| author | bulwahn |
| Tue, 31 Aug 2010 08:00:53 +0200 | |
| changeset 38950 | 62578950e748 |
| parent 37936 | 1e4c5015a72e |
| child 38964 | b1a7bef0907a |
| 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 |
|
| 13956 | 11 |
header{*Theory of Events for Security Protocols*}
|
12 |
||
| 16417 | 13 |
theory Event imports Message begin |
| 11104 | 14 |
|
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
15 |
consts (*Initial states of agents -- parameter of the construction*) |
| 11104 | 16 |
initState :: "agent => msg set" |
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
17 |
|
|
6399
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
paulson
parents:
6308
diff
changeset
|
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 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
24 |
bad :: "agent set" -- {* compromised agents *}
|
| 11104 | 25 |
knows :: "agent => event list => msg set" |
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
26 |
|
|
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
27 |
|
|
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
28 |
text{*The constant "spies" is retained for compatibility's sake*}
|
| 20768 | 29 |
|
30 |
abbreviation (input) |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
31 |
spies :: "event list => msg set" where |
| 20768 | 32 |
"spies == knows Spy" |
|
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
33 |
|
|
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
34 |
text{*Spy has access to his own key for spoof messages, but Server is secure*}
|
|
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
35 |
specification (bad) |
|
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
14126
diff
changeset
|
36 |
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
|
37 |
Server_not_bad [iff]: "Server \<notin> bad" |
|
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
38 |
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
|
39 |
|
| 5183 | 40 |
primrec |
| 11104 | 41 |
knows_Nil: "knows A [] = initState A" |
42 |
knows_Cons: |
|
|
6399
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
paulson
parents:
6308
diff
changeset
|
43 |
"knows A (ev # evs) = |
|
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
paulson
parents:
6308
diff
changeset
|
44 |
(if A = Spy then |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
45 |
(case ev of |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
46 |
Says A' B X => insert X (knows Spy evs) |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
47 |
| Gets A' X => knows Spy evs |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
48 |
| Notes A' X => |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
49 |
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
|
50 |
else |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
51 |
(case ev of |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
52 |
Says A' B X => |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
53 |
if A'=A then insert X (knows A evs) else knows A evs |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
54 |
| Gets A' X => |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
55 |
if A'=A then insert X (knows A evs) else knows A evs |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
56 |
| Notes A' X => |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
57 |
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
|
58 |
|
|
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
59 |
(* |
|
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
60 |
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
|
61 |
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
|
62 |
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
|
63 |
*) |
| 3678 | 64 |
|
| 3683 | 65 |
consts |
66 |
(*Set of items that might be visible to somebody: |
|
67 |
complement of the set of fresh items*) |
|
| 11104 | 68 |
used :: "event list => msg set" |
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
69 |
|
| 5183 | 70 |
primrec |
| 11104 | 71 |
used_Nil: "used [] = (UN B. parts (initState B))" |
72 |
used_Cons: "used (ev # evs) = |
|
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
73 |
(case ev of |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
74 |
Says A B X => parts {X} \<union> used evs
|
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
75 |
| Gets A X => used evs |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
76 |
| Notes A X => parts {X} \<union> used evs)"
|
| 13935 | 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"}. *}
|
|
|
6308
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents:
5183
diff
changeset
|
80 |
|
| 13926 | 81 |
lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs --> X \<in> used evs" |
82 |
apply (induct_tac evs) |
|
| 11463 | 83 |
apply (auto split: event.split) |
84 |
done |
|
85 |
||
| 13926 | 86 |
lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs --> X \<in> used evs" |
87 |
apply (induct_tac evs) |
|
| 11463 | 88 |
apply (auto split: event.split) |
89 |
done |
|
90 |
||
| 13926 | 91 |
|
92 |
subsection{*Function @{term knows}*}
|
|
93 |
||
| 13956 | 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.*) |
|
| 13935 | 97 |
lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard] |
| 13926 | 98 |
|
99 |
lemma knows_Spy_Says [simp]: |
|
100 |
"knows Spy (Says A B X # evs) = insert X (knows Spy evs)" |
|
101 |
by simp |
|
102 |
||
|
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
14126
diff
changeset
|
103 |
text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
|
|
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
14126
diff
changeset
|
104 |
on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
|
| 13926 | 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: |
|
| 13935 | 114 |
"knows Spy evs \<subseteq> knows Spy (Says A B X # evs)" |
| 13926 | 115 |
by (simp add: subset_insertI) |
116 |
||
117 |
lemma knows_Spy_subset_knows_Spy_Notes: |
|
| 13935 | 118 |
"knows Spy evs \<subseteq> knows Spy (Notes A X # evs)" |
| 13926 | 119 |
by force |
120 |
||
121 |
lemma knows_Spy_subset_knows_Spy_Gets: |
|
| 13935 | 122 |
"knows Spy evs \<subseteq> knows Spy (Gets A X # evs)" |
| 13926 | 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*} |
|
| 32404 | 141 |
lemmas Says_imp_parts_knows_Spy = |
142 |
Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] |
|
143 |
||
| 13926 | 144 |
lemmas knows_Spy_partsEs = |
| 32404 | 145 |
Says_imp_parts_knows_Spy parts.Body [THEN revcut_rl, standard] |
| 13926 | 146 |
|
| 18570 | 147 |
lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj] |
148 |
||
| 13926 | 149 |
text{*Compatibility for the old "spies" function*}
|
150 |
lemmas spies_partsEs = knows_Spy_partsEs |
|
151 |
lemmas Says_imp_spies = Says_imp_knows_Spy |
|
| 13935 | 152 |
lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy] |
| 13926 | 153 |
|
154 |
||
155 |
subsection{*Knowledge of Agents*}
|
|
156 |
||
157 |
lemma knows_Says: "knows A (Says A B X # evs) = insert X (knows A evs)" |
|
158 |
by simp |
|
159 |
||
160 |
lemma knows_Notes: "knows A (Notes A X # evs) = insert X (knows A evs)" |
|
161 |
by simp |
|
162 |
||
163 |
lemma knows_Gets: |
|
164 |
"A \<noteq> Spy --> knows A (Gets A X # evs) = insert X (knows A evs)" |
|
165 |
by simp |
|
166 |
||
167 |
||
| 13935 | 168 |
lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)" |
169 |
by (simp add: subset_insertI) |
|
| 13926 | 170 |
|
| 13935 | 171 |
lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)" |
172 |
by (simp add: subset_insertI) |
|
| 13926 | 173 |
|
| 13935 | 174 |
lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)" |
175 |
by (simp add: subset_insertI) |
|
| 13926 | 176 |
|
177 |
text{*Agents know what they say*}
|
|
178 |
lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs --> X \<in> knows A evs" |
|
179 |
apply (induct_tac "evs") |
|
180 |
apply (simp_all (no_asm_simp) split add: event.split) |
|
181 |
apply blast |
|
182 |
done |
|
183 |
||
184 |
text{*Agents know what they note*}
|
|
185 |
lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs --> X \<in> knows A evs" |
|
186 |
apply (induct_tac "evs") |
|
187 |
apply (simp_all (no_asm_simp) split add: event.split) |
|
188 |
apply blast |
|
189 |
done |
|
190 |
||
191 |
text{*Agents know what they receive*}
|
|
192 |
lemma Gets_imp_knows_agents [rule_format]: |
|
193 |
"A \<noteq> Spy --> Gets A X \<in> set evs --> X \<in> knows A evs" |
|
194 |
apply (induct_tac "evs") |
|
195 |
apply (simp_all (no_asm_simp) split add: event.split) |
|
196 |
done |
|
197 |
||
198 |
||
199 |
text{*What agents DIFFERENT FROM Spy know
|
|
200 |
was either said, or noted, or got, or known initially*} |
|
201 |
lemma knows_imp_Says_Gets_Notes_initState [rule_format]: |
|
202 |
"[| X \<in> knows A evs; A \<noteq> Spy |] ==> EX B. |
|
203 |
Says A B X \<in> set evs | Gets A X \<in> set evs | Notes A X \<in> set evs | X \<in> initState A" |
|
204 |
apply (erule rev_mp) |
|
205 |
apply (induct_tac "evs") |
|
206 |
apply (simp_all (no_asm_simp) split add: event.split) |
|
207 |
apply blast |
|
208 |
done |
|
209 |
||
210 |
text{*What the Spy knows -- for the time being --
|
|
211 |
was either said or noted, or known initially*} |
|
212 |
lemma knows_Spy_imp_Says_Notes_initState [rule_format]: |
|
213 |
"[| X \<in> knows Spy evs |] ==> EX A B. |
|
214 |
Says A B X \<in> set evs | Notes A X \<in> set evs | X \<in> initState Spy" |
|
215 |
apply (erule rev_mp) |
|
216 |
apply (induct_tac "evs") |
|
217 |
apply (simp_all (no_asm_simp) split add: event.split) |
|
218 |
apply blast |
|
219 |
done |
|
220 |
||
| 13935 | 221 |
lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs" |
222 |
apply (induct_tac "evs", force) |
|
223 |
apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) |
|
| 13926 | 224 |
done |
225 |
||
226 |
lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] |
|
227 |
||
228 |
lemma initState_into_used: "X \<in> parts (initState B) ==> X \<in> used evs" |
|
229 |
apply (induct_tac "evs") |
|
| 13935 | 230 |
apply (simp_all add: parts_insert_knows_A split add: event.split, blast) |
| 13926 | 231 |
done |
232 |
||
233 |
lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
|
|
234 |
by simp |
|
235 |
||
236 |
lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
|
|
237 |
by simp |
|
238 |
||
239 |
lemma used_Gets [simp]: "used (Gets A X # evs) = used evs" |
|
240 |
by simp |
|
241 |
||
| 13935 | 242 |
lemma used_nil_subset: "used [] \<subseteq> used evs" |
243 |
apply simp |
|
| 13926 | 244 |
apply (blast intro: initState_into_used) |
245 |
done |
|
246 |
||
247 |
text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
|
|
| 13935 | 248 |
declare knows_Cons [simp del] |
249 |
used_Nil [simp del] used_Cons [simp del] |
|
| 13926 | 250 |
|
251 |
||
252 |
text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
|
|
253 |
New events added by induction to "evs" are discarded. Provided |
|
254 |
this information isn't needed, the proof will be much shorter, since |
|
255 |
it will omit complicated reasoning about @{term analz}.*}
|
|
256 |
||
257 |
lemmas analz_mono_contra = |
|
258 |
knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD] |
|
259 |
knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD] |
|
260 |
knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD] |
|
261 |
||
| 11104 | 262 |
|
| 13922 | 263 |
lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)" |
264 |
by (induct e, auto simp: knows_Cons) |
|
265 |
||
| 13935 | 266 |
lemma initState_subset_knows: "initState A \<subseteq> knows A evs" |
| 13926 | 267 |
apply (induct_tac evs, simp) |
| 13922 | 268 |
apply (blast intro: knows_subset_knows_Cons [THEN subsetD]) |
269 |
done |
|
270 |
||
271 |
||
| 13926 | 272 |
text{*For proving @{text new_keys_not_used}*}
|
| 13922 | 273 |
lemma keysFor_parts_insert: |
| 13926 | 274 |
"[| K \<in> keysFor (parts (insert X G)); X \<in> synth (analz H) |] |
275 |
==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"; |
|
| 13922 | 276 |
by (force |
277 |
dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD] |
|
278 |
analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] |
|
279 |
intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD]) |
|
280 |
||
| 24122 | 281 |
|
| 27225 | 282 |
lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)", standard] |
283 |
||
| 24122 | 284 |
ML |
285 |
{*
|
|
286 |
val analz_mono_contra_tac = |
|
| 27225 | 287 |
rtac @{thm analz_impI} THEN'
|
288 |
REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
|
|
289 |
THEN' mp_tac |
|
| 24122 | 290 |
*} |
291 |
||
| 11104 | 292 |
method_setup analz_mono_contra = {*
|
| 30549 | 293 |
Scan.succeed (K (SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac))) *} |
| 13922 | 294 |
"for proving theorems of the form X \<notin> analz (knows Spy evs) --> P" |
295 |
||
296 |
subsubsection{*Useful for case analysis on whether a hash is a spoof or not*}
|
|
297 |
||
| 27225 | 298 |
lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))", standard] |
299 |
||
| 13922 | 300 |
ML |
301 |
{*
|
|
302 |
val synth_analz_mono_contra_tac = |
|
| 27225 | 303 |
rtac @{thm syan_impI} THEN'
|
304 |
REPEAT1 o |
|
305 |
(dresolve_tac |
|
306 |
[@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
|
|
307 |
@{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
|
|
308 |
@{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
|
|
309 |
THEN' |
|
310 |
mp_tac |
|
| 13922 | 311 |
*} |
312 |
||
313 |
method_setup synth_analz_mono_contra = {*
|
|
| 30549 | 314 |
Scan.succeed (K (SIMPLE_METHOD (REPEAT_FIRST synth_analz_mono_contra_tac))) *} |
| 13922 | 315 |
"for proving theorems of the form X \<notin> synth (analz (knows Spy evs)) --> P" |
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
316 |
|
|
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff
changeset
|
317 |
end |