1(*  Title:      HOL/Auth/Event.thy
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3    Copyright   1996  University of Cambridge
4
5Datatype 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
11section\<open>Theory of Events for Security Protocols\<close>
12
13theory Event imports Message begin
14
15consts  (*Initial states of agents -- parameter of the construction*)
16  initState :: "agent \<Rightarrow> msg set"
17
18datatype
19  event = Says  agent agent msg
20        | Gets  agent       msg
21        | Notes agent       msg
22       
23consts 
24  bad    :: "agent set"                         \<comment> \<open>compromised agents\<close>
25
26text\<open>Spy has access to his own key for spoof messages, but Server is secure\<close>
27specification (bad)
28  Spy_in_bad     [iff]: "Spy \<in> bad"
29  Server_not_bad [iff]: "Server \<notin> bad"
30    by (rule exI [of _ "{Spy}"], simp)
31
32primrec knows :: "agent \<Rightarrow> event list \<Rightarrow> msg set"
33where
34  knows_Nil:   "knows A [] = initState A"
35| knows_Cons:
36    "knows A (ev # evs) =
37       (if A = Spy then 
38        (case ev of
39           Says A' B X \<Rightarrow> insert X (knows Spy evs)
40         | Gets A' X \<Rightarrow> knows Spy evs
41         | Notes A' X  \<Rightarrow> 
42             if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
43        else
44        (case ev of
45           Says A' B X \<Rightarrow> 
46             if A'=A then insert X (knows A evs) else knows A evs
47         | Gets A' X    \<Rightarrow> 
48             if A'=A then insert X (knows A evs) else knows A evs
49         | Notes A' X    \<Rightarrow> 
50             if A'=A then insert X (knows A evs) else knows A evs))"
51(*
52  Case A=Spy on the Gets event
53  enforces the fact that if a message is received then it must have been sent,
54  therefore the oops case must use Notes
55*)
56
57text\<open>The constant "spies" is retained for compatibility's sake\<close>
58
59abbreviation (input)
60  spies  :: "event list \<Rightarrow> msg set" where
61  "spies == knows Spy"
62
63
64(*Set of items that might be visible to somebody:
65    complement of the set of fresh items*)
66
67primrec used :: "event list \<Rightarrow> msg set"
68where
69  used_Nil:   "used []         = (UN B. parts (initState B))"
70| used_Cons:  "used (ev # evs) =
71                     (case ev of
72                        Says A B X \<Rightarrow> parts {X} \<union> used evs
73                      | Gets A X   \<Rightarrow> used evs
74                      | Notes A X  \<Rightarrow> parts {X} \<union> used evs)"
75    \<comment> \<open>The case for \<^term>\<open>Gets\<close> seems anomalous, but \<^term>\<open>Gets\<close> always
76        follows \<^term>\<open>Says\<close> in real protocols.  Seems difficult to change.
77        See \<open>Gets_correct\<close> in theory \<open>Guard/Extensions.thy\<close>.\<close>
78
79lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs \<longrightarrow> X \<in> used evs"
80apply (induct_tac evs)
81apply (auto split: event.split) 
82done
83
84lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs \<longrightarrow> X \<in> used evs"
85apply (induct_tac evs)
86apply (auto split: event.split) 
87done
88
89
90subsection\<open>Function \<^term>\<open>knows\<close>\<close>
91
92(*Simplifying   
93 parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
94  This version won't loop with the simplifier.*)
95lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs
96
97lemma knows_Spy_Says [simp]:
98     "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
99by simp
100
101text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits
102      on whether \<^term>\<open>A=Spy\<close> and whether \<^term>\<open>A\<in>bad\<close>\<close>
103lemma knows_Spy_Notes [simp]:
104     "knows Spy (Notes A X # evs) =  
105          (if A\<in>bad then insert X (knows Spy evs) else knows Spy evs)"
106by simp
107
108lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
109by simp
110
111lemma knows_Spy_subset_knows_Spy_Says:
112     "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)"
113by (simp add: subset_insertI)
114
115lemma knows_Spy_subset_knows_Spy_Notes:
116     "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)"
117by force
118
119lemma knows_Spy_subset_knows_Spy_Gets:
120     "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)"
121by (simp add: subset_insertI)
122
123text\<open>Spy sees what is sent on the traffic\<close>
124lemma Says_imp_knows_Spy [rule_format]:
125     "Says A B X \<in> set evs \<longrightarrow> X \<in> knows Spy evs"
126apply (induct_tac "evs")
127apply (simp_all (no_asm_simp) split: event.split)
128done
129
130lemma Notes_imp_knows_Spy [rule_format]:
131     "Notes A X \<in> set evs \<longrightarrow> A \<in> bad \<longrightarrow> X \<in> knows Spy evs"
132apply (induct_tac "evs")
133apply (simp_all (no_asm_simp) split: event.split)
134done
135
136
137text\<open>Elimination rules: derive contradictions from old Says events containing
138  items known to be fresh\<close>
139lemmas Says_imp_parts_knows_Spy = 
140       Says_imp_knows_Spy [THEN parts.Inj, elim_format] 
141
142lemmas knows_Spy_partsEs =
143     Says_imp_parts_knows_Spy parts.Body [elim_format]
144
145lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj]
146
147text\<open>Compatibility for the old "spies" function\<close>
148lemmas spies_partsEs = knows_Spy_partsEs
149lemmas Says_imp_spies = Says_imp_knows_Spy
150lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy]
151
152
153subsection\<open>Knowledge of Agents\<close>
154
155lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)"
156by (simp add: subset_insertI)
157
158lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)"
159by (simp add: subset_insertI)
160
161lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)"
162by (simp add: subset_insertI)
163
164text\<open>Agents know what they say\<close>
165lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs \<longrightarrow> X \<in> knows A evs"
166apply (induct_tac "evs")
167apply (simp_all (no_asm_simp) split: event.split)
168apply blast
169done
170
171text\<open>Agents know what they note\<close>
172lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
173apply (induct_tac "evs")
174apply (simp_all (no_asm_simp) split: event.split)
175apply blast
176done
177
178text\<open>Agents know what they receive\<close>
179lemma Gets_imp_knows_agents [rule_format]:
180     "A \<noteq> Spy \<longrightarrow> Gets A X \<in> set evs \<longrightarrow> X \<in> knows A evs"
181apply (induct_tac "evs")
182apply (simp_all (no_asm_simp) split: event.split)
183done
184
185
186text\<open>What agents DIFFERENT FROM Spy know 
187  was either said, or noted, or got, or known initially\<close>
188lemma knows_imp_Says_Gets_Notes_initState [rule_format]:
189     "[| X \<in> knows A evs; A \<noteq> Spy |] ==> \<exists> B.
190  Says A B X \<in> set evs \<or> Gets A X \<in> set evs \<or> Notes A X \<in> set evs \<or> X \<in> initState A"
191apply (erule rev_mp)
192apply (induct_tac "evs")
193apply (simp_all (no_asm_simp) split: event.split)
194apply blast
195done
196
197text\<open>What the Spy knows -- for the time being --
198  was either said or noted, or known initially\<close>
199lemma knows_Spy_imp_Says_Notes_initState [rule_format]:
200     "X \<in> knows Spy evs \<Longrightarrow> \<exists>A B.
201  Says A B X \<in> set evs \<or> Notes A X \<in> set evs \<or> X \<in> initState Spy"
202apply (erule rev_mp)
203apply (induct_tac "evs")
204apply (simp_all (no_asm_simp) split: event.split)
205apply blast
206done
207
208lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs"
209apply (induct_tac "evs", force)  
210apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) 
211done
212
213lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
214
215lemma initState_into_used: "X \<in> parts (initState B) \<Longrightarrow> X \<in> used evs"
216apply (induct_tac "evs")
217apply (simp_all add: parts_insert_knows_A split: event.split, blast)
218done
219
220lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
221by simp
222
223lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
224by simp
225
226lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
227by simp
228
229lemma used_nil_subset: "used [] \<subseteq> used evs"
230apply simp
231apply (blast intro: initState_into_used)
232done
233
234text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close>
235declare knows_Cons [simp del]
236        used_Nil [simp del] used_Cons [simp del]
237
238
239text\<open>For proving theorems of the form \<^term>\<open>X \<notin> analz (knows Spy evs) \<longrightarrow> P\<close>
240  New events added by induction to "evs" are discarded.  Provided 
241  this information isn't needed, the proof will be much shorter, since
242  it will omit complicated reasoning about \<^term>\<open>analz\<close>.\<close>
243
244lemmas analz_mono_contra =
245       knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
246       knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
247       knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]
248
249
250lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)"
251by (cases e, auto simp: knows_Cons)
252
253lemma initState_subset_knows: "initState A \<subseteq> knows A evs"
254apply (induct_tac evs, simp) 
255apply (blast intro: knows_subset_knows_Cons [THEN subsetD])
256done
257
258
259text\<open>For proving \<open>new_keys_not_used\<close>\<close>
260lemma keysFor_parts_insert:
261     "[| K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) |] 
262      ==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"
263by (force 
264    dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
265           analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD]
266    intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD])
267
268
269lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs
270
271ML
272\<open>
273fun analz_mono_contra_tac ctxt = 
274  resolve_tac ctxt @{thms analz_impI} THEN' 
275  REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
276  THEN' (mp_tac ctxt)
277\<close>
278
279method_setup analz_mono_contra = \<open>
280    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close>
281    "for proving theorems of the form X \<notin> analz (knows Spy evs) \<longrightarrow> P"
282
283subsubsection\<open>Useful for case analysis on whether a hash is a spoof or not\<close>
284
285lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))"] for Y evs
286
287ML
288\<open>
289fun synth_analz_mono_contra_tac ctxt = 
290  resolve_tac ctxt @{thms syan_impI} THEN'
291  REPEAT1 o 
292    (dresolve_tac ctxt
293     [@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
294      @{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
295      @{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
296  THEN'
297  mp_tac ctxt
298\<close>
299
300method_setup synth_analz_mono_contra = \<open>
301    Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))\<close>
302    "for proving theorems of the form X \<notin> synth (analz (knows Spy evs)) \<longrightarrow> P"
303
304end
305