src/HOL/Auth/NS_Public.thy
author blanchet
Wed, 12 Dec 2012 03:47:02 +0100
changeset 50486 d5dc28fafd9d
parent 37936 1e4c5015a72e
child 56681 e8d5d60d655e
permissions -rw-r--r--
made MaSh evaluation driver work with SMT solvers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 32960
diff changeset
     1
(*  Title:      HOL/Auth/NS_Public.thy
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     3
    Copyright   1996  University of Cambridge
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     4
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     5
Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
2538
c55f68761a8d Mended spelling error
paulson
parents: 2516
diff changeset
     6
Version incorporating Lowe's fix (inclusion of B's identity in round 2).
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     7
*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     8
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13926
diff changeset
     9
header{*Verifying the Needham-Schroeder-Lowe Public-Key Protocol*}
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13926
diff changeset
    10
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14207
diff changeset
    11
theory NS_Public imports Public begin
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    12
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    13
inductive_set ns_public :: "event list set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    14
  where 
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    15
         (*Initial trace is empty*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    16
   Nil:  "[] \<in> ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    17
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    18
         (*The spy MAY say anything he CAN say.  We do not expect him to
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    19
           invent new nonces here, but he can also use NS1.  Common to
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    20
           all similar protocols.*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    21
 | Fake: "\<lbrakk>evsf \<in> ns_public;  X \<in> synth (analz (spies evsf))\<rbrakk>
11366
b42287eb20cf renaming of evs in the Fake rule
paulson
parents: 11230
diff changeset
    22
          \<Longrightarrow> Says Spy B X  # evsf \<in> ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    23
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    24
         (*Alice initiates a protocol run, sending a nonce to Bob*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    25
 | NS1:  "\<lbrakk>evs1 \<in> ns_public;  Nonce NA \<notin> used evs1\<rbrakk>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    26
          \<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    27
                 # evs1  \<in>  ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    28
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    29
         (*Bob responds to Alice's message with a further nonce*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    30
 | NS2:  "\<lbrakk>evs2 \<in> ns_public;  Nonce NB \<notin> used evs2;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    31
           Says A' B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    32
          \<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    33
                # evs2  \<in>  ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    34
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    35
         (*Alice proves her existence by sending NB back to Bob.*)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    36
 | NS3:  "\<lbrakk>evs3 \<in> ns_public;
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    37
           Says A  B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    38
           Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    39
              \<in> set evs3\<rbrakk>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    40
          \<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) # evs3 \<in> ns_public"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    41
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    42
declare knows_Spy_partsEs [elim]
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
    43
declare knows_Spy_partsEs [elim]
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
    44
declare analz_into_parts [dest]
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
    45
declare Fake_parts_insert_in_Un  [dest]
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    46
declare image_eq_UN [simp]  (*accelerates proofs involving nested images*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    47
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    48
(*A "possibility property": there are traces that reach the end*)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    49
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    50
apply (intro exI bexI)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    51
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, 
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    52
                                   THEN ns_public.NS3], possibility)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    53
done
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    54
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    55
(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    56
    sends messages containing X! **)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    57
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    58
(*Spy never sees another agent's private key! (unless it's bad at start)*)
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    59
lemma Spy_see_priEK [simp]: 
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    60
      "evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> parts (spies evs)) = (A \<in> bad)"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    61
by (erule ns_public.induct, auto)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    62
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    63
lemma Spy_analz_priEK [simp]: 
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    64
      "evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> analz (spies evs)) = (A \<in> bad)"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    65
by auto
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    66
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
    67
subsection{*Authenticity properties obtained from NS2*}
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    68
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    69
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    70
(*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    71
  is secret.  (Honest users generate fresh nonces.)*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    72
lemma no_nonce_NS1_NS2 [rule_format]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    73
      "evs \<in> ns_public 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    74
       \<Longrightarrow> Crypt (pubEK C) \<lbrace>NA', Nonce NA, Agent D\<rbrace> \<in> parts (spies evs) \<longrightarrow>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    75
           Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>  
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    76
           Nonce NA \<in> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    77
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    78
apply (blast intro: analz_insertI)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    79
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    80
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    81
(*Unicity for NS1: nonce NA identifies agents A and B*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    82
lemma unique_NA: 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    83
     "\<lbrakk>Crypt(pubEK B)  \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(spies evs);  
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    84
       Crypt(pubEK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(spies evs);  
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    85
       Nonce NA \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    86
      \<Longrightarrow> A=A' \<and> B=B'"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    87
apply (erule rev_mp, erule rev_mp, erule rev_mp)   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    88
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    89
(*Fake, NS1*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    90
apply (blast intro: analz_insertI)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    91
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    92
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    93
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    94
(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    95
  The major premise "Says A B ..." makes it a dest-rule, so we use
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    96
  (erule rev_mp) rather than rule_format. *)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    97
theorem Spy_not_see_NA: 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
    98
      "\<lbrakk>Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    99
        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   100
       \<Longrightarrow> Nonce NA \<notin> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   101
apply (erule rev_mp)   
13507
febb8e5d2a9d tidying of Isar scripts
paulson
parents: 11366
diff changeset
   102
apply (erule ns_public.induct, simp_all, spy_analz)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   103
apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   104
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   105
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   106
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   107
(*Authentication for A: if she receives message 2 and has used NA
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   108
  to start a run, then B has sent message 2.*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   109
lemma A_trusts_NS2_lemma [rule_format]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   110
   "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   111
    \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   112
        Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   113
        Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   114
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   115
(*Fake, NS1*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   116
apply (blast dest: Spy_not_see_NA)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   117
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   118
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   119
theorem A_trusts_NS2: 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   120
     "\<lbrakk>Says A  B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;   
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   121
       Says B' A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   122
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   123
      \<Longrightarrow> Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   124
by (blast intro: A_trusts_NS2_lemma)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   125
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   126
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   127
(*If the encrypted message appears then it originated with Alice in NS1*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   128
lemma B_trusts_NS1 [rule_format]:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   129
     "evs \<in> ns_public                                         
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   130
      \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   131
          Nonce NA \<notin> analz (spies evs) \<longrightarrow>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 23746
diff changeset
   132
          Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   133
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   134
(*Fake*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   135
apply (blast intro!: analz_insertI)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   136
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   137
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   138
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
   139
subsection{*Authenticity properties obtained from NS2*}
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   140
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   141
(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   142
  [unicity of B makes Lowe's fix work]
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   143
  [proof closely follows that for unique_NA] *)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   144
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   145
lemma unique_NB [dest]: 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   146
     "\<lbrakk>Crypt(pubEK A)  \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts(spies evs);
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   147
       Crypt(pubEK A') \<lbrace>Nonce NA', Nonce NB, Agent B'\<rbrace> \<in> parts(spies evs);
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   148
       Nonce NB \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   149
      \<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B'"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   150
apply (erule rev_mp, erule rev_mp, erule rev_mp)   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   151
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   152
(*Fake, NS2*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   153
apply (blast intro: analz_insertI)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   154
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   155
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   156
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   157
(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   158
theorem Spy_not_see_NB [dest]:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   159
     "\<lbrakk>Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   160
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   161
      \<Longrightarrow> Nonce NB \<notin> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   162
apply (erule rev_mp)
13507
febb8e5d2a9d tidying of Isar scripts
paulson
parents: 11366
diff changeset
   163
apply (erule ns_public.induct, simp_all, spy_analz)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   164
apply (blast intro: no_nonce_NS1_NS2)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   165
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   166
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   167
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   168
(*Authentication for B: if he receives message 3 and has used NB
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   169
  in message 2, then A has sent message 3.*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   170
lemma B_trusts_NS3_lemma [rule_format]:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   171
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   172
      Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   173
      Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   174
      Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   175
by (erule ns_public.induct, auto)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   176
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   177
theorem B_trusts_NS3:
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   178
     "\<lbrakk>Says B A  (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   179
       Says A' B (Crypt (pubEK B) (Nonce NB)) \<in> set evs;             
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   180
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                    
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   181
      \<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   182
by (blast intro: B_trusts_NS3_lemma)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   183
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13956
diff changeset
   184
subsection{*Overall guarantee for B*}
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   185
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   186
(*If NS3 has been sent and the nonce NB agrees with the nonce B joined with
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   187
  NA, then A initiated the run using NA.*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   188
theorem B_trusts_protocol:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   189
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk> \<Longrightarrow>
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   190
      Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   191
      Says B A  (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 13507
diff changeset
   192
      Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   193
by (erule ns_public.induct, auto)
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   194
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   195
end