src/Doc/Tutorial/Protocol/Event.thy
changeset 48985 5386df44a037
parent 46471 2289a3869c88
child 49322 fbb320d02420
equal deleted inserted replaced
48984:f51d4a302962 48985:5386df44a037
       
     1 (*  Title:      HOL/Auth/Event
       
     2     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     3     Copyright   1996  University of Cambridge
       
     4 
       
     5 Datatype of events; function "spies"; freshness
       
     6 
       
     7 "bad" agents have been broken by the Spy; their private keys and internal
       
     8     stores are visible to him
       
     9 *)(*<*)
       
    10 
       
    11 header{*Theory of Events for Security Protocols*}
       
    12 
       
    13 theory Event imports Message begin
       
    14 
       
    15 consts  (*Initial states of agents -- parameter of the construction*)
       
    16   initState :: "agent => msg set"
       
    17 
       
    18 datatype
       
    19   event = Says  agent agent msg
       
    20         | Gets  agent       msg
       
    21         | Notes agent       msg
       
    22        
       
    23 consts 
       
    24   bad    :: "agent set"                         -- {* compromised agents *}
       
    25 
       
    26 
       
    27 text{*The constant "spies" is retained for compatibility's sake*}
       
    28 
       
    29 primrec
       
    30   knows :: "agent => event list => msg set"
       
    31 where
       
    32   knows_Nil:   "knows A [] = initState A"
       
    33 | knows_Cons:
       
    34     "knows A (ev # evs) =
       
    35        (if A = Spy then 
       
    36         (case ev of
       
    37            Says A' B X => insert X (knows Spy evs)
       
    38          | Gets A' X => knows Spy evs
       
    39          | Notes A' X  => 
       
    40              if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
       
    41         else
       
    42         (case ev of
       
    43            Says A' B X => 
       
    44              if A'=A then insert X (knows A evs) else knows A evs
       
    45          | Gets A' X    => 
       
    46              if A'=A then insert X (knows A evs) else knows A evs
       
    47          | Notes A' X    => 
       
    48              if A'=A then insert X (knows A evs) else knows A evs))"
       
    49 
       
    50 abbreviation (input)
       
    51   spies  :: "event list => msg set" where
       
    52   "spies == knows Spy"
       
    53 
       
    54 text{*Spy has access to his own key for spoof messages, but Server is secure*}
       
    55 specification (bad)
       
    56   Spy_in_bad     [iff]: "Spy \<in> bad"
       
    57   Server_not_bad [iff]: "Server \<notin> bad"
       
    58     by (rule exI [of _ "{Spy}"], simp)
       
    59 
       
    60 (*
       
    61   Case A=Spy on the Gets event
       
    62   enforces the fact that if a message is received then it must have been sent,
       
    63   therefore the oops case must use Notes
       
    64 *)
       
    65 
       
    66 primrec
       
    67   (*Set of items that might be visible to somebody:
       
    68     complement of the set of fresh items*)
       
    69   used :: "event list => msg set"
       
    70 where
       
    71   used_Nil:   "used []         = (UN B. parts (initState B))"
       
    72 | used_Cons:  "used (ev # evs) =
       
    73                      (case ev of
       
    74                         Says A B X => parts {X} \<union> used evs
       
    75                       | Gets A X   => used evs
       
    76                       | Notes A X  => parts {X} \<union> used evs)"
       
    77     --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
       
    78         follows @{term Says} in real protocols.  Seems difficult to change.
       
    79         See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
       
    80 
       
    81 lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs --> X \<in> used evs"
       
    82 apply (induct_tac evs)
       
    83 apply (auto split: event.split) 
       
    84 done
       
    85 
       
    86 lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs --> X \<in> used evs"
       
    87 apply (induct_tac evs)
       
    88 apply (auto split: event.split) 
       
    89 done
       
    90 
       
    91 
       
    92 subsection{*Function @{term knows}*}
       
    93 
       
    94 (*Simplifying   
       
    95  parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
       
    96   This version won't loop with the simplifier.*)
       
    97 lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard]
       
    98 
       
    99 lemma knows_Spy_Says [simp]:
       
   100      "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
       
   101 by simp
       
   102 
       
   103 text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
       
   104       on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
       
   105 lemma knows_Spy_Notes [simp]:
       
   106      "knows Spy (Notes A X # evs) =  
       
   107           (if A:bad then insert X (knows Spy evs) else knows Spy evs)"
       
   108 by simp
       
   109 
       
   110 lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
       
   111 by simp
       
   112 
       
   113 lemma knows_Spy_subset_knows_Spy_Says:
       
   114      "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)"
       
   115 by (simp add: subset_insertI)
       
   116 
       
   117 lemma knows_Spy_subset_knows_Spy_Notes:
       
   118      "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)"
       
   119 by force
       
   120 
       
   121 lemma knows_Spy_subset_knows_Spy_Gets:
       
   122      "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)"
       
   123 by (simp add: subset_insertI)
       
   124 
       
   125 text{*Spy sees what is sent on the traffic*}
       
   126 lemma Says_imp_knows_Spy [rule_format]:
       
   127      "Says A B X \<in> set evs --> X \<in> knows Spy evs"
       
   128 apply (induct_tac "evs")
       
   129 apply (simp_all (no_asm_simp) split add: event.split)
       
   130 done
       
   131 
       
   132 lemma Notes_imp_knows_Spy [rule_format]:
       
   133      "Notes A X \<in> set evs --> A: bad --> X \<in> knows Spy evs"
       
   134 apply (induct_tac "evs")
       
   135 apply (simp_all (no_asm_simp) split add: event.split)
       
   136 done
       
   137 
       
   138 
       
   139 text{*Elimination rules: derive contradictions from old Says events containing
       
   140   items known to be fresh*}
       
   141 lemmas knows_Spy_partsEs =
       
   142      Says_imp_knows_Spy [THEN parts.Inj, elim_format] 
       
   143      parts.Body [elim_format]
       
   144 
       
   145 lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj]
       
   146 
       
   147 text{*Compatibility for the old "spies" function*}
       
   148 lemmas spies_partsEs = knows_Spy_partsEs
       
   149 lemmas Says_imp_spies = Says_imp_knows_Spy
       
   150 lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy]
       
   151 
       
   152 
       
   153 subsection{*Knowledge of Agents*}
       
   154 
       
   155 lemma knows_Says: "knows A (Says A B X # evs) = insert X (knows A evs)"
       
   156 by simp
       
   157 
       
   158 lemma knows_Notes: "knows A (Notes A X # evs) = insert X (knows A evs)"
       
   159 by simp
       
   160 
       
   161 lemma knows_Gets:
       
   162      "A \<noteq> Spy --> knows A (Gets A X # evs) = insert X (knows A evs)"
       
   163 by simp
       
   164 
       
   165 
       
   166 lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)"
       
   167 by (simp add: subset_insertI)
       
   168 
       
   169 lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)"
       
   170 by (simp add: subset_insertI)
       
   171 
       
   172 lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)"
       
   173 by (simp add: subset_insertI)
       
   174 
       
   175 text{*Agents know what they say*}
       
   176 lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs --> X \<in> knows A evs"
       
   177 apply (induct_tac "evs")
       
   178 apply (simp_all (no_asm_simp) split add: event.split)
       
   179 apply blast
       
   180 done
       
   181 
       
   182 text{*Agents know what they note*}
       
   183 lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs --> X \<in> knows A evs"
       
   184 apply (induct_tac "evs")
       
   185 apply (simp_all (no_asm_simp) split add: event.split)
       
   186 apply blast
       
   187 done
       
   188 
       
   189 text{*Agents know what they receive*}
       
   190 lemma Gets_imp_knows_agents [rule_format]:
       
   191      "A \<noteq> Spy --> Gets A X \<in> set evs --> X \<in> knows A evs"
       
   192 apply (induct_tac "evs")
       
   193 apply (simp_all (no_asm_simp) split add: event.split)
       
   194 done
       
   195 
       
   196 
       
   197 text{*What agents DIFFERENT FROM Spy know 
       
   198   was either said, or noted, or got, or known initially*}
       
   199 lemma knows_imp_Says_Gets_Notes_initState [rule_format]:
       
   200      "[| X \<in> knows A evs; A \<noteq> Spy |] ==> EX B.  
       
   201   Says A B X \<in> set evs | Gets A X \<in> set evs | Notes A X \<in> set evs | X \<in> initState A"
       
   202 apply (erule rev_mp)
       
   203 apply (induct_tac "evs")
       
   204 apply (simp_all (no_asm_simp) split add: event.split)
       
   205 apply blast
       
   206 done
       
   207 
       
   208 text{*What the Spy knows -- for the time being --
       
   209   was either said or noted, or known initially*}
       
   210 lemma knows_Spy_imp_Says_Notes_initState [rule_format]:
       
   211      "[| X \<in> knows Spy evs |] ==> EX A B.  
       
   212   Says A B X \<in> set evs | Notes A X \<in> set evs | X \<in> initState Spy"
       
   213 apply (erule rev_mp)
       
   214 apply (induct_tac "evs")
       
   215 apply (simp_all (no_asm_simp) split add: event.split)
       
   216 apply blast
       
   217 done
       
   218 
       
   219 lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs"
       
   220 apply (induct_tac "evs", force)  
       
   221 apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) 
       
   222 done
       
   223 
       
   224 lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
       
   225 
       
   226 lemma initState_into_used: "X \<in> parts (initState B) ==> X \<in> used evs"
       
   227 apply (induct_tac "evs")
       
   228 apply (simp_all add: parts_insert_knows_A split add: event.split, blast)
       
   229 done
       
   230 
       
   231 lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
       
   232 by simp
       
   233 
       
   234 lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
       
   235 by simp
       
   236 
       
   237 lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
       
   238 by simp
       
   239 
       
   240 lemma used_nil_subset: "used [] \<subseteq> used evs"
       
   241 apply simp
       
   242 apply (blast intro: initState_into_used)
       
   243 done
       
   244 
       
   245 text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
       
   246 declare knows_Cons [simp del]
       
   247         used_Nil [simp del] used_Cons [simp del]
       
   248 
       
   249 
       
   250 text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
       
   251   New events added by induction to "evs" are discarded.  Provided 
       
   252   this information isn't needed, the proof will be much shorter, since
       
   253   it will omit complicated reasoning about @{term analz}.*}
       
   254 
       
   255 lemmas analz_mono_contra =
       
   256        knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
       
   257        knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
       
   258        knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]
       
   259 
       
   260 lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)", standard]
       
   261 
       
   262 ML
       
   263 {*
       
   264 val analz_mono_contra_tac = 
       
   265   rtac @{thm analz_impI} THEN' 
       
   266   REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
       
   267   THEN' mp_tac
       
   268 *}
       
   269 
       
   270 lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)"
       
   271 by (induct e, auto simp: knows_Cons)
       
   272 
       
   273 lemma initState_subset_knows: "initState A \<subseteq> knows A evs"
       
   274 apply (induct_tac evs, simp) 
       
   275 apply (blast intro: knows_subset_knows_Cons [THEN subsetD])
       
   276 done
       
   277 
       
   278 
       
   279 text{*For proving @{text new_keys_not_used}*}
       
   280 lemma keysFor_parts_insert:
       
   281      "[| K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) |] 
       
   282       ==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"; 
       
   283 by (force 
       
   284     dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
       
   285            analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD]
       
   286     intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD])
       
   287 
       
   288 method_setup analz_mono_contra = {*
       
   289     Scan.succeed (K (SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac))) *}
       
   290     "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P"
       
   291 
       
   292 subsubsection{*Useful for case analysis on whether a hash is a spoof or not*}
       
   293 
       
   294 lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))", standard]
       
   295 
       
   296 ML
       
   297 {*
       
   298 val knows_Cons = @{thm knows_Cons};
       
   299 val used_Nil = @{thm used_Nil};
       
   300 val used_Cons = @{thm used_Cons};
       
   301 
       
   302 val Notes_imp_used = @{thm Notes_imp_used};
       
   303 val Says_imp_used = @{thm Says_imp_used};
       
   304 val Says_imp_knows_Spy = @{thm Says_imp_knows_Spy};
       
   305 val Notes_imp_knows_Spy = @{thm Notes_imp_knows_Spy};
       
   306 val knows_Spy_partsEs = @{thms knows_Spy_partsEs};
       
   307 val spies_partsEs = @{thms spies_partsEs};
       
   308 val Says_imp_spies = @{thm Says_imp_spies};
       
   309 val parts_insert_spies = @{thm parts_insert_spies};
       
   310 val Says_imp_knows = @{thm Says_imp_knows};
       
   311 val Notes_imp_knows = @{thm Notes_imp_knows};
       
   312 val Gets_imp_knows_agents = @{thm Gets_imp_knows_agents};
       
   313 val knows_imp_Says_Gets_Notes_initState = @{thm knows_imp_Says_Gets_Notes_initState};
       
   314 val knows_Spy_imp_Says_Notes_initState = @{thm knows_Spy_imp_Says_Notes_initState};
       
   315 val usedI = @{thm usedI};
       
   316 val initState_into_used = @{thm initState_into_used};
       
   317 val used_Says = @{thm used_Says};
       
   318 val used_Notes = @{thm used_Notes};
       
   319 val used_Gets = @{thm used_Gets};
       
   320 val used_nil_subset = @{thm used_nil_subset};
       
   321 val analz_mono_contra = @{thms analz_mono_contra};
       
   322 val knows_subset_knows_Cons = @{thm knows_subset_knows_Cons};
       
   323 val initState_subset_knows = @{thm initState_subset_knows};
       
   324 val keysFor_parts_insert = @{thm keysFor_parts_insert};
       
   325 
       
   326 
       
   327 val synth_analz_mono = @{thm synth_analz_mono};
       
   328 
       
   329 val knows_Spy_subset_knows_Spy_Says = @{thm knows_Spy_subset_knows_Spy_Says};
       
   330 val knows_Spy_subset_knows_Spy_Notes = @{thm knows_Spy_subset_knows_Spy_Notes};
       
   331 val knows_Spy_subset_knows_Spy_Gets = @{thm knows_Spy_subset_knows_Spy_Gets};
       
   332 
       
   333 
       
   334 val synth_analz_mono_contra_tac = 
       
   335   rtac @{thm syan_impI} THEN'
       
   336   REPEAT1 o 
       
   337     (dresolve_tac 
       
   338      [@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
       
   339       @{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
       
   340       @{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
       
   341   THEN'
       
   342   mp_tac
       
   343 *}
       
   344 
       
   345 method_setup synth_analz_mono_contra = {*
       
   346     Scan.succeed (K (SIMPLE_METHOD (REPEAT_FIRST synth_analz_mono_contra_tac))) *}
       
   347     "for proving theorems of the form X \<notin> synth (analz (knows Spy evs)) --> P"
       
   348 (*>*)
       
   349 
       
   350 section{* Event Traces \label{sec:events} *}
       
   351 
       
   352 text {*
       
   353 The system's behaviour is formalized as a set of traces of
       
   354 \emph{events}.  The most important event, @{text "Says A B X"}, expresses
       
   355 $A\to B : X$, which is the attempt by~$A$ to send~$B$ the message~$X$.
       
   356 A trace is simply a list, constructed in reverse
       
   357 using~@{text "#"}.  Other event types include reception of messages (when
       
   358 we want to make it explicit) and an agent's storing a fact.
       
   359 
       
   360 Sometimes the protocol requires an agent to generate a new nonce. The
       
   361 probability that a 20-byte random number has appeared before is effectively
       
   362 zero.  To formalize this important property, the set @{term "used evs"}
       
   363 denotes the set of all items mentioned in the trace~@{text evs}.
       
   364 The function @{text used} has a straightforward
       
   365 recursive definition.  Here is the case for @{text Says} event:
       
   366 @{thm [display,indent=5] used_Says [no_vars]}
       
   367 
       
   368 The function @{text knows} formalizes an agent's knowledge.  Mostly we only
       
   369 care about the spy's knowledge, and @{term "knows Spy evs"} is the set of items
       
   370 available to the spy in the trace~@{text evs}.  Already in the empty trace,
       
   371 the spy starts with some secrets at his disposal, such as the private keys
       
   372 of compromised users.  After each @{text Says} event, the spy learns the
       
   373 message that was sent:
       
   374 @{thm [display,indent=5] knows_Spy_Says [no_vars]}
       
   375 Combinations of functions express other important
       
   376 sets of messages derived from~@{text evs}:
       
   377 \begin{itemize}
       
   378 \item @{term "analz (knows Spy evs)"} is everything that the spy could
       
   379 learn by decryption
       
   380 \item @{term "synth (analz (knows Spy evs))"} is everything that the spy
       
   381 could generate
       
   382 \end{itemize}
       
   383 *}
       
   384 
       
   385 (*<*)
       
   386 end
       
   387 (*>*)