| author | wenzelm | 
| Mon, 10 Mar 2025 23:30:15 +0100 | |
| changeset 82258 | 9cbc848d263f | 
| parent 76299 | 0ad6f6508274 | 
| permissions | -rw-r--r-- | 
| 37936 | 1 | (* Title: HOL/Auth/OtwayRees.thy | 
| 11251 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1996 University of Cambridge | |
| 1941 | 4 | *) | 
| 5 | ||
| 61830 | 6 | section\<open>The Original Otway-Rees Protocol\<close> | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 7 | |
| 16417 | 8 | theory OtwayRees imports Public begin | 
| 13907 | 9 | |
| 61830 | 10 | text\<open>From page 244 of | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 11 | Burrows, Abadi and Needham (1989). A Logic of Authentication. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 12 | Proc. Royal Soc. 426 | 
| 1941 | 13 | |
| 61830 | 14 | This is the original version, which encrypts Nonce NB.\<close> | 
| 6308 
76f3865a2b1d
Added Bella's "Gets" model for Otway_Rees.  Also affects some other theories.
 paulson parents: 
5434diff
changeset | 15 | |
| 23746 | 16 | inductive_set otway :: "event list set" | 
| 17 | where | |
| 11251 | 18 | Nil: "[] \<in> otway" | 
| 76299 | 19 | \<comment> \<open>Initial trace is empty\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 20 | | Fake: "\<lbrakk>evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) \<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 21 | \<Longrightarrow> Says Spy B X # evsf \<in> otway" | 
| 76299 | 22 | \<comment> \<open>The spy can say almost anything.\<close> | 
| 23 | | Reception: "\<lbrakk>evsr \<in> otway; Says A B X \<in>set evsr\<rbrakk> \<Longrightarrow> Gets B X # evsr \<in> otway" | |
| 24 | \<comment> \<open>A message that has been sent can be received by the intended recipient.\<close> | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 25 | | OR1: "\<lbrakk>evs1 \<in> otway; Nonce NA \<notin> used evs1\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 26 | \<Longrightarrow> Says A B \<lbrace>Nonce NA, Agent A, Agent B, | 
| 61956 | 27 | Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace> \<rbrace> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 28 | # evs1 \<in> otway" | 
| 76299 | 29 | \<comment> \<open>Alice initiates a protocol run\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 30 | | OR2: "\<lbrakk>evs2 \<in> otway; Nonce NB \<notin> used evs2; | 
| 67613 | 31 | Gets B \<lbrace>Nonce NA, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 32 | \<Longrightarrow> Says B Server | 
| 61956 | 33 | \<lbrace>Nonce NA, Agent A, Agent B, X, | 
| 2451 
ce85a2aafc7a
Extensive tidying and simplification, largely stemming from
 paulson parents: 
2378diff
changeset | 34 | Crypt (shrK B) | 
| 61956 | 35 | \<lbrace>Nonce NA, Nonce NB, Agent A, Agent B\<rbrace>\<rbrace> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 36 | # evs2 \<in> otway" | 
| 76299 | 37 | \<comment> \<open>Bob's response to Alice's message. Note that NB is encrypted.\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 38 | | OR3: "\<lbrakk>evs3 \<in> otway; Key KAB \<notin> used evs3; | 
| 11251 | 39 | Gets Server | 
| 61956 | 40 | \<lbrace>Nonce NA, Agent A, Agent B, | 
| 41 | Crypt (shrK A) \<lbrace>Nonce NA, Agent A, Agent B\<rbrace>, | |
| 42 | Crypt (shrK B) \<lbrace>Nonce NA, Nonce NB, Agent A, Agent B\<rbrace>\<rbrace> | |
| 67613 | 43 | \<in> set evs3\<rbrakk> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 44 | \<Longrightarrow> Says Server B | 
| 61956 | 45 | \<lbrace>Nonce NA, | 
| 46 | Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>, | |
| 47 | Crypt (shrK B) \<lbrace>Nonce NB, Key KAB\<rbrace>\<rbrace> | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 48 | # evs3 \<in> otway" | 
| 76299 | 49 | \<comment> \<open>The Server receives Bob's message and checks that the three NAs | 
| 50 | match. Then he sends a new session key to Bob with a packet for forwarding to Alice\<close> | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 51 | | OR4: "\<lbrakk>evs4 \<in> otway; B \<noteq> Server; | 
| 61956 | 52 | Says B Server \<lbrace>Nonce NA, Agent A, Agent B, X', | 
| 2284 
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
 paulson parents: 
