| author | paulson <lp15@cam.ac.uk> | 
| Mon, 28 Aug 2017 16:30:51 +0100 | |
| changeset 66535 | 64035d9161d3 | 
| parent 62390 | 842917225d56 | 
| child 67613 | ce654b0e6d69 | 
| permissions | -rw-r--r-- | 
| 41775 | 1 | (* Title: HOL/Auth/Guard/Guard_Shared.thy | 
| 2 | Author: Frederic Blanqui, University of Cambridge Computer Laboratory | |
| 3 | Copyright 2002 University of Cambridge | |
| 4 | *) | |
| 13508 | 5 | |
| 61830 | 6 | section\<open>lemmas on guarded messages for protocols with symmetric keys\<close> | 
| 13508 | 7 | |
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
35416diff
changeset | 8 | theory Guard_Shared imports Guard GuardK "../Shared" begin | 
| 13508 | 9 | |
| 61830 | 10 | subsection\<open>Extensions to Theory \<open>Shared\<close>\<close> | 
| 13508 | 11 | |
| 12 | declare initState.simps [simp del] | |
| 13 | ||
| 61830 | 14 | subsubsection\<open>a little abbreviation\<close> | 
| 13508 | 15 | |
| 19363 | 16 | abbreviation | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19363diff
changeset | 17 | Ciph :: "agent => msg => msg" where | 
| 19363 | 18 | "Ciph A X == Crypt (shrK A) X" | 
| 13508 | 19 | |
| 61830 | 20 | subsubsection\<open>agent associated to a key\<close> | 
| 13508 | 21 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
21404diff
changeset | 22 | definition agt :: "key => agent" where | 
| 13508 | 23 | "agt K == @A. K = shrK A" | 
| 24 | ||
| 25 | lemma agt_shrK [simp]: "agt (shrK A) = A" | |
| 26 | by (simp add: agt_def) | |
| 27 | ||
| 61830 | 28 | subsubsection\<open>basic facts about @{term initState}\<close>
 | 
| 13508 | 29 | |
| 30 | lemma no_Crypt_in_parts_init [simp]: "Crypt K X ~:parts (initState A)" | |
| 31 | by (cases A, auto simp: initState.simps) | |
| 32 | ||
| 33 | lemma no_Crypt_in_analz_init [simp]: "Crypt K X ~:analz (initState A)" | |
| 34 | by auto | |
| 35 | ||
| 36 | lemma no_shrK_in_analz_init [simp]: "A ~:bad | |
| 37 | ==> Key (shrK A) ~:analz (initState Spy)" | |
| 38 | by (auto simp: initState.simps) | |
| 39 | ||
| 40 | lemma shrK_notin_initState_Friend [simp]: "A ~= Friend C | |
| 41 | ==> Key (shrK A) ~: parts (initState (Friend C))" | |
| 42 | by (auto simp: initState.simps) | |
| 43 | ||
| 44 | lemma keyset_init [iff]: "keyset (initState A)" | |
| 45 | by (cases A, auto simp: keyset_def initState.simps) | |
| 46 | ||
| 61830 | 47 | subsubsection\<open>sets of symmetric keys\<close> | 
| 13508 | 48 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
21404diff
changeset | 49 | definition shrK_set :: "key set => bool" where | 
| 13508 | 50 | "shrK_set Ks == ALL K. K:Ks --> (EX A. K = shrK A)" | 
| 51 | ||
| 52 | lemma in_shrK_set: "[| shrK_set Ks; K:Ks |] ==> EX A. K = shrK A" | |
| 53 | by (simp add: shrK_set_def) | |
| 54 | ||
| 55 | lemma shrK_set1 [iff]: "shrK_set {shrK A}"
 | |
| 56 | by (simp add: shrK_set_def) | |
| 57 | ||
| 58 | lemma shrK_set2 [iff]: "shrK_set {shrK A, shrK B}"
 | |
| 59 | by (simp add: shrK_set_def) | |
| 60 | ||
| 61830 | 61 | subsubsection\<open>sets of good keys\<close> | 
| 13508 | 62 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
21404diff
changeset | 63 | definition good :: "key set => bool" where | 
| 13508 | 64 | "good Ks == ALL K. K:Ks --> agt K ~:bad" | 
| 65 | ||
| 66 | lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad" | |
| 67 | by (simp add: good_def) | |
| 68 | ||
| 69 | lemma good1 [simp]: "A ~:bad ==> good {shrK A}"
 | |
