| author | huffman | 
| Tue, 11 Oct 2005 23:47:29 +0200 | |
| changeset 17836 | 5d9c9e284d16 | 
| parent 16417 | 9bc16273c2d4 | 
| child 17990 | 86d462f305e0 | 
| permissions | -rw-r--r-- | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 1 | (* Title: HOL/Auth/Event | 
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 4 | Copyright 1996 University of Cambridge | 
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 5 | |
| 3683 | 6 | Datatype of events; function "spies"; freshness | 
| 3678 | 7 | |
| 3683 | 8 | "bad" agents have been broken by the Spy; their private keys and internal | 
| 3678 | 9 | stores are visible to him | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 10 | *) | 
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 11 | |
| 13956 | 12 | header{*Theory of Events for Security Protocols*}
 | 
| 13 | ||
| 16417 | 14 | theory Event imports Message begin | 
| 11104 | 15 | |
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 16 | consts (*Initial states of agents -- parameter of the construction*) | 
| 11104 | 17 | initState :: "agent => msg set" | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 18 | |
| 6399 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 19 | datatype | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 20 | event = Says agent agent msg | 
| 6399 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 21 | | Gets agent msg | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 22 | | Notes agent msg | 
| 6308 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 23 | |
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 24 | consts | 
| 11104 | 25 | bad :: "agent set" (*compromised agents*) | 
| 26 | knows :: "agent => event list => msg set" | |
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 27 | |
| 6308 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 28 | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 29 | text{*The constant "spies" is retained for compatibility's sake*}
 | 
| 6308 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 30 | syntax | 
| 11104 | 31 | spies :: "event list => msg set" | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3512diff
changeset | 32 | |
| 6308 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 33 | translations | 
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 34 | "spies" => "knows Spy" | 
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 35 | |
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 36 | 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: 
13956diff
changeset | 37 | specification (bad) | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 38 | Spy_in_bad [iff]: "Spy \<in> bad" | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 39 | Server_not_bad [iff]: "Server \<notin> bad" | 
| 14126 
28824746d046
Tidying and replacement of some axioms by specifications
 paulson parents: 
13956diff
changeset | 40 |     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 | 41 | |
| 5183 | 42 | primrec | 
| 11104 | 43 | knows_Nil: "knows A [] = initState A" | 
| 44 | knows_Cons: | |
| 6399 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 45 | "knows A (ev # evs) = | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 46 | (if A = Spy then | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 47 | (case ev of | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 48 | Says A' B X => insert X (knows Spy evs) | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 49 | | Gets A' X => knows Spy evs | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 50 | | Notes A' X => | 
| 13922 | 51 | if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs) | 
| 6399 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 52 | else | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 53 | (case ev of | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 54 | Says A' B X => | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 55 | if A'=A then insert X (knows A evs) else knows A evs | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 56 | | Gets A' X => | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 57 | if A'=A then insert X (knows A evs) else knows A evs | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 58 | | Notes A' X => | 
| 
4a9040b85e2e
exchanged the order of Gets and Notes in datatype event
 paulson parents: 
6308diff
changeset | 59 | 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 | 60 | |
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 61 | (* | 
| 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5183diff
changeset | 62 | 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 | 63 | 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 | 64 | 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 | 65 | *) | 
| 3678 | 66 | |
| 3683 | 67 | consts | 
| 68 | (*Set of items that might be visible to somebody: | |
| 69 | complement of the set of fresh items*) | |
| 11104 | 70 | used :: "event list => msg set" | 
| 3512 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 71 | |
| 5183 | 72 | primrec | 
| 11104 | 73 | used_Nil: "used [] = (UN B. parts (initState B))" | 
| 74 | used_Cons: "used (ev # evs) = | |
| 75 | (case ev of | |
| 13935 | 76 | 			Says A B X => parts {X} \<union> used evs
 | 
| 11104 | 77 | | Gets A X => used evs | 
| 13935 | 78 | 		      | Notes A X  => parts {X} \<union> used evs)"
 | 
| 79 |     --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
 | |
| 80 |         follows @{term Says} in real protocols.  Seems difficult to change.
 | |
| 81 |         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: 
5183diff
changeset | 82 | |
| 13926 | 83 | lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs --> X \<in> used evs" | 
| 84 | apply (induct_tac evs) | |
| 11463 | 85 | apply (auto split: event.split) | 
| 86 | done | |
| 87 | ||
| 13926 | 88 | lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs --> X \<in> used evs" | 
| 89 | apply (induct_tac evs) | |
| 11463 | 90 | apply (auto split: event.split) | 
| 91 | done | |
| 92 | ||
| 93 | lemma MPair_used [rule_format]: | |
| 13926 | 94 | "MPair X Y \<in> used evs --> X \<in> used evs & Y \<in> used evs" | 
| 95 | apply (induct_tac evs) | |
| 11463 | 96 | apply (auto split: event.split) | 
| 97 | done | |
| 98 | ||
| 13926 | 99 | |
| 100 | subsection{*Function @{term knows}*}
 | |
| 101 | ||
| 13956 | 102 | (*Simplifying | 
| 103 |  parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
 | |
| 104 | This version won't loop with the simplifier.*) | |
| 13935 | 105 | lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard] | 
| 13926 | 106 | |
| 107 | lemma knows_Spy_Says [simp]: | |
| 108 | "knows Spy (Says A B X # evs) = insert X (knows Spy evs)" | |
| 109 | by simp | |
| 110 | ||
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14126diff
changeset | 111 | 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: 
14126diff
changeset | 112 |       on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
 | 
| 13926 | 113 | lemma knows_Spy_Notes [simp]: | 
| 114 | "knows Spy (Notes A X # evs) = | |
| 115 | (if A:bad then insert X (knows Spy evs) else knows Spy evs)" | |
| 116 | by simp | |
| 117 | ||
| 118 | lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" | |
| 119 | by simp | |
| 120 | ||
| 121 | lemma knows_Spy_subset_knows_Spy_Says: | |
| 13935 | 122 | "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)" | 
| 13926 | 123 | by (simp add: subset_insertI) | 
| 124 | ||
| 125 | lemma knows_Spy_subset_knows_Spy_Notes: | |
| 13935 | 126 | "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)" | 
| 13926 | 127 | by force | 
| 128 | ||
| 129 | lemma knows_Spy_subset_knows_Spy_Gets: | |
| 13935 | 130 | "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)" | 
| 13926 | 131 | by (simp add: subset_insertI) | 
| 132 | ||
| 133 | text{*Spy sees what is sent on the traffic*}
 | |
| 134 | lemma Says_imp_knows_Spy [rule_format]: | |
| 135 | "Says A B X \<in> set evs --> X \<in> knows Spy evs" | |
| 136 | apply (induct_tac "evs") | |
| 137 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 138 | done | |
| 139 | ||
| 140 | lemma Notes_imp_knows_Spy [rule_format]: | |
| 141 | "Notes A X \<in> set evs --> A: bad --> X \<in> knows Spy evs" | |
| 142 | apply (induct_tac "evs") | |
| 143 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 144 | done | |
| 145 | ||
| 146 | ||
| 147 | text{*Elimination rules: derive contradictions from old Says events containing
 | |
| 148 | items known to be fresh*} | |
| 149 | lemmas knows_Spy_partsEs = | |
| 150 | Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] | |
| 151 | parts.Body [THEN revcut_rl, standard] | |
| 152 | ||
| 153 | text{*Compatibility for the old "spies" function*}
 | |
