| author | wenzelm | 
| Sat, 04 Jul 2009 11:46:51 +0200 | |
| changeset 31929 | ecfc667cac53 | 
| parent 23746 | a455e69c31cc | 
| child 37936 | 1e4c5015a72e | 
| permissions | -rw-r--r-- | 
| 2274 | 1 | (* Title: HOL/Auth/WooLam | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 7 | header{*The Woo-Lam Protocol*}
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 8 | |
| 16417 | 9 | theory WooLam imports Public begin | 
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 10 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 11 | text{*Simplified version from page 11 of
 | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 12 | Abadi and Needham (1996). | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 13 | Prudent Engineering Practice for Cryptographic Protocols. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 14 | IEEE Trans. S.E. 22(1), pages 6-15. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 15 | |
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 16 | Note: this differs from the Woo-Lam protocol discussed by Lowe (1996): | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 17 | Some New Attacks upon Security Protocols. | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 18 | Computer Security Foundations Workshop | 
| 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 19 | *} | 
| 2274 | 20 | |
| 23746 | 21 | inductive_set woolam :: "event list set" | 
| 22 | where | |
| 2274 | 23 | (*Initial trace is empty*) | 
| 11251 | 24 | Nil: "[] \<in> woolam" | 
| 2274 | 25 | |
| 5434 
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
 paulson parents: 
3683diff
changeset | 26 | (** These rules allow agents to send messages to themselves **) | 
| 
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
 paulson parents: 
3683diff
changeset | 27 | |
| 2274 | 28 | (*The spy MAY say anything he CAN say. We do not expect him to | 
| 29 | invent new nonces here, but he can also use NS1. Common to | |
| 30 | all similar protocols.*) | |
| 23746 | 31 | | Fake: "[| evsf \<in> woolam; X \<in> synth (analz (spies evsf)) |] | 
| 11251 | 32 | ==> Says Spy B X # evsf \<in> woolam" | 
| 2274 | 33 | |
| 34 | (*Alice initiates a protocol run*) | |
| 23746 | 35 | | WL1: "evs1 \<in> woolam ==> Says A B (Agent A) # evs1 \<in> woolam" | 
| 2274 | 36 | |
| 37 | (*Bob responds to Alice's message with a challenge.*) | |
| 23746 | 38 | | WL2: "[| evs2 \<in> woolam; Says A' B (Agent A) \<in> set evs2 |] | 
| 11251 | 39 | ==> Says B A (Nonce NB) # evs2 \<in> woolam" | 
| 2274 | 40 | |
| 41 | (*Alice responds to Bob's challenge by encrypting NB with her key. | |
| 42 | B is *not* properly determined -- Alice essentially broadcasts | |
| 43 | her reply.*) | |
| 23746 | 44 | | WL3: "[| evs3 \<in> woolam; | 
| 11251 | 45 | Says A B (Agent A) \<in> set evs3; | 
| 46 | Says B' A (Nonce NB) \<in> set evs3 |] | |
| 47 | ==> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \<in> woolam" | |
| 2274 | 48 | |
| 3470 | 49 | (*Bob forwards Alice's response to the Server. NOTE: usually | 
| 50 | the messages are shown in chronological order, for clarity. | |
| 51 | But here, exchanging the two events would cause the lemma | |
| 3683 | 52 | WL4_analz_spies to pick up the wrong assumption!*) | 
| 23746 | 53 | | WL4: "[| evs4 \<in> woolam; | 
| 11251 | 54 | Says A' B X \<in> set evs4; | 
| 55 | Says A'' B (Agent A) \<in> set evs4 |] | |
| 56 |           ==> Says B Server {|Agent A, Agent B, X|} # evs4 \<in> woolam"
 | |
| 2274 | 57 | |
| 58 | (*Server decrypts Alice's response for Bob.*) | |
| 23746 | 59 | | WL5: "[| evs5 \<in> woolam; | 
| 2283 
68829cf138fc
Swapping arguments of Crypt; removing argument lost
 paulson parents: 
2274diff
changeset | 60 |              Says B' Server {|Agent A, Agent B, Crypt (shrK A) (Nonce NB)|}
 | 
| 11251 | 61 | \<in> set evs5 |] | 
| 2283 
68829cf138fc
Swapping arguments of Crypt; removing argument lost
 paulson parents: 
2274diff
changeset | 62 |           ==> Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|})
 | 
| 11251 | 63 | # evs5 \<in> woolam" | 
| 64 | ||
| 65 | ||
| 66 | declare Says_imp_knows_Spy [THEN analz.Inj, dest] | |
| 67 | declare parts.Body [dest] | |
| 68 | declare analz_into_parts [dest] | |
| 69 | declare Fake_parts_insert_in_Un [dest] | |
| 70 | ||
| 71 | ||
| 72 | (*A "possibility property": there are traces that reach the end*) | |
| 73 | lemma "\<exists>NB. \<exists>evs \<in> woolam. | |
| 74 |              Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|}) \<in> set evs"
 | |
