| author | haftmann | 
| Tue, 20 Jun 2017 13:07:49 +0200 | |
| changeset 66149 | 4bf16fb7c14d | 
| parent 63648 | f9f3006a5579 | 
| child 67443 | 3abf6a722518 | 
| 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 | |
| 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 | |
| 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: 
6308diff
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: 
5183diff
changeset | 22 | |
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
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: 
5183diff
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: 
13956diff
changeset | 27 | specification (bad) | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 28 | Spy_in_bad [iff]: "Spy \<in> bad" | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 29 | Server_not_bad [iff]: "Server \<notin> bad" | 
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
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: 
5183diff
changeset | 31 | |
| 38964 | 32 | primrec knows :: "agent => event list => msg set" | 
| 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: 
6308diff
changeset | 36 | "knows A (ev # evs) = | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 37 | (if A = Spy then | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 38 | (case ev of | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 39 | Says A' B X => insert X (knows Spy evs) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 40 | | Gets A' X => knows Spy evs | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 41 | | Notes A' X => | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
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: 
32404diff
changeset | 43 | else | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 44 | (case ev of | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 45 | Says A' B X => | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 46 | 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: 
32404diff
changeset | 47 | | Gets A' X => | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 48 | 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: 
32404diff
changeset | 49 | | Notes A' X => | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
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: 
5183diff
changeset | 51 | (* | 
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
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: 
5183diff
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: 
5183diff
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: 
5183diff
changeset | 55 | *) | 
| 3678 | 56 | |
| 61830 | 57 | text\<open>The constant "spies" is retained for compatibility's sake\<close> | 
| 38964 | 58 | |
| 59 | abbreviation (input) | |
| 60 | spies :: "event list => msg set" where | |
| 61 | "spies == knows Spy" | |
| 62 | ||
| 63 | ||
| 64 | (*Set of items that might be visible to somebody: | |
| 3683 | 65 | complement of the set of fresh items*) | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 66 | |
| 38964 | 67 | primrec used :: "event list => msg set" | 
| 68 | where | |
| 11104 | 69 | used_Nil: "used [] = (UN B. parts (initState B))" | 
| 38964 | 70 | | used_Cons: "used (ev # evs) = | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 71 | (case ev of | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 72 |                         Says A B X => parts {X} \<union> used evs
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 73 | | Gets A X => used evs | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32404diff
changeset | 74 |                       | Notes A X  => parts {X} \<union> used evs)"
 | 
| 61830 | 75 |     \<comment>\<open>The case for @{term Gets} seems anomalous, but @{term Gets} always
 | 
| 13935 | 76 |         follows @{term Says} in real protocols.  Seems difficult to change.
 | 
| 61830 | 77 | 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: 
5183diff
changeset | 78 | |
| 13926 | 79 | lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs --> X \<in> used evs" | 
| 80 | apply (induct_tac evs) | |
| 11463 | 81 | apply (auto split: event.split) | 
| 82 | done | |
| 83 | ||
| 13926 | 84 | lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs --> X \<in> used evs" | 
| 85 | apply (induct_tac evs) | |
| 11463 | 86 | apply (auto split: event.split) | 
| 87 | done | |
| 88 | ||
| 13926 | 89 | |
| 61830 | 90 | subsection\<open>Function @{term knows}\<close>
 | 
| 13926 | 91 | |
| 13956 | 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.*) | |
| 45605 | 95 | lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs | 
| 13926 | 96 | |
| 97 | lemma knows_Spy_Says [simp]: | |
| 98 | "knows Spy (Says A B X # evs) = insert X (knows Spy evs)" | |
| 99 | by simp | |
| 100 | ||
| 61830 | 101 | text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits | 
| 102 |       on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}\<close>
 | |