2135diff
changeset | 53 | Crypt (shrK B) | 
| 61956 | 54 | \<lbrace>Nonce NA, Nonce NB, Agent A, Agent B\<rbrace>\<rbrace> | 
| 67613 | 55 | \<in> set evs4; | 
| 61956 | 56 | Gets B \<lbrace>Nonce NA, X, Crypt (shrK B) \<lbrace>Nonce NB, Key K\<rbrace>\<rbrace> | 
| 67613 | 57 | \<in> set evs4\<rbrakk> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 58 | \<Longrightarrow> Says B A \<lbrace>Nonce NA, X\<rbrace> # evs4 \<in> otway" | 
| 76299 | 59 | \<comment> \<open>Bob receives the Server's (?) message and compares the Nonces with | 
| 60 | those in the message he previously sent the Server. | |
| 61 |        Need @{term"B \<noteq> Server"} because we allow messages to self.\<close>
 | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 62 | | Oops: "\<lbrakk>evso \<in> otway; | 
| 61956 | 63 | Says Server B \<lbrace>Nonce NA, X, Crypt (shrK B) \<lbrace>Nonce NB, Key K\<rbrace>\<rbrace> | 
| 67613 | 64 | \<in> set evso\<rbrakk> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 65 | \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> otway" | 
| 76299 | 66 | \<comment> \<open>This message models possible leaks of session keys. The nonces identify the protocol run\<close> | 
| 11251 | 67 | |
| 18570 | 68 | declare Says_imp_analz_Spy [dest] | 
| 11251 | 69 | declare parts.Body [dest] | 
| 70 | declare analz_into_parts [dest] | |
| 71 | declare Fake_parts_insert_in_Un [dest] | |
| 72 | ||
| 73 | ||
| 61830 | 74 | text\<open>A "possibility property": there are traces that reach the end\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 75 | lemma "\<lbrakk>B \<noteq> Server; Key K \<notin> used []\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 76 | \<Longrightarrow> \<exists>evs \<in> otway. | 
| 61956 | 77 | Says B A \<lbrace>Nonce NA, Crypt (shrK A) \<lbrace>Nonce NA, Key K\<rbrace>\<rbrace> | 
| 11251 | 78 | \<in> set evs" | 
| 79 | apply (intro exI bexI) | |
| 80 | apply (rule_tac [2] otway.Nil | |
| 81 | [THEN otway.OR1, THEN otway.Reception, | |
| 82 | THEN otway.OR2, THEN otway.Reception, | |
| 14200 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13907diff
changeset | 83 | THEN otway.OR3, THEN otway.Reception, THEN otway.OR4]) | 
| 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 paulson parents: 
13907diff
changeset | 84 | apply (possibility, simp add: used_Cons) | 
| 11251 | 85 | done | 
| 86 | ||
| 87 | lemma Gets_imp_Says [dest!]: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 88 | "\<lbrakk>Gets B X \<in> set evs; evs \<in> otway\<rbrakk> \<Longrightarrow> \<exists>A. Says A B X \<in> set evs" | 
| 11251 | 89 | apply (erule rev_mp) | 
| 13507 | 90 | apply (erule otway.induct, auto) | 
| 11251 | 91 | done | 
| 92 | ||
| 93 | ||
| 94 | (** For reasoning about the encrypted portion of messages **) | |
| 95 | ||
| 96 | lemma OR2_analz_knows_Spy: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 97 | "\<lbrakk>Gets B \<lbrace>N, Agent A, Agent B, X\<rbrace> \<in> set evs; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 98 | \<Longrightarrow> X \<in> analz (knows Spy evs)" | 
| 76299 | 99 | by blast | 
| 11251 | 100 | |
| 101 | lemma OR4_analz_knows_Spy: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 102 | "\<lbrakk>Gets B \<lbrace>N, X, Crypt (shrK B) X'\<rbrace> \<in> set evs; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 103 | \<Longrightarrow> X \<in> analz (knows Spy evs)" | 
| 76299 | 104 | by blast | 
| 11251 | 105 | |
| 106 | (*These lemmas assist simplification by removing forwarded X-variables. | |
| 107 | We can replace them by rewriting with parts_insert2 and proving using | |
| 108 | dest: parts_cut, but the proofs become more difficult.*) | |
| 109 | lemmas OR2_parts_knows_Spy = | |
| 45605 | 110 | OR2_analz_knows_Spy [THEN analz_into_parts] | 
| 11251 | 111 | |
| 112 | (*There could be OR4_parts_knows_Spy and Oops_parts_knows_Spy, but for | |
| 113 | some reason proofs work without them!*) | |
| 114 | ||
| 115 | ||
| 69597 | 116 | text\<open>Theorems of the form \<^term>\<open>X \<notin> parts (spies evs)\<close> imply that | 
| 61830 | 117 | NOBODY sends messages containing X!\<close> | 
| 11251 | 118 | |
| 61830 | 119 | text\<open>Spy never sees a good agent's shared key!\<close> | 
| 11251 | 120 | lemma Spy_see_shrK [simp]: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 121 | "evs \<in> otway \<Longrightarrow> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)" | 
| 13907 | 122 | by (erule otway.induct, force, | 
| 123 | drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+) | |
| 124 | ||
| 11251 | 125 | |
| 126 | lemma Spy_analz_shrK [simp]: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 127 | "evs \<in> otway \<Longrightarrow> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)" | 
| 11251 | 128 | by auto | 
| 129 | ||
| 130 | lemma Spy_see_shrK_D [dest!]: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 131 | "\<lbrakk>Key (shrK A) \<in> parts (knows Spy evs); evs \<in> otway\<rbrakk> \<Longrightarrow> A \<in> bad" | 
| 11251 | 132 | by (blast dest: Spy_see_shrK) | 
| 133 | ||
| 134 | ||
| 69597 | 135 | subsection\<open>Towards Secrecy: Proofs Involving \<^term>\<open>analz\<close>\<close> | 
| 11251 | 136 | |
| 76297 | 137 | text \<open>Describes the form of K and NA when the Server sends this message. Also | 
| 138 | for Oops case.\<close> | |
| 11251 | 139 | lemma Says_Server_message_form: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 140 | "\<lbrakk>Says Server B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs; | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 141 | evs \<in> otway\<rbrakk> | 
| 67613 | 142 | \<Longrightarrow> K \<notin> range shrK \<and> (\<exists>i. NA = Nonce i) \<and> (\<exists>j. NB = Nonce j)" | 
| 17778 | 143 | by (erule rev_mp, erule otway.induct, simp_all) | 
| 11251 | 144 | |
| 145 | ||
| 146 | (**** | |
| 147 | The following is to prove theorems of the form | |
| 148 | ||
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 149 | Key K \<in> analz (insert (Key KAB) (knows Spy evs)) \<Longrightarrow> | 
| 11251 | 150 | Key K \<in> analz (knows Spy evs) | 
| 151 | ||
| 152 | A more general formula must be proved inductively. | |
| 153 | ****) | |
| 154 | ||
| 155 | ||
| 61830 | 156 | text\<open>Session keys are not used to encrypt other session keys\<close> | 
| 11251 | 157 | |
| 61830 | 158 | text\<open>The equality makes the induction hypothesis easier to apply\<close> | 
| 11251 | 159 | lemma analz_image_freshK [rule_format]: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 160 | "evs \<in> otway \<Longrightarrow> | 
| 67613 | 161 | \<forall>K KK. KK \<subseteq> -(range shrK) \<longrightarrow> | 
| 162 | (Key K \<in> analz (Key`KK \<union> (knows Spy evs))) = | |
| 11251 | 163 | (K \<in> KK | Key K \<in> analz (knows Spy evs))" | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 164 | apply (erule otway.induct) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 165 | apply (frule_tac [8] Says_Server_message_form) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 166 | apply (drule_tac [7] OR4_analz_knows_Spy) | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
14200diff
changeset | 167 | apply (drule_tac [5] OR2_analz_knows_Spy, analz_freshK, spy_analz, auto) | 
| 11251 | 168 | done | 
| 169 | ||
| 170 | lemma analz_insert_freshK: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 171 | "\<lbrakk>evs \<in> otway; KAB \<notin> range shrK\<rbrakk> \<Longrightarrow> | 
| 11655 | 172 | (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) = | 
| 11251 | 173 | (K = KAB | Key K \<in> analz (knows Spy evs))" | 
| 174 | by (simp only: analz_image_freshK analz_image_freshK_simps) | |
| 175 | ||
| 176 | ||
| 61830 | 177 | text\<open>The Key K uniquely identifies the Server's message.\<close> | 
| 11251 | 178 | lemma unique_session_keys: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 179 | "\<lbrakk>Says Server B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, K\<rbrace>\<rbrace> \<in> set evs; | 
| 61956 | 180 | Says Server B' \<lbrace>NA',X',Crypt (shrK B') \<lbrace>NB',K\<rbrace>\<rbrace> \<in> set evs; | 
| 67613 | 181 | evs \<in> otway\<rbrakk> \<Longrightarrow> X=X' \<and> B=B' \<and> NA=NA' \<and> NB=NB'" | 
| 11251 | 182 | apply (erule rev_mp) | 
| 183 | apply (erule rev_mp) | |
| 184 | apply (erule otway.induct, simp_all) | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 185 | apply blast+ \<comment> \<open>OR3 and OR4\<close> | 
| 11251 | 186 | done | 
| 187 | ||
| 188 | ||
| 61830 | 189 | subsection\<open>Authenticity properties relating to NA\<close> | 
| 11251 | 190 | |
| 61830 | 191 | text\<open>Only OR1 can have caused such a part of a message to appear.\<close> | 
| 11251 | 192 | lemma Crypt_imp_OR1 [rule_format]: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 193 | "\<lbrakk>A \<notin> bad; evs \<in> otway\<rbrakk> | 
| 67613 | 194 | \<Longrightarrow> Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow> | 
| 61956 | 195 | Says A B \<lbrace>NA, Agent A, Agent B, | 
| 196 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> | |
| 11251 | 197 | \<in> set evs" | 
| 14225 | 198 | by (erule otway.induct, force, | 
| 199 | drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+) | |
| 11251 | 200 | |
| 201 | lemma Crypt_imp_OR1_Gets: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 202 | "\<lbrakk>Gets B \<lbrace>NA, Agent A, Agent B, | 
| 61956 | 203 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs; | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 204 | A \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 205 | \<Longrightarrow> Says A B \<lbrace>NA, Agent A, Agent B, | 
| 61956 | 206 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> | 
| 11251 | 207 | \<in> set evs" | 
| 208 | by (blast dest: Crypt_imp_OR1) | |
| 209 | ||
| 210 | ||
| 61830 | 211 | text\<open>The Nonce NA uniquely identifies A's message\<close> | 
| 11251 | 212 | lemma unique_NA: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 213 | "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs); | 
| 61956 | 214 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent C\<rbrace> \<in> parts (knows Spy evs); | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 215 | evs \<in> otway; A \<notin> bad\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 216 | \<Longrightarrow> B = C" | 
| 11251 | 217 | apply (erule rev_mp, erule rev_mp) | 
| 218 | apply (erule otway.induct, force, | |
| 13507 | 219 | drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+) | 
| 11251 | 220 | done | 
| 221 | ||
| 222 | ||
| 61830 | 223 | text\<open>It is impossible to re-use a nonce in both OR1 and OR2. This holds because | 
| 11251 | 224 | OR2 encrypts Nonce NB. It prevents the attack that can occur in the | 
| 61830 | 225 | over-simplified version of this protocol: see \<open>OtwayRees_Bad\<close>.\<close> | 
| 11251 | 226 | lemma no_nonce_OR1_OR2: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 227 | "\<lbrakk>Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs); | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 228 | A \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 229 | \<Longrightarrow> Crypt (shrK A) \<lbrace>NA', NA, Agent A', Agent A\<rbrace> \<notin> parts (knows Spy evs)" | 
| 11251 | 230 | apply (erule rev_mp) | 
| 231 | apply (erule otway.induct, force, | |
| 13507 | 232 | drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+) | 
| 11251 | 233 | done | 
| 234 | ||
| 61830 | 235 | text\<open>Crucial property: If the encrypted message appears, and A has used NA | 
| 236 | to start a run, then it originated with the Server!\<close> | |
| 11251 | 237 | lemma NA_Crypt_imp_Server_msg [rule_format]: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 238 | "\<lbrakk>A \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 239 | \<Longrightarrow> Says A B \<lbrace>NA, Agent A, Agent B, | 
| 67613 | 240 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs \<longrightarrow> | 
| 61956 | 241 | Crypt (shrK A) \<lbrace>NA, Key K\<rbrace> \<in> parts (knows Spy evs) | 
| 67613 | 242 | \<longrightarrow> (\<exists>NB. Says Server B | 
| 61956 | 243 | \<lbrace>NA, | 
| 244 | Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | |
| 245 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs)" | |
| 11251 | 246 | apply (erule otway.induct, force, | 
| 13507 | 247 | drule_tac [4] OR2_parts_knows_Spy, simp_all, blast) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 248 | subgoal \<comment> \<open>OR1: by freshness\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 249 | by blast | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 250 | subgoal \<comment> \<open>OR3\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 251 | by (blast dest!: no_nonce_OR1_OR2 intro: unique_NA) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 252 | subgoal \<comment> \<open>OR4\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 253 | by (blast intro!: Crypt_imp_OR1) | 
| 11251 | 254 | done | 
| 255 | ||
| 256 | ||
| 61830 | 257 | text\<open>Corollary: if A receives B's OR4 message and the nonce NA agrees | 
| 11251 | 258 | then the key really did come from the Server! CANNOT prove this of the | 
| 259 | bad form of this protocol, even though we can prove | |
| 61830 | 260 | \<open>Spy_not_see_encrypted_key\<close>\<close> | 
| 11251 | 261 | lemma A_trusts_OR4: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 262 | "\<lbrakk>Says A B \<lbrace>NA, Agent A, Agent B, | 
| 61956 | 263 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs; | 
| 264 | Says B' A \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>\<rbrace> \<in> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 265 | A \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 266 | \<Longrightarrow> \<exists>NB. Says Server B | 
| 61956 | 267 | \<lbrace>NA, | 
| 268 | Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | |
| 269 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> | |
| 11251 | 270 | \<in> set evs" | 
| 271 | by (blast intro!: NA_Crypt_imp_Server_msg) | |
| 272 | ||
| 273 | ||
| 61830 | 274 | text\<open>Crucial secrecy property: Spy does not see the keys sent in msg OR3 | 
| 11251 | 275 | Does not in itself guarantee security: an attack could violate | 
| 69597 | 276 | the premises, e.g. by having \<^term>\<open>A=Spy\<close>\<close> | 
| 11251 | 277 | lemma secrecy_lemma: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 278 | "\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 279 | \<Longrightarrow> Says Server B | 
| 61956 | 280 | \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | 
| 67613 | 281 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs \<longrightarrow> | 
| 282 | Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs \<longrightarrow> | |
| 11251 | 283 | Key K \<notin> analz (knows Spy evs)" | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 284 | apply (erule otway.induct, force, simp_all) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 285 | subgoal \<comment> \<open>Fake\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 286 | by spy_analz | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 287 | subgoal \<comment> \<open>OR2\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 288 | by (drule OR2_analz_knows_Spy) (auto simp: analz_insert_eq) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 289 | subgoal \<comment> \<open>OR3\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 290 | by (auto simp add: analz_insert_freshK pushes) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 291 | subgoal \<comment> \<open>OR4\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 292 | by (drule OR4_analz_knows_Spy) (auto simp: analz_insert_eq) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 293 | subgoal \<comment> \<open>Oops\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 294 | by (auto simp add: Says_Server_message_form analz_insert_freshK unique_session_keys) | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 295 | done | 
| 11251 | 296 | |
| 13907 | 297 | theorem Spy_not_see_encrypted_key: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 298 | "\<lbrakk>Says Server B | 
| 61956 | 299 | \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | 
| 300 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs; | |
| 301 | Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 302 | A \<notin> bad; B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 303 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs)" | 
| 11251 | 304 | by (blast dest: Says_Server_message_form secrecy_lemma) | 
| 305 | ||
| 61830 | 306 | text\<open>This form is an immediate consequence of the previous result. It is | 
| 13907 | 307 | similar to the assertions established by other methods. It is equivalent | 
| 69597 | 308 | to the previous result in that the Spy already has \<^term>\<open>analz\<close> and | 
| 309 | \<^term>\<open>synth\<close> at his disposal. However, the conclusion | |
| 310 | \<^term>\<open>Key K \<notin> knows Spy evs\<close> appears not to be inductive: all the cases | |
| 13907 | 311 | other than Fake are trivial, while Fake requires | 
| 69597 | 312 | \<^term>\<open>Key K \<notin> analz (knows Spy evs)\<close>.\<close> | 
| 13907 | 313 | lemma Spy_not_know_encrypted_key: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 314 | "\<lbrakk>Says Server B | 
| 61956 | 315 | \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | 
| 316 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs; | |
| 317 | Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 318 | A \<notin> bad; B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 319 | \<Longrightarrow> Key K \<notin> knows Spy evs" | 
| 13907 | 320 | by (blast dest: Spy_not_see_encrypted_key) | 
| 321 | ||
| 11251 | 322 | |
| 61830 | 323 | text\<open>A's guarantee. The Oops premise quantifies over NB because A cannot know | 
| 324 | what it is.\<close> | |
| 11251 | 325 | lemma A_gets_good_key: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 326 | "\<lbrakk>Says A B \<lbrace>NA, Agent A, Agent B, | 
| 61956 | 327 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs; | 
| 328 | Says B' A \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>\<rbrace> \<in> set evs; | |
| 329 | \<forall>NB. Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 330 | A \<notin> bad; B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 331 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs)" | 
| 11251 | 332 | by (blast dest!: A_trusts_OR4 Spy_not_see_encrypted_key) | 
| 333 | ||
| 334 | ||
| 61830 | 335 | subsection\<open>Authenticity properties relating to NB\<close> | 
| 11251 | 336 | |
| 61830 | 337 | text\<open>Only OR2 can have caused such a part of a message to appear. We do not | 
| 338 | know anything about X: it does NOT have to have the right form.\<close> | |
| 11251 | 339 | lemma Crypt_imp_OR2: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 340 | "\<lbrakk>Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace> \<in> parts (knows Spy evs); | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 341 | B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 342 | \<Longrightarrow> \<exists>X. Says B Server | 
| 61956 | 343 | \<lbrace>NA, Agent A, Agent B, X, | 
| 344 | Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace>\<rbrace> | |
| 11251 | 345 | \<in> set evs" | 
| 346 | apply (erule rev_mp) | |
| 347 | apply (erule otway.induct, force, | |
| 13507 | 348 | drule_tac [4] OR2_parts_knows_Spy, simp_all, blast+) | 
| 11251 | 349 | done | 
| 350 | ||
| 351 | ||
| 61830 | 352 | text\<open>The Nonce NB uniquely identifies B's message\<close> | 
| 11251 | 353 | lemma unique_NB: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 354 | "\<lbrakk>Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace> \<in> parts(knows Spy evs); | 
| 61956 | 355 | Crypt (shrK B) \<lbrace>NC, NB, Agent C, Agent B\<rbrace> \<in> parts(knows Spy evs); | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 356 | evs \<in> otway; B \<notin> bad\<rbrakk> | 
| 67613 | 357 | \<Longrightarrow> NC = NA \<and> C = A" | 
| 11251 | 358 | apply (erule rev_mp, erule rev_mp) | 
| 359 | apply (erule otway.induct, force, | |
| 360 | drule_tac [4] OR2_parts_knows_Spy, simp_all) | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 361 | apply blast+ \<comment> \<open>Fake, OR2\<close> | 
| 11251 | 362 | done | 
| 363 | ||
| 61830 | 364 | text\<open>If the encrypted message appears, and B has used Nonce NB, | 
| 365 | then it originated with the Server! Quite messy proof.\<close> | |
| 11251 | 366 | lemma NB_Crypt_imp_Server_msg [rule_format]: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 367 | "\<lbrakk>B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 368 | \<Longrightarrow> Crypt (shrK B) \<lbrace>NB, Key K\<rbrace> \<in> parts (knows Spy evs) | 
| 67613 | 369 | \<longrightarrow> (\<forall>X'. Says B Server | 
| 61956 | 370 | \<lbrace>NA, Agent A, Agent B, X', | 
| 371 | Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace>\<rbrace> | |
| 11251 | 372 | \<in> set evs | 
| 67613 | 373 | \<longrightarrow> Says Server B | 
| 61956 | 374 | \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | 
| 375 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> | |
| 11251 | 376 | \<in> set evs)" | 
| 377 | apply simp | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 378 | apply (erule otway.induct, force, simp_all) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 379 | subgoal \<comment> \<open>Fake\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 380 | by blast | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 381 | subgoal \<comment> \<open>OR2\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 382 | by (force dest!: OR2_parts_knows_Spy) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 383 | subgoal \<comment> \<open>OR3\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 384 | by (blast dest: unique_NB dest!: no_nonce_OR1_OR2) \<comment> \<open>OR3\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63975diff
changeset | 385 | subgoal \<comment> \<open>OR4\<close> | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 386 | by (blast dest!: Crypt_imp_OR2) | 
| 11251 | 387 | done | 
| 388 | ||
| 389 | ||
| 61830 | 390 | text\<open>Guarantee for B: if it gets a message with matching NB then the Server | 
| 391 | has sent the correct message.\<close> | |
| 13907 | 392 | theorem B_trusts_OR3: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 393 | "\<lbrakk>Says B Server \<lbrace>NA, Agent A, Agent B, X', | 
| 61956 | 394 | Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace>\<rbrace> | 
| 11251 | 395 | \<in> set evs; | 
| 61956 | 396 | Gets B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs; | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 397 | B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 398 | \<Longrightarrow> Says Server B | 
| 61956 | 399 | \<lbrace>NA, | 
| 400 | Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | |
| 401 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> | |
| 11251 | 402 | \<in> set evs" | 
| 403 | by (blast intro!: NB_Crypt_imp_Server_msg) | |
| 404 | ||
| 405 | ||
| 61830 | 406 | text\<open>The obvious combination of \<open>B_trusts_OR3\<close> with | 
| 407 | \<open>Spy_not_see_encrypted_key\<close>\<close> | |
| 11251 | 408 | lemma B_gets_good_key: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 409 | "\<lbrakk>Says B Server \<lbrace>NA, Agent A, Agent B, X', | 
| 61956 | 410 | Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace>\<rbrace> | 
| 11251 | 411 | \<in> set evs; | 
| 61956 | 412 | Gets B \<lbrace>NA, X, Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs; | 
| 413 | Notes Spy \<lbrace>NA, NB, Key K\<rbrace> \<notin> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 414 | A \<notin> bad; B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 415 | \<Longrightarrow> Key K \<notin> analz (knows Spy evs)" | 
| 11251 | 416 | by (blast dest!: B_trusts_OR3 Spy_not_see_encrypted_key) | 
| 417 | ||
| 418 | ||
| 419 | lemma OR3_imp_OR2: | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 420 | "\<lbrakk>Says Server B | 
| 61956 | 421 | \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>, | 
| 422 | Crypt (shrK B) \<lbrace>NB, Key K\<rbrace>\<rbrace> \<in> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 423 | B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 424 | \<Longrightarrow> \<exists>X. Says B Server \<lbrace>NA, Agent A, Agent B, X, | 
| 61956 | 425 | Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace>\<rbrace> | 
| 11251 | 426 | \<in> set evs" | 
| 427 | apply (erule rev_mp) | |
| 428 | apply (erule otway.induct, simp_all) | |
| 429 | apply (blast dest!: Crypt_imp_OR2)+ | |
| 430 | done | |
| 431 | ||
| 432 | ||
| 61830 | 433 | text\<open>After getting and checking OR4, agent A can trust that B has been active. | 
| 11251 | 434 | We could probably prove that X has the expected form, but that is not | 
| 61830 | 435 | strictly necessary for authentication.\<close> | 
| 13907 | 436 | theorem A_auths_B: | 
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 437 | "\<lbrakk>Says B' A \<lbrace>NA, Crypt (shrK A) \<lbrace>NA, Key K\<rbrace>\<rbrace> \<in> set evs; | 
| 61956 | 438 | Says A B \<lbrace>NA, Agent A, Agent B, | 
| 439 | Crypt (shrK A) \<lbrace>NA, Agent A, Agent B\<rbrace>\<rbrace> \<in> set evs; | |
| 63975 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 440 | A \<notin> bad; B \<notin> bad; evs \<in> otway\<rbrakk> | 
| 
6728b5007ad0
Trying out "subgoal", and no more [| |]
 paulson <lp15@cam.ac.uk> parents: 
61956diff
changeset | 441 | \<Longrightarrow> \<exists>NB X. Says B Server \<lbrace>NA, Agent A, Agent B, X, | 
| 61956 | 442 | Crypt (shrK B) \<lbrace>NA, NB, Agent A, Agent B\<rbrace>\<rbrace> | 
| 11251 | 443 | \<in> set evs" | 
| 444 | by (blast dest!: A_trusts_OR4 OR3_imp_OR2) | |
| 445 | ||
| 1941 | 446 | end |