src/HOL/Auth/NS_Public_Bad.thy
author paulson
Tue, 13 Feb 2001 13:16:27 +0100
changeset 11104 f2024fed9f0c
parent 5434 9b4bed3f394c
child 11150 67387142225e
permissions -rw-r--r--
partial conversion to Isar script style simplified unicity proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/NS_Public_Bad
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     2
    ID:         $Id$
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     5
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     6
Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol.
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     7
Flawed version, vulnerable to Lowe's attack.
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     8
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     9
From page 260 of
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    10
  Burrows, Abadi and Needham.  A Logic of Authentication.
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    11
  Proc. Royal Soc. 426 (1989)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    12
*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    13
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    14
theory NS_Public_Bad = Public:
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    15
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    16
consts  ns_public  :: "event list set"
2549
0c723635b9e6 Cosmetic improvements
paulson
parents: 2516
diff changeset
    17
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    18
inductive ns_public
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    19
  intros 
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    20
         (*Initial trace is empty*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    21
   Nil:  "[] \<in> ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    22
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    23
         (*The spy MAY say anything he CAN say.  We do not expect him to
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    24
           invent new nonces here, but he can also use NS1.  Common to
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    25
           all similar protocols.*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    26
   Fake: "\<lbrakk>evs \<in> ns_public;  X \<in> synth (analz (spies evs))\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    27
          \<Longrightarrow> Says Spy B X  # evs \<in> ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    28
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    29
         (*Alice initiates a protocol run, sending a nonce to Bob*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    30
   NS1:  "\<lbrakk>evs1 \<in> ns_public;  Nonce NA \<notin> used evs1\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    31
          \<Longrightarrow> Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    32
                # evs1  \<in>  ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    33
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    34
         (*Bob responds to Alice's message with a further nonce*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    35
   NS2:  "\<lbrakk>evs2 \<in> ns_public;  Nonce NB \<notin> used evs2;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    36
           Says A' B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    37
          \<Longrightarrow> Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    38
                # evs2  \<in>  ns_public"
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    39
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    40
         (*Alice proves her existence by sending NB back to Bob.*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    41
   NS3:  "\<lbrakk>evs3 \<in> ns_public;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    42
           Says A  B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    43
           Says B' A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    44
          \<Longrightarrow> Says A B (Crypt (pubK B) (Nonce NB)) # evs3 \<in> ns_public"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    45
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    46
  (*No Oops message.  Should there be one?*)
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
declare knows_Spy_partsEs [elim]
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    49
declare analz_subset_parts [THEN subsetD, dest]
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    50
declare Fake_parts_insert [THEN subsetD, dest]
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    51
declare image_eq_UN [simp]  (*accelerates proofs involving nested images*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    52
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    53
(*A "possibility property": there are traces that reach the end*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    54
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubK B) (Nonce NB)) \<in> set evs"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    55
apply (intro exI bexI)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    56
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    57
                                   THEN ns_public.NS3])
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    58
by possibility
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    59
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    60
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    61
(**** Inductive proofs about ns_public ****)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    62
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    63
(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    64
    sends messages containing X! **)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    65
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    66
(*Spy never sees another agent's private key! (unless it's bad at start)*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    67
lemma Spy_see_priK [simp]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    68
      "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    69
by (erule ns_public.induct, auto)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    70
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    71
lemma Spy_analz_priK [simp]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    72
      "evs \<in> ns_public \<Longrightarrow> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    73
by auto
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    74
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    75
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    76
(*** Authenticity properties obtained from NS2 ***)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    77
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    78
(*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
    79
  is secret.  (Honest users generate fresh nonces.)*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    80
lemma no_nonce_NS1_NS2 [rule_format]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    81
      "evs \<in> ns_public 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    82
       \<Longrightarrow> Crypt (pubK C) \<lbrace>NA', Nonce NA\<rbrace> \<in> parts (spies evs) \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    83
           Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>  
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    84
           Nonce NA \<in> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    85
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    86
apply (blast intro: analz_insertI)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    87
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    88
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    89
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    90
(*Unicity for NS1: nonce NA identifies agents A and B*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    91
lemma unique_NA: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    92
     "\<lbrakk>Crypt(pubK B)  \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(spies evs);  
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    93
       Crypt(pubK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(spies evs);  
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    94
       Nonce NA \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    95
      \<Longrightarrow> A=A' \<and> B=B'"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    96
apply (erule rev_mp, erule rev_mp, erule rev_mp)   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    97
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    98
(*Fake, NS1*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
    99
apply (blast intro!: analz_insertI)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   100
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   101
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   102
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   103
(*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
   104
  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
   105
  (erule rev_mp) rather than rule_format. *)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   106
theorem Spy_not_see_NA: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   107
      "\<lbrakk>Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   108
        A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   109
       \<Longrightarrow> Nonce NA \<notin> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   110
apply (erule rev_mp)   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   111
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   112
apply spy_analz
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   113
apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   114
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   115
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   116
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   117
(*Authentication for A: if she receives message 2 and has used NA
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   118
  to start a run, then B has sent message 2.*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   119
lemma A_trusts_NS2_lemma [rule_format]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   120
   "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   121
    \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   122
	Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   123
	Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   124
apply (erule ns_public.induct)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   125
apply (auto dest: Spy_not_see_NA unique_NA)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   126
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   127
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   128
theorem A_trusts_NS2: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   129
     "\<lbrakk>Says A  B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   130
       Says B' A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   131
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                     
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   132
      \<Longrightarrow> Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   133
by (blast intro: A_trusts_NS2_lemma)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   134
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   135
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   136
(*If the encrypted message appears then it originated with Alice in NS1*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   137
lemma B_trusts_NS1 [rule_format]:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   138
     "evs \<in> ns_public                                         
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   139
      \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   140
	  Nonce NA \<notin> analz (spies evs) \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   141
	  Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   142
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   143
(*Fake*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   144
apply (blast intro!: analz_insertI)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   145
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   146
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   147
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   148
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   149
(*** Authenticity properties obtained from NS2 ***)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   150
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   151
(*Unicity for NS2: nonce NB identifies nonce NA and agent A
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   152
  [proof closely follows that for unique_NA] *)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   153
lemma unique_NB [dest]: 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   154
     "\<lbrakk>Crypt(pubK A)  \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts(spies evs);
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   155
       Crypt(pubK A') \<lbrace>Nonce NA', Nonce NB\<rbrace> \<in> parts(spies evs);
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   156
       Nonce NB \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   157
     \<Longrightarrow> A=A' \<and> NA=NA'"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   158
apply (erule rev_mp, erule rev_mp, erule rev_mp)   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   159
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   160
(*Fake, NS2*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   161
apply (blast intro!: analz_insertI)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   162
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   163
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   164
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   165
(*NB remains secret PROVIDED Alice never responds with round 3*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   166
theorem Spy_not_see_NB [dest]:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   167
     "\<lbrakk>Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs;   
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   168
       \<forall>C. Says A C (Crypt (pubK C) (Nonce NB)) \<notin> set evs;       
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   169
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                      
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   170
     \<Longrightarrow> Nonce NB \<notin> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   171
apply (erule rev_mp, erule rev_mp)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   172
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   173
apply spy_analz
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   174
apply (simp_all add: all_conj_distrib) (*speeds up the next step*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   175
apply (blast intro: no_nonce_NS1_NS2)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   176
done
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   177
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   178
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   179
(*Authentication for B: if he receives message 3 and has used NB
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   180
  in message 2, then A has sent message 3--to somebody....*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   181
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   182
lemma B_trusts_NS3_lemma [rule_format]:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   183
     "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                    
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   184
      \<Longrightarrow> Crypt (pubK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   185
          Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   186
          (\<exists>C. Says A C (Crypt (pubK C) (Nonce NB)) \<in> set evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   187
apply (erule ns_public.induct, auto)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   188
by (blast intro: no_nonce_NS1_NS2)+
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   189
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   190
theorem B_trusts_NS3:
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   191
     "\<lbrakk>Says B A  (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   192
       Says A' B (Crypt (pubK B) (Nonce NB)) \<in> set evs;             
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   193
       A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>                    
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   194
      \<Longrightarrow> \<exists>C. Says A C (Crypt (pubK C) (Nonce NB)) \<in> set evs"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   195
by (blast intro: B_trusts_NS3_lemma)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   196
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   197
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   198
(*Can we strengthen the secrecy theorem Spy_not_see_NB?  NO*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   199
lemma "\<lbrakk>A \<notin> bad;  B \<notin> bad;  evs \<in> ns_public\<rbrakk>            
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   200
       \<Longrightarrow> Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs  
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   201
           \<longrightarrow> Nonce NB \<notin> analz (spies evs)"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   202
apply (erule ns_public.induct, simp_all)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   203
apply spy_analz
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   204
(*NS1: by freshness*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   205
apply (blast)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   206
(*NS2: by freshness and unicity of NB*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   207
apply (blast intro: no_nonce_NS1_NS2)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   208
(*NS3: unicity of NB identifies A and NA, but not B*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   209
apply clarify
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   210
apply (frule_tac A' = "A" in 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   211
       Says_imp_knows_Spy [THEN parts.Inj, THEN unique_NB])
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   212
apply auto
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   213
apply (rename_tac C B' evs3)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   214
oops
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   215
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   216
(*
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   217
THIS IS THE ATTACK!
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   218
Level 8
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   219
!!evs. \<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   220
       \<Longrightarrow> Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs \<longrightarrow>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   221
           Nonce NB \<notin> analz (spies evs)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   222
 1. !!C B' evs3.
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   223
       \<lbrakk>A \<notin> bad; B \<notin> bad; evs3 \<in> ns_public
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   224
        Says A C (Crypt (pubK C) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   225
        Says B' A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3; 
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   226
        C \<in> bad;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   227
        Says B A (Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3;
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   228
        Nonce NB \<notin> analz (spies evs3)\<rbrakk>
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   229
       \<Longrightarrow> False
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 5434
diff changeset
   230
*)
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   231
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   232
end