| 70 | by (simp add: good_def) | |
| 71 | ||
| 72 | lemma good2 [simp]: "[| A ~:bad; B ~:bad |] ==> good {shrK A, shrK B}"
 | |
| 73 | by (simp add: good_def) | |
| 74 | ||
| 75 | ||
| 61830 | 76 | subsection\<open>Proofs About Guarded Messages\<close> | 
| 13508 | 77 | |
| 61830 | 78 | subsubsection\<open>small hack\<close> | 
| 13508 | 79 | |
| 80 | lemma shrK_is_invKey_shrK: "shrK A = invKey (shrK A)" | |
| 81 | by simp | |
| 82 | ||
| 83 | lemmas shrK_is_invKey_shrK_substI = shrK_is_invKey_shrK [THEN ssubst] | |
| 84 | ||
| 85 | lemmas invKey_invKey_substI = invKey [THEN ssubst] | |
| 86 | ||
| 87 | lemma "Nonce n:parts {X} ==> Crypt (shrK A) X:guard n {shrK A}"
 | |
| 88 | apply (rule shrK_is_invKey_shrK_substI, rule invKey_invKey_substI) | |
| 89 | by (rule Guard_Nonce, simp+) | |
| 90 | ||
| 61830 | 91 | subsubsection\<open>guardedness results on nonces\<close> | 
| 13508 | 92 | |
| 93 | lemma guard_ciph [simp]: "shrK A:Ks ==> Ciph A X:guard n Ks" | |
| 94 | by (rule Guard_Nonce, simp) | |
| 95 | ||
| 13523 | 96 | lemma guardK_ciph [simp]: "shrK A:Ks ==> Ciph A X:guardK n Ks" | 
| 13508 | 97 | by (rule Guard_Key, simp) | 
| 98 | ||
| 99 | lemma Guard_init [iff]: "Guard n Ks (initState B)" | |
| 100 | by (induct B, auto simp: Guard_def initState.simps) | |
| 101 | ||
| 102 | lemma Guard_knows_max': "Guard n Ks (knows_max' C evs) | |
| 103 | ==> Guard n Ks (knows_max C evs)" | |
| 104 | by (simp add: knows_max_def) | |
| 105 | ||
| 106 | lemma Nonce_not_used_Guard_spies [dest]: "Nonce n ~:used evs | |
| 107 | ==> Guard n Ks (spies evs)" | |
| 108 | by (auto simp: Guard_def dest: not_used_not_known parts_sub) | |
| 109 | ||
| 110 | lemma Nonce_not_used_Guard [dest]: "[| evs:p; Nonce n ~:used evs; | |
| 111 | Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" | |
| 112 | by (auto simp: Guard_def dest: known_used parts_trans) | |
| 113 | ||
| 114 | lemma Nonce_not_used_Guard_max [dest]: "[| evs:p; Nonce n ~:used evs; | |
| 115 | Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" | |
| 116 | by (auto simp: Guard_def dest: known_max_used parts_trans) | |
| 117 | ||
| 118 | lemma Nonce_not_used_Guard_max' [dest]: "[| evs:p; Nonce n ~:used evs; | |
| 119 | Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" | |
| 120 | apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) | |
| 121 | by (auto simp: knows_max_def) | |
| 122 | ||
| 61830 | 123 | subsubsection\<open>guardedness results on keys\<close> | 
| 13508 | 124 | |
| 125 | lemma GuardK_init [simp]: "n ~:range shrK ==> GuardK n Ks (initState B)" | |
| 126 | by (induct B, auto simp: GuardK_def initState.simps) | |
| 127 | ||
| 128 | lemma GuardK_knows_max': "[| GuardK n A (knows_max' C evs); n ~:range shrK |] | |
| 129 | ==> GuardK n A (knows_max C evs)" | |
| 130 | by (simp add: knows_max_def) | |
| 131 | ||
| 132 | lemma Key_not_used_GuardK_spies [dest]: "Key n ~:used evs | |
| 133 | ==> GuardK n A (spies evs)" | |
| 134 | by (auto simp: GuardK_def dest: not_used_not_known parts_sub) | |
| 135 | ||
| 136 | lemma Key_not_used_GuardK [dest]: "[| evs:p; Key n ~:used evs; | |
| 137 | Gets_correct p; one_step p |] ==> GuardK n A (knows (Friend C) evs)" | |
| 138 | by (auto simp: GuardK_def dest: known_used parts_trans) | |
| 139 | ||
| 140 | lemma Key_not_used_GuardK_max [dest]: "[| evs:p; Key n ~:used evs; | |
| 141 | Gets_correct p; one_step p |] ==> GuardK n A (knows_max (Friend C) evs)" | |
| 142 | by (auto simp: GuardK_def dest: known_max_used parts_trans) | |
| 143 | ||
| 144 | lemma Key_not_used_GuardK_max' [dest]: "[| evs:p; Key n ~:used evs; | |
| 145 | Gets_correct p; one_step p |] ==> GuardK n A (knows_max' (Friend C) evs)" | |
| 146 | apply (rule_tac H="knows_max (Friend C) evs" in GuardK_mono) | |
| 147 | by (auto simp: knows_max_def) | |
| 148 | ||
| 61830 | 149 | subsubsection\<open>regular protocols\<close> | 
| 13508 | 150 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
21404diff
changeset | 151 | definition regular :: "event list set => bool" where | 
| 13508 | 152 | "regular p == ALL evs A. evs:p --> (Key (shrK A):parts (spies evs)) = (A:bad)" | 
| 153 | ||
| 154 | lemma shrK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==> | |
| 155 | (Key (shrK A):parts (spies evs)) = (A:bad)" | |
| 156 | by (auto simp: regular_def) | |
| 157 | ||
| 158 | lemma shrK_analz_iff_bad [simp]: "[| evs:p; regular p |] ==> | |
| 159 | (Key (shrK A):analz (spies evs)) = (A:bad)" | |
| 160 | by auto | |
| 161 | ||
| 162 | lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs:p; | |
| 163 | shrK_set Ks; good Ks; regular p |] ==> Nonce n ~:analz (spies evs)" | |
| 164 | apply (clarify, simp only: knows_decomp) | |
| 165 | apply (drule Guard_invKey_keyset, simp+, safe) | |
| 166 | apply (drule in_good, simp) | |
| 167 | apply (drule in_shrK_set, simp+, clarify) | |
| 168 | apply (frule_tac A=A in shrK_analz_iff_bad) | |
| 169 | by (simp add: knows_decomp)+ | |
| 170 | ||
| 61928 | 171 | lemma GuardK_Key_analz: | 
| 172 | assumes "GuardK n Ks (spies evs)" "evs \<in> p" "shrK_set Ks" | |
| 173 | "good Ks" "regular p" "n \<notin> range shrK" | |
| 174 | shows "Key n \<notin> analz (spies evs)" | |
| 175 | proof (rule ccontr) | |
| 176 | assume "\<not> Key n \<notin> analz (knows Spy evs)" | |
| 177 | then have *: "Key n \<in> analz (spies' evs \<union> initState Spy)" | |
| 178 | by (simp add: knows_decomp) | |
| 179 | from \<open>GuardK n Ks (spies evs)\<close> | |
| 180 | have "GuardK n Ks (spies' evs \<union> initState Spy)" | |
| 181 | by (simp add: knows_decomp) | |
| 182 | then have "GuardK n Ks (spies' evs)" | |
| 183 | and "finite (spies' evs)" "keyset (initState Spy)" | |
| 184 | by simp_all | |
| 185 | moreover have "Key n \<notin> initState Spy" | |
| 186 | using \<open>n \<notin> range shrK\<close> by (simp add: image_iff initState_Spy) | |
| 187 | ultimately obtain K | |
| 188 | where "K \<in> Ks" and **: "Key K \<in> analz (spies' evs \<union> initState Spy)" | |
| 189 | using * by (auto dest: GuardK_invKey_keyset) | |
| 190 | from \<open>K \<in> Ks\<close> and \<open>good Ks\<close> have "agt K \<notin> bad" | |
| 191 | by (auto dest: in_good) | |
| 192 | from \<open>K \<in> Ks\<close> \<open>shrK_set Ks\<close> obtain A | |
| 193 | where "K = shrK A" | |
| 194 | by (auto dest: in_shrK_set) | |
| 195 | then have "agt K \<in> bad" | |
| 196 | using ** \<open>evs \<in> p\<close> \<open>regular p\<close> shrK_analz_iff_bad [of evs p "agt K"] | |
| 197 | by (simp add: knows_decomp) | |
| 198 | with \<open>agt K \<notin> bad\<close> show False by simp | |
| 199 | qed | |
| 13508 | 200 | |
| 62390 | 201 | end |