doc-src/TutorialI/Protocol/NS_Public.thy
author wenzelm
Fri, 21 May 2004 21:24:22 +0200
changeset 14787 724ce6e574e3
parent 11269 4095353bd0d7
child 16417 9bc16273c2d4
permissions -rw-r--r--
adapted tsig interface;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/NS_Public
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     2
    ID:         $Id$
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     5
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     6
Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     7
Version incorporating Lowe's fix (inclusion of B's identity in round 2).
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     8
*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
     9
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    10
theory NS_Public = Public:
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    11
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    12
consts  ns_public  :: "event list set"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    13
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    14
inductive ns_public
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    15
  intros
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    16
         (*Initial trace is empty*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    17
   Nil:  "[] \<in> ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    18
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    19
         (*The spy MAY say anything he CAN say.  We do not expect him to
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    20
           invent new nonces here, but he can also use NS1.  Common to
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    21
           all similar protocols.*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    22
   Fake: "\<lbrakk>evs \<in> ns_public;  X \<in> synth (analz (knows Spy evs))\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    23
          \<Longrightarrow> Says Spy B X  # evs \<in> ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    24
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    25
         (*Alice initiates a protocol run, sending a nonce to Bob*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    26
   NS1:  "\<lbrakk>evs1 \<in> ns_public;  Nonce NA \<notin> used evs1\<rbrakk>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    27
          \<Longrightarrow> Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    28
                 # evs1  \<in>  ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    29
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    30
         (*Bob responds to Alice's message with a further nonce*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    31
   NS2:  "\<lbrakk>evs2 \<in> ns_public;  Nonce NB \<notin> used evs2;
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    32
           Says A' B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    33
          \<Longrightarrow> Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    34
                # evs2  \<in>  ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    35
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    36
         (*Alice proves her existence by sending NB back to Bob.*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    37
   NS3:  "\<lbrakk>evs3 \<in> ns_public;
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    38
           Says A  B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    39
           Says B' A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    40
              \<in> set evs3\<rbrakk>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    41
          \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) # evs3 \<in> ns_public"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    42
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    43
declare knows_Spy_partsEs [elim]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    44
declare analz_subset_parts [THEN subsetD, dest]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    45
declare Fake_parts_insert [THEN subsetD, dest]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    46
declare image_eq_UN [simp]  (*accelerates proofs involving nested images*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    47
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    48
(*A "possibility property": there are traces that reach the end*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    49
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    50
apply (intro exI bexI)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    51
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2,
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    52
                                   THEN ns_public.NS3])
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    53
by possibility
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    54
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    55
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    56
(**** Inductive proofs about ns_public ****)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    57
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    58
(** Theorems of the form X \<notin> parts (knows Spy evs) imply that NOBODY
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    59
    sends messages containing X! **)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    60
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    61
(*Spy never sees another agent's private key! (unless it's bad at start)*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    62
lemma Spy_see_priK [simp]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    63
      "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    64
by (erule ns_public.induct, auto)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    65
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    66
lemma Spy_analz_priK [simp]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    67
      "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    68
by auto
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    69
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    70
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    71
(*** Authenticity properties obtained from NS2 ***)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    72
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    73
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    74
(*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    75
  is secret.  (Honest users generate fresh nonces.)*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    76
lemma no_nonce_NS1_NS2:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    77
      "\<lbrakk>Crypt (pubK C) \<lbrace>NA', Nonce NA, Agent D\<rbrace> \<in> parts (knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    78
        Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    79
        evs \<in> ns_public\<rbrakk>
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    80
       \<Longrightarrow> Nonce NA \<in> analz (knows Spy evs)"
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    81
apply (erule rev_mp, erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    82
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    83
apply (blast intro: analz_insertI)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    84
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    85
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    86
(*Unicity for NS1: nonce NA identifies agents A and B*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    87
lemma unique_NA:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    88
     "\<lbrakk>Crypt(pubK B)  \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    89
       Crypt(pubK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    90
       Nonce NA \<notin> analz (knows Spy evs); evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    91
      \<Longrightarrow> A=A' \<and> B=B'"
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
    92
apply (erule rev_mp, erule rev_mp, erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    93
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    94
(*Fake, NS1*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    95
apply (blast intro: analz_insertI)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    96
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    97
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    98
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
    99
(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   100
  The major premise "Says A B ..." makes it a dest-rule, so we use
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   101
  (erule rev_mp) rather than rule_format. *)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   102
theorem Spy_not_see_NA:
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   103
      "\<lbrakk>Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   104
        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   105
       \<Longrightarrow> Nonce NA \<notin> analz (knows Spy evs)"
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   106
apply (erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   107
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   108
apply spy_analz
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   109
apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   110
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   111
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   112
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   113
(*Authentication for A: if she receives message 2 and has used NA
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   114
  to start a run, then B has sent message 2.*)
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   115
lemma A_trusts_NS2_lemma [rule_format]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   116
   "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   117
    \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   118
	Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   119
	Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   120
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   121
(*Fake, NS1*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   122
apply (blast dest: Spy_not_see_NA)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   123
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   124
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   125
theorem A_trusts_NS2:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   126
     "\<lbrakk>Says A  B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   127
       Says B' A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   128
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   129
      \<Longrightarrow> Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   130
by (blast intro: A_trusts_NS2_lemma)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   131
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   132
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   133
(*If the encrypted message appears then it originated with Alice in NS1*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   134
lemma B_trusts_NS1 [rule_format]:
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   135
     "evs \<in> ns_public
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   136
      \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   137
	  Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   138
	  Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   139
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   140
(*Fake*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   141
apply (blast intro!: analz_insertI)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   142
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   143
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   144
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   145
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   146
(*** Authenticity properties obtained from NS2 ***)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   147
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   148
(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   149
  [unicity of B makes Lowe's fix work]
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   150
  [proof closely follows that for unique_NA] *)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   151
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   152
lemma unique_NB [dest]:
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   153
     "\<lbrakk>Crypt(pubK A)  \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   154
       Crypt(pubK A') \<lbrace>Nonce NA', Nonce NB, Agent B'\<rbrace> \<in> parts(knows Spy evs);
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   155
       Nonce NB \<notin> analz (knows Spy evs); evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   156
      \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B'"
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   157
apply (erule rev_mp, erule rev_mp, erule rev_mp)
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   158
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   159
(*Fake, NS2*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   160
apply (blast intro: analz_insertI)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   161
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   162
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   163
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   164
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   165
text{*
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   166
@{thm[display] analz_Crypt_if[no_vars]}
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   167
\rulename{analz_Crypt_if}
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   168
*}
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   169
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   170
(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   171
theorem Spy_not_see_NB [dest]:
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   172
     "\<lbrakk>Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   173
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   174
      \<Longrightarrow> Nonce NB \<notin> analz (knows Spy evs)"
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   175
apply (erule rev_mp)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   176
apply (erule ns_public.induct, simp_all)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   177
apply spy_analz
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   178
apply (blast intro: no_nonce_NS1_NS2)+
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   179
done
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   180
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   181
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   182
(*Authentication for B: if he receives message 3 and has used NB
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   183
  in message 2, then A has sent message 3.*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   184
lemma B_trusts_NS3_lemma [rule_format]:
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   185
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   186
      Crypt (pubK B) (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   187
      Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   188
      Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   189
by (erule ns_public.induct, auto)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   190
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   191
theorem B_trusts_NS3:
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   192
     "\<lbrakk>Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   193
       Says A' B (Crypt (pubK B) (Nonce NB)) \<in> set evs;
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   194
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   195
      \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   196
by (blast intro: B_trusts_NS3_lemma)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   197
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   198
(*** Overall guarantee for B ***)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   199
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   200
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   201
(*If NS3 has been sent and the nonce NB agrees with the nonce B joined with
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   202
  NA, then A initiated the run using NA.*)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   203
theorem B_trusts_protocol:
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   204
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
11269
4095353bd0d7 changes specifically for the book version
paulson
parents: 11250
diff changeset
   205
      Crypt (pubK B) (Nonce NB) \<in> parts (knows Spy evs) \<longrightarrow>
11250
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   206
      Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   207
      Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   208
by (erule ns_public.induct, auto)
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   209
c8bbf4c4bc2d symlinks to ../../../HOL/Auth. Fingers crossed...
paulson
parents:
diff changeset
   210
end