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