| author | wenzelm | 
| Sun, 18 Jan 2015 21:35:54 +0100 | |
| changeset 59397 | fc909f7e7ce5 | 
| parent 58889 | 5b7a9633cfa8 | 
| child 61830 | 4f5ab843cf5b | 
| permissions | -rw-r--r-- | 
| 37936 | 1 | (* Title: HOL/Auth/CertifiedEmail.thy | 
| 13922 | 2 | Author: Giampaolo Bella, Christiano Longo and Lawrence C Paulson | 
| 13956 | 3 | *) | 
| 13922 | 4 | |
| 58889 | 5 | section{*The Certified Electronic Mail Protocol by Abadi et al.*}
 | 
| 13922 | 6 | |
| 16417 | 7 | theory CertifiedEmail imports Public begin | 
| 13922 | 8 | |
| 20768 | 9 | abbreviation | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 10 | TTP :: agent where | 
| 20768 | 11 | "TTP == Server" | 
| 13922 | 12 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 13 | abbreviation | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20768diff
changeset | 14 | RPwd :: "agent => key" where | 
| 20768 | 15 | "RPwd == shrK" | 
| 13922 | 16 | |
| 17 | ||
| 18 | (*FIXME: the four options should be represented by pairs of 0 or 1. | |
| 19 | Right now only BothAuth is modelled.*) | |
| 20 | consts | |
| 21 | NoAuth :: nat | |
| 22 | TTPAuth :: nat | |
| 23 | SAuth :: nat | |
| 24 | BothAuth :: nat | |
| 25 | ||
| 26 | text{*We formalize a fixed way of computing responses.  Could be better.*}
 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 27 | definition "response" :: "agent => agent => nat => msg" where | 
| 13922 | 28 |    "response S R q == Hash {|Agent S, Key (shrK R), Nonce q|}"
 | 
| 29 | ||
| 30 | ||
| 23746 | 31 | inductive_set certified_mail :: "event list set" | 
| 32 | where | |
| 13922 | 33 | |
| 23746 | 34 |   Nil: --{*The empty trace*}
 | 
| 13922 | 35 | "[] \<in> certified_mail" | 
| 36 | ||
| 23746 | 37 | | Fake: --{*The Spy may say anything he can say.  The sender field is correct,
 | 
| 13922 | 38 | but agents don't use that information.*} | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 39 | "[| evsf \<in> certified_mail; X \<in> synth(analz(spies evsf))|] | 
| 13956 | 40 | ==> Says Spy B X # evsf \<in> certified_mail" | 
| 13922 | 41 | |
| 23746 | 42 | | FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
 | 
| 13922 | 43 | equipped with the necessary credentials to serve as an SSL server.*} | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 44 | "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|] | 
| 13922 | 45 |           ==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
 | 
| 46 | ||
| 23746 | 47 | | CM1: --{*The sender approaches the recipient.  The message is a number.*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 48 | "[|evs1 \<in> certified_mail; | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 49 | Key K \<notin> used evs1; | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 50 | K \<in> symKeys; | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 51 | Nonce q \<notin> used evs1; | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 52 |     hs = Hash{|Number cleartext, Nonce q, response S R q, Crypt K (Number m)|};
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 53 |     S2TTP = Crypt(pubEK TTP) {|Agent S, Number BothAuth, Key K, Agent R, hs|}|]
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 54 |   ==> Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number BothAuth, 
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 55 | Number cleartext, Nonce q, S2TTP|} # evs1 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 56 | \<in> certified_mail" | 
| 13922 | 57 | |
| 23746 | 58 | | CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
 | 
| 13922 | 59 |      password to @{term TTP} over an SSL channel.*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 60 | "[|evs2 \<in> certified_mail; | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 61 |     Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext, 
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 62 | Nonce q, S2TTP|} \<in> set evs2; | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 63 | TTP \<noteq> R; | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 64 |     hr = Hash {|Number cleartext, Nonce q, response S R q, em|} |]
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 65 | ==> | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 66 |    Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} # evs2
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 67 | \<in> certified_mail" | 
| 13922 | 68 | |
| 23746 | 69 | | CM3: --{*@{term TTP} simultaneously reveals the key to the recipient and gives
 | 
| 13922 | 70 | a receipt to the sender. The SSL channel does not authenticate | 
| 14735 | 71 |          the client (@{term R}), but @{term TTP} accepts the message only 
 | 
| 72 |          if the given password is that of the claimed sender, @{term R}.
 | |
| 13922 | 73 | He replies over the established SSL channel.*} | 
| 74 | "[|evs3 \<in> certified_mail; | |
| 14735 | 75 |     Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} \<in> set evs3;
 | 