| 75 | apply (intro exI bexI) | |
| 76 | apply (rule_tac [2] woolam.Nil | |
| 77 | [THEN woolam.WL1, THEN woolam.WL2, THEN woolam.WL3, | |
| 13507 | 78 | THEN woolam.WL4, THEN woolam.WL5], possibility) | 
| 11251 | 79 | done | 
| 80 | ||
| 81 | (*Could prove forwarding lemmas for WL4, but we do not need them!*) | |
| 82 | ||
| 83 | (**** Inductive proofs about woolam ****) | |
| 84 | ||
| 85 | (** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY | |
| 86 | sends messages containing X! **) | |
| 87 | ||
| 88 | (*Spy never sees a good agent's shared key!*) | |
| 89 | lemma Spy_see_shrK [simp]: | |
| 90 | "evs \<in> woolam ==> (Key (shrK A) \<in> parts (spies evs)) = (A \<in> bad)" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 91 | by (erule woolam.induct, force, simp_all, blast+) | 
| 11251 | 92 | |
| 93 | lemma Spy_analz_shrK [simp]: | |
| 94 | "evs \<in> woolam ==> (Key (shrK A) \<in> analz (spies evs)) = (A \<in> bad)" | |
| 95 | by auto | |
| 96 | ||
| 97 | lemma Spy_see_shrK_D [dest!]: | |
| 98 | "[|Key (shrK A) \<in> parts (knows Spy evs); evs \<in> woolam|] ==> A \<in> bad" | |
| 99 | by (blast dest: Spy_see_shrK) | |
| 100 | ||
| 101 | ||
| 102 | (**** Autheticity properties for Woo-Lam ****) | |
| 103 | ||
| 104 | (*** WL4 ***) | |
| 105 | ||
| 106 | (*If the encrypted message appears then it originated with Alice*) | |
| 107 | lemma NB_Crypt_imp_Alice_msg: | |
| 108 | "[| Crypt (shrK A) (Nonce NB) \<in> parts (spies evs); | |
| 109 | A \<notin> bad; evs \<in> woolam |] | |
| 110 | ==> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 111 | by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) | 
| 11251 | 112 | |
| 113 | (*Guarantee for Server: if it gets a message containing a certificate from | |
| 114 | Alice, then she originated that certificate. But we DO NOT know that B | |
| 115 | ever saw it: the Spy may have rerouted the message to the Server.*) | |
| 116 | lemma Server_trusts_WL4 [dest]: | |
| 117 |      "[| Says B' Server {|Agent A, Agent B, Crypt (shrK A) (Nonce NB)|}
 | |
| 118 | \<in> set evs; | |
| 119 | A \<notin> bad; evs \<in> woolam |] | |
| 120 | ==> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs" | |
| 121 | by (blast intro!: NB_Crypt_imp_Alice_msg) | |
| 122 | ||
| 123 | ||
| 124 | (*** WL5 ***) | |
| 125 | ||
| 126 | (*Server sent WL5 only if it received the right sort of message*) | |
| 127 | lemma Server_sent_WL5 [dest]: | |
| 128 |      "[| Says Server B (Crypt (shrK B) {|Agent A, NB|}) \<in> set evs;
 | |
| 129 | evs \<in> woolam |] | |
| 130 |       ==> \<exists>B'. Says B' Server {|Agent A, Agent B, Crypt (shrK A) NB|}
 | |
| 131 | \<in> set evs" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 132 | by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) | 
| 11251 | 133 | |
| 134 | (*If the encrypted message appears then it originated with the Server!*) | |
| 135 | lemma NB_Crypt_imp_Server_msg [rule_format]: | |
| 136 |      "[| Crypt (shrK B) {|Agent A, NB|} \<in> parts (spies evs);
 | |
| 137 | B \<notin> bad; evs \<in> woolam |] | |
| 138 |       ==> Says Server B (Crypt (shrK B) {|Agent A, NB|}) \<in> set evs"
 | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 139 | by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) | 
| 11251 | 140 | |
| 141 | (*Guarantee for B. If B gets the Server's certificate then A has encrypted | |
| 142 | the nonce using her key. This event can be no older than the nonce itself. | |
| 143 | But A may have sent the nonce to some other agent and it could have reached | |
| 144 | the Server via the Spy.*) | |
| 145 | lemma B_trusts_WL5: | |
| 146 |      "[| Says S B (Crypt (shrK B) {|Agent A, Nonce NB|}): set evs;
 | |
| 147 | A \<notin> bad; B \<notin> bad; evs \<in> woolam |] | |
| 148 | ==> \<exists>B. Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs" | |
| 149 | by (blast dest!: NB_Crypt_imp_Server_msg) | |
| 150 | ||
| 151 | ||
| 152 | (*B only issues challenges in response to WL1. Not used.*) | |
| 153 | lemma B_said_WL2: | |
| 154 | "[| Says B A (Nonce NB) \<in> set evs; B \<noteq> Spy; evs \<in> woolam |] | |
| 155 | ==> \<exists>A'. Says A' B (Agent A) \<in> set evs" | |
| 14207 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 paulson parents: 
13507diff
changeset | 156 | by (erule rev_mp, erule woolam.induct, force, simp_all, blast+) | 
| 11251 | 157 | |
| 158 | ||
| 159 | (**CANNOT be proved because A doesn't know where challenges come from...*) | |
| 160 | lemma "[| A \<notin> bad; B \<noteq> Spy; evs \<in> woolam |] | |
| 161 | ==> Crypt (shrK A) (Nonce NB) \<in> parts (spies evs) & | |
| 162 | Says B A (Nonce NB) \<in> set evs | |
| 163 | --> Says A B (Crypt (shrK A) (Nonce NB)) \<in> set evs" | |
| 13507 | 164 | apply (erule rev_mp, erule woolam.induct, force, simp_all, blast, auto) | 
| 11251 | 165 | oops | 
| 2274 | 166 | |
| 167 | end |