author | paulson |
Mon, 14 Jul 1997 12:47:21 +0200 | |
changeset 3519 | ab0a9fbed4c0 |
parent 3471 | cd37ec057028 |
child 3683 | aafe719dff14 |
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 |
Inductive relation "woolam" for the Woo-Lam protocol. |
|
7 |
||
8 |
Simplified version from page 11 of |
|
9 |
Abadi and Needham. Prudent Engineering Practice for Cryptographic Protocols. |
|
10 |
IEEE Trans. S.E. 22(1), 1996, pages 6-15. |
|
11 |
*) |
|
12 |
||
13 |
open WooLam; |
|
14 |
||
15 |
proof_timing:=true; |
|
16 |
HOL_quantifiers := false; |
|
17 |
||
18 |
||
2321 | 19 |
(*A "possibility property": there are traces that reach the end*) |
2274 | 20 |
goal thy |
21 |
"!!A B. [| A ~= B; A ~= Server; B ~= Server |] \ |
|
2283
68829cf138fc
Swapping arguments of Crypt; removing argument lost
paulson
parents:
2274
diff
changeset
|
22 |
\ ==> EX NB. EX evs: woolam. \ |
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
23 |
\ Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|}) : set evs"; |
2274 | 24 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
25 |
by (rtac (woolam.Nil RS woolam.WL1 RS woolam.WL2 RS woolam.WL3 RS |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2451
diff
changeset
|
26 |
woolam.WL4 RS woolam.WL5) 2); |
3471
cd37ec057028
Now the possibility proof calls the appropriate tactic
paulson
parents:
3466
diff
changeset
|
27 |
by possibility_tac; |
2274 | 28 |
result(); |
29 |
||
30 |
||
31 |
(**** Inductive proofs about woolam ****) |
|
32 |
||
33 |
(*Nobody sends themselves messages*) |
|
3465 | 34 |
goal thy "!!evs. evs : woolam ==> ALL A X. Says A A X ~: set evs"; |
2274 | 35 |
by (etac woolam.induct 1); |
36 |
by (Auto_tac()); |
|
37 |
qed_spec_mp "not_Says_to_self"; |
|
38 |
Addsimps [not_Says_to_self]; |
|
39 |
AddSEs [not_Says_to_self RSN (2, rev_notE)]; |
|
40 |
||
41 |
||
42 |
(** For reasoning about the encrypted portion of messages **) |
|
43 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
44 |
goal thy "!!evs. Says A' B X : set evs ==> X : analz (sees Spy evs)"; |
2274 | 45 |
by (etac (Says_imp_sees_Spy RS analz.Inj) 1); |
46 |
qed "WL4_analz_sees_Spy"; |
|
47 |
||
48 |
bind_thm ("WL4_parts_sees_Spy", |
|
49 |
WL4_analz_sees_Spy RS (impOfSubs analz_subset_parts)); |
|
50 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
51 |
(*For proving the easier theorems about X ~: parts (sees Spy evs) *) |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
52 |
fun parts_induct_tac i = |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
53 |
etac woolam.induct i THEN |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
54 |
forward_tac [WL4_parts_sees_Spy] (i+5) THEN |
3471
cd37ec057028
Now the possibility proof calls the appropriate tactic
paulson
parents:
3466
diff
changeset
|
55 |
prove_simple_subgoals_tac 1; |
2274 | 56 |
|
57 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
58 |
(** Theorems of the form X ~: parts (sees Spy evs) imply that NOBODY |
2274 | 59 |
sends messages containing X! **) |
60 |
||
61 |
(*Spy never sees another agent's shared key! (unless it's lost at start)*) |
|
62 |
goal thy |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
63 |
"!!evs. evs : woolam ==> (Key (shrK A) : parts (sees Spy evs)) = (A : lost)"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
64 |
by (parts_induct_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
65 |
by (Fake_parts_insert_tac 1); |
2274 | 66 |
qed "Spy_see_shrK"; |
67 |
Addsimps [Spy_see_shrK]; |
|
68 |
||
69 |
goal thy |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
70 |
"!!evs. evs : woolam ==> (Key (shrK A) : analz (sees Spy evs)) = (A : lost)"; |
2274 | 71 |
by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset)); |
72 |
qed "Spy_analz_shrK"; |
|
73 |
Addsimps [Spy_analz_shrK]; |
|
74 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
75 |
goal thy "!!A. [| Key (shrK A) : parts (sees Spy evs); \ |
2283
68829cf138fc
Swapping arguments of Crypt; removing argument lost
paulson
parents:
2274
diff
changeset
|
76 |
\ evs : woolam |] ==> A:lost"; |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
77 |
by (blast_tac (!claset addDs [Spy_see_shrK]) 1); |
2274 | 78 |
qed "Spy_see_shrK_D"; |
79 |
||
80 |
bind_thm ("Spy_analz_shrK_D", analz_subset_parts RS subsetD RS Spy_see_shrK_D); |
|
81 |
AddSDs [Spy_see_shrK_D, Spy_analz_shrK_D]; |
|
82 |
||
83 |
||
84 |
(**** Autheticity properties for Woo-Lam ****) |
|
85 |
||
86 |
||
87 |
(*** WL4 ***) |
|
88 |
||
89 |
(*If the encrypted message appears then it originated with Alice*) |
|
90 |
goal thy |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
91 |
"!!evs. [| A ~: lost; evs : woolam |] \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
92 |
\ ==> Crypt (shrK A) (Nonce NB) : parts (sees Spy evs) \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
93 |
\ --> (EX B. Says A B (Crypt (shrK A) (Nonce NB)) : set evs)"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
94 |
by (parts_induct_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
95 |
by (Fake_parts_insert_tac 1); |
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
96 |
by (Blast_tac 1); |
2274 | 97 |
qed_spec_mp "NB_Crypt_imp_Alice_msg"; |
98 |
||
99 |
(*Guarantee for Server: if it gets a message containing a certificate from |
|
100 |
Alice, then she originated that certificate. But we DO NOT know that B |
|
101 |
ever saw it: the Spy may have rerouted the message to the Server.*) |
|
102 |
goal thy |
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
103 |
"!!evs. [| A ~: lost; evs : woolam; \ |
2283
68829cf138fc
Swapping arguments of Crypt; removing argument lost
paulson
parents:
2274
diff
changeset
|
104 |
\ Says B' Server {|Agent A, Agent B, Crypt (shrK A) (Nonce NB)|} \ |
3465 | 105 |
\ : set evs |] \ |
106 |
\ ==> EX B. Says A B (Crypt (shrK A) (Nonce NB)) : set evs"; |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
107 |
by (blast_tac (!claset addSIs [NB_Crypt_imp_Alice_msg] |
2274 | 108 |
addSEs [MPair_parts] |
109 |
addDs [Says_imp_sees_Spy RS parts.Inj]) 1); |
|
2321 | 110 |
qed "Server_trusts_WL4"; |
2274 | 111 |
|
112 |
||
113 |
(*** WL5 ***) |
|
114 |
||
115 |
(*Server sent WL5 only if it received the right sort of message*) |
|
116 |
goal thy |
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
117 |
"!!evs. evs : woolam ==> \ |
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
118 |
\ Says Server B (Crypt (shrK B) {|Agent A, NB|}) : set evs \ |
2283
68829cf138fc
Swapping arguments of Crypt; removing argument lost
paulson
parents:
2274
diff
changeset
|
119 |
\ --> (EX B'. Says B' Server {|Agent A, Agent B, Crypt (shrK A) NB|} \ |
3465 | 120 |
\ : set evs)"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
121 |
by (parts_induct_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
122 |
by (Fake_parts_insert_tac 1); |
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
123 |
by (ALLGOALS Blast_tac); |
2274 | 124 |
bind_thm ("Server_sent_WL5", result() RSN (2, rev_mp)); |
125 |
||
126 |
||
127 |
(*If the encrypted message appears then it originated with the Server!*) |
|
128 |
goal thy |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
129 |
"!!evs. [| B ~: lost; evs : woolam |] \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
130 |
\ ==> Crypt (shrK B) {|Agent A, NB|} : parts (sees Spy evs) \ |
3465 | 131 |
\ --> Says Server B (Crypt (shrK B) {|Agent A, NB|}) : set evs"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
132 |
by (parts_induct_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
133 |
by (Fake_parts_insert_tac 1); |
2274 | 134 |
qed_spec_mp "NB_Crypt_imp_Server_msg"; |
135 |
||
136 |
(*Partial guarantee for B: if it gets a message of correct form then the Server |
|
137 |
sent the same message.*) |
|
138 |
goal thy |
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
139 |
"!!evs. [| Says S B (Crypt (shrK B) {|Agent A, NB|}) : set evs; \ |
2283
68829cf138fc
Swapping arguments of Crypt; removing argument lost
paulson
parents:
2274
diff
changeset
|
140 |
\ B ~: lost; evs : woolam |] \ |
3465 | 141 |
\ ==> Says Server B (Crypt (shrK B) {|Agent A, NB|}) : set evs"; |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
142 |
by (blast_tac (!claset addSIs [NB_Crypt_imp_Server_msg] |
2274 | 143 |
addDs [Says_imp_sees_Spy RS parts.Inj]) 1); |
144 |
qed "B_got_WL5"; |
|
145 |
||
146 |
(*Guarantee for B. If B gets the Server's certificate then A has encrypted |
|
147 |
the nonce using her key. This event can be no older than the nonce itself. |
|
148 |
But A may have sent the nonce to some other agent and it could have reached |
|
149 |
the Server via the Spy.*) |
|
150 |
goal thy |
|
3465 | 151 |
"!!evs. [| Says S B (Crypt (shrK B) {|Agent A, Nonce NB|}): set evs; \ |
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
152 |
\ A ~: lost; B ~: lost; evs : woolam |] \ |
3465 | 153 |
\ ==> EX B. Says A B (Crypt (shrK A) (Nonce NB)) : set evs"; |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
154 |
by (blast_tac (!claset addIs [Server_trusts_WL4] |
2274 | 155 |
addSDs [B_got_WL5 RS Server_sent_WL5]) 1); |
2321 | 156 |
qed "B_trusts_WL5"; |
2274 | 157 |
|
158 |
||
159 |
(*B only issues challenges in response to WL1. Useful??*) |
|
160 |
goal thy |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
161 |
"!!evs. [| B ~= Spy; evs : woolam |] \ |
3465 | 162 |
\ ==> Says B A (Nonce NB) : set evs \ |
163 |
\ --> (EX A'. Says A' B (Agent A) : set evs)"; |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
164 |
by (parts_induct_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
165 |
by (Fake_parts_insert_tac 1); |
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
166 |
by (ALLGOALS Blast_tac); |
2274 | 167 |
bind_thm ("B_said_WL2", result() RSN (2, rev_mp)); |
168 |
||
169 |
||
170 |
(**CANNOT be proved because A doesn't know where challenges come from... |
|
171 |
goal thy |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
172 |
"!!evs. [| A ~: lost; B ~= Spy; evs : woolam |] \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
173 |
\ ==> Crypt (shrK A) (Nonce NB) : parts (sees Spy evs) & \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
174 |
\ Says B A (Nonce NB) : set evs \ |
3465 | 175 |
\ --> Says A B (Crypt (shrK A) (Nonce NB)) : set evs"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3471
diff
changeset
|
176 |
by (parts_induct_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
177 |
by (Fake_parts_insert_tac 1); |
2274 | 178 |
by (Step_tac 1); |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
179 |
by (blast_tac (!claset addSEs partsEs) 1); |
2274 | 180 |
**) |