| author | wenzelm | 
| Mon, 12 Apr 2021 22:57:39 +0200 | |
| changeset 73576 | b50f8cc8c08e | 
| parent 69597 | ff784d5a5bfb | 
| child 76287 | cdc14f94c754 | 
| permissions | -rw-r--r-- | 
| 41775 | 1 | (* Title: HOL/Auth/Guard/Guard.thy | 
| 2 | Author: Frederic Blanqui, University of Cambridge Computer Laboratory | |
| 3 | Copyright 2002 University of Cambridge | |
| 4 | *) | |
| 13508 | 5 | |
| 61830 | 6 | section\<open>Protocol-Independent Confidentiality Theorem on Nonces\<close> | 
| 13508 | 7 | |
| 16417 | 8 | theory Guard imports Analz Extensions begin | 
| 13508 | 9 | |
| 10 | (****************************************************************************** | |
| 11 | messages where all the occurrences of Nonce n are | |
| 12 | in a sub-message of the form Crypt (invKey K) X with K:Ks | |
| 13 | ******************************************************************************) | |
| 14 | ||
| 23746 | 15 | inductive_set | 
| 67613 | 16 | guard :: "nat \<Rightarrow> key set \<Rightarrow> msg set" | 
| 23746 | 17 | for n :: nat and Ks :: "key set" | 
| 18 | where | |
| 67613 | 19 |   No_Nonce [intro]: "Nonce n \<notin> parts {X} \<Longrightarrow> X \<in> guard n Ks"
 | 
| 20 | | Guard_Nonce [intro]: "invKey K \<in> Ks \<Longrightarrow> Crypt K X \<in> guard n Ks" | |
| 21 | | Crypt [intro]: "X \<in> guard n Ks \<Longrightarrow> Crypt K X \<in> guard n Ks" | |
| 22 | | Pair [intro]: "[| X \<in> guard n Ks; Y \<in> guard n Ks |] ==> \<lbrace>X,Y\<rbrace> \<in> guard n Ks" | |
| 13508 | 23 | |
| 69597 | 24 | subsection\<open>basic facts about \<^term>\<open>guard\<close>\<close> | 
| 13508 | 25 | |
| 67613 | 26 | lemma Key_is_guard [iff]: "Key K \<in> guard n Ks" | 
| 13508 | 27 | by auto | 
| 28 | ||
| 67613 | 29 | lemma Agent_is_guard [iff]: "Agent A \<in> guard n Ks" | 
| 13508 | 30 | by auto | 
| 31 | ||
| 67613 | 32 | lemma Number_is_guard [iff]: "Number r \<in> guard n Ks" | 
| 13508 | 33 | by auto | 
| 34 | ||
| 67613 | 35 | lemma Nonce_notin_guard: "X \<in> guard n Ks \<Longrightarrow> X \<noteq> Nonce n" | 
| 13508 | 36 | by (erule guard.induct, auto) | 
| 37 | ||
| 67613 | 38 | lemma Nonce_notin_guard_iff [iff]: "Nonce n \<notin> guard n Ks" | 
| 13508 | 39 | by (auto dest: Nonce_notin_guard) | 
| 40 | ||
| 67613 | 41 | lemma guard_has_Crypt [rule_format]: "X \<in> guard n Ks ==> Nonce n \<in> parts {X}
 | 
| 42 | \<longrightarrow> (\<exists>K Y. Crypt K Y \<in> kparts {X} \<and> Nonce n \<in> parts {Y})"
 | |
| 13508 | 43 | by (erule guard.induct, auto) | 
| 44 | ||
| 67613 | 45 | lemma Nonce_notin_kparts_msg: "X \<in> guard n Ks \<Longrightarrow> Nonce n \<notin> kparts {X}"
 | 