| 13926 | 103 | lemma knows_Spy_Notes [simp]: | 
| 104 | "knows Spy (Notes A X # evs) = | |
| 105 | (if A:bad then insert X (knows Spy evs) else knows Spy evs)" | |
| 106 | by simp | |
| 107 | ||
| 108 | lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" | |
| 109 | by simp | |
| 110 | ||
| 111 | lemma knows_Spy_subset_knows_Spy_Says: | |
| 13935 | 112 | "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)" | 
| 13926 | 113 | by (simp add: subset_insertI) | 
| 114 | ||
| 115 | lemma knows_Spy_subset_knows_Spy_Notes: | |
| 13935 | 116 | "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)" | 
| 13926 | 117 | by force | 
| 118 | ||
| 119 | lemma knows_Spy_subset_knows_Spy_Gets: | |
| 13935 | 120 | "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)" | 
| 13926 | 121 | by (simp add: subset_insertI) | 
| 122 | ||
| 61830 | 123 | text\<open>Spy sees what is sent on the traffic\<close> | 
| 13926 | 124 | lemma Says_imp_knows_Spy [rule_format]: | 
| 125 | "Says A B X \<in> set evs --> X \<in> knows Spy evs" | |
| 126 | apply (induct_tac "evs") | |
| 63648 | 127 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 128 | done | 
| 129 | ||
| 130 | lemma Notes_imp_knows_Spy [rule_format]: | |
| 131 | "Notes A X \<in> set evs --> A: bad --> X \<in> knows Spy evs" | |
| 132 | apply (induct_tac "evs") | |
| 63648 | 133 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 134 | done | 
| 135 | ||
| 136 | ||
| 61830 | 137 | text\<open>Elimination rules: derive contradictions from old Says events containing | 
| 138 | items known to be fresh\<close> | |
| 32404 | 139 | lemmas Says_imp_parts_knows_Spy = | 
| 46471 | 140 | Says_imp_knows_Spy [THEN parts.Inj, elim_format] | 
| 32404 | 141 | |
| 13926 | 142 | lemmas knows_Spy_partsEs = | 
| 46471 | 143 | Says_imp_parts_knows_Spy parts.Body [elim_format] | 
| 13926 | 144 | |
| 18570 | 145 | lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj] | 
| 146 | ||
| 61830 | 147 | text\<open>Compatibility for the old "spies" function\<close> | 
| 13926 | 148 | lemmas spies_partsEs = knows_Spy_partsEs | 
| 149 | lemmas Says_imp_spies = Says_imp_knows_Spy | |
| 13935 | 150 | lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy] | 
| 13926 | 151 | |
| 152 | ||
| 61830 | 153 | subsection\<open>Knowledge of Agents\<close> | 
| 13926 | 154 | |
| 13935 | 155 | lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)" | 
| 156 | by (simp add: subset_insertI) | |
| 13926 | 157 | |
| 13935 | 158 | lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)" | 
| 159 | by (simp add: subset_insertI) | |
| 13926 | 160 | |
| 13935 | 161 | lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)" | 
| 162 | by (simp add: subset_insertI) | |
| 13926 | 163 | |
| 61830 | 164 | text\<open>Agents know what they say\<close> | 
| 13926 | 165 | lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs --> X \<in> knows A evs" | 
| 166 | apply (induct_tac "evs") | |
| 63648 | 167 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 168 | apply blast | 
| 169 | done | |
| 170 | ||
| 61830 | 171 | text\<open>Agents know what they note\<close> | 
| 13926 | 172 | lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs --> X \<in> knows A evs" | 
| 173 | apply (induct_tac "evs") | |
| 63648 | 174 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 175 | apply blast | 
| 176 | done | |
| 177 | ||
| 61830 | 178 | text\<open>Agents know what they receive\<close> | 
| 13926 | 179 | lemma Gets_imp_knows_agents [rule_format]: | 
| 180 | "A \<noteq> Spy --> Gets A X \<in> set evs --> X \<in> knows A evs" | |
| 181 | apply (induct_tac "evs") | |
| 63648 | 182 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 183 | done | 
| 184 | ||
| 185 | ||
| 61830 | 186 | text\<open>What agents DIFFERENT FROM Spy know | 
| 187 | was either said, or noted, or got, or known initially\<close> | |
| 13926 | 188 | lemma knows_imp_Says_Gets_Notes_initState [rule_format]: | 
| 189 | "[| X \<in> knows A evs; A \<noteq> Spy |] ==> EX B. | |
| 190 | Says A B X \<in> set evs | Gets A X \<in> set evs | Notes A X \<in> set evs | X \<in> initState A" | |
| 191 | apply (erule rev_mp) | |
| 192 | apply (induct_tac "evs") | |
| 63648 | 193 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 194 | apply blast | 
| 195 | done | |
| 196 | ||
| 61830 | 197 | text\<open>What the Spy knows -- for the time being -- | 
| 198 | was either said or noted, or known initially\<close> | |
| 13926 | 199 | lemma knows_Spy_imp_Says_Notes_initState [rule_format]: | 
| 200 | "[| X \<in> knows Spy evs |] ==> EX A B. | |
| 201 | Says A B X \<in> set evs | Notes A X \<in> set evs | X \<in> initState Spy" | |
| 202 | apply (erule rev_mp) | |
| 203 | apply (induct_tac "evs") | |
| 63648 | 204 | apply (simp_all (no_asm_simp) split: event.split) | 
| 13926 | 205 | apply blast | 
| 206 | done | |
| 207 | ||
| 13935 | 208 | lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs" | 
| 209 | apply (induct_tac "evs", force) | |
| 210 | apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) | |
| 13926 | 211 | done | 
| 212 | ||
| 213 | lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] | |
| 214 | ||
| 215 | lemma initState_into_used: "X \<in> parts (initState B) ==> X \<in> used evs" | |
| 216 | apply (induct_tac "evs") | |
| 63648 | 217 | apply (simp_all add: parts_insert_knows_A split: event.split, blast) | 
| 13926 | 218 | done | 
| 219 | ||
| 220 | lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
 | |
