author | wenzelm |
Mon, 11 Sep 2023 19:30:48 +0200 | |
changeset 78659 | b5f3d1051b13 |
parent 76341 | d72a8cdca1ab |
permissions | -rw-r--r-- |
37936 | 1 |
(* Title: HOL/Auth/Event.thy |
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
3 |
Copyright 1996 University of Cambridge |
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
4 |
|
3683 | 5 |
Datatype of events; function "spies"; freshness |
3678 | 6 |
|
3683 | 7 |
"bad" agents have been broken by the Spy; their private keys and internal |
3678 | 8 |
stores are visible to him |
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
9 |
*) |
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
10 |
|
61830 | 11 |
section\<open>Theory of Events for Security Protocols\<close> |
13956 | 12 |
|
16417 | 13 |
theory Event imports Message begin |
11104 | 14 |
|
76341 | 15 |
consts \<comment> \<open>Initial states of agents --- a parameter of the construction\<close> |
67613 | 16 |
initState :: "agent \<Rightarrow> msg set" |
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
17 |
|
58310 | 18 |
datatype |
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
19 |
event = Says agent agent msg |
6399
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
paulson
parents:
6308
diff
changeset
|
20 |
| Gets agent msg |
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
21 |
| Notes agent msg |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
22 |
|
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
23 |
consts |
61830 | 24 |
bad :: "agent set" \<comment> \<open>compromised agents\<close> |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
25 |
|
61830 | 26 |
text\<open>Spy has access to his own key for spoof messages, but Server is secure\<close> |
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
27 |
specification (bad) |
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
14126
diff
changeset
|
28 |
Spy_in_bad [iff]: "Spy \<in> bad" |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
14126
diff
changeset
|
29 |
Server_not_bad [iff]: "Server \<notin> bad" |
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
30 |
by (rule exI [of _ "{Spy}"], simp) |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
31 |
|
67613 | 32 |
primrec knows :: "agent \<Rightarrow> event list \<Rightarrow> msg set" |
38964 | 33 |
where |
11104 | 34 |
knows_Nil: "knows A [] = initState A" |
38964 | 35 |
| knows_Cons: |
6399
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
paulson
parents:
6308
diff
changeset
|
36 |
"knows A (ev # evs) = |
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
paulson
parents:
6308
diff
changeset
|
37 |
(if A = Spy then |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
38 |
(case ev of |
67613 | 39 |
Says A' B X \<Rightarrow> insert X (knows Spy evs) |
40 |
| Gets A' X \<Rightarrow> knows Spy evs |
|
41 |
| Notes A' X \<Rightarrow> |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
42 |
if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
43 |
else |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
44 |
(case ev of |
67613 | 45 |
Says A' B X \<Rightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
46 |
if A'=A then insert X (knows A evs) else knows A evs |
67613 | 47 |
| Gets A' X \<Rightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
48 |
if A'=A then insert X (knows A evs) else knows A evs |
67613 | 49 |
| Notes A' X \<Rightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
50 |
if A'=A then insert X (knows A evs) else knows A evs))" |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
51 |
(* |
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
52 |
Case A=Spy on the Gets event |
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
53 |
enforces the fact that if a message is received then it must have been sent, |
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
54 |
therefore the oops case must use Notes |
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
55 |
*) |
3678 | 56 |
|
61830 | 57 |
text\<open>The constant "spies" is retained for compatibility's sake\<close> |
38964 | 58 |
|
59 |
abbreviation (input) |
|
67613 | 60 |
spies :: "event list \<Rightarrow> msg set" where |
76341 | 61 |
"spies \<equiv> knows Spy" |
38964 | 62 |
|
63 |
||
76341 | 64 |
text \<open>Set of items that might be visible to somebody: complement of the set of fresh items\<close> |
67613 | 65 |
primrec used :: "event list \<Rightarrow> msg set" |
38964 | 66 |
where |
11104 | 67 |
used_Nil: "used [] = (UN B. parts (initState B))" |
38964 | 68 |
| used_Cons: "used (ev # evs) = |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
69 |
(case ev of |
67613 | 70 |
Says A B X \<Rightarrow> parts {X} \<union> used evs |
71 |
| Gets A X \<Rightarrow> used evs |
|
72 |
| Notes A X \<Rightarrow> parts {X} \<union> used evs)" |
|
69597 | 73 |
\<comment> \<open>The case for \<^term>\<open>Gets\<close> seems anomalous, but \<^term>\<open>Gets\<close> always |
74 |
follows \<^term>\<open>Says\<close> in real protocols. Seems difficult to change. |
|
61830 | 75 |
See \<open>Gets_correct\<close> in theory \<open>Guard/Extensions.thy\<close>.\<close> |
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
76 |
|
76341 | 77 |
lemma Notes_imp_used: "Notes A X \<in> set evs \<Longrightarrow> X \<in> used evs" |
78 |
by (induct evs) (auto split: event.split) |
|
11463 | 79 |
|
76341 | 80 |
lemma Says_imp_used: "Says A B X \<in> set evs \<Longrightarrow> X \<in> used evs" |
81 |
by (induct evs) (auto split: event.split) |
|
11463 | 82 |
|
13926 | 83 |
|
69597 | 84 |
subsection\<open>Function \<^term>\<open>knows\<close>\<close> |
13926 | 85 |
|
13956 | 86 |
(*Simplifying |
87 |
parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs). |
|
88 |
This version won't loop with the simplifier.*) |
|
45605 | 89 |
lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs |
13926 | 90 |
|
91 |
lemma knows_Spy_Says [simp]: |
|
76341 | 92 |
"knows Spy (Says A B X # evs) = insert X (knows Spy evs)" |
93 |
by simp |
|
13926 | 94 |
|
61830 | 95 |
text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits |
69597 | 96 |
on whether \<^term>\<open>A=Spy\<close> and whether \<^term>\<open>A\<in>bad\<close>\<close> |
13926 | 97 |
lemma knows_Spy_Notes [simp]: |
76341 | 98 |
"knows Spy (Notes A X # evs) = |
67613 | 99 |
(if A\<in>bad then insert X (knows Spy evs) else knows Spy evs)" |
76341 | 100 |
by simp |
13926 | 101 |
|
102 |
lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" |
|
76341 | 103 |
by simp |
13926 | 104 |
|
105 |
lemma knows_Spy_subset_knows_Spy_Says: |
|
76341 | 106 |
"knows Spy evs \<subseteq> knows Spy (Says A B X # evs)" |
107 |
by (simp add: subset_insertI) |
|
13926 | 108 |
|
109 |
lemma knows_Spy_subset_knows_Spy_Notes: |
|
76341 | 110 |
"knows Spy evs \<subseteq> knows Spy (Notes A X # evs)" |
111 |
by force |
|
13926 | 112 |
|
113 |
lemma knows_Spy_subset_knows_Spy_Gets: |
|
76341 | 114 |
"knows Spy evs \<subseteq> knows Spy (Gets A X # evs)" |
115 |
by (simp add: subset_insertI) |
|
13926 | 116 |
|
61830 | 117 |
text\<open>Spy sees what is sent on the traffic\<close> |
76341 | 118 |
lemma Says_imp_knows_Spy: |
119 |
"Says A B X \<in> set evs \<Longrightarrow> X \<in> knows Spy evs" |
|
120 |
by (induct evs) (auto split: event.split) |
|
13926 | 121 |
|
122 |
lemma Notes_imp_knows_Spy [rule_format]: |
|
76341 | 123 |
"Notes A X \<in> set evs \<Longrightarrow> A \<in> bad \<Longrightarrow> X \<in> knows Spy evs" |
124 |
by (induct evs) (auto split: event.split) |
|
13926 | 125 |
|
126 |
||
61830 | 127 |
text\<open>Elimination rules: derive contradictions from old Says events containing |
128 |
items known to be fresh\<close> |
|
32404 | 129 |
lemmas Says_imp_parts_knows_Spy = |
46471 | 130 |
Says_imp_knows_Spy [THEN parts.Inj, elim_format] |
32404 | 131 |
|
13926 | 132 |
lemmas knows_Spy_partsEs = |
46471 | 133 |
Says_imp_parts_knows_Spy parts.Body [elim_format] |
13926 | 134 |
|
18570 | 135 |
lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj] |
136 |
||
61830 | 137 |
text\<open>Compatibility for the old "spies" function\<close> |
13926 | 138 |
lemmas spies_partsEs = knows_Spy_partsEs |
139 |
lemmas Says_imp_spies = Says_imp_knows_Spy |
|
13935 | 140 |
lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy] |
13926 | 141 |
|
142 |
||
61830 | 143 |
subsection\<open>Knowledge of Agents\<close> |
13926 | 144 |
|
13935 | 145 |
lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)" |
76341 | 146 |
by (simp add: subset_insertI) |
13926 | 147 |
|
13935 | 148 |
lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)" |
76341 | 149 |
by (simp add: subset_insertI) |
13926 | 150 |
|
13935 | 151 |
lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)" |
76341 | 152 |
by (simp add: subset_insertI) |
13926 | 153 |
|
61830 | 154 |
text\<open>Agents know what they say\<close> |
76341 | 155 |
lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs \<Longrightarrow> X \<in> knows A evs" |
156 |
by (induct evs) (force split: event.split)+ |
|
13926 | 157 |
|
61830 | 158 |
text\<open>Agents know what they note\<close> |
76341 | 159 |
lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs \<Longrightarrow> X \<in> knows A evs" |
160 |
by (induct evs) (force split: event.split)+ |
|
13926 | 161 |
|
61830 | 162 |
text\<open>Agents know what they receive\<close> |
13926 | 163 |
lemma Gets_imp_knows_agents [rule_format]: |
76341 | 164 |
"A \<noteq> Spy \<Longrightarrow> Gets A X \<in> set evs \<Longrightarrow> X \<in> knows A evs" |
165 |
by (induct evs) (force split: event.split)+ |
|
13926 | 166 |
|
61830 | 167 |
text\<open>What agents DIFFERENT FROM Spy know |
168 |
was either said, or noted, or got, or known initially\<close> |
|
76341 | 169 |
lemma knows_imp_Says_Gets_Notes_initState: |
170 |
"\<lbrakk>X \<in> knows A evs; A \<noteq> Spy\<rbrakk> \<Longrightarrow> |
|
171 |
\<exists>B. 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" |
|
172 |
by(induct evs) (auto split: event.split_asm if_split_asm) |
|
13926 | 173 |
|
61830 | 174 |
text\<open>What the Spy knows -- for the time being -- |
175 |
was either said or noted, or known initially\<close> |
|
76341 | 176 |
lemma knows_Spy_imp_Says_Notes_initState: |
177 |
"X \<in> knows Spy evs \<Longrightarrow> |
|
178 |
\<exists>A B. Says A B X \<in> set evs \<or> Notes A X \<in> set evs \<or> X \<in> initState Spy" |
|
179 |
by(induct evs) (auto split: event.split_asm if_split_asm) |
|
13926 | 180 |
|
13935 | 181 |
lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs" |
76341 | 182 |
by (induct evs) (auto simp: parts_insert_knows_A split: event.split_asm if_split_asm) |
13926 | 183 |
|
184 |
lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] |
|
185 |
||
67613 | 186 |
lemma initState_into_used: "X \<in> parts (initState B) \<Longrightarrow> X \<in> used evs" |
76341 | 187 |
by (induct evs) (auto simp: parts_insert_knows_A split: event.split) |
188 |
||
189 |
text \<open>New simprules to replace the primitive ones for @{term used} and @{term knows}\<close> |
|
13926 | 190 |
|
191 |
lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs" |
|
76341 | 192 |
by simp |
13926 | 193 |
|
194 |
lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs" |
|
76341 | 195 |
by simp |
13926 | 196 |
|
197 |
lemma used_Gets [simp]: "used (Gets A X # evs) = used evs" |
|
76341 | 198 |
by simp |
13926 | 199 |
|
13935 | 200 |
lemma used_nil_subset: "used [] \<subseteq> used evs" |
76341 | 201 |
using initState_into_used by auto |
13926 | 202 |
|
76341 | 203 |
text\<open>NOTE REMOVAL: the laws above are cleaner, as they don't involve "case"\<close> |
13935 | 204 |
declare knows_Cons [simp del] |
205 |
used_Nil [simp del] used_Cons [simp del] |
|
13926 | 206 |
|
207 |
||
69597 | 208 |
text\<open>For proving theorems of the form \<^term>\<open>X \<notin> analz (knows Spy evs) \<longrightarrow> P\<close> |
13926 | 209 |
New events added by induction to "evs" are discarded. Provided |
210 |
this information isn't needed, the proof will be much shorter, since |
|
69597 | 211 |
it will omit complicated reasoning about \<^term>\<open>analz\<close>.\<close> |
13926 | 212 |
|
213 |
lemmas analz_mono_contra = |
|
214 |
knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD] |
|
215 |
knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD] |
|
216 |
knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD] |
|
217 |
||
11104 | 218 |
|
13922 | 219 |
lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)" |
76341 | 220 |
by (cases e, auto simp: knows_Cons) |
13922 | 221 |
|
13935 | 222 |
lemma initState_subset_knows: "initState A \<subseteq> knows A evs" |
76341 | 223 |
by (induct evs) (use knows_subset_knows_Cons in fastforce)+ |
13922 | 224 |
|
61830 | 225 |
text\<open>For proving \<open>new_keys_not_used\<close>\<close> |
13922 | 226 |
lemma keysFor_parts_insert: |
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
227 |
"\<lbrakk>K \<in> keysFor (parts (insert X G)); X \<in> synth (analz H)\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
228 |
\<Longrightarrow> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H" |
13922 | 229 |
by (force |
230 |
dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD] |
|
231 |
analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] |
|
232 |
intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD]) |
|
233 |
||
24122 | 234 |
|
45605 | 235 |
lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs |
27225 | 236 |
|
24122 | 237 |
ML |
61830 | 238 |
\<open> |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
239 |
fun analz_mono_contra_tac ctxt = |
60754 | 240 |
resolve_tac ctxt @{thms analz_impI} THEN' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
241 |
REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra}) |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
242 |
THEN' (mp_tac ctxt) |
61830 | 243 |
\<close> |
24122 | 244 |
|
61830 | 245 |
method_setup analz_mono_contra = \<open> |
246 |
Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close> |
|
67613 | 247 |
"for proving theorems of the form X \<notin> analz (knows Spy evs) \<longrightarrow> P" |
13922 | 248 |
|
76299 | 249 |
text\<open>Useful for case analysis on whether a hash is a spoof or not\<close> |
45605 | 250 |
lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))"] for Y evs |
27225 | 251 |
|
13922 | 252 |
ML |
61830 | 253 |
\<open> |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
254 |
fun synth_analz_mono_contra_tac ctxt = |
60754 | 255 |
resolve_tac ctxt @{thms syan_impI} THEN' |
27225 | 256 |
REPEAT1 o |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
257 |
(dresolve_tac ctxt |
27225 | 258 |
[@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}, |
259 |
@{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}, |
|
260 |
@{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}]) |
|
261 |
THEN' |
|
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
262 |
mp_tac ctxt |
61830 | 263 |
\<close> |
13922 | 264 |
|
61830 | 265 |
method_setup synth_analz_mono_contra = \<open> |
266 |
Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))\<close> |
|
67613 | 267 |
"for proving theorems of the form X \<notin> synth (analz (knows Spy evs)) \<longrightarrow> P" |
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
268 |
|
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
269 |
end |