| 154 | lemmas spies_partsEs = knows_Spy_partsEs | |
| 155 | lemmas Says_imp_spies = Says_imp_knows_Spy | |
| 13935 | 156 | lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy] | 
| 13926 | 157 | |
| 158 | ||
| 159 | subsection{*Knowledge of Agents*}
 | |
| 160 | ||
| 161 | lemma knows_Says: "knows A (Says A B X # evs) = insert X (knows A evs)" | |
| 162 | by simp | |
| 163 | ||
| 164 | lemma knows_Notes: "knows A (Notes A X # evs) = insert X (knows A evs)" | |
| 165 | by simp | |
| 166 | ||
| 167 | lemma knows_Gets: | |
| 168 | "A \<noteq> Spy --> knows A (Gets A X # evs) = insert X (knows A evs)" | |
| 169 | by simp | |
| 170 | ||
| 171 | ||
| 13935 | 172 | lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)" | 
| 173 | by (simp add: subset_insertI) | |
| 13926 | 174 | |
| 13935 | 175 | lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)" | 
| 176 | by (simp add: subset_insertI) | |
| 13926 | 177 | |
| 13935 | 178 | lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)" | 
| 179 | by (simp add: subset_insertI) | |
| 13926 | 180 | |
| 181 | text{*Agents know what they say*}
 | |
| 182 | lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs --> X \<in> knows A evs" | |
| 183 | apply (induct_tac "evs") | |
| 184 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 185 | apply blast | |
| 186 | done | |
| 187 | ||
| 188 | text{*Agents know what they note*}
 | |
| 189 | lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs --> X \<in> knows A evs" | |
| 190 | apply (induct_tac "evs") | |
| 191 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 192 | apply blast | |
| 193 | done | |
| 194 | ||
| 195 | text{*Agents know what they receive*}
 | |
| 196 | lemma Gets_imp_knows_agents [rule_format]: | |
| 197 | "A \<noteq> Spy --> Gets A X \<in> set evs --> X \<in> knows A evs" | |
| 198 | apply (induct_tac "evs") | |
| 199 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 200 | done | |
| 201 | ||
| 202 | ||
| 203 | text{*What agents DIFFERENT FROM Spy know 
 | |
| 204 | was either said, or noted, or got, or known initially*} | |
| 205 | lemma knows_imp_Says_Gets_Notes_initState [rule_format]: | |
| 206 | "[| X \<in> knows A evs; A \<noteq> Spy |] ==> EX B. | |
| 207 | Says A B X \<in> set evs | Gets A X \<in> set evs | Notes A X \<in> set evs | X \<in> initState A" | |
| 208 | apply (erule rev_mp) | |
| 209 | apply (induct_tac "evs") | |
| 210 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 211 | apply blast | |
| 212 | done | |
| 213 | ||
| 214 | text{*What the Spy knows -- for the time being --
 | |
| 215 | was either said or noted, or known initially*} | |
| 216 | lemma knows_Spy_imp_Says_Notes_initState [rule_format]: | |
| 217 | "[| X \<in> knows Spy evs |] ==> EX A B. | |
| 218 | Says A B X \<in> set evs | Notes A X \<in> set evs | X \<in> initState Spy" | |
| 219 | apply (erule rev_mp) | |
| 220 | apply (induct_tac "evs") | |
| 221 | apply (simp_all (no_asm_simp) split add: event.split) | |
| 222 | apply blast | |
| 223 | done | |
| 224 | ||
| 13935 | 225 | lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs" | 
| 226 | apply (induct_tac "evs", force) | |
| 227 | apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) | |
| 13926 | 228 | done | 
| 229 | ||
| 230 | lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] | |
| 231 | ||
| 232 | lemma initState_into_used: "X \<in> parts (initState B) ==> X \<in> used evs" | |
| 233 | apply (induct_tac "evs") | |
| 13935 | 234 | apply (simp_all add: parts_insert_knows_A split add: event.split, blast) | 
| 13926 | 235 | done | 
| 236 | ||
| 237 | lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
 | |
| 238 | by simp | |
| 239 | ||
| 240 | lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
 | |
| 241 | by simp | |
| 242 | ||
| 243 | lemma used_Gets [simp]: "used (Gets A X # evs) = used evs" | |
| 244 | by simp | |
| 245 | ||
| 13935 | 246 | lemma used_nil_subset: "used [] \<subseteq> used evs" | 
| 247 | apply simp | |
| 13926 | 248 | apply (blast intro: initState_into_used) | 
| 249 | done | |
| 250 | ||
| 251 | text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
 | |
| 13935 | 252 | declare knows_Cons [simp del] | 
| 253 | used_Nil [simp del] used_Cons [simp del] | |
| 13926 | 254 | |
| 255 | ||
| 256 | text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
 | |
| 257 | New events added by induction to "evs" are discarded. Provided | |
| 258 | this information isn't needed, the proof will be much shorter, since | |
| 259 |   it will omit complicated reasoning about @{term analz}.*}
 | |
| 260 | ||
| 261 | lemmas analz_mono_contra = | |
| 262 | knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD] | |
| 263 | knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD] | |
| 264 | knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD] | |
| 265 | ||
| 266 | ML | |
| 267 | {*
 | |
| 268 | val analz_mono_contra_tac = | |
| 269 | let val analz_impI = inst "P" "?Y \<notin> analz (knows Spy ?evs)" impI | |
| 270 | in | |
| 271 | rtac analz_impI THEN' | |
| 272 | REPEAT1 o | |
| 273 | (dresolve_tac (thms"analz_mono_contra")) | |
| 274 | THEN' mp_tac | |
| 275 | end | |
| 276 | *} | |
| 277 | ||
| 11104 | 278 | |
| 13922 | 279 | lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)" | 
| 280 | by (induct e, auto simp: knows_Cons) | |
| 281 | ||
| 13935 | 282 | lemma initState_subset_knows: "initState A \<subseteq> knows A evs" | 
| 13926 | 283 | apply (induct_tac evs, simp) | 
| 13922 | 284 | apply (blast intro: knows_subset_knows_Cons [THEN subsetD]) | 
| 285 | done | |
| 286 | ||
| 287 | ||
| 13926 | 288 | text{*For proving @{text new_keys_not_used}*}
 | 
| 13922 | 289 | lemma keysFor_parts_insert: | 
| 13926 | 290 | "[| K \<in> keysFor (parts (insert X G)); X \<in> synth (analz H) |] | 
| 291 | ==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"; | |
| 13922 | 292 | by (force | 
| 293 | dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD] | |
| 294 | analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] | |
| 295 | intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD]) | |
| 296 | ||
| 11104 | 297 | method_setup analz_mono_contra = {*
 | 
| 298 | Method.no_args | |
| 299 | (Method.METHOD (fn facts => REPEAT_FIRST analz_mono_contra_tac)) *} | |
| 13922 | 300 | "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P" | 
| 301 | ||
| 302 | subsubsection{*Useful for case analysis on whether a hash is a spoof or not*}
 | |
| 303 | ||
| 304 | ML | |
| 305 | {*
 | |
| 13926 | 306 | val knows_Cons = thm "knows_Cons" | 
| 307 | val used_Nil = thm "used_Nil" | |
| 308 | val used_Cons = thm "used_Cons" | |
| 309 | ||
| 310 | val Notes_imp_used = thm "Notes_imp_used"; | |
| 311 | val Says_imp_used = thm "Says_imp_used"; | |
| 312 | val MPair_used = thm "MPair_used"; | |
| 313 | val Says_imp_knows_Spy = thm "Says_imp_knows_Spy"; | |
| 314 | val Notes_imp_knows_Spy = thm "Notes_imp_knows_Spy"; | |
| 315 | val knows_Spy_partsEs = thms "knows_Spy_partsEs"; | |
| 316 | val spies_partsEs = thms "spies_partsEs"; | |
| 317 | val Says_imp_spies = thm "Says_imp_spies"; | |
| 318 | val parts_insert_spies = thm "parts_insert_spies"; | |
| 319 | val Says_imp_knows = thm "Says_imp_knows"; | |
| 320 | val Notes_imp_knows = thm "Notes_imp_knows"; | |
| 321 | val Gets_imp_knows_agents = thm "Gets_imp_knows_agents"; | |
| 322 | val knows_imp_Says_Gets_Notes_initState = thm "knows_imp_Says_Gets_Notes_initState"; | |
| 323 | val knows_Spy_imp_Says_Notes_initState = thm "knows_Spy_imp_Says_Notes_initState"; | |
| 324 | val usedI = thm "usedI"; | |
| 325 | val initState_into_used = thm "initState_into_used"; | |
| 326 | val used_Says = thm "used_Says"; | |
| 327 | val used_Notes = thm "used_Notes"; | |
| 328 | val used_Gets = thm "used_Gets"; | |
| 329 | val used_nil_subset = thm "used_nil_subset"; | |
| 330 | val analz_mono_contra = thms "analz_mono_contra"; | |
| 331 | val knows_subset_knows_Cons = thm "knows_subset_knows_Cons"; | |
| 332 | val initState_subset_knows = thm "initState_subset_knows"; | |
| 333 | val keysFor_parts_insert = thm "keysFor_parts_insert"; | |
| 334 | ||
| 335 | ||
| 13922 | 336 | val synth_analz_mono = thm "synth_analz_mono"; | 
| 337 | ||
| 13935 | 338 | val knows_Spy_subset_knows_Spy_Says = thm "knows_Spy_subset_knows_Spy_Says"; | 
| 339 | val knows_Spy_subset_knows_Spy_Notes = thm "knows_Spy_subset_knows_Spy_Notes"; | |
| 340 | val knows_Spy_subset_knows_Spy_Gets = thm "knows_Spy_subset_knows_Spy_Gets"; | |
| 341 | ||
| 342 | ||
| 13922 | 343 | val synth_analz_mono_contra_tac = | 
| 13926 | 344 | let val syan_impI = inst "P" "?Y \<notin> synth (analz (knows Spy ?evs))" impI | 
| 13922 | 345 | in | 
| 346 | rtac syan_impI THEN' | |
| 347 | REPEAT1 o | |
| 348 | (dresolve_tac | |
| 349 | [knows_Spy_subset_knows_Spy_Says RS synth_analz_mono RS contra_subsetD, | |
| 350 | knows_Spy_subset_knows_Spy_Notes RS synth_analz_mono RS contra_subsetD, | |
| 351 | knows_Spy_subset_knows_Spy_Gets RS synth_analz_mono RS contra_subsetD]) | |
| 352 | THEN' | |
| 353 | mp_tac | |
| 354 | end; | |
| 355 | *} | |
| 356 | ||
| 357 | method_setup synth_analz_mono_contra = {*
 | |
| 358 | Method.no_args | |
| 359 | (Method.METHOD (fn facts => REPEAT_FIRST synth_analz_mono_contra_tac)) *} | |
| 360 | "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 | 361 | |
| 
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
 paulson parents: diff
changeset | 362 | end |