| 221 | by simp | |
| 222 | ||
| 223 | lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
 | |
| 224 | by simp | |
| 225 | ||
| 226 | lemma used_Gets [simp]: "used (Gets A X # evs) = used evs" | |
| 227 | by simp | |
| 228 | ||
| 13935 | 229 | lemma used_nil_subset: "used [] \<subseteq> used evs" | 
| 230 | apply simp | |
| 13926 | 231 | apply (blast intro: initState_into_used) | 
| 232 | done | |
| 233 | ||
| 61830 | 234 | text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close> | 
| 13935 | 235 | declare knows_Cons [simp del] | 
| 236 | used_Nil [simp del] used_Cons [simp del] | |
| 13926 | 237 | |
| 238 | ||
| 61830 | 239 | text\<open>For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
 | 
| 13926 | 240 | New events added by induction to "evs" are discarded. Provided | 
| 241 | this information isn't needed, the proof will be much shorter, since | |
| 61830 | 242 |   it will omit complicated reasoning about @{term analz}.\<close>
 | 
| 13926 | 243 | |
| 244 | lemmas 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 | ||
| 11104 | 249 | |
| 13922 | 250 | lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)" | 
| 53428 | 251 | by (cases e, auto simp: knows_Cons) | 
| 13922 | 252 | |
| 13935 | 253 | lemma initState_subset_knows: "initState A \<subseteq> knows A evs" | 
| 13926 | 254 | apply (induct_tac evs, simp) | 
| 13922 | 255 | apply (blast intro: knows_subset_knows_Cons [THEN subsetD]) | 
| 256 | done | |
| 257 | ||
| 258 | ||
| 61830 | 259 | text\<open>For proving \<open>new_keys_not_used\<close>\<close> | 
| 13922 | 260 | lemma keysFor_parts_insert: | 
| 13926 | 261 | "[| K \<in> keysFor (parts (insert X G)); X \<in> synth (analz H) |] | 
| 41693 
47532fe9e075
Introduction of metis calls and other cosmetic modifications.
 paulson parents: 
38964diff
changeset | 262 | ==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H" | 
| 13922 | 263 | by (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 | ||
| 24122 | 268 | |
| 45605 | 269 | lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs | 
| 27225 | 270 | |
| 24122 | 271 | ML | 
| 61830 | 272 | \<open> | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58889diff
changeset | 273 | fun analz_mono_contra_tac ctxt = | 
| 60754 | 274 |   resolve_tac ctxt @{thms analz_impI} THEN' 
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 275 |   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: 
58889diff
changeset | 276 | THEN' (mp_tac ctxt) | 
| 61830 | 277 | \<close> | 
| 24122 | 278 | |
| 61830 | 279 | method_setup analz_mono_contra = \<open> | 
| 280 | Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close> | |
| 13922 | 281 | "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P" | 
| 282 | ||
| 61830 | 283 | subsubsection\<open>Useful for case analysis on whether a hash is a spoof or not\<close> | 
| 13922 | 284 | |
| 45605 | 285 | lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))"] for Y evs | 
| 27225 | 286 | |
| 13922 | 287 | ML | 
| 61830 | 288 | \<open> | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58889diff
changeset | 289 | fun synth_analz_mono_contra_tac ctxt = | 
| 60754 | 290 |   resolve_tac ctxt @{thms syan_impI} THEN'
 | 
| 27225 | 291 | REPEAT1 o | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 292 | (dresolve_tac ctxt | 
| 27225 | 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' | |
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58889diff
changeset | 297 | mp_tac ctxt | 
| 61830 | 298 | \<close> | 
| 13922 | 299 | |
| 61830 | 300 | method_setup synth_analz_mono_contra = \<open> | 
| 301 | Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))\<close> | |
| 13922 | 302 | "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 | 303 | |
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 304 | end |