| 76 | S2TTP = Crypt (pubEK TTP) | |
| 77 |                      {|Agent S, Number BothAuth, Key k, Agent R, hs|};
 | |
| 78 | TTP \<noteq> R; hs = hr; k \<in> symKeys|] | |
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 79 | ==> | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 80 |    Notes R {|Agent TTP, Agent R, Key k, hr|} # 
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 81 | Gets S (Crypt (priSK TTP) S2TTP) # | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
14735diff
changeset | 82 | Says TTP S (Crypt (priSK TTP) S2TTP) # evs3 \<in> certified_mail" | 
| 13922 | 83 | |
| 23746 | 84 | | Reception: | 
| 13922 | 85 | "[|evsr \<in> certified_mail; Says A B X \<in> set evsr|] | 
| 86 | ==> Gets B X#evsr \<in> certified_mail" | |
| 87 | ||
| 88 | ||
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 89 | declare Says_imp_knows_Spy [THEN analz.Inj, dest] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 90 | declare analz_into_parts [dest] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 91 | |
| 13922 | 92 | (*A "possibility property": there are traces that reach the end*) | 
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14145diff
changeset | 93 | lemma "[| Key K \<notin> used []; K \<in> symKeys |] ==> | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14145diff
changeset | 94 | \<exists>S2TTP. \<exists>evs \<in> certified_mail. | 
| 13956 | 95 | Says TTP S (Crypt (priSK TTP) S2TTP) \<in> set evs" | 
| 13922 | 96 | apply (intro exI bexI) | 
| 97 | apply (rule_tac [2] certified_mail.Nil | |
| 98 | [THEN certified_mail.CM1, THEN certified_mail.Reception, | |
| 99 | THEN certified_mail.CM2, | |
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14145diff
changeset | 100 | THEN certified_mail.CM3]) | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
14145diff
changeset | 101 | apply (possibility, auto) | 
| 13922 | 102 | done | 
| 103 | ||
| 104 | ||
| 105 | lemma Gets_imp_Says: | |
| 106 | "[| Gets B X \<in> set evs; evs \<in> certified_mail |] ==> \<exists>A. Says A B X \<in> set evs" | |
| 107 | apply (erule rev_mp) | |
| 108 | apply (erule certified_mail.induct, auto) | |
| 109 | done | |
| 110 | ||
| 111 | ||
| 112 | lemma Gets_imp_parts_knows_Spy: | |
| 113 | "[|Gets A X \<in> set evs; evs \<in> certified_mail|] ==> X \<in> parts(spies evs)" | |
| 114 | apply (drule Gets_imp_Says, simp) | |
| 115 | apply (blast dest: Says_imp_knows_Spy parts.Inj) | |
| 116 | done | |
| 117 | ||
| 118 | lemma CM2_S2TTP_analz_knows_Spy: | |
| 119 |  "[|Gets R {|Agent A, Agent B, em, Number AO, Number cleartext, 
 | |
| 120 | Nonce q, S2TTP|} \<in> set evs; | |
| 121 | evs \<in> certified_mail|] | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 122 | ==> S2TTP \<in> analz(spies evs)" | 
| 13922 | 123 | apply (drule Gets_imp_Says, simp) | 
| 124 | apply (blast dest: Says_imp_knows_Spy analz.Inj) | |
| 125 | done | |
| 126 | ||
| 127 | lemmas CM2_S2TTP_parts_knows_Spy = | |
| 128 | CM2_S2TTP_analz_knows_Spy [THEN analz_subset_parts [THEN subsetD]] | |
| 129 | ||
| 130 | lemma hr_form_lemma [rule_format]: | |
| 131 | "evs \<in> certified_mail | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 132 | ==> hr \<notin> synth (analz (spies evs)) --> | 
| 13922 | 133 |       (\<forall>S2TTP. Notes TTP {|Agent R, Agent TTP, S2TTP, pwd, hr|}
 | 
| 134 | \<in> set evs --> | |
| 135 |       (\<exists>clt q S em. hr = Hash {|Number clt, Nonce q, response S R q, em|}))"
 | |
| 136 | apply (erule certified_mail.induct) | |
| 137 | apply (synth_analz_mono_contra, simp_all, blast+) | |
| 138 | done | |
| 139 | ||
| 140 | text{*Cannot strengthen the first disjunct to @{term "R\<noteq>Spy"} because
 | |
| 141 | the fakessl rule allows Spy to spoof the sender's name. Maybe can | |
| 142 | strengthen the second disjunct with @{term "R\<noteq>Spy"}.*}
 | |
| 143 | lemma hr_form: | |
| 144 |  "[|Notes TTP {|Agent R, Agent TTP, S2TTP, pwd, hr|} \<in> set evs;
 | |
| 145 | evs \<in> certified_mail|] | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 146 | ==> hr \<in> synth (analz (spies evs)) | | 
| 13922 | 147 |       (\<exists>clt q S em. hr = Hash {|Number clt, Nonce q, response S R q, em|})"
 | 
| 148 | by (blast intro: hr_form_lemma) | |
| 149 | ||
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 150 | lemma Spy_dont_know_private_keys [dest!]: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 151 | "[|Key (privateKey b A) \<in> parts (spies evs); evs \<in> certified_mail|] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 152 | ==> A \<in> bad" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 153 | apply (erule rev_mp) | 
| 13922 | 154 | apply (erule certified_mail.induct, simp_all) | 
| 155 | txt{*Fake*}
 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 156 | apply (blast dest: Fake_parts_insert_in_Un) | 
| 13922 | 157 | txt{*Message 1*}
 | 
| 158 | apply blast | |
| 159 | txt{*Message 3*}
 | |
| 160 | apply (frule_tac hr_form, assumption) | |
| 161 | apply (elim disjE exE) | |
| 162 | apply (simp_all add: parts_insert2) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 163 | apply (force dest!: parts_insert_subset_Un [THEN [2] rev_subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 164 | analz_subset_parts [THEN subsetD], blast) | 
| 13922 | 165 | done | 
| 166 | ||
| 167 | lemma Spy_know_private_keys_iff [simp]: | |
| 168 | "evs \<in> certified_mail | |
| 169 | ==> (Key (privateKey b A) \<in> parts (spies evs)) = (A \<in> bad)" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 170 | by blast | 
| 13922 | 171 | |
| 13956 | 172 | lemma Spy_dont_know_TTPKey_parts [simp]: | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 173 | "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> parts(spies evs)" | 
| 13956 | 174 | by simp | 
| 13934 | 175 | |
| 13956 | 176 | lemma Spy_dont_know_TTPKey_analz [simp]: | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 177 | "evs \<in> certified_mail ==> Key (privateKey b TTP) \<notin> analz(spies evs)" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 178 | by auto | 
| 13934 | 179 | |
| 180 | text{*Thus, prove any goal that assumes that @{term Spy} knows a private key
 | |
| 181 | belonging to @{term TTP}*}
 | |
| 13956 | 182 | declare Spy_dont_know_TTPKey_parts [THEN [2] rev_notE, elim!] | 
| 13934 | 183 | |
| 13922 | 184 | |
| 185 | lemma CM3_k_parts_knows_Spy: | |
| 186 | "[| evs \<in> certified_mail; | |
| 187 |      Notes TTP {|Agent A, Agent TTP,
 | |
| 13956 | 188 |                  Crypt (pubEK TTP) {|Agent S, Number AO, Key K, 
 | 
| 13922 | 189 | Agent R, hs|}, Key (RPwd R), hs|} \<in> set evs|] | 
| 190 | ==> Key K \<in> parts(spies evs)" | |
| 191 | apply (rotate_tac 1) | |
| 192 | apply (erule rev_mp) | |
| 193 | apply (erule certified_mail.induct, simp_all) | |
| 194 | apply (blast intro:parts_insertI) | |
| 195 | txt{*Fake SSL*}
 | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 196 | apply (blast dest: parts.Body) | 
| 13922 | 197 | txt{*Message 2*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 198 | apply (blast dest!: Gets_imp_Says elim!: knows_Spy_partsEs) | 
| 13922 | 199 | txt{*Message 3*}
 | 
| 43584 | 200 | apply (metis parts_insertI) | 
| 13922 | 201 | done | 
| 202 | ||
| 203 | lemma Spy_dont_know_RPwd [rule_format]: | |
| 204 | "evs \<in> certified_mail ==> Key (RPwd A) \<in> parts(spies evs) --> A \<in> bad" | |
| 205 | apply (erule certified_mail.induct, simp_all) | |
| 206 | txt{*Fake*}
 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 207 | apply (blast dest: Fake_parts_insert_in_Un) | 
| 13922 | 208 | txt{*Message 1*}
 | 
| 209 | apply blast | |
| 210 | txt{*Message 3*}
 | |
| 211 | apply (frule CM3_k_parts_knows_Spy, assumption) | |
| 212 | apply (frule_tac hr_form, assumption) | |
| 213 | apply (elim disjE exE) | |
| 214 | apply (simp_all add: parts_insert2) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 215 | apply (force dest!: parts_insert_subset_Un [THEN [2] rev_subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 216 | analz_subset_parts [THEN subsetD]) | 
| 13922 | 217 | done | 
| 218 | ||
| 219 | ||
| 220 | lemma Spy_know_RPwd_iff [simp]: | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 221 | "evs \<in> certified_mail ==> (Key (RPwd A) \<in> parts(spies evs)) = (A\<in>bad)" | 
| 13922 | 222 | by (auto simp add: Spy_dont_know_RPwd) | 
| 223 | ||
| 224 | lemma Spy_analz_RPwd_iff [simp]: | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 225 | "evs \<in> certified_mail ==> (Key (RPwd A) \<in> analz(spies evs)) = (A\<in>bad)" | 
| 43584 | 226 | by (metis Spy_know_RPwd_iff Spy_spies_bad_shrK analz.Inj analz_into_parts) | 
| 13922 | 227 | |
| 228 | text{*Unused, but a guarantee of sorts*}
 | |
| 229 | theorem CertAutenticity: | |
| 13956 | 230 | "[|Crypt (priSK TTP) X \<in> parts (spies evs); evs \<in> certified_mail|] | 
| 231 | ==> \<exists>A. Says TTP A (Crypt (priSK TTP) X) \<in> set evs" | |
| 13922 | 232 | apply (erule rev_mp) | 
| 233 | apply (erule certified_mail.induct, simp_all) | |
| 234 | txt{*Fake*}
 | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 235 | apply (blast dest: Spy_dont_know_private_keys Fake_parts_insert_in_Un) | 
| 13922 | 236 | txt{*Message 1*}
 | 
| 237 | apply blast | |
| 238 | txt{*Message 3*}
 | |
| 239 | apply (frule_tac hr_form, assumption) | |
| 240 | apply (elim disjE exE) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 241 | apply (simp_all add: parts_insert2 parts_insert_knows_A) | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 242 | apply (blast dest!: Fake_parts_sing_imp_Un, blast) | 
| 13922 | 243 | done | 
| 244 | ||
| 245 | ||
| 246 | subsection{*Proving Confidentiality Results*}
 | |
| 247 | ||
| 248 | lemma analz_image_freshK [rule_format]: | |
| 249 | "evs \<in> certified_mail ==> | |
| 13956 | 250 | \<forall>K KK. invKey (pubEK TTP) \<notin> KK --> | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 251 | (Key K \<in> analz (Key`KK Un (spies evs))) = | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 252 | (K \<in> KK | Key K \<in> analz (spies evs))" | 
| 13922 | 253 | apply (erule certified_mail.induct) | 
| 254 | apply (drule_tac [6] A=TTP in symKey_neq_priEK) | |
| 255 | apply (erule_tac [6] disjE [OF hr_form]) | |
| 256 | apply (drule_tac [5] CM2_S2TTP_analz_knows_Spy) | |
| 257 | prefer 9 | |
| 258 | apply (elim exE) | |
| 259 | apply (simp_all add: synth_analz_insert_eq | |
| 260 | subset_trans [OF _ subset_insertI] | |
| 261 | subset_trans [OF _ Un_upper2] | |
| 262 | del: image_insert image_Un add: analz_image_freshK_simps) | |
| 263 | done | |
| 264 | ||
| 265 | ||
| 266 | lemma analz_insert_freshK: | |
| 13956 | 267 | "[| evs \<in> certified_mail; KAB \<noteq> invKey (pubEK TTP) |] ==> | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 268 | (Key K \<in> analz (insert (Key KAB) (spies evs))) = | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 269 | (K = KAB | Key K \<in> analz (spies evs))" | 
| 13922 | 270 | by (simp only: analz_image_freshK analz_image_freshK_simps) | 
| 271 | ||
| 272 | text{*@{term S2TTP} must have originated from a valid sender
 | |
| 273 |     provided @{term K} is secure.  Proof is surprisingly hard.*}
 | |
| 274 | ||
| 275 | lemma Notes_SSL_imp_used: | |
| 276 |      "[|Notes B {|Agent A, Agent B, X|} \<in> set evs|] ==> X \<in> used evs"
 | |
| 277 | by (blast dest!: Notes_imp_used) | |
| 278 | ||
| 279 | ||
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 280 | (*The weaker version, replacing "used evs" by "parts (spies evs)", | 
| 13922 | 281 | isn't inductive: message 3 case can't be proved *) | 
| 282 | lemma S2TTP_sender_lemma [rule_format]: | |
| 283 | "evs \<in> certified_mail ==> | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 284 | Key K \<notin> analz (spies evs) --> | 
| 13956 | 285 | (\<forall>AO. Crypt (pubEK TTP) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 286 |            {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
 | 
| 13922 | 287 | (\<exists>m ctxt q. | 
| 288 |         hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 289 | Says S R | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 290 |            {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 291 | Number ctxt, Nonce q, | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 292 | Crypt (pubEK TTP) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 293 |               {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))" 
 | 
| 13922 | 294 | apply (erule certified_mail.induct, analz_mono_contra) | 
| 295 | apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp) | |
| 296 | apply (simp add: used_Nil Crypt_notin_initState, simp_all) | |
| 297 | txt{*Fake*}
 | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 298 | apply (blast dest: Fake_parts_sing [THEN subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 299 | dest!: analz_subset_parts [THEN subsetD]) | 
| 13922 | 300 | txt{*Fake SSL*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 301 | apply (blast dest: Fake_parts_sing [THEN subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 302 | dest: analz_subset_parts [THEN subsetD]) | 
| 13922 | 303 | txt{*Message 1*}
 | 
| 13956 | 304 | apply (clarsimp, blast) | 
| 13922 | 305 | txt{*Message 2*}
 | 
| 306 | apply (simp add: parts_insert2, clarify) | |
| 43584 | 307 | apply (metis parts_cut Un_empty_left usedI) | 
| 13922 | 308 | txt{*Message 3*} 
 | 
| 309 | apply (blast dest: Notes_SSL_imp_used used_parts_subset_parts) | |
| 310 | done | |
| 311 | ||
| 312 | lemma S2TTP_sender: | |
| 13956 | 313 |  "[|Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs;
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 314 | Key K \<notin> analz (spies evs); | 
| 13922 | 315 | evs \<in> certified_mail|] | 
| 316 | ==> \<exists>m ctxt q. | |
| 317 |         hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 318 | Says S R | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 319 |            {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 320 | Number ctxt, Nonce q, | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 321 | Crypt (pubEK TTP) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 322 |               {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs" 
 | 
| 13922 | 323 | by (blast intro: S2TTP_sender_lemma) | 
| 324 | ||
| 325 | ||
| 326 | text{*Nobody can have used non-existent keys!*}
 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 327 | lemma new_keys_not_used [simp]: | 
| 13922 | 328 | "[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> certified_mail|] | 
| 329 | ==> K \<notin> keysFor (parts (spies evs))" | |
| 330 | apply (erule rev_mp) | |
| 331 | apply (erule certified_mail.induct, simp_all) | |
| 332 | txt{*Fake*}
 | |
| 333 | apply (force dest!: keysFor_parts_insert) | |
| 334 | txt{*Message 1*}
 | |
| 335 | apply blast | |
| 336 | txt{*Message 3*}
 | |
| 337 | apply (frule CM3_k_parts_knows_Spy, assumption) | |
| 338 | apply (frule_tac hr_form, assumption) | |
| 339 | apply (force dest!: keysFor_parts_insert) | |
| 340 | done | |
| 341 | ||
| 342 | ||
| 343 | text{*Less easy to prove @{term "m'=m"}.  Maybe needs a separate unicity
 | |
| 13926 | 344 | theorem for ciphertexts of the form @{term "Crypt K (Number m)"}, 
 | 
| 13922 | 345 | where @{term K} is secure.*}
 | 
| 346 | lemma Key_unique_lemma [rule_format]: | |
| 347 | "evs \<in> certified_mail ==> | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 348 | Key K \<notin> analz (spies evs) --> | 
| 13922 | 349 | (\<forall>m cleartext q hs. | 
| 350 | Says S R | |
| 351 |            {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
 | |
| 352 | Number cleartext, Nonce q, | |
| 13956 | 353 |              Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|}|}
 | 
| 13922 | 354 | \<in> set evs --> | 
| 355 | (\<forall>m' cleartext' q' hs'. | |
| 356 | Says S' R' | |
| 357 |            {|Agent S', Agent TTP, Crypt K (Number m'), Number AO',
 | |
| 358 | Number cleartext', Nonce q', | |
| 13956 | 359 |              Crypt (pubEK TTP) {|Agent S', Number AO', Key K, Agent R', hs'|}|}
 | 
| 13922 | 360 | \<in> set evs --> R' = R & S' = S & AO' = AO & hs' = hs))" | 
| 361 | apply (erule certified_mail.induct, analz_mono_contra, simp_all) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 362 | prefer 2 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 363 |  txt{*Message 1*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 364 | apply (blast dest!: Says_imp_knows_Spy [THEN parts.Inj] new_keys_not_used Crypt_imp_keysFor) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 365 | txt{*Fake*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 366 | apply (auto dest!: usedI S2TTP_sender analz_subset_parts [THEN subsetD]) | 
| 13922 | 367 | done | 
| 368 | ||
| 369 | text{*The key determines the sender, recipient and protocol options.*}
 | |
| 13926 | 370 | lemma Key_unique: | 
| 13922 | 371 | "[|Says S R | 
| 372 |            {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
 | |
| 373 | Number cleartext, Nonce q, | |
| 13956 | 374 |              Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|}|}
 | 
| 13922 | 375 | \<in> set evs; | 
| 376 | Says S' R' | |
| 377 |            {|Agent S', Agent TTP, Crypt K (Number m'), Number AO',
 | |
| 378 | Number cleartext', Nonce q', | |
| 13956 | 379 |              Crypt (pubEK TTP) {|Agent S', Number AO', Key K, Agent R', hs'|}|}
 | 
| 13922 | 380 | \<in> set evs; | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 381 | Key K \<notin> analz (spies evs); | 
| 13922 | 382 | evs \<in> certified_mail|] | 
| 383 | ==> R' = R & S' = S & AO' = AO & hs' = hs" | |
| 384 | by (rule Key_unique_lemma, assumption+) | |
| 385 | ||
| 13934 | 386 | |
| 13926 | 387 | subsection{*The Guarantees for Sender and Recipient*}
 | 
| 388 | ||
| 13934 | 389 | text{*A Sender's guarantee:
 | 
| 390 |       If Spy gets the key then @{term R} is bad and @{term S} moreover
 | |
| 13922 | 391 | gets his return receipt (and therefore has no grounds for complaint).*} | 
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 392 | theorem S_fairness_bad_R: | 
| 13922 | 393 |       "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
 | 
| 394 | Number cleartext, Nonce q, S2TTP|} \<in> set evs; | |
| 13956 | 395 |          S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
 | 
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 396 | Key K \<in> analz (spies evs); | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 397 | evs \<in> certified_mail; | 
| 13922 | 398 | S\<noteq>Spy|] | 
| 13956 | 399 | ==> R \<in> bad & Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs" | 
| 13922 | 400 | apply (erule rev_mp) | 
| 401 | apply (erule ssubst) | |
| 13956 | 402 | apply (erule rev_mp) | 
| 13922 | 403 | apply (erule certified_mail.induct, simp_all) | 
| 404 | txt{*Fake*}
 | |
| 405 | apply spy_analz | |
| 406 | txt{*Fake SSL*}
 | |
| 407 | apply spy_analz | |
| 408 | txt{*Message 3*}
 | |
| 409 | apply (frule_tac hr_form, assumption) | |
| 410 | apply (elim disjE exE) | |
| 411 | apply (simp_all add: synth_analz_insert_eq | |
| 412 | subset_trans [OF _ subset_insertI] | |
| 413 | subset_trans [OF _ Un_upper2] | |
| 414 | del: image_insert image_Un add: analz_image_freshK_simps) | |
| 415 | apply (simp_all add: symKey_neq_priEK analz_insert_freshK) | |
| 416 | apply (blast dest: Notes_SSL_imp_used S2TTP_sender Key_unique)+ | |
| 417 | done | |
| 418 | ||
| 419 | text{*Confidentially for the symmetric key*}
 | |
| 420 | theorem Spy_not_see_encrypted_key: | |
| 421 |       "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
 | |
| 422 | Number cleartext, Nonce q, S2TTP|} \<in> set evs; | |
| 13956 | 423 |          S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 424 | evs \<in> certified_mail; | 
| 13922 | 425 | S\<noteq>Spy; R \<notin> bad|] | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 426 | ==> Key K \<notin> analz(spies evs)" | 
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 427 | by (blast dest: S_fairness_bad_R) | 
| 13922 | 428 | |
| 429 | ||
| 430 | text{*Agent @{term R}, who may be the Spy, doesn't receive the key
 | |
| 431 |  until @{term S} has access to the return receipt.*} 
 | |
| 432 | theorem S_guarantee: | |
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 433 |      "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO, 
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 434 | Number cleartext, Nonce q, S2TTP|} \<in> set evs; | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 435 |         S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 436 |         Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23746diff
changeset | 437 | S\<noteq>Spy; evs \<in> certified_mail|] | 
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 438 | ==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs" | 
| 13922 | 439 | apply (erule rev_mp) | 
| 440 | apply (erule ssubst) | |
| 13956 | 441 | apply (erule rev_mp) | 
| 13922 | 442 | apply (erule certified_mail.induct, simp_all) | 
| 443 | txt{*Message 1*}
 | |
| 444 | apply (blast dest: Notes_imp_used) | |
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 445 | txt{*Message 3*}
 | 
| 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 446 | apply (blast dest: Notes_SSL_imp_used S2TTP_sender Key_unique S_fairness_bad_R) | 
| 13922 | 447 | done | 
| 448 | ||
| 449 | ||
| 17689 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 450 | text{*If @{term R} sends message 2, and a delivery certificate exists, 
 | 
| 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 451 |  then @{term R} receives the necessary key.  This result is also important
 | 
| 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 452 |  to @{term S}, as it confirms the validity of the return receipt.*}
 | 
| 
a04b5b43625e
streamlined theory; conformance to recent publication
 paulson parents: 
16417diff
changeset | 453 | theorem RR_validity: | 
| 13956 | 454 | "[|Crypt (priSK TTP) S2TTP \<in> used evs; | 
| 455 | S2TTP = Crypt (pubEK TTP) | |
| 13934 | 456 |                {|Agent S, Number AO, Key K, Agent R, 
 | 
| 457 |                  Hash {|Number cleartext, Nonce q, r, em|}|};
 | |
| 13922 | 458 |      hr = Hash {|Number cleartext, Nonce q, r, em|};
 | 
| 459 | R\<noteq>Spy; evs \<in> certified_mail|] | |
| 460 |   ==> Notes R {|Agent TTP, Agent R, Key K, hr|} \<in> set evs"
 | |
| 461 | apply (erule rev_mp) | |
| 462 | apply (erule ssubst) | |
| 463 | apply (erule ssubst) | |
| 464 | apply (erule certified_mail.induct, simp_all) | |
| 13934 | 465 | txt{*Fake*} 
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 466 | apply (blast dest: Fake_parts_sing [THEN subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 467 | dest!: analz_subset_parts [THEN subsetD]) | 
| 13934 | 468 | txt{*Fake SSL*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 469 | apply (blast dest: Fake_parts_sing [THEN subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 470 | dest!: analz_subset_parts [THEN subsetD]) | 
| 13922 | 471 | txt{*Message 2*}
 | 
| 13934 | 472 | apply (drule CM2_S2TTP_parts_knows_Spy, assumption) | 
| 473 | apply (force dest: parts_cut) | |
| 474 | txt{*Message 3*}
 | |
| 475 | apply (frule_tac hr_form, assumption) | |
| 476 | apply (elim disjE exE, simp_all) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 477 | apply (blast dest: Fake_parts_sing [THEN subsetD] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: 
13956diff
changeset | 478 | dest!: analz_subset_parts [THEN subsetD]) | 
| 13922 | 479 | done | 
| 480 | ||
| 481 | end |