| 13508 | 46 | by (erule guard.induct, auto) | 
| 47 | ||
| 67613 | 48 | lemma Nonce_in_kparts_imp_no_guard: "Nonce n \<in> kparts H | 
| 49 | \<Longrightarrow> \<exists>X. X \<in> H \<and> X \<notin> guard n Ks" | |
| 13508 | 50 | apply (drule in_kparts, clarify) | 
| 51 | apply (rule_tac x=X in exI, clarify) | |
| 52 | by (auto dest: Nonce_notin_kparts_msg) | |
| 53 | ||
| 67613 | 54 | lemma guard_kparts [rule_format]: "X \<in> guard n Ks \<Longrightarrow> | 
| 55 | Y \<in> kparts {X} \<longrightarrow> Y \<in> guard n Ks"
 | |
| 13508 | 56 | by (erule guard.induct, auto) | 
| 57 | ||
| 67613 | 58 | lemma guard_Crypt: "[| Crypt K Y \<in> guard n Ks; K \<notin> invKey`Ks |] ==> Y \<in> guard n Ks" | 
| 59 | by (ind_cases "Crypt K Y \<in> guard n Ks") (auto intro!: image_eqI) | |
| 13508 | 60 | |
| 61956 | 61 | lemma guard_MPair [iff]: "(\<lbrace>X,Y\<rbrace> \<in> guard n Ks) = (X \<in> guard n Ks \<and> Y \<in> guard n Ks)" | 
| 62 | by (auto, (ind_cases "\<lbrace>X,Y\<rbrace> \<in> guard n Ks", auto)+) | |
| 13508 | 63 | |
| 67613 | 64 | lemma guard_not_guard [rule_format]: "X \<in> guard n Ks \<Longrightarrow> | 
| 65 | Crypt K Y \<in> kparts {X} \<longrightarrow> Nonce n \<in> kparts {Y} \<longrightarrow> Y \<notin> guard n Ks"
 | |
| 13508 | 66 | by (erule guard.induct, auto dest: guard_kparts) | 
| 67 | ||
| 67613 | 68 | lemma guard_extand: "[| X \<in> guard n Ks; Ks \<subseteq> Ks' |] ==> X \<in> guard n Ks'" | 
| 13508 | 69 | by (erule guard.induct, auto) | 
| 70 | ||
| 61830 | 71 | subsection\<open>guarded sets\<close> | 
| 13508 | 72 | |
| 67613 | 73 | definition Guard :: "nat \<Rightarrow> key set \<Rightarrow> msg set \<Rightarrow> bool" where | 
| 74 | "Guard n Ks H \<equiv> \<forall>X. X \<in> H \<longrightarrow> X \<in> guard n Ks" | |
| 13508 | 75 | |
| 69597 | 76 | subsection\<open>basic facts about \<^term>\<open>Guard\<close>\<close> | 
| 13508 | 77 | |
| 78 | lemma Guard_empty [iff]: "Guard n Ks {}"
 | |
| 79 | by (simp add: Guard_def) | |
| 80 | ||
| 67613 | 81 | lemma notin_parts_Guard [intro]: "Nonce n \<notin> parts G \<Longrightarrow> Guard n Ks G" | 
| 13508 | 82 | apply (unfold Guard_def, clarify) | 
| 67613 | 83 | apply (subgoal_tac "Nonce n \<notin> parts {X}")
 | 
