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