| 
18886
 | 
     1  | 
(*  ID:         $Id$
  | 
| 
 | 
     2  | 
    Author:     Giampaolo Bella, Catania University
  | 
| 
 | 
     3  | 
*)
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
header{*Bella's version of the Otway-Rees protocol*}
 | 
| 
 | 
     6  | 
  | 
| 
 | 
     7  | 
  | 
| 
 | 
     8  | 
theory OtwayReesBella imports Public begin
  | 
| 
 | 
     9  | 
  | 
| 
 | 
    10  | 
text{*Bella's modifications to a version of the Otway-Rees protocol taken from
 | 
| 
 | 
    11  | 
the BAN paper only concern message 7. The updated protocol makes the goal of
  | 
| 
 | 
    12  | 
key distribution of the session key available to A. Investigating the
  | 
| 
 | 
    13  | 
principle of Goal Availability undermines the BAN claim about the original
  | 
| 
 | 
    14  | 
protocol, that "this protocol does not make use of Kab as an encryption key,
  | 
| 
 | 
    15  | 
so neither principal can know whether the key is known to the other". The
  | 
| 
 | 
    16  | 
updated protocol makes no use of the session key to encrypt but informs A that
  | 
| 
 | 
    17  | 
B knows it.*}
  | 
| 
 | 
    18  | 
  | 
| 
 | 
    19  | 
consts  orb   :: "event list set"
  | 
| 
 | 
    20  | 
inductive "orb"
  | 
| 
 | 
    21  | 
 intros 
  | 
| 
 | 
    22  | 
  | 
| 
 | 
    23  | 
  Nil:  "[]\<in> orb"
  | 
| 
 | 
    24  | 
  | 
| 
 | 
    25  | 
  Fake: "\<lbrakk>evsa\<in> orb;  X\<in> synth (analz (knows Spy evsa))\<rbrakk>
  | 
| 
 | 
    26  | 
 	 \<Longrightarrow> Says Spy B X  # evsa \<in> orb"
  | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
  Reception: "\<lbrakk>evsr\<in> orb;  Says A B X \<in> set evsr\<rbrakk>
  | 
| 
 | 
    29  | 
	      \<Longrightarrow> Gets B X # evsr \<in> orb"
  | 
| 
 | 
    30  | 
  | 
| 
 | 
    31  | 
  OR1:  "\<lbrakk>evs1\<in> orb;  Nonce NA \<notin> used evs1\<rbrakk>
  | 
| 
 | 
    32  | 
	 \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, 
  | 
| 
 | 
    33  | 
		   Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
  | 
| 
 | 
    34  | 
	       # evs1 \<in> orb"
  | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
  OR2:  "\<lbrakk>evs2\<in> orb;  Nonce NB \<notin> used evs2;
  | 
| 
 | 
    37  | 
	   Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
  | 
| 
 | 
    38  | 
	\<Longrightarrow> Says B Server 
  | 
| 
 | 
    39  | 
		\<lbrace>Nonce M, Agent A, Agent B, X, 
  | 
| 
 | 
    40  | 
	   Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
  | 
| 
 | 
    41  | 
	       # evs2 \<in> orb"
  | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
  OR3:  "\<lbrakk>evs3\<in> orb;  Key KAB \<notin> used evs3;
  | 
| 
 | 
    44  | 
	  Gets Server 
  | 
| 
 | 
    45  | 
	     \<lbrace>Nonce M, Agent A, Agent B, 
  | 
| 
 | 
    46  | 
	       Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>, 
  | 
| 
 | 
    47  | 
	       Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
  | 
| 
 | 
    48  | 
	  \<in> set evs3\<rbrakk>
  | 
| 
 | 
    49  | 
	\<Longrightarrow> Says Server B \<lbrace>Nonce M,
  | 
| 
 | 
    50  | 
		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
  | 
| 
 | 
    51  | 
				      Nonce NB, Key KAB\<rbrace>\<rbrace>
  | 
| 
 | 
    52  | 
	       # evs3 \<in> orb"
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
  (*B can only check that the message he is bouncing is a ciphertext*)
  | 
| 
 | 
    55  | 
  (*Sending M back is omitted*)   
  | 
| 
 | 
    56  | 
  OR4:  "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>; 
  | 
| 
 | 
    57  | 
	  Says B Server \<lbrace>Nonce M, Agent A, Agent B, X', 
  | 
| 
 | 
    58  | 
		Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
  | 
| 
 | 
    59  | 
	    \<in> set evs4;
  | 
| 
 | 
    60  | 
	  Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
  | 
| 
 | 
    61  | 
	    \<in> set evs4\<rbrakk>
  | 
| 
 | 
    62  | 
	\<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
  | 
| 
 | 
    63  | 
  | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
  Oops: "\<lbrakk>evso\<in> orb;  
  | 
| 
 | 
    66  | 
	   Says Server B \<lbrace>Nonce M,
  | 
| 
 | 
    67  | 
		    Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
  | 
| 
 | 
    68  | 
				      Nonce NB, Key KAB\<rbrace>\<rbrace> 
  | 
| 
 | 
    69  | 
	     \<in> set evso\<rbrakk>
  | 
| 
 | 
    70  | 
 \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso 
  | 
| 
 | 
    71  | 
     \<in> orb"
  | 
| 
 | 
    72  | 
  | 
| 
 | 
    73  | 
  | 
| 
 | 
    74  | 
  | 
| 
 | 
    75  | 
declare knows_Spy_partsEs [elim]
  | 
| 
 | 
    76  | 
declare analz_into_parts [dest]
  | 
| 
 | 
    77  | 
declare Fake_parts_insert_in_Un  [dest]
  | 
| 
 | 
    78  | 
  | 
| 
 | 
    79  | 
  | 
| 
 | 
    80  | 
text{*Fragile proof, with backtracking in the possibility call.*}
 | 
| 
 | 
    81  | 
lemma possibility_thm: "\<lbrakk>A \<noteq> Server; B \<noteq> Server; Key K \<notin> used[]\<rbrakk>    
  | 
| 
 | 
    82  | 
      \<Longrightarrow>   \<exists> evs \<in> orb.           
  | 
| 
 | 
    83  | 
     Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"
  | 
| 
 | 
    84  | 
apply (intro exI bexI)
  | 
| 
 | 
    85  | 
apply (rule_tac [2] orb.Nil
  | 
| 
 | 
    86  | 
                    [THEN orb.OR1, THEN orb.Reception,
  | 
| 
 | 
    87  | 
                     THEN orb.OR2, THEN orb.Reception,
  | 
| 
 | 
    88  | 
                     THEN orb.OR3, THEN orb.Reception, THEN orb.OR4]) 
  | 
| 
 | 
    89  | 
apply (possibility, simp add: used_Cons)  
  | 
| 
 | 
    90  | 
done
  | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
  | 
| 
 | 
    93  | 
lemma Gets_imp_Says :
  | 
| 
 | 
    94  | 
     "\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs"
  | 
| 
 | 
    95  | 
apply (erule rev_mp)
  | 
| 
 | 
    96  | 
apply (erule orb.induct)
  | 
| 
 | 
    97  | 
apply auto
  | 
| 
 | 
    98  | 
done
  | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
lemma Gets_imp_knows_Spy: 
  | 
| 
 | 
   101  | 
     "\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk>  \<Longrightarrow> X \<in> knows Spy evs"
  | 
| 
 | 
   102  | 
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy)
  | 