| 13508 | 84 | by (auto dest: parts_sub) | 
| 85 | ||
| 67613 | 86 | lemma Nonce_notin_kparts [simplified]: "Guard n Ks H \<Longrightarrow> Nonce n \<notin> kparts H" | 
| 13508 | 87 | by (auto simp: Guard_def dest: in_kparts Nonce_notin_kparts_msg) | 
| 88 | ||
| 67613 | 89 | lemma Guard_must_decrypt: "[| Guard n Ks H; Nonce n \<in> analz H |] ==> | 
| 90 | \<exists>K Y. Crypt K Y \<in> kparts H \<and> Key (invKey K) \<in> kparts H" | |
| 91 | apply (drule_tac P="\<lambda>G. Nonce n \<in> G" in analz_pparts_kparts_substD, simp) | |
| 13508 | 92 | by (drule must_decrypt, auto dest: Nonce_notin_kparts) | 
| 93 | ||
| 94 | lemma Guard_kparts [intro]: "Guard n Ks H ==> Guard n Ks (kparts H)" | |
| 95 | by (auto simp: Guard_def dest: in_kparts guard_kparts) | |
| 96 | ||
| 97 | lemma Guard_mono: "[| Guard n Ks H; G <= H |] ==> Guard n Ks G" | |
| 98 | by (auto simp: Guard_def) | |
| 99 | ||
| 100 | lemma Guard_insert [iff]: "Guard n Ks (insert X H) | |
| 67613 | 101 | = (Guard n Ks H \<and> X \<in> guard n Ks)" | 
| 13508 | 102 | by (auto simp: Guard_def) | 
| 103 | ||
| 104 | lemma Guard_Un [iff]: "Guard n Ks (G Un H) = (Guard n Ks G & Guard n Ks H)" | |
| 105 | by (auto simp: Guard_def) | |
| 106 | ||
| 107 | lemma Guard_synth [intro]: "Guard n Ks G ==> Guard n Ks (synth G)" | |
| 108 | by (auto simp: Guard_def, erule synth.induct, auto) | |
| 109 | ||
| 67613 | 110 | lemma Guard_analz [intro]: "[| Guard n Ks G; \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |] | 
| 13508 | 111 | ==> Guard n Ks (analz G)" | 
| 112 | apply (auto simp: Guard_def) | |
| 113 | apply (erule analz.induct, auto) | |
| 67613 | 114 | by (ind_cases "Crypt K Xa \<in> guard n Ks" for K Xa, auto) | 
| 13508 | 115 | |
| 67613 | 116 | lemma in_Guard [dest]: "[| X \<in> G; Guard n Ks G |] ==> X \<in> guard n Ks" | 
| 13508 | 117 | by (auto simp: Guard_def) | 
| 118 | ||
| 67613 | 119 | lemma in_synth_Guard: "[| X \<in> synth G; Guard n Ks G |] ==> X \<in> guard n Ks" | 
| 13508 | 120 | by (drule Guard_synth, auto) | 
| 121 | ||
| 67613 | 122 | lemma in_analz_Guard: "[| X \<in> analz G; Guard n Ks G; | 
| 123 | \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |] ==> X \<in> guard n Ks" | |
| 13508 | 124 | by (drule Guard_analz, auto) | 
| 125 | ||
| 126 | lemma Guard_keyset [simp]: "keyset G ==> Guard n Ks G" | |
| 127 | by (auto simp: Guard_def) | |
| 128 | ||
| 67613 | 129 | lemma Guard_Un_keyset: "[| Guard n Ks G; keyset H |] ==> Guard n Ks (G \<union> H)" | 
| 13508 | 130 | by auto | 
| 131 | ||
| 67613 | 132 | lemma in_Guard_kparts: "[| X \<in> G; Guard n Ks G; Y \<in> kparts {X} |] ==> Y \<in> guard n Ks"
 | 
| 13508 | 133 | by blast | 
| 134 | ||
| 67613 | 135 | lemma in_Guard_kparts_neq: "[| X \<in> G; Guard n Ks G; Nonce n' \<in> kparts {X} |]
 | 
| 136 | ==> n \<noteq> n'" | |
| 13508 | 137 | by (blast dest: in_Guard_kparts) | 
| 138 | ||
| 67613 | 139 | lemma in_Guard_kparts_Crypt: "[| X \<in> G; Guard n Ks G; is_MPair X; | 
| 140 | Crypt K Y \<in> kparts {X}; Nonce n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
 | |
| 13508 | 141 | apply (drule in_Guard, simp) | 
| 142 | apply (frule guard_not_guard, simp+) | |
| 143 | apply (drule guard_kparts, simp) | |
| 67613 | 144 | by (ind_cases "Crypt K Y \<in> guard n Ks", auto) | 
| 13508 | 145 | |
| 67613 | 146 | lemma Guard_extand: "[| Guard n Ks G; Ks \<subseteq> Ks' |] ==> Guard n Ks' G" | 
| 13508 | 147 | by (auto simp: Guard_def dest: guard_extand) | 
| 148 | ||
| 67613 | 149 | lemma guard_invKey [rule_format]: "[| X \<in> guard n Ks; Nonce n \<in> kparts {Y} |] ==>
 | 
| 150 | Crypt K Y \<in> kparts {X} \<longrightarrow> invKey K \<in> Ks"
 | |
| 13508 | 151 | by (erule guard.induct, auto) | 
| 152 | ||
| 67613 | 153 | lemma Crypt_guard_invKey [rule_format]: "[| Crypt K Y \<in> guard n Ks; | 
| 154 | Nonce n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
 | |
| 13508 | 155 | by (auto dest: guard_invKey) | 
| 156 | ||
| 61830 | 157 | subsection\<open>set obtained by decrypting a message\<close> | 
| 13508 | 158 | |
| 20768 | 159 | abbreviation (input) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 160 | decrypt :: "msg set => key => msg => msg set" where | 
| 20768 | 161 |   "decrypt H K Y == insert Y (H - {Crypt K Y})"
 | 
| 13508 | 162 | |
| 67613 | 163 | lemma analz_decrypt: "[| Crypt K Y \<in> H; Key (invKey K) \<in> H; Nonce n \<in> analz H |] | 
| 164 | ==> Nonce n \<in> analz (decrypt H K Y)" | |
| 165 | apply (drule_tac P="\<lambda>H. Nonce n \<in> analz H" in ssubst [OF insert_Diff]) | |
| 14307 | 166 | apply assumption | 
| 167 | apply (simp only: analz_Crypt_if, simp) | |
| 168 | done | |
| 13508 | 169 | |
| 67613 | 170 | lemma parts_decrypt: "[| Crypt K Y \<in> H; X \<in> parts (decrypt H K Y) |] ==> X \<in> parts H" | 
| 13508 | 171 | by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body) | 
| 172 | ||
| 61830 | 173 | subsection\<open>number of Crypt's in a message\<close> | 
| 13508 | 174 | |
| 35418 | 175 | fun crypt_nb :: "msg => nat" | 
| 176 | where | |
| 177 | "crypt_nb (Crypt K X) = Suc (crypt_nb X)" | |
| 61956 | 178 | | "crypt_nb \<lbrace>X,Y\<rbrace> = crypt_nb X + crypt_nb Y" | 
| 35418 | 179 | | "crypt_nb X = 0" (* otherwise *) | 
| 13508 | 180 | |
| 69597 | 181 | subsection\<open>basic facts about \<^term>\<open>crypt_nb\<close>\<close> | 
| 13508 | 182 | |
| 67613 | 183 | lemma non_empty_crypt_msg: "Crypt K Y \<in> parts {X} \<Longrightarrow> crypt_nb X \<noteq> 0"
 | 
| 13508 | 184 | by (induct X, simp_all, safe, simp_all) | 
| 185 | ||
| 61830 | 186 | subsection\<open>number of Crypt's in a message list\<close> | 
| 13508 | 187 | |
| 35418 | 188 | primrec cnb :: "msg list => nat" | 
| 189 | where | |
| 190 | "cnb [] = 0" | |
| 191 | | "cnb (X#l) = crypt_nb X + cnb l" | |
| 13508 | 192 | |
| 69597 | 193 | subsection\<open>basic facts about \<^term>\<open>cnb\<close>\<close> | 
| 13508 | 194 | |
| 195 | lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'" | |
| 196 | by (induct l, auto) | |
| 197 | ||
| 37596 | 198 | lemma mem_cnb_minus: "x \<in> set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)" | 
| 199 | by (induct l) auto | |
| 13508 | 200 | |
| 201 | lemmas mem_cnb_minus_substI = mem_cnb_minus [THEN ssubst] | |
| 202 | ||
| 37596 | 203 | lemma cnb_minus [simp]: "x \<in> set l ==> cnb (remove l x) = cnb l - crypt_nb x" | 
| 13508 | 204 | apply (induct l, auto) | 
| 45600 
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
 wenzelm parents: 
41775diff
changeset | 205 | apply (erule_tac l=l and x=x in mem_cnb_minus_substI) | 
| 37596 | 206 | apply simp | 
| 207 | done | |
| 13508 | 208 | |
| 67613 | 209 | lemma parts_cnb: "Z \<in> parts (set l) \<Longrightarrow> | 
| 13508 | 210 | cnb l = (cnb l - crypt_nb Z) + crypt_nb Z" | 
| 211 | by (erule parts.induct, auto simp: in_set_conv_decomp) | |
| 212 | ||
| 67613 | 213 | lemma non_empty_crypt: "Crypt K Y \<in> parts (set l) \<Longrightarrow> cnb l \<noteq> 0" | 
| 13508 | 214 | by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD) | 
| 215 | ||
| 61830 | 216 | subsection\<open>list of kparts\<close> | 
| 13508 | 217 | |
| 67613 | 218 | lemma kparts_msg_set: "\<exists>l. kparts {X} = set l \<and> cnb l = crypt_nb X"
 | 
| 13508 | 219 | apply (induct X, simp_all) | 
| 58250 | 220 | apply (rename_tac agent, rule_tac x="[Agent agent]" in exI, simp) | 
| 221 | apply (rename_tac nat, rule_tac x="[Number nat]" in exI, simp) | |
| 222 | apply (rename_tac nat, rule_tac x="[Nonce nat]" in exI, simp) | |
| 223 | apply (rename_tac nat, rule_tac x="[Key nat]" in exI, simp) | |
| 224 | apply (rename_tac X, rule_tac x="[Hash X]" in exI, simp) | |
| 13508 | 225 | apply (clarify, rule_tac x="l@la" in exI, simp) | 
| 58250 | 226 | by (clarify, rename_tac nat X y, rule_tac x="[Crypt nat X]" in exI, simp) | 
| 13508 | 227 | |
| 67613 | 228 | lemma kparts_set: "\<exists>l'. kparts (set l) = set l' \<and> cnb l' = cnb l" | 
| 13508 | 229 | apply (induct l) | 
| 230 | apply (rule_tac x="[]" in exI, simp, clarsimp) | |
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
45600diff
changeset | 231 | apply (rename_tac a b l') | 
| 67613 | 232 | apply (subgoal_tac "\<exists>l''.  kparts {a} = set l'' \<and> cnb l'' = crypt_nb a", clarify)
 | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
14307diff
changeset | 233 | apply (rule_tac x="l''@l'" in exI, simp) | 
| 13508 | 234 | apply (rule kparts_insert_substI, simp) | 
| 235 | by (rule kparts_msg_set) | |
| 236 | ||
| 61830 | 237 | subsection\<open>list corresponding to "decrypt"\<close> | 
| 13508 | 238 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
25134diff
changeset | 239 | definition decrypt' :: "msg list => key => msg => msg list" where | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17087diff
changeset | 240 | "decrypt' l K Y == Y # remove l (Crypt K Y)" | 
| 13508 | 241 | |
| 242 | declare decrypt'_def [simp] | |
| 243 | ||
| 69597 | 244 | subsection\<open>basic facts about \<^term>\<open>decrypt'\<close>\<close> | 
| 13508 | 245 | |
| 246 | lemma decrypt_minus: "decrypt (set l) K Y <= set (decrypt' l K Y)" | |
| 247 | by (induct l, auto) | |
| 248 | ||
| 61830 | 249 | subsection\<open>if the analyse of a finite guarded set gives n then it must also gives | 
| 250 | one of the keys of Ks\<close> | |
| 13508 | 251 | |
| 67613 | 252 | lemma Guard_invKey_by_list [rule_format]: "\<forall>l. cnb l = p | 
| 253 | \<longrightarrow> Guard n Ks (set l) \<longrightarrow> Nonce n \<in> analz (set l) | |
| 254 | \<longrightarrow> (\<exists>K. K \<in> Ks \<and> Key K \<in> analz (set l))" | |
| 13508 | 255 | apply (induct p) | 
| 256 | (* case p=0 *) | |
| 257 | apply (clarify, drule Guard_must_decrypt, simp, clarify) | |
| 258 | apply (drule kparts_parts, drule non_empty_crypt, simp) | |
| 259 | (* case p>0 *) | |
| 260 | apply (clarify, frule Guard_must_decrypt, simp, clarify) | |
| 67613 | 261 | apply (drule_tac P="\<lambda>G. Nonce n \<in> G" in analz_pparts_kparts_substD, simp) | 
| 13508 | 262 | apply (frule analz_decrypt, simp_all) | 
| 67613 | 263 | apply (subgoal_tac "\<exists>l'. kparts (set l) = set l' \<and> cnb l' = cnb l", clarsimp) | 
| 13508 | 264 | apply (drule_tac G="insert Y (set l' - {Crypt K Y})"
 | 
| 265 | and H="set (decrypt' l' K Y)" in analz_sub, rule decrypt_minus) | |
| 266 | apply (rule_tac analz_pparts_kparts_substI, simp) | |
| 67613 | 267 | apply (case_tac "K \<in> invKey`Ks") | 
| 13508 | 268 | (* K:invKey`Ks *) | 
| 269 | apply (clarsimp, blast) | |
| 270 | (* K ~:invKey`Ks *) | |
| 271 | apply (subgoal_tac "Guard n Ks (set (decrypt' l' K Y))") | |
| 37596 | 272 | apply (drule_tac x="decrypt' l' K Y" in spec, simp) | 
| 67613 | 273 | apply (subgoal_tac "Crypt K Y \<in> parts (set l)") | 
| 13508 | 274 | apply (drule parts_cnb, rotate_tac -1, simp) | 
| 275 | apply (clarify, drule_tac X="Key Ka" and H="insert Y (set l')" in analz_sub) | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17087diff
changeset | 276 | apply (rule insert_mono, rule set_remove) | 
| 13508 | 277 | apply (simp add: analz_insertD, blast) | 
| 278 | (* Crypt K Y:parts (set l) *) | |
| 279 | apply (blast dest: kparts_parts) | |
| 280 | (* Guard n Ks (set (decrypt' l' K Y)) *) | |
| 281 | apply (rule_tac H="insert Y (set l')" in Guard_mono) | |
| 282 | apply (subgoal_tac "Guard n Ks (set l')", simp) | |
| 283 | apply (rule_tac K=K in guard_Crypt, simp add: Guard_def, simp) | |
| 284 | apply (drule_tac t="set l'" in sym, simp) | |
| 285 | apply (rule Guard_kparts, simp, simp) | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17087diff
changeset | 286 | apply (rule_tac B="set l'" in subset_trans, rule set_remove, blast) | 
| 13508 | 287 | by (rule kparts_set) | 
| 288 | ||
| 67613 | 289 | lemma Guard_invKey_finite: "[| Nonce n \<in> analz G; Guard n Ks G; finite G |] | 
| 290 | ==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G" | |
| 13508 | 291 | apply (drule finite_list, clarify) | 
| 292 | by (rule Guard_invKey_by_list, auto) | |
| 293 | ||
| 67613 | 294 | lemma Guard_invKey: "[| Nonce n \<in> analz G; Guard n Ks G |] | 
| 295 | ==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G" | |
| 13508 | 296 | by (auto dest: analz_needs_only_finite Guard_invKey_finite) | 
| 297 | ||
| 61830 | 298 | subsection\<open>if the analyse of a finite guarded set and a (possibly infinite) set of keys | 
| 299 | gives n then it must also gives Ks\<close> | |
| 13508 | 300 | |
| 67613 | 301 | lemma Guard_invKey_keyset: "[| Nonce n \<in> analz (G \<union> H); Guard n Ks G; finite G; | 
| 302 | keyset H |] ==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz (G \<union> H)" | |
| 303 | apply (frule_tac P="\<lambda>G. Nonce n \<in> G" and G=G in analz_keyset_substD, simp_all) | |
| 13508 | 304 | apply (drule_tac G="G Un (H Int keysfor G)" in Guard_invKey_finite) | 
| 305 | by (auto simp: Guard_def intro: analz_sub) | |
| 306 | ||
| 62390 | 307 | end |