| 
 | 
   103  | 
  | 
| 
 | 
   104  | 
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
  | 
| 
 | 
   105  | 
  | 
| 
 | 
   106  | 
lemma Gets_imp_knows:
  | 
| 
 | 
   107  | 
     "\<lbrakk>Gets B X \<in> set evs; evs \<in> orb\<rbrakk>  \<Longrightarrow> X \<in> knows B evs"
  | 
| 
 | 
   108  | 
apply (case_tac "B = Spy")
  | 
| 
 | 
   109  | 
apply (blast dest!: Gets_imp_knows_Spy)
  | 
| 
 | 
   110  | 
apply (blast dest!: Gets_imp_knows_agents)
  | 
| 
 | 
   111  | 
done
  | 
| 
 | 
   112  | 
  | 
| 
 | 
   113  | 
  | 
| 
 | 
   114  | 
lemma OR2_analz_knows_Spy: 
  | 
| 
 | 
   115  | 
   "\<lbrakk>Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs; evs \<in> orb\<rbrakk>   
  | 
| 
 | 
   116  | 
    \<Longrightarrow> X \<in> analz (knows Spy evs)"
  | 
| 
 | 
   117  | 
by (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
  | 
| 
 | 
   118  | 
  | 
| 
 | 
   119  | 
lemma OR4_parts_knows_Spy: 
  | 
| 
 | 
   120  | 
   "\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key Kab\<rbrace>\<rbrace>  \<in> set evs; 
  | 
| 
 | 
   121  | 
      evs \<in> orb\<rbrakk>   \<Longrightarrow> X \<in> parts (knows Spy evs)"
  | 
| 
 | 
   122  | 
by blast
  | 
| 
 | 
   123  | 
  | 
| 
 | 
   124  | 
lemma Oops_parts_knows_Spy: 
  | 
| 
 | 
   125  | 
    "Says Server B \<lbrace>Nonce M, Crypt K' \<lbrace>X, Nonce Nb, K\<rbrace>\<rbrace> \<in> set evs  
  | 
| 
 | 
   126  | 
     \<Longrightarrow> K \<in> parts (knows Spy evs)"
  | 
| 
 | 
   127  | 
by blast
  | 
| 
 | 
   128  | 
  | 
| 
 | 
   129  | 
lemmas OR2_parts_knows_Spy =
  | 
| 
 | 
   130  | 
    OR2_analz_knows_Spy [THEN analz_into_parts, standard]
  | 
| 
 | 
   131  | 
  | 
| 
 | 
   132  | 
ML
  | 
| 
 | 
   133  | 
{*
 | 
| 
 | 
   134  | 
val Oops_parts_knows_Spy = thm "Oops_parts_knows_Spy"
  | 
| 
 | 
   135  | 
val OR4_parts_knows_Spy = thm "OR4_parts_knows_Spy"
  | 
| 
 | 
   136  | 
val OR2_parts_knows_Spy = thm "OR2_parts_knows_Spy"
  | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
fun parts_explicit_tac i = 
  | 
| 
 | 
   139  | 
    forward_tac [Oops_parts_knows_Spy] (i+7) THEN
  | 
| 
 | 
   140  | 
    forward_tac [OR4_parts_knows_Spy]  (i+6) THEN
  | 
| 
 | 
   141  | 
    forward_tac [OR2_parts_knows_Spy]  (i+4)
  | 
| 
 | 
   142  | 
*}
  | 
| 
 | 
   143  | 
 
  | 
| 
 | 
   144  | 
method_setup parts_explicit = {*
 | 
| 
21588
 | 
   145  | 
    Method.no_args (Method.SIMPLE_METHOD' parts_explicit_tac) *}
  | 
| 
18886
 | 
   146  | 
  "to explicitly state that some message components belong to parts knows Spy"
  | 
| 
 | 
   147  | 
  | 
| 
 | 
   148  | 
  | 
| 
 | 
   149  | 
lemma Spy_see_shrK [simp]: 
  | 
| 
 | 
   150  | 
    "evs \<in> orb \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
  | 
| 
 | 
   151  | 
by (erule orb.induct, parts_explicit, simp_all, blast+)
  | 
| 
 | 
   152  | 
  | 
| 
 | 
   153  | 
lemma Spy_analz_shrK [simp]: 
  | 
| 
 | 
   154  | 
"evs \<in> orb \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
  | 
| 
 | 
   155  | 
by auto
  | 
| 
 | 
   156  | 
  | 
| 
 | 
   157  | 
lemma Spy_see_shrK_D [dest!]:
  | 
| 
 | 
   158  | 
     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> orb|] ==> A \<in> bad"
  | 
| 
 | 
   159  | 
by (blast dest: Spy_see_shrK)
  | 
| 
 | 
   160  | 
  | 
| 
 | 
   161  | 
lemma new_keys_not_used [simp]:
  | 
| 
 | 
   162  | 
   "\<lbrakk>Key K \<notin> used evs; K \<in> symKeys; evs \<in> orb\<rbrakk>  \<Longrightarrow> K \<notin> keysFor (parts (knows Spy evs))"
  | 
| 
 | 
   163  | 
apply (erule rev_mp)
  | 
| 
 | 
   164  | 
apply (erule orb.induct, parts_explicit, simp_all)
  | 
| 
 | 
   165  | 
apply (force dest!: keysFor_parts_insert)
  | 
| 
 | 
   166  | 
apply (blast+)
  | 
| 
 | 
   167  | 
done
  | 
| 
 | 
   168  | 
  | 
| 
 | 
   169  | 
  | 
| 
 | 
   170  | 
  | 
| 
 | 
   171  | 
subsection{* Proofs involving analz *}
 | 
| 
 | 
   172  | 
  | 
| 
 | 
   173  | 
text{*Describes the form of K and NA when the Server sends this message.  Also
 | 
| 
 | 
   174  | 
  for Oops case.*}
  | 
| 
 | 
   175  | 
lemma Says_Server_message_form: 
  | 
| 
 | 
   176  | 
"\<lbrakk>Says Server B  \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  
  | 
| 
 | 
   177  | 
     evs \<in> orb\<rbrakk>                                            
  | 
| 
 | 
   178  | 
 \<Longrightarrow> K \<notin> range shrK & (\<exists> A Na. X=(Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>))"
  | 
| 
 | 
   179  | 
by (erule rev_mp, erule orb.induct, simp_all)
  | 
| 
 | 
   180  | 
  | 
| 
 | 
   181  | 
lemma Says_Server_imp_Gets: 
  | 
| 
 | 
   182  | 
 "\<lbrakk>Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>,
  | 
| 
 | 
   183  | 
                                             Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;
  | 
| 
 | 
   184  | 
    evs \<in> orb\<rbrakk>
  | 
| 
 | 
   185  | 
  \<Longrightarrow>  Gets Server \<lbrace>Nonce M, Agent A, Agent B, 
  | 
| 
 | 
   186  | 
                   Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>, 
  | 
| 
 | 
   187  | 
               Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
  | 
| 
 | 
   188  | 
         \<in> set evs"
  | 
| 
 | 
   189  | 
by (erule rev_mp, erule orb.induct, simp_all)
  | 
| 
 | 
   190  | 
  | 
| 
 | 
   191  | 
  | 
| 
 | 
   192  | 
lemma A_trusts_OR1: 
  | 
| 
 | 
   193  | 
"\<lbrakk>Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs);  
  | 
| 
 | 
   194  | 
    A \<notin> bad; evs \<in> orb\<rbrakk>                   
  | 
| 
 | 
   195  | 
 \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs"
  | 
| 
 | 
   196  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
  | 
| 
 | 
   197  | 
apply (blast)
  | 
| 
 | 
   198  | 
done
  | 
| 
 | 
   199  | 
  | 
| 
 | 
   200  | 
  | 
| 
 | 
   201  | 
lemma B_trusts_OR2:
  | 
| 
 | 
   202  | 
 "\<lbrakk>Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>  
  | 
| 
 | 
   203  | 
      \<in> parts (knows Spy evs);  B \<notin> bad; evs \<in> orb\<rbrakk>                   
  | 
| 
 | 
   204  | 
  \<Longrightarrow> (\<exists> X. Says B Server \<lbrace>Nonce M, Agent A, Agent B, X,  
  | 
| 
 | 
   205  | 
              Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> 
  | 
| 
 | 
   206  | 
          \<in> set evs)"
  | 
| 
 | 
   207  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
  | 
| 
 | 
   208  | 
apply (blast+)
  | 
| 
 | 
   209  | 
done
  | 
| 
 | 
   210  | 
  | 
| 
 | 
   211  | 
  | 
| 
 | 
   212  | 
lemma B_trusts_OR3: 
  | 
| 
 | 
   213  | 
"\<lbrakk>Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace> \<in> parts (knows Spy evs);  
  | 
| 
 | 
   214  | 
   B \<notin> bad; evs \<in> orb\<rbrakk>                   
  | 
| 
 | 
   215  | 
\<Longrightarrow> \<exists> M. Says Server B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> 
  | 
| 
 | 
   216  | 
         \<in> set evs"
  | 
| 
 | 
   217  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
  | 
| 
 | 
   218  | 
apply (blast+)
  | 
| 
 | 
   219  | 
done
  | 
| 
 | 
   220  | 
  | 
| 
 | 
   221  | 
lemma Gets_Server_message_form: 
  | 
| 
 | 
   222  | 
"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs;  
  | 
| 
 | 
   223  | 
    evs \<in> orb\<rbrakk>                                              
  | 
| 
 | 
   224  | 
 \<Longrightarrow> (K \<notin> range shrK & (\<exists> A Na. X = (Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>)))    
  | 
| 
 | 
   225  | 
             | X \<in> analz (knows Spy evs)"
  | 
| 
 | 
   226  | 
apply (case_tac "B \<in> bad")
  | 
| 
 | 
   227  | 
apply (drule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, 
  | 
| 
 | 
   228  | 
                                 THEN analz.Decrypt, THEN analz.Fst])
  | 
| 
 | 
   229  | 
prefer 3 apply blast
  | 
| 
 | 
   230  | 
prefer 3 apply (blast dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN 
  | 
| 
 | 
   231  | 
                                                 parts.Snd, THEN B_trusts_OR3]
  | 
| 
 | 
   232  | 
                             Says_Server_message_form)
  | 
| 
 | 
   233  | 
apply simp_all                                    
  | 
| 
 | 
   234  | 
done
  | 
| 
 | 
   235  | 
  | 
| 
 | 
   236  | 
lemma unique_Na: "\<lbrakk>Says A B  \<lbrace>Nonce M, Agent A, Agent B, Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs;   
  | 
| 
 | 
   237  | 
         Says A B' \<lbrace>Nonce M', Agent A, Agent B', Crypt (shrK A) \<lbrace>Nonce Na, Nonce M', Agent A, Agent B'\<rbrace>\<rbrace> \<in> set evs;  
  | 
| 
 | 
   238  | 
    A \<notin> bad; evs \<in> orb\<rbrakk> \<Longrightarrow> B=B' & M=M'"
  | 
| 
 | 
   239  | 
by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+)
  | 
| 
 | 
   240  | 
  | 
| 
 | 
   241  | 
lemma unique_Nb: "\<lbrakk>Says B Server \<lbrace>Nonce M, Agent A, Agent B, X, Crypt (shrK B) \<lbrace>Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs;   
  | 
| 
 | 
   242  | 
         Says B Server \<lbrace>Nonce M', Agent A', Agent B, X', Crypt (shrK B) \<lbrace>Nonce Nb,Nonce M', Nonce M', Agent A', Agent B\<rbrace>\<rbrace> \<in> set evs;   
  | 
| 
 | 
   243  | 
    B \<notin> bad; evs \<in> orb\<rbrakk> \<Longrightarrow>   M=M' & A=A' & X=X'"
  | 
| 
 | 
   244  | 
by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+)
  | 
| 
 | 
   245  | 
  | 
| 
 | 
   246  | 
lemma analz_image_freshCryptK_lemma:
  | 
| 
 | 
   247  | 
"(Crypt K X \<in> analz (Key`nE \<union> H)) \<longrightarrow> (Crypt K X \<in> analz H) \<Longrightarrow>  
  | 
| 
 | 
   248  | 
        (Crypt K X \<in> analz (Key`nE \<union> H)) = (Crypt K X \<in> analz H)";
  | 
| 
 | 
   249  | 
by (blast intro: analz_mono [THEN [2] rev_subsetD])
  | 
| 
 | 
   250  | 
  | 
| 
 | 
   251  | 
ML
  | 
| 
 | 
   252  | 
{*
 | 
| 
 | 
   253  | 
val analz_image_freshCryptK_lemma = thm "analz_image_freshCryptK_lemma";
  | 
| 
 | 
   254  | 
val analz_image_freshK_simps = thms "analz_image_freshK_simps";
  | 
| 
 | 
   255  | 
  | 
| 
 | 
   256  | 
val analz_image_freshK_ss = 
  | 
| 
 | 
   257  | 
     simpset() delsimps [image_insert, image_Un]
  | 
| 
 | 
   258  | 
	       delsimps [imp_disjL]    (*reduces blow-up*)
  | 
| 
 | 
   259  | 
	       addsimps thms "analz_image_freshK_simps"
  | 
| 
 | 
   260  | 
*}
  | 
| 
 | 
   261  | 
  | 
| 
 | 
   262  | 
method_setup analz_freshCryptK = {*
 | 
| 
20048
 | 
   263  | 
    Method.ctxt_args (fn ctxt =>
  | 
| 
21588
 | 
   264  | 
     (Method.SIMPLE_METHOD
  | 
| 
 | 
   265  | 
      (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
  | 
| 
18886
 | 
   266  | 
                          REPEAT_FIRST (rtac analz_image_freshCryptK_lemma),
  | 
| 
20048
 | 
   267  | 
                          ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss))]))) *}
  | 
| 
18886
 | 
   268  | 
  "for proving useful rewrite rule"
  | 
| 
 | 
   269  | 
  | 
| 
 | 
   270  | 
  | 
| 
 | 
   271  | 
method_setup disentangle = {*
 | 
| 
 | 
   272  | 
    Method.no_args
  | 
| 
21588
 | 
   273  | 
     (Method.SIMPLE_METHOD
  | 
| 
 | 
   274  | 
      (REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE] 
  | 
| 
18886
 | 
   275  | 
                   ORELSE' hyp_subst_tac))) *}
  | 
| 
 | 
   276  | 
  "for eliminating conjunctions, disjunctions and the like"
  | 
| 
 | 
   277  | 
  | 
| 
 | 
   278  | 
  | 
| 
 | 
   279  | 
  | 
| 
 | 
   280  | 
lemma analz_image_freshCryptK [rule_format]: 
  | 
| 
 | 
   281  | 
"evs \<in> orb \<Longrightarrow>                              
  | 
| 
 | 
   282  | 
     Key K \<notin> analz (knows Spy evs) \<longrightarrow>  
  | 
| 
 | 
   283  | 
       (\<forall> KK. KK \<subseteq> - (range shrK) \<longrightarrow>                  
  | 
| 
 | 
   284  | 
             (Crypt K X \<in> analz (Key`KK \<union> (knows Spy evs))) =   
  | 
| 
 | 
   285  | 
             (Crypt K X \<in> analz (knows Spy evs)))"
  | 
| 
 | 
   286  | 
apply (erule orb.induct)
  | 
| 
 | 
   287  | 
apply (analz_mono_contra)
  | 
| 
 | 
   288  | 
apply (frule_tac [7] Gets_Server_message_form)
  | 
| 
 | 
   289  | 
apply (frule_tac [9] Says_Server_message_form)
  | 
| 
 | 
   290  | 
apply disentangle
  | 
| 
 | 
   291  | 
apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN  analz.Snd])
  | 
| 
 | 
   292  | 
prefer 8 apply clarify
  | 
| 
 | 
   293  | 
apply (analz_freshCryptK, spy_analz, fastsimp)
  | 
| 
 | 
   294  | 
done
  | 
| 
 | 
   295  | 
  | 
| 
 | 
   296  | 
  | 
| 
 | 
   297  | 
  | 
| 
 | 
   298  | 
lemma analz_insert_freshCryptK: 
  | 
| 
 | 
   299  | 
"\<lbrakk>evs \<in> orb;  Key K \<notin> analz (knows Spy evs);  
  | 
| 
 | 
   300  | 
         Seskey \<notin> range shrK\<rbrakk> \<Longrightarrow>   
  | 
| 
 | 
   301  | 
         (Crypt K X \<in> analz (insert (Key Seskey) (knows Spy evs))) =  
  | 
| 
 | 
   302  | 
         (Crypt K X \<in> analz (knows Spy evs))"
  | 
| 
 | 
   303  | 
by (simp only: analz_image_freshCryptK analz_image_freshK_simps)
  | 
| 
 | 
   304  | 
  | 
| 
 | 
   305  | 
  | 
| 
 | 
   306  | 
lemma analz_hard: 
  | 
| 
 | 
   307  | 
"\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B,  
  | 
| 
 | 
   308  | 
             Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in>set evs; 
  | 
| 
 | 
   309  | 
   Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace> \<in> analz (knows Spy evs);  
  | 
| 
 | 
   310  | 
   A \<notin> bad; B \<notin> bad; evs \<in> orb\<rbrakk>                   
  | 
| 
 | 
   311  | 
 \<Longrightarrow>  Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs"
  | 
| 
 | 
   312  | 
apply (erule rev_mp)
  | 
| 
 | 
   313  | 
apply (erule rev_mp)
  | 
| 
 | 
   314  | 
apply (erule orb.induct)
  | 
| 
 | 
   315  | 
apply (frule_tac [7] Gets_Server_message_form)
  | 
| 
 | 
   316  | 
apply (frule_tac [9] Says_Server_message_form)
  | 
| 
 | 
   317  | 
apply disentangle
  | 
| 
 | 
   318  | 
txt{*letting the simplifier solve OR2*}
 | 
| 
 | 
   319  | 
apply (drule_tac [5] Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Snd, THEN analz.Snd])
  | 
| 
 | 
   320  | 
apply (simp_all (no_asm_simp) add: analz_insert_eq pushes split_ifs)
  | 
| 
 | 
   321  | 
apply (spy_analz)
  | 
| 
 | 
   322  | 
txt{*OR1*}
 | 
| 
 | 
   323  | 
apply blast
  | 
| 
 | 
   324  | 
txt{*Oops*}
 | 
| 
 | 
   325  | 
prefer 4 apply (blast dest: analz_insert_freshCryptK)
  | 
| 
 | 
   326  | 
txt{*OR4 - ii*}
 | 
| 
 | 
   327  | 
prefer 3 apply (blast)
  | 
| 
 | 
   328  | 
txt{*OR3*}
 | 
| 
 | 
   329  | 
(*adding Gets_imp_ and Says_imp_ for efficiency*)
  | 
| 
 | 
   330  | 
apply (blast dest: 
  | 
| 
 | 
   331  | 
       A_trusts_OR1 unique_Na Key_not_used analz_insert_freshCryptK)
  | 
| 
 | 
   332  | 
txt{*OR4 - i *}
 | 
| 
 | 
   333  | 
apply clarify
  | 
| 
 | 
   334  | 
apply (simp add: pushes split_ifs)
  | 
| 
 | 
   335  | 
apply (case_tac "Aaa\<in>bad")
  | 
| 
 | 
   336  | 
apply (blast dest: analz_insert_freshCryptK)
  | 
| 
 | 
   337  | 
apply clarify
  | 
| 
 | 
   338  | 
apply simp
  | 
| 
 | 
   339  | 
apply (case_tac "Ba\<in>bad")
  | 
| 
 | 
   340  | 
apply (frule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd, THEN analz.Decrypt, THEN analz.Fst] , assumption)
  | 
| 
 | 
   341  | 
apply (simp (no_asm_simp))
  | 
| 
 | 
   342  | 
apply clarify
  | 
| 
 | 
   343  | 
apply (frule Gets_imp_knows_Spy 
  | 
| 
 | 
   344  | 
             [THEN parts.Inj, THEN parts.Snd, THEN B_trusts_OR3],  
  | 
| 
 | 
   345  | 
       assumption, assumption, assumption, erule exE)
  | 
| 
 | 
   346  | 
apply (frule Says_Server_imp_Gets 
  | 
| 
 | 
   347  | 
            [THEN Gets_imp_knows_Spy, THEN parts.Inj, THEN parts.Snd, 
  | 
| 
 | 
   348  | 
            THEN parts.Snd, THEN parts.Snd, THEN parts.Fst, THEN A_trusts_OR1],
  | 
| 
 | 
   349  | 
       assumption, assumption, assumption, assumption)
  | 
| 
 | 
   350  | 
apply (blast dest: Says_Server_imp_Gets B_trusts_OR2 unique_Na unique_Nb)
  | 
| 
 | 
   351  | 
done
  | 
| 
 | 
   352  | 
  | 
| 
 | 
   353  | 
  | 
| 
 | 
   354  | 
lemma Gets_Server_message_form': 
  | 
| 
 | 
   355  | 
"\<lbrakk>Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce Nb, Key K\<rbrace>\<rbrace>  \<in> set evs;  
  | 
| 
 | 
   356  | 
   B \<notin> bad; evs \<in> orb\<rbrakk>                              
  | 
| 
 | 
   357  | 
  \<Longrightarrow> K \<notin> range shrK & (\<exists> A Na. X = (Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>))"
  | 
| 
 | 
   358  | 
by (blast dest!: B_trusts_OR3 Says_Server_message_form)
  | 
| 
 | 
   359  | 
  | 
| 
 | 
   360  | 
  | 
| 
 | 
   361  | 
lemma OR4_imp_Gets: 
  | 
| 
 | 
   362  | 
"\<lbrakk>Says B A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs;   
  | 
| 
 | 
   363  | 
   B \<notin> bad; evs \<in> orb\<rbrakk>  
  | 
| 
 | 
   364  | 
 \<Longrightarrow> (\<exists> Nb. Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>,
  | 
| 
 | 
   365  | 
                                             Nonce Nb, Key K\<rbrace>\<rbrace> \<in> set evs)"
  | 
| 
 | 
   366  | 
apply (erule rev_mp, erule orb.induct, parts_explicit, simp_all)
  | 
| 
 | 
   367  | 
prefer 3 apply (blast dest: Gets_Server_message_form')
  | 
| 
 | 
   368  | 
apply blast+
  | 
| 
 | 
   369  | 
done
  | 
| 
 | 
   370  | 
  | 
| 
 | 
   371  | 
  | 
| 
 | 
   372  | 
lemma A_keydist_to_B: 
  | 
| 
 | 
   373  | 
"\<lbrakk>Says A B \<lbrace>Nonce M, Agent A, Agent B,  
  | 
| 
 | 
   374  | 
            Crypt (shrK A) \<lbrace>Nonce Na, Nonce M, Agent A, Agent B\<rbrace>\<rbrace> \<in>set evs; 
  | 
| 
 | 
   375  | 
   Gets A \<lbrace>Nonce M, Crypt (shrK A) \<lbrace>Nonce Na, Key K\<rbrace>\<rbrace> \<in> set evs;    
  | 
| 
 | 
   376  | 
   A \<notin> bad; B \<notin> bad; evs \<in> orb\<rbrakk>  
  | 
| 
 | 
   377  | 
  \<Longrightarrow> Key K \<in> analz (knows B evs)"
  | 
| 
 | 
   378  | 
apply (drule Gets_imp_knows_Spy [THEN analz.Inj, THEN analz.Snd], assumption)
  | 
| 
 | 
   379  | 
apply (drule analz_hard, assumption, assumption, assumption, assumption)
  | 
| 
 | 
   380  | 
apply (drule OR4_imp_Gets, assumption, assumption)
  | 
| 
 | 
   381  | 
apply (erule exE)
  | 
| 
 | 
   382  | 
(*blast doesn't do because it can't infer that Key (shrK P) \<in> (knows P evs)*)
  | 
| 
 | 
   383  | 
apply (fastsimp dest!: Gets_imp_knows [THEN analz.Inj] analz.Decrypt)
  | 
| 
 | 
   384  | 
done
  | 
| 
 | 
   385  | 
  | 
| 
 | 
   386  | 
  | 
| 
 | 
   387  | 
text{*Other properties as for the original protocol*}
 | 
| 
 | 
   388  | 
  | 
| 
 | 
   389  | 
  | 
| 
 | 
   390  | 
end
  |