| author | paulson |
| Sun, 13 Jun 1999 13:52:26 +0200 | |
| changeset 6821 | 350f27e2d649 |
| parent 6141 | a6922171b396 |
| child 6915 | 4ab8e31a8421 |
| permissions | -rw-r--r-- |
| 1839 | 1 |
(* Title: HOL/Auth/Message |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
Datatypes of agents and messages; |
|
| 1913 | 7 |
Inductive relations "parts", "analz" and "synth" |
| 1839 | 8 |
*) |
9 |
||
| 3702 | 10 |
|
11 |
(*Eliminates a commonly-occurring expression*) |
|
12 |
goal HOL.thy "~ (ALL x. x~=y)"; |
|
13 |
by (Blast_tac 1); |
|
14 |
Addsimps [result()]; |
|
15 |
||
| 3668 | 16 |
AddIffs atomic.inject; |
17 |
AddIffs msg.inject; |
|
| 1839 | 18 |
|
|
4422
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4157
diff
changeset
|
19 |
(*Equations hold because constructors are injective; cannot prove for all f*) |
| 5076 | 20 |
Goal "(Friend x : Friend``A) = (x:A)"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
21 |
by Auto_tac; |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
22 |
qed "Friend_image_eq"; |
|
4422
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4157
diff
changeset
|
23 |
|
| 5076 | 24 |
Goal "(Key x : Key``A) = (x:A)"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
25 |
by Auto_tac; |
|
4422
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4157
diff
changeset
|
26 |
qed "Key_image_eq"; |
|
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4157
diff
changeset
|
27 |
|
| 5076 | 28 |
Goal "(Nonce x ~: Key``A)"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
29 |
by Auto_tac; |
|
4422
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4157
diff
changeset
|
30 |
qed "Nonce_Key_image_eq"; |
|
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4157
diff
changeset
|
31 |
Addsimps [Friend_image_eq, Key_image_eq, Nonce_Key_image_eq]; |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
32 |
|
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
33 |
|
| 1839 | 34 |
(** Inverse of keys **) |
35 |
||
| 5493 | 36 |
Goal "(invKey K = invKey K') = (K=K')"; |
| 3730 | 37 |
by Safe_tac; |
| 2032 | 38 |
by (rtac box_equals 1); |
| 1839 | 39 |
by (REPEAT (rtac invKey 2)); |
40 |
by (Asm_simp_tac 1); |
|
41 |
qed "invKey_eq"; |
|
42 |
||
43 |
Addsimps [invKey, invKey_eq]; |
|
44 |
||
45 |
||
46 |
(**** keysFor operator ****) |
|
47 |
||
| 5076 | 48 |
Goalw [keysFor_def] "keysFor {} = {}";
|
| 2891 | 49 |
by (Blast_tac 1); |
| 1839 | 50 |
qed "keysFor_empty"; |
51 |
||
| 5076 | 52 |
Goalw [keysFor_def] "keysFor (H Un H') = keysFor H Un keysFor H'"; |
| 2891 | 53 |
by (Blast_tac 1); |
| 1839 | 54 |
qed "keysFor_Un"; |
55 |
||
| 5076 | 56 |
Goalw [keysFor_def] "keysFor (UN i:A. H i) = (UN i:A. keysFor (H i))"; |
| 2891 | 57 |
by (Blast_tac 1); |
| 4157 | 58 |
qed "keysFor_UN"; |
| 1839 | 59 |
|
60 |
(*Monotonicity*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
61 |
Goalw [keysFor_def] "G<=H ==> keysFor(G) <= keysFor(H)"; |
| 2891 | 62 |
by (Blast_tac 1); |
| 1839 | 63 |
qed "keysFor_mono"; |
64 |
||
| 5076 | 65 |
Goalw [keysFor_def] "keysFor (insert (Agent A) H) = keysFor H"; |
| 3102 | 66 |
by (Blast_tac 1); |
| 1839 | 67 |
qed "keysFor_insert_Agent"; |
68 |
||
| 5076 | 69 |
Goalw [keysFor_def] "keysFor (insert (Nonce N) H) = keysFor H"; |
| 3102 | 70 |
by (Blast_tac 1); |
| 1839 | 71 |
qed "keysFor_insert_Nonce"; |
72 |
||
| 5076 | 73 |
Goalw [keysFor_def] "keysFor (insert (Number N) H) = keysFor H"; |
| 3668 | 74 |
by (Blast_tac 1); |
75 |
qed "keysFor_insert_Number"; |
|
76 |
||
| 5076 | 77 |
Goalw [keysFor_def] "keysFor (insert (Key K) H) = keysFor H"; |
| 3102 | 78 |
by (Blast_tac 1); |
| 1839 | 79 |
qed "keysFor_insert_Key"; |
80 |
||
| 5076 | 81 |
Goalw [keysFor_def] "keysFor (insert (Hash X) H) = keysFor H"; |
| 3102 | 82 |
by (Blast_tac 1); |
| 2373 | 83 |
qed "keysFor_insert_Hash"; |
84 |
||
| 5076 | 85 |
Goalw [keysFor_def] "keysFor (insert {|X,Y|} H) = keysFor H";
|
| 3102 | 86 |
by (Blast_tac 1); |
| 1839 | 87 |
qed "keysFor_insert_MPair"; |
88 |
||
| 5076 | 89 |
Goalw [keysFor_def] |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
90 |
"keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
91 |
by Auto_tac; |
| 1839 | 92 |
qed "keysFor_insert_Crypt"; |
93 |
||
| 4157 | 94 |
Addsimps [keysFor_empty, keysFor_Un, keysFor_UN, |
| 3668 | 95 |
keysFor_insert_Agent, keysFor_insert_Nonce, |
96 |
keysFor_insert_Number, keysFor_insert_Key, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
97 |
keysFor_insert_Hash, keysFor_insert_MPair, keysFor_insert_Crypt]; |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
98 |
AddSEs [keysFor_Un RS equalityD1 RS subsetD RS UnE, |
| 4157 | 99 |
keysFor_UN RS equalityD1 RS subsetD RS UN_E]; |
| 1839 | 100 |
|
| 5076 | 101 |
Goalw [keysFor_def] "keysFor (Key``E) = {}";
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
102 |
by Auto_tac; |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
103 |
qed "keysFor_image_Key"; |
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
104 |
Addsimps [keysFor_image_Key]; |
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
105 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
106 |
Goalw [keysFor_def] "Crypt K X : H ==> invKey K : keysFor H"; |
| 2891 | 107 |
by (Blast_tac 1); |
| 2068 | 108 |
qed "Crypt_imp_invKey_keysFor"; |
109 |
||
| 1839 | 110 |
|
111 |
(**** Inductive relation "parts" ****) |
|
112 |
||
113 |
val major::prems = |
|
| 5076 | 114 |
Goal "[| {|X,Y|} : parts H; \
|
| 1839 | 115 |
\ [| X : parts H; Y : parts H |] ==> P \ |
116 |
\ |] ==> P"; |
|
117 |
by (cut_facts_tac [major] 1); |
|
| 2032 | 118 |
by (resolve_tac prems 1); |
| 1839 | 119 |
by (REPEAT (eresolve_tac [asm_rl, parts.Fst, parts.Snd] 1)); |
120 |
qed "MPair_parts"; |
|
121 |
||
122 |
AddIs [parts.Inj]; |
|
|
1929
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
123 |
|
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
124 |
val partsEs = [MPair_parts, make_elim parts.Body]; |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
125 |
|
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
126 |
AddSEs partsEs; |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
127 |
(*NB These two rules are UNSAFE in the formal sense, as they discard the |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
128 |
compound message. They work well on THIS FILE, perhaps because its |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
129 |
proofs concern only atomic messages.*) |
| 1839 | 130 |
|
| 5076 | 131 |
Goal "H <= parts(H)"; |
| 2891 | 132 |
by (Blast_tac 1); |
| 1839 | 133 |
qed "parts_increasing"; |
134 |
||
135 |
(*Monotonicity*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
136 |
Goalw parts.defs "G<=H ==> parts(G) <= parts(H)"; |
| 1839 | 137 |
by (rtac lfp_mono 1); |
138 |
by (REPEAT (ares_tac basic_monos 1)); |
|
139 |
qed "parts_mono"; |
|
140 |
||
| 2373 | 141 |
val parts_insertI = impOfSubs (subset_insertI RS parts_mono); |
142 |
||
| 5076 | 143 |
Goal "parts{} = {}";
|
| 3730 | 144 |
by Safe_tac; |
| 2032 | 145 |
by (etac parts.induct 1); |
| 2891 | 146 |
by (ALLGOALS Blast_tac); |
| 1839 | 147 |
qed "parts_empty"; |
148 |
Addsimps [parts_empty]; |
|
149 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
150 |
Goal "X: parts{} ==> P";
|
| 1839 | 151 |
by (Asm_full_simp_tac 1); |
152 |
qed "parts_emptyE"; |
|
153 |
AddSEs [parts_emptyE]; |
|
154 |
||
| 1893 | 155 |
(*WARNING: loops if H = {Y}, therefore must not be repeated!*)
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
156 |
Goal "X: parts H ==> EX Y:H. X: parts {Y}";
|
| 2032 | 157 |
by (etac parts.induct 1); |
| 2891 | 158 |
by (ALLGOALS Blast_tac); |
| 1893 | 159 |
qed "parts_singleton"; |
160 |
||
| 1839 | 161 |
|
162 |
(** Unions **) |
|
163 |
||
| 5076 | 164 |
Goal "parts(G) Un parts(H) <= parts(G Un H)"; |
| 1839 | 165 |
by (REPEAT (ares_tac [Un_least, parts_mono, Un_upper1, Un_upper2] 1)); |
166 |
val parts_Un_subset1 = result(); |
|
167 |
||
| 5076 | 168 |
Goal "parts(G Un H) <= parts(G) Un parts(H)"; |
| 2032 | 169 |
by (rtac subsetI 1); |
170 |
by (etac parts.induct 1); |
|
| 2891 | 171 |
by (ALLGOALS Blast_tac); |
| 1839 | 172 |
val parts_Un_subset2 = result(); |
173 |
||
| 5076 | 174 |
Goal "parts(G Un H) = parts(G) Un parts(H)"; |
| 1839 | 175 |
by (REPEAT (ares_tac [equalityI, parts_Un_subset1, parts_Un_subset2] 1)); |
176 |
qed "parts_Un"; |
|
177 |
||
| 5076 | 178 |
Goal "parts (insert X H) = parts {X} Un parts H";
|
| 1852 | 179 |
by (stac (read_instantiate [("A","H")] insert_is_Un) 1);
|
| 2011 | 180 |
by (simp_tac (HOL_ss addsimps [parts_Un]) 1); |
181 |
qed "parts_insert"; |
|
182 |
||
183 |
(*TWO inserts to avoid looping. This rewrite is better than nothing. |
|
184 |
Not suitable for Addsimps: its behaviour can be strange.*) |
|
| 5076 | 185 |
Goal "parts (insert X (insert Y H)) = parts {X} Un parts {Y} Un parts H";
|
| 4091 | 186 |
by (simp_tac (simpset() addsimps [Un_assoc]) 1); |
187 |
by (simp_tac (simpset() addsimps [parts_insert RS sym]) 1); |
|
| 1852 | 188 |
qed "parts_insert2"; |
189 |
||
| 5076 | 190 |
Goal "(UN x:A. parts(H x)) <= parts(UN x:A. H x)"; |
| 1839 | 191 |
by (REPEAT (ares_tac [UN_least, parts_mono, UN_upper] 1)); |
192 |
val parts_UN_subset1 = result(); |
|
193 |
||
| 5076 | 194 |
Goal "parts(UN x:A. H x) <= (UN x:A. parts(H x))"; |
| 2032 | 195 |
by (rtac subsetI 1); |
196 |
by (etac parts.induct 1); |
|
| 2891 | 197 |
by (ALLGOALS Blast_tac); |
| 1839 | 198 |
val parts_UN_subset2 = result(); |
199 |
||
| 5076 | 200 |
Goal "parts(UN x:A. H x) = (UN x:A. parts(H x))"; |
| 1839 | 201 |
by (REPEAT (ares_tac [equalityI, parts_UN_subset1, parts_UN_subset2] 1)); |
202 |
qed "parts_UN"; |
|
203 |
||
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
204 |
(*Added to simplify arguments to parts, analz and synth. |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
205 |
NOTE: the UN versions are no longer used!*) |
| 4157 | 206 |
Addsimps [parts_Un, parts_UN]; |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
207 |
AddSEs [parts_Un RS equalityD1 RS subsetD RS UnE, |
| 4157 | 208 |
parts_UN RS equalityD1 RS subsetD RS UN_E]; |
| 1839 | 209 |
|
| 5076 | 210 |
Goal "insert X (parts H) <= parts(insert X H)"; |
| 4091 | 211 |
by (blast_tac (claset() addIs [impOfSubs parts_mono]) 1); |
| 1839 | 212 |
qed "parts_insert_subset"; |
213 |
||
214 |
(** Idempotence and transitivity **) |
|
215 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
216 |
Goal "X: parts (parts H) ==> X: parts H"; |
| 2032 | 217 |
by (etac parts.induct 1); |
| 2891 | 218 |
by (ALLGOALS Blast_tac); |
| 2922 | 219 |
qed "parts_partsD"; |
220 |
AddSDs [parts_partsD]; |
|
| 1839 | 221 |
|
| 5076 | 222 |
Goal "parts (parts H) = parts H"; |
| 2891 | 223 |
by (Blast_tac 1); |
| 1839 | 224 |
qed "parts_idem"; |
225 |
Addsimps [parts_idem]; |
|
226 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
227 |
Goal "[| X: parts G; G <= parts H |] ==> X: parts H"; |
| 1839 | 228 |
by (dtac parts_mono 1); |
| 2891 | 229 |
by (Blast_tac 1); |
| 1839 | 230 |
qed "parts_trans"; |
231 |
||
232 |
(*Cut*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
233 |
Goal "[| Y: parts (insert X G); X: parts H |] \ |
| 2373 | 234 |
\ ==> Y: parts (G Un H)"; |
| 2032 | 235 |
by (etac parts_trans 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
236 |
by Auto_tac; |
| 1839 | 237 |
qed "parts_cut"; |
238 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
239 |
Goal "X: parts H ==> parts (insert X H) = parts H"; |
| 4091 | 240 |
by (fast_tac (claset() addSDs [parts_cut] |
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
241 |
addIs [parts_insertI] |
|
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
242 |
addss (simpset())) 1); |
|
1929
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
243 |
qed "parts_cut_eq"; |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
244 |
|
|
2028
738bb98d65ec
Last working version prior to addition of "lost" component
paulson
parents:
2026
diff
changeset
|
245 |
Addsimps [parts_cut_eq]; |
|
738bb98d65ec
Last working version prior to addition of "lost" component
paulson
parents:
2026
diff
changeset
|
246 |
|
| 1839 | 247 |
|
248 |
(** Rewrite rules for pulling out atomic messages **) |
|
249 |
||
| 2373 | 250 |
fun parts_tac i = |
251 |
EVERY [rtac ([subsetI, parts_insert_subset] MRS equalityI) i, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
252 |
etac parts.induct i, |
| 3102 | 253 |
REPEAT (Blast_tac i)]; |
| 2373 | 254 |
|
| 5076 | 255 |
Goal "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"; |
| 2373 | 256 |
by (parts_tac 1); |
| 1839 | 257 |
qed "parts_insert_Agent"; |
258 |
||
| 5076 | 259 |
Goal "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"; |
| 2373 | 260 |
by (parts_tac 1); |
| 1839 | 261 |
qed "parts_insert_Nonce"; |
262 |
||
| 5076 | 263 |
Goal "parts (insert (Number N) H) = insert (Number N) (parts H)"; |
| 3668 | 264 |
by (parts_tac 1); |
265 |
qed "parts_insert_Number"; |
|
266 |
||
| 5076 | 267 |
Goal "parts (insert (Key K) H) = insert (Key K) (parts H)"; |
| 2373 | 268 |
by (parts_tac 1); |
| 1839 | 269 |
qed "parts_insert_Key"; |
270 |
||
| 5076 | 271 |
Goal "parts (insert (Hash X) H) = insert (Hash X) (parts H)"; |
| 2373 | 272 |
by (parts_tac 1); |
273 |
qed "parts_insert_Hash"; |
|
274 |
||
| 5076 | 275 |
Goal "parts (insert (Crypt K X) H) = \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
276 |
\ insert (Crypt K X) (parts (insert X H))"; |
| 2032 | 277 |
by (rtac equalityI 1); |
278 |
by (rtac subsetI 1); |
|
279 |
by (etac parts.induct 1); |
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
280 |
by Auto_tac; |
| 2032 | 281 |
by (etac parts.induct 1); |
| 4091 | 282 |
by (ALLGOALS (blast_tac (claset() addIs [parts.Body]))); |
| 1839 | 283 |
qed "parts_insert_Crypt"; |
284 |
||
| 5076 | 285 |
Goal "parts (insert {|X,Y|} H) = \
|
| 1839 | 286 |
\ insert {|X,Y|} (parts (insert X (insert Y H)))";
|
| 2032 | 287 |
by (rtac equalityI 1); |
288 |
by (rtac subsetI 1); |
|
289 |
by (etac parts.induct 1); |
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
290 |
by Auto_tac; |
| 2032 | 291 |
by (etac parts.induct 1); |
| 4091 | 292 |
by (ALLGOALS (blast_tac (claset() addIs [parts.Fst, parts.Snd]))); |
| 1839 | 293 |
qed "parts_insert_MPair"; |
294 |
||
| 3668 | 295 |
Addsimps [parts_insert_Agent, parts_insert_Nonce, |
296 |
parts_insert_Number, parts_insert_Key, |
|
| 2373 | 297 |
parts_insert_Hash, parts_insert_Crypt, parts_insert_MPair]; |
| 1839 | 298 |
|
299 |
||
| 5076 | 300 |
Goal "parts (Key``N) = Key``N"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
301 |
by Auto_tac; |
| 2032 | 302 |
by (etac parts.induct 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
303 |
by Auto_tac; |
|
2026
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
304 |
qed "parts_image_Key"; |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
305 |
Addsimps [parts_image_Key]; |
|
2026
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
306 |
|
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
307 |
|
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
308 |
(*In any message, there is an upper bound N on its greatest nonce.*) |
| 5076 | 309 |
Goal "EX N. ALL n. N<=n --> Nonce n ~: parts {msg}";
|
| 3668 | 310 |
by (induct_tac "msg" 1); |
311 |
by (induct_tac "atomic" 1); |
|
| 4091 | 312 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [exI, parts_insert2]))); |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
313 |
(*MPair case: blast_tac works out the necessary sum itself!*) |
| 4091 | 314 |
by (blast_tac (claset() addSEs [add_leE]) 2); |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
315 |
(*Nonce case*) |
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
316 |
by (res_inst_tac [("x","N + Suc nat")] exI 1);
|
| 5983 | 317 |
by (auto_tac (claset() addSEs [add_leE], simpset())); |
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
318 |
qed "msg_Nonce_supply"; |
|
2026
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
319 |
|
|
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
320 |
|
| 1913 | 321 |
(**** Inductive relation "analz" ****) |
| 1839 | 322 |
|
323 |
val major::prems = |
|
| 5076 | 324 |
Goal "[| {|X,Y|} : analz H; \
|
| 1913 | 325 |
\ [| X : analz H; Y : analz H |] ==> P \ |
| 1839 | 326 |
\ |] ==> P"; |
327 |
by (cut_facts_tac [major] 1); |
|
| 2032 | 328 |
by (resolve_tac prems 1); |
| 1913 | 329 |
by (REPEAT (eresolve_tac [asm_rl, analz.Fst, analz.Snd] 1)); |
330 |
qed "MPair_analz"; |
|
| 1839 | 331 |
|
| 1913 | 332 |
AddIs [analz.Inj]; |
| 2011 | 333 |
AddSEs [MPair_analz]; (*Perhaps it should NOT be deemed safe!*) |
| 1913 | 334 |
AddDs [analz.Decrypt]; |
| 1839 | 335 |
|
| 1913 | 336 |
Addsimps [analz.Inj]; |
| 1885 | 337 |
|
| 5076 | 338 |
Goal "H <= analz(H)"; |
| 2891 | 339 |
by (Blast_tac 1); |
| 1913 | 340 |
qed "analz_increasing"; |
| 1839 | 341 |
|
| 5076 | 342 |
Goal "analz H <= parts H"; |
| 1839 | 343 |
by (rtac subsetI 1); |
| 2032 | 344 |
by (etac analz.induct 1); |
| 2891 | 345 |
by (ALLGOALS Blast_tac); |
| 1913 | 346 |
qed "analz_subset_parts"; |
| 1839 | 347 |
|
| 1913 | 348 |
bind_thm ("not_parts_not_analz", analz_subset_parts RS contra_subsetD);
|
| 1839 | 349 |
|
350 |
||
| 5076 | 351 |
Goal "parts (analz H) = parts H"; |
| 2032 | 352 |
by (rtac equalityI 1); |
353 |
by (rtac (analz_subset_parts RS parts_mono RS subset_trans) 1); |
|
| 1839 | 354 |
by (Simp_tac 1); |
| 4091 | 355 |
by (blast_tac (claset() addIs [analz_increasing RS parts_mono RS subsetD]) 1); |
| 1913 | 356 |
qed "parts_analz"; |
357 |
Addsimps [parts_analz]; |
|
| 1839 | 358 |
|
| 5076 | 359 |
Goal "analz (parts H) = parts H"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
360 |
by Auto_tac; |
| 2032 | 361 |
by (etac analz.induct 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
362 |
by Auto_tac; |
| 1913 | 363 |
qed "analz_parts"; |
364 |
Addsimps [analz_parts]; |
|
| 1885 | 365 |
|
| 1839 | 366 |
(*Monotonicity; Lemma 1 of Lowe*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
367 |
Goalw analz.defs "G<=H ==> analz(G) <= analz(H)"; |
| 1839 | 368 |
by (rtac lfp_mono 1); |
369 |
by (REPEAT (ares_tac basic_monos 1)); |
|
| 1913 | 370 |
qed "analz_mono"; |
| 1839 | 371 |
|
| 2373 | 372 |
val analz_insertI = impOfSubs (subset_insertI RS analz_mono); |
373 |
||
| 1839 | 374 |
(** General equational properties **) |
375 |
||
| 5076 | 376 |
Goal "analz{} = {}";
|
| 3730 | 377 |
by Safe_tac; |
| 2032 | 378 |
by (etac analz.induct 1); |
| 2891 | 379 |
by (ALLGOALS Blast_tac); |
| 1913 | 380 |
qed "analz_empty"; |
381 |
Addsimps [analz_empty]; |
|
| 1839 | 382 |
|
| 1913 | 383 |
(*Converse fails: we can analz more from the union than from the |
| 1839 | 384 |
separate parts, as a key in one might decrypt a message in the other*) |
| 5076 | 385 |
Goal "analz(G) Un analz(H) <= analz(G Un H)"; |
| 1913 | 386 |
by (REPEAT (ares_tac [Un_least, analz_mono, Un_upper1, Un_upper2] 1)); |
387 |
qed "analz_Un"; |
|
| 1839 | 388 |
|
| 5076 | 389 |
Goal "insert X (analz H) <= analz(insert X H)"; |
| 4091 | 390 |
by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1); |
| 1913 | 391 |
qed "analz_insert"; |
| 1839 | 392 |
|
393 |
(** Rewrite rules for pulling out atomic messages **) |
|
394 |
||
| 2373 | 395 |
fun analz_tac i = |
396 |
EVERY [rtac ([subsetI, analz_insert] MRS equalityI) i, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
397 |
etac analz.induct i, |
| 3102 | 398 |
REPEAT (Blast_tac i)]; |
| 2373 | 399 |
|
| 5076 | 400 |
Goal "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"; |
| 2373 | 401 |
by (analz_tac 1); |
| 1913 | 402 |
qed "analz_insert_Agent"; |
| 1839 | 403 |
|
| 5076 | 404 |
Goal "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"; |
| 2373 | 405 |
by (analz_tac 1); |
| 1913 | 406 |
qed "analz_insert_Nonce"; |
| 1839 | 407 |
|
| 5076 | 408 |
Goal "analz (insert (Number N) H) = insert (Number N) (analz H)"; |
| 3668 | 409 |
by (analz_tac 1); |
410 |
qed "analz_insert_Number"; |
|
411 |
||
| 5076 | 412 |
Goal "analz (insert (Hash X) H) = insert (Hash X) (analz H)"; |
| 2373 | 413 |
by (analz_tac 1); |
414 |
qed "analz_insert_Hash"; |
|
415 |
||
| 1839 | 416 |
(*Can only pull out Keys if they are not needed to decrypt the rest*) |
| 5076 | 417 |
Goalw [keysFor_def] |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
418 |
"K ~: keysFor (analz H) ==> \ |
| 1913 | 419 |
\ analz (insert (Key K) H) = insert (Key K) (analz H)"; |
| 2373 | 420 |
by (analz_tac 1); |
| 1913 | 421 |
qed "analz_insert_Key"; |
| 1839 | 422 |
|
| 5076 | 423 |
Goal "analz (insert {|X,Y|} H) = \
|
| 1913 | 424 |
\ insert {|X,Y|} (analz (insert X (insert Y H)))";
|
| 2032 | 425 |
by (rtac equalityI 1); |
426 |
by (rtac subsetI 1); |
|
427 |
by (etac analz.induct 1); |
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
428 |
by Auto_tac; |
| 2032 | 429 |
by (etac analz.induct 1); |
| 4091 | 430 |
by (ALLGOALS (blast_tac (claset() addIs [analz.Fst, analz.Snd]))); |
| 1913 | 431 |
qed "analz_insert_MPair"; |
| 1885 | 432 |
|
433 |
(*Can pull out enCrypted message if the Key is not known*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
434 |
Goal "Key (invKey K) ~: analz H ==> \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
435 |
\ analz (insert (Crypt K X) H) = \ |
|
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
436 |
\ insert (Crypt K X) (analz H)"; |
| 2373 | 437 |
by (analz_tac 1); |
| 1913 | 438 |
qed "analz_insert_Crypt"; |
| 1839 | 439 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
440 |
Goal "Key (invKey K) : analz H ==> \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
441 |
\ analz (insert (Crypt K X) H) <= \ |
|
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
442 |
\ insert (Crypt K X) (analz (insert X H))"; |
| 2032 | 443 |
by (rtac subsetI 1); |
| 5102 | 444 |
by (eres_inst_tac [("xa","x")] analz.induct 1);
|
| 3102 | 445 |
by (ALLGOALS (Blast_tac)); |
| 1839 | 446 |
val lemma1 = result(); |
447 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
448 |
Goal "Key (invKey K) : analz H ==> \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
449 |
\ insert (Crypt K X) (analz (insert X H)) <= \ |
|
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
450 |
\ analz (insert (Crypt K X) H)"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
451 |
by Auto_tac; |
| 5102 | 452 |
by (eres_inst_tac [("xa","x")] analz.induct 1);
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
453 |
by Auto_tac; |
| 4091 | 454 |
by (blast_tac (claset() addIs [analz_insertI, analz.Decrypt]) 1); |
| 1839 | 455 |
val lemma2 = result(); |
456 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
457 |
Goal "Key (invKey K) : analz H ==> \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
458 |
\ analz (insert (Crypt K X) H) = \ |
|
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
459 |
\ insert (Crypt K X) (analz (insert X H))"; |
| 1839 | 460 |
by (REPEAT (ares_tac [equalityI, lemma1, lemma2] 1)); |
| 1913 | 461 |
qed "analz_insert_Decrypt"; |
| 1839 | 462 |
|
| 1885 | 463 |
(*Case analysis: either the message is secure, or it is not! |
| 1946 | 464 |
Effective, but can cause subgoals to blow up! |
| 5055 | 465 |
Use with split_if; apparently split_tac does not cope with patterns |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
466 |
such as "analz (insert (Crypt K X) H)" *) |
| 5076 | 467 |
Goal "analz (insert (Crypt K X) H) = \ |
| 2154 | 468 |
\ (if (Key (invKey K) : analz H) \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
469 |
\ then insert (Crypt K X) (analz (insert X H)) \ |
|
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
470 |
\ else insert (Crypt K X) (analz H))"; |
| 2102 | 471 |
by (case_tac "Key (invKey K) : analz H " 1); |
| 4091 | 472 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [analz_insert_Crypt, |
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
473 |
analz_insert_Decrypt]))); |
| 1913 | 474 |
qed "analz_Crypt_if"; |
| 1885 | 475 |
|
| 3668 | 476 |
Addsimps [analz_insert_Agent, analz_insert_Nonce, |
477 |
analz_insert_Number, analz_insert_Key, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
478 |
analz_insert_Hash, analz_insert_MPair, analz_Crypt_if]; |
| 1839 | 479 |
|
480 |
(*This rule supposes "for the sake of argument" that we have the key.*) |
|
| 5076 | 481 |
Goal "analz (insert (Crypt K X) H) <= \ |
|
2284
80ebd1a213fd
Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents:
2170
diff
changeset
|
482 |
\ insert (Crypt K X) (analz (insert X H))"; |
| 2032 | 483 |
by (rtac subsetI 1); |
484 |
by (etac analz.induct 1); |
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
485 |
by Auto_tac; |
| 1913 | 486 |
qed "analz_insert_Crypt_subset"; |
| 1839 | 487 |
|
488 |
||
| 5076 | 489 |
Goal "analz (Key``N) = Key``N"; |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
490 |
by Auto_tac; |
| 2032 | 491 |
by (etac analz.induct 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
492 |
by Auto_tac; |
|
2026
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
493 |
qed "analz_image_Key"; |
|
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
494 |
|
|
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
495 |
Addsimps [analz_image_Key]; |
|
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
496 |
|
|
0df5a96bf77e
Last working version prior to introduction of "lost"
paulson
parents:
2011
diff
changeset
|
497 |
|
| 1839 | 498 |
(** Idempotence and transitivity **) |
499 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
500 |
Goal "X: analz (analz H) ==> X: analz H"; |
| 2032 | 501 |
by (etac analz.induct 1); |
| 2891 | 502 |
by (ALLGOALS Blast_tac); |
| 2922 | 503 |
qed "analz_analzD"; |
504 |
AddSDs [analz_analzD]; |
|
| 1839 | 505 |
|
| 5076 | 506 |
Goal "analz (analz H) = analz H"; |
| 2891 | 507 |
by (Blast_tac 1); |
| 1913 | 508 |
qed "analz_idem"; |
509 |
Addsimps [analz_idem]; |
|
| 1839 | 510 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
511 |
Goal "[| X: analz G; G <= analz H |] ==> X: analz H"; |
| 1913 | 512 |
by (dtac analz_mono 1); |
| 2891 | 513 |
by (Blast_tac 1); |
| 1913 | 514 |
qed "analz_trans"; |
| 1839 | 515 |
|
516 |
(*Cut; Lemma 2 of Lowe*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
517 |
Goal "[| Y: analz (insert X H); X: analz H |] ==> Y: analz H"; |
| 2032 | 518 |
by (etac analz_trans 1); |
| 2891 | 519 |
by (Blast_tac 1); |
| 1913 | 520 |
qed "analz_cut"; |
| 1839 | 521 |
|
522 |
(*Cut can be proved easily by induction on |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
523 |
"Y: analz (insert X H) ==> X: analz H --> Y: analz H" |
| 1839 | 524 |
*) |
525 |
||
| 3449 | 526 |
(*This rewrite rule helps in the simplification of messages that involve |
527 |
the forwarding of unknown components (X). Without it, removing occurrences |
|
528 |
of X can be very complicated. *) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
529 |
Goal "X: analz H ==> analz (insert X H) = analz H"; |
| 4091 | 530 |
by (blast_tac (claset() addIs [analz_cut, analz_insertI]) 1); |
| 3431 | 531 |
qed "analz_insert_eq"; |
532 |
||
| 1885 | 533 |
|
| 1913 | 534 |
(** A congruence rule for "analz" **) |
| 1885 | 535 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
536 |
Goal "[| analz G <= analz G'; analz H <= analz H' \ |
| 1913 | 537 |
\ |] ==> analz (G Un H) <= analz (G' Un H')"; |
| 3714 | 538 |
by (Clarify_tac 1); |
| 2032 | 539 |
by (etac analz.induct 1); |
| 4091 | 540 |
by (ALLGOALS (best_tac (claset() addIs [analz_mono RS subsetD]))); |
| 1913 | 541 |
qed "analz_subset_cong"; |
| 1885 | 542 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
543 |
Goal "[| analz G = analz G'; analz H = analz H' \ |
| 1913 | 544 |
\ |] ==> analz (G Un H) = analz (G' Un H')"; |
545 |
by (REPEAT_FIRST (ares_tac [equalityI, analz_subset_cong] |
|
| 2032 | 546 |
ORELSE' etac equalityE)); |
| 1913 | 547 |
qed "analz_cong"; |
| 1885 | 548 |
|
549 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
550 |
Goal "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"; |
| 4091 | 551 |
by (asm_simp_tac (simpset() addsimps [insert_def] delsimps [singleton_conv] |
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
552 |
setloop (rtac analz_cong)) 1); |
| 1913 | 553 |
qed "analz_insert_cong"; |
| 1885 | 554 |
|
| 1913 | 555 |
(*If there are no pairs or encryptions then analz does nothing*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
556 |
Goal "[| ALL X Y. {|X,Y|} ~: H; ALL X K. Crypt K X ~: H |] ==> \
|
| 1913 | 557 |
\ analz H = H"; |
| 3730 | 558 |
by Safe_tac; |
| 2032 | 559 |
by (etac analz.induct 1); |
| 2891 | 560 |
by (ALLGOALS Blast_tac); |
| 1913 | 561 |
qed "analz_trivial"; |
| 1839 | 562 |
|
| 4157 | 563 |
(*These two are obsolete (with a single Spy) but cost little to prove...*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
564 |
Goal "X: analz (UN i:A. analz (H i)) ==> X: analz (UN i:A. H i)"; |
| 2032 | 565 |
by (etac analz.induct 1); |
| 4091 | 566 |
by (ALLGOALS (blast_tac (claset() addIs [impOfSubs analz_mono]))); |
| 1839 | 567 |
val lemma = result(); |
568 |
||
| 5076 | 569 |
Goal "analz (UN i:A. analz (H i)) = analz (UN i:A. H i)"; |
| 4091 | 570 |
by (blast_tac (claset() addIs [lemma, impOfSubs analz_mono]) 1); |
| 1913 | 571 |
qed "analz_UN_analz"; |
572 |
Addsimps [analz_UN_analz]; |
|
| 1839 | 573 |
|
574 |
||
| 1913 | 575 |
(**** Inductive relation "synth" ****) |
| 1839 | 576 |
|
| 1913 | 577 |
AddIs synth.intrs; |
| 1839 | 578 |
|
| 3668 | 579 |
(*NO Agent_synth, as any Agent name can be synthesized. Ditto for Number*) |
| 6141 | 580 |
val Nonce_synth = synth.mk_cases "Nonce n : synth H"; |
581 |
val Key_synth = synth.mk_cases "Key K : synth H"; |
|
582 |
val Hash_synth = synth.mk_cases "Hash X : synth H"; |
|
583 |
val MPair_synth = synth.mk_cases "{|X,Y|} : synth H";
|
|
584 |
val Crypt_synth = synth.mk_cases "Crypt K X : synth H"; |
|
| 2011 | 585 |
|
| 2373 | 586 |
AddSEs [Nonce_synth, Key_synth, Hash_synth, MPair_synth, Crypt_synth]; |
| 2011 | 587 |
|
| 5076 | 588 |
Goal "H <= synth(H)"; |
| 2891 | 589 |
by (Blast_tac 1); |
| 1913 | 590 |
qed "synth_increasing"; |
| 1839 | 591 |
|
592 |
(*Monotonicity*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
593 |
Goalw synth.defs "G<=H ==> synth(G) <= synth(H)"; |
| 1839 | 594 |
by (rtac lfp_mono 1); |
595 |
by (REPEAT (ares_tac basic_monos 1)); |
|
| 1913 | 596 |
qed "synth_mono"; |
| 1839 | 597 |
|
598 |
(** Unions **) |
|
599 |
||
| 1913 | 600 |
(*Converse fails: we can synth more from the union than from the |
| 1839 | 601 |
separate parts, building a compound message using elements of each.*) |
| 5076 | 602 |
Goal "synth(G) Un synth(H) <= synth(G Un H)"; |
| 1913 | 603 |
by (REPEAT (ares_tac [Un_least, synth_mono, Un_upper1, Un_upper2] 1)); |
604 |
qed "synth_Un"; |
|
| 1839 | 605 |
|
| 5076 | 606 |
Goal "insert X (synth H) <= synth(insert X H)"; |
| 4091 | 607 |
by (blast_tac (claset() addIs [impOfSubs synth_mono]) 1); |
| 1913 | 608 |
qed "synth_insert"; |
| 1885 | 609 |
|
| 1839 | 610 |
(** Idempotence and transitivity **) |
611 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
612 |
Goal "X: synth (synth H) ==> X: synth H"; |
| 2032 | 613 |
by (etac synth.induct 1); |
| 2891 | 614 |
by (ALLGOALS Blast_tac); |
| 2922 | 615 |
qed "synth_synthD"; |
616 |
AddSDs [synth_synthD]; |
|
| 1839 | 617 |
|
| 5076 | 618 |
Goal "synth (synth H) = synth H"; |
| 2891 | 619 |
by (Blast_tac 1); |
| 1913 | 620 |
qed "synth_idem"; |
| 1839 | 621 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
622 |
Goal "[| X: synth G; G <= synth H |] ==> X: synth H"; |
| 1913 | 623 |
by (dtac synth_mono 1); |
| 2891 | 624 |
by (Blast_tac 1); |
| 1913 | 625 |
qed "synth_trans"; |
| 1839 | 626 |
|
627 |
(*Cut; Lemma 2 of Lowe*) |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
628 |
Goal "[| Y: synth (insert X H); X: synth H |] ==> Y: synth H"; |
| 2032 | 629 |
by (etac synth_trans 1); |
| 2891 | 630 |
by (Blast_tac 1); |
| 1913 | 631 |
qed "synth_cut"; |
| 1839 | 632 |
|
| 5076 | 633 |
Goal "Agent A : synth H"; |
| 2891 | 634 |
by (Blast_tac 1); |
| 1946 | 635 |
qed "Agent_synth"; |
636 |
||
| 5076 | 637 |
Goal "Number n : synth H"; |
| 3668 | 638 |
by (Blast_tac 1); |
639 |
qed "Number_synth"; |
|
640 |
||
| 5076 | 641 |
Goal "(Nonce N : synth H) = (Nonce N : H)"; |
| 2891 | 642 |
by (Blast_tac 1); |
| 1913 | 643 |
qed "Nonce_synth_eq"; |
| 1839 | 644 |
|
| 5076 | 645 |
Goal "(Key K : synth H) = (Key K : H)"; |
| 2891 | 646 |
by (Blast_tac 1); |
| 1913 | 647 |
qed "Key_synth_eq"; |
| 1839 | 648 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
649 |
Goal "Key K ~: H ==> (Crypt K X : synth H) = (Crypt K X : H)"; |
| 2891 | 650 |
by (Blast_tac 1); |
| 2011 | 651 |
qed "Crypt_synth_eq"; |
652 |
||
| 3668 | 653 |
Addsimps [Agent_synth, Number_synth, |
654 |
Nonce_synth_eq, Key_synth_eq, Crypt_synth_eq]; |
|
| 1839 | 655 |
|
656 |
||
| 5076 | 657 |
Goalw [keysFor_def] |
| 1913 | 658 |
"keysFor (synth H) = keysFor H Un invKey``{K. Key K : H}";
|
| 2891 | 659 |
by (Blast_tac 1); |
| 1913 | 660 |
qed "keysFor_synth"; |
661 |
Addsimps [keysFor_synth]; |
|
| 1839 | 662 |
|
663 |
||
| 1913 | 664 |
(*** Combinations of parts, analz and synth ***) |
| 1839 | 665 |
|
| 5076 | 666 |
Goal "parts (synth H) = parts H Un synth H"; |
| 2032 | 667 |
by (rtac equalityI 1); |
668 |
by (rtac subsetI 1); |
|
669 |
by (etac parts.induct 1); |
|
| 1839 | 670 |
by (ALLGOALS |
| 4091 | 671 |
(blast_tac (claset() addIs ((synth_increasing RS parts_mono RS subsetD) |
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
672 |
::parts.intrs)))); |
| 1913 | 673 |
qed "parts_synth"; |
674 |
Addsimps [parts_synth]; |
|
| 1839 | 675 |
|
| 5076 | 676 |
Goal "analz (analz G Un H) = analz (G Un H)"; |
| 2373 | 677 |
by (REPEAT_FIRST (resolve_tac [equalityI, analz_subset_cong])); |
678 |
by (ALLGOALS Simp_tac); |
|
679 |
qed "analz_analz_Un"; |
|
680 |
||
| 5076 | 681 |
Goal "analz (synth G Un H) = analz (G Un H) Un synth G"; |
| 2032 | 682 |
by (rtac equalityI 1); |
683 |
by (rtac subsetI 1); |
|
684 |
by (etac analz.induct 1); |
|
| 4091 | 685 |
by (blast_tac (claset() addIs [impOfSubs analz_mono]) 5); |
686 |
by (ALLGOALS (blast_tac (claset() addIs analz.intrs))); |
|
| 2373 | 687 |
qed "analz_synth_Un"; |
688 |
||
| 5076 | 689 |
Goal "analz (synth H) = analz H Un synth H"; |
| 2373 | 690 |
by (cut_inst_tac [("H","{}")] analz_synth_Un 1);
|
691 |
by (Full_simp_tac 1); |
|
| 1913 | 692 |
qed "analz_synth"; |
| 2373 | 693 |
Addsimps [analz_analz_Un, analz_synth_Un, analz_synth]; |
| 1839 | 694 |
|
| 1946 | 695 |
|
696 |
(** For reasoning about the Fake rule in traces **) |
|
697 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
698 |
Goal "X: G ==> parts(insert X H) <= parts G Un parts H"; |
| 2032 | 699 |
by (rtac ([parts_mono, parts_Un_subset2] MRS subset_trans) 1); |
| 2891 | 700 |
by (Blast_tac 1); |
|
1929
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
701 |
qed "parts_insert_subset_Un"; |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
702 |
|
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
703 |
(*More specifically for Fake. Very occasionally we could do with a version |
|
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
704 |
of the form parts{X} <= synth (analz H) Un parts H *)
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
705 |
Goal "X: synth (analz H) ==> \ |
|
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
706 |
\ parts (insert X H) <= synth (analz H) Un parts H"; |
| 2032 | 707 |
by (dtac parts_insert_subset_Un 1); |
| 1946 | 708 |
by (Full_simp_tac 1); |
| 2891 | 709 |
by (Blast_tac 1); |
| 1946 | 710 |
qed "Fake_parts_insert"; |
711 |
||
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
712 |
(*H is sometimes (Key `` KK Un spies evs), so can't put G=H*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
713 |
Goal "X: synth (analz G) ==> \ |
|
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
714 |
\ analz (insert X H) <= synth (analz G) Un analz (G Un H)"; |
| 2373 | 715 |
by (rtac subsetI 1); |
716 |
by (subgoal_tac "x : analz (synth (analz G) Un H)" 1); |
|
| 4091 | 717 |
by (blast_tac (claset() addIs [impOfSubs analz_mono, |
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
718 |
impOfSubs (analz_mono RS synth_mono)]) 2); |
| 2373 | 719 |
by (Full_simp_tac 1); |
| 2891 | 720 |
by (Blast_tac 1); |
| 2373 | 721 |
qed "Fake_analz_insert"; |
722 |
||
| 5076 | 723 |
Goal "(X: analz H & X: parts H) = (X: analz H)"; |
| 4091 | 724 |
by (blast_tac (claset() addIs [impOfSubs analz_subset_parts]) 1); |
| 2011 | 725 |
val analz_conj_parts = result(); |
726 |
||
| 5076 | 727 |
Goal "(X: analz H | X: parts H) = (X: parts H)"; |
| 4091 | 728 |
by (blast_tac (claset() addIs [impOfSubs analz_subset_parts]) 1); |
| 2011 | 729 |
val analz_disj_parts = result(); |
730 |
||
731 |
AddIffs [analz_conj_parts, analz_disj_parts]; |
|
732 |
||
|
1998
f8230821f1e8
Reordering of premises for cut theorems, and new law MPair_synth_analz
paulson
parents:
1994
diff
changeset
|
733 |
(*Without this equation, other rules for synth and analz would yield |
|
f8230821f1e8
Reordering of premises for cut theorems, and new law MPair_synth_analz
paulson
parents:
1994
diff
changeset
|
734 |
redundant cases*) |
| 5076 | 735 |
Goal "({|X,Y|} : synth (analz H)) = \
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
736 |
\ (X : synth (analz H) & Y : synth (analz H))"; |
| 2891 | 737 |
by (Blast_tac 1); |
|
1998
f8230821f1e8
Reordering of premises for cut theorems, and new law MPair_synth_analz
paulson
parents:
1994
diff
changeset
|
738 |
qed "MPair_synth_analz"; |
|
f8230821f1e8
Reordering of premises for cut theorems, and new law MPair_synth_analz
paulson
parents:
1994
diff
changeset
|
739 |
|
|
f8230821f1e8
Reordering of premises for cut theorems, and new law MPair_synth_analz
paulson
parents:
1994
diff
changeset
|
740 |
AddIffs [MPair_synth_analz]; |
|
1929
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
741 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
742 |
Goal "[| Key K : analz H; Key (invKey K) : analz H |] \ |
|
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
743 |
\ ==> (Crypt K X : synth (analz H)) = (X : synth (analz H))"; |
| 2891 | 744 |
by (Blast_tac 1); |
| 2154 | 745 |
qed "Crypt_synth_analz"; |
746 |
||
|
1929
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
747 |
|
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
748 |
Goal "X ~: synth (analz H) \ |
|
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
749 |
\ ==> (Hash{|X,Y|} : synth (analz H)) = (Hash{|X,Y|} : analz H)";
|
| 2891 | 750 |
by (Blast_tac 1); |
| 2373 | 751 |
qed "Hash_synth_analz"; |
752 |
Addsimps [Hash_synth_analz]; |
|
753 |
||
754 |
||
| 2484 | 755 |
(**** HPair: a combination of Hash and MPair ****) |
756 |
||
757 |
(*** Freeness ***) |
|
758 |
||
| 5076 | 759 |
Goalw [HPair_def] "Agent A ~= Hash[X] Y"; |
| 2484 | 760 |
by (Simp_tac 1); |
761 |
qed "Agent_neq_HPair"; |
|
762 |
||
| 5076 | 763 |
Goalw [HPair_def] "Nonce N ~= Hash[X] Y"; |
| 2484 | 764 |
by (Simp_tac 1); |
765 |
qed "Nonce_neq_HPair"; |
|
766 |
||
| 5076 | 767 |
Goalw [HPair_def] "Number N ~= Hash[X] Y"; |
| 3668 | 768 |
by (Simp_tac 1); |
769 |
qed "Number_neq_HPair"; |
|
770 |
||
| 5076 | 771 |
Goalw [HPair_def] "Key K ~= Hash[X] Y"; |
| 2484 | 772 |
by (Simp_tac 1); |
773 |
qed "Key_neq_HPair"; |
|
774 |
||
| 5076 | 775 |
Goalw [HPair_def] "Hash Z ~= Hash[X] Y"; |
| 2484 | 776 |
by (Simp_tac 1); |
777 |
qed "Hash_neq_HPair"; |
|
778 |
||
| 5076 | 779 |
Goalw [HPair_def] "Crypt K X' ~= Hash[X] Y"; |
| 2484 | 780 |
by (Simp_tac 1); |
781 |
qed "Crypt_neq_HPair"; |
|
782 |
||
| 3668 | 783 |
val HPair_neqs = [Agent_neq_HPair, Nonce_neq_HPair, Number_neq_HPair, |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
784 |
Key_neq_HPair, Hash_neq_HPair, Crypt_neq_HPair]; |
| 2484 | 785 |
|
786 |
AddIffs HPair_neqs; |
|
787 |
AddIffs (HPair_neqs RL [not_sym]); |
|
788 |
||
| 5076 | 789 |
Goalw [HPair_def] "(Hash[X'] Y' = Hash[X] Y) = (X' = X & Y'=Y)"; |
| 2484 | 790 |
by (Simp_tac 1); |
791 |
qed "HPair_eq"; |
|
792 |
||
| 5076 | 793 |
Goalw [HPair_def] "({|X',Y'|} = Hash[X] Y) = (X' = Hash{|X,Y|} & Y'=Y)";
|
| 2484 | 794 |
by (Simp_tac 1); |
795 |
qed "MPair_eq_HPair"; |
|
796 |
||
| 5076 | 797 |
Goalw [HPair_def] "(Hash[X] Y = {|X',Y'|}) = (X' = Hash{|X,Y|} & Y'=Y)";
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4422
diff
changeset
|
798 |
by Auto_tac; |
| 2484 | 799 |
qed "HPair_eq_MPair"; |
800 |
||
801 |
AddIffs [HPair_eq, MPair_eq_HPair, HPair_eq_MPair]; |
|
802 |
||
803 |
||
804 |
(*** Specialized laws, proved in terms of those for Hash and MPair ***) |
|
805 |
||
| 5076 | 806 |
Goalw [HPair_def] "keysFor (insert (Hash[X] Y) H) = keysFor H"; |
| 2484 | 807 |
by (Simp_tac 1); |
808 |
qed "keysFor_insert_HPair"; |
|
809 |
||
| 5076 | 810 |
Goalw [HPair_def] |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
811 |
"parts (insert (Hash[X] Y) H) = \ |
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
812 |
\ insert (Hash[X] Y) (insert (Hash{|X,Y|}) (parts (insert Y H)))";
|
| 2484 | 813 |
by (Simp_tac 1); |
814 |
qed "parts_insert_HPair"; |
|
815 |
||
| 5076 | 816 |
Goalw [HPair_def] |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
817 |
"analz (insert (Hash[X] Y) H) = \ |
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
818 |
\ insert (Hash[X] Y) (insert (Hash{|X,Y|}) (analz (insert Y H)))";
|
| 2484 | 819 |
by (Simp_tac 1); |
820 |
qed "analz_insert_HPair"; |
|
821 |
||
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5102
diff
changeset
|
822 |
Goalw [HPair_def] "X ~: synth (analz H) \ |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
823 |
\ ==> (Hash[X] Y : synth (analz H)) = \ |
| 2484 | 824 |
\ (Hash {|X, Y|} : analz H & Y : synth (analz H))";
|
825 |
by (Simp_tac 1); |
|
| 2891 | 826 |
by (Blast_tac 1); |
| 2484 | 827 |
qed "HPair_synth_analz"; |
828 |
||
829 |
Addsimps [keysFor_insert_HPair, parts_insert_HPair, analz_insert_HPair, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
830 |
HPair_synth_analz, HPair_synth_analz]; |
| 2484 | 831 |
|
832 |
||
|
1929
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
833 |
(*We do NOT want Crypt... messages broken up in protocols!!*) |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
834 |
Delrules partsEs; |
|
f0839bab4b00
Working version of NS, messages 1-3, WITH INTERLEAVING
paulson
parents:
1913
diff
changeset
|
835 |
|
| 2327 | 836 |
|
837 |
(** Rewrites to push in Key and Crypt messages, so that other messages can |
|
838 |
be pulled out using the analz_insert rules **) |
|
839 |
||
840 |
fun insComm thy x y = read_instantiate_sg (sign_of thy) [("x",x), ("y",y)]
|
|
841 |
insert_commute; |
|
842 |
||
843 |
val pushKeys = map (insComm thy "Key ?K") |
|
| 3668 | 844 |
["Agent ?C", "Nonce ?N", "Number ?N", |
845 |
"Hash ?X", "MPair ?X ?Y", "Crypt ?X ?K'"]; |
|
| 2327 | 846 |
|
847 |
val pushCrypts = map (insComm thy "Crypt ?X ?K") |
|
| 3668 | 848 |
["Agent ?C", "Nonce ?N", "Number ?N", |
849 |
"Hash ?X'", "MPair ?X' ?Y"]; |
|
| 2327 | 850 |
|
851 |
(*Cannot be added with Addsimps -- we don't always want to re-order messages*) |
|
852 |
val pushes = pushKeys@pushCrypts; |
|
853 |
||
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
854 |
|
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
855 |
(*** Tactics useful for many protocol proofs ***) |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
856 |
|
| 3470 | 857 |
(*Prove base case (subgoal i) and simplify others. A typical base case |
| 3683 | 858 |
concerns Crypt K X ~: Key``shrK``bad and cannot be proved by rewriting |
| 3470 | 859 |
alone.*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
860 |
fun prove_simple_subgoals_tac i = |
| 5421 | 861 |
Force_tac i THEN ALLGOALS Asm_simp_tac; |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
862 |
|
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
863 |
fun Fake_parts_insert_tac i = |
|
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
864 |
blast_tac (claset() addIs [parts_insertI] |
|
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
865 |
addDs [impOfSubs analz_subset_parts, |
|
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
866 |
impOfSubs Fake_parts_insert]) i; |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
867 |
|
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
868 |
(*Apply rules to break down assumptions of the form |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
869 |
Y : parts(insert X H) and Y : analz(insert X H) |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
3102
diff
changeset
|
870 |
*) |
| 2373 | 871 |
val Fake_insert_tac = |
872 |
dresolve_tac [impOfSubs Fake_analz_insert, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
873 |
impOfSubs Fake_parts_insert] THEN' |
| 2373 | 874 |
eresolve_tac [asm_rl, synth.Inj]; |
875 |
||
| 3702 | 876 |
fun Fake_insert_simp_tac i = |
877 |
REPEAT (Fake_insert_tac i) THEN Asm_full_simp_tac i; |
|
878 |
||
879 |
||
| 3449 | 880 |
(*Analysis of Fake cases. Also works for messages that forward unknown parts, |
881 |
but this application is no longer necessary if analz_insert_eq is used. |
|
| 2327 | 882 |
Abstraction over i is ESSENTIAL: it delays the dereferencing of claset |
883 |
DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) |
|
|
4735
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
884 |
|
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
885 |
val atomic_spy_analz_tac = SELECT_GOAL |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
886 |
(Fake_insert_simp_tac 1 |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
887 |
THEN |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
888 |
IF_UNSOLVED (Blast.depth_tac |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
889 |
(claset() addIs [analz_insertI, |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
890 |
impOfSubs analz_subset_parts]) 4 1)); |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
891 |
|
| 2327 | 892 |
fun spy_analz_tac i = |
| 2373 | 893 |
DETERM |
894 |
(SELECT_GOAL |
|
895 |
(EVERY |
|
896 |
[ (*push in occurrences of X...*) |
|
897 |
(REPEAT o CHANGED) |
|
898 |
(res_inst_tac [("x1","X")] (insert_commute RS ssubst) 1),
|
|
899 |
(*...allowing further simplifications*) |
|
| 4686 | 900 |
Simp_tac 1, |
|
3476
1be4fee7606b
spy_analz_tac: Restored iffI to the list of rules used to break down
paulson
parents:
3470
diff
changeset
|
901 |
REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])), |
|
4735
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
902 |
DEPTH_SOLVE (atomic_spy_analz_tac 1)]) i); |
|
6d544b44a70e
spy_analz_tac now handles individual conjuncts properly
paulson
parents:
4686
diff
changeset
|
903 |
|
| 2327 | 904 |
|
| 2415 | 905 |
(** Useful in many uniqueness proofs **) |
| 2327 | 906 |
fun ex_strip_tac i = REPEAT (swap_res_tac [exI, conjI] i) THEN |
907 |
assume_tac (i+1); |
|
908 |
||
| 2415 | 909 |
(*Apply the EX-ALL quantifification to prove uniqueness theorems in |
910 |
their standard form*) |
|
911 |
fun prove_unique_tac lemma = |
|
912 |
EVERY' [dtac lemma, |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
913 |
REPEAT o (mp_tac ORELSE' eresolve_tac [asm_rl,exE]), |
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
914 |
(*Duplicate the assumption*) |
|
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2484
diff
changeset
|
915 |
forw_inst_tac [("psi", "ALL C.?P(C)")] asm_rl,
|
| 4091 | 916 |
Blast.depth_tac (claset() addSDs [spec]) 0]; |
| 2415 | 917 |
|
| 2373 | 918 |
|
919 |
(*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*) |
|
920 |
goal Set.thy "A Un (B Un A) = B Un A"; |
|
| 2891 | 921 |
by (Blast_tac 1); |
| 2373 | 922 |
val Un_absorb3 = result(); |
923 |
Addsimps [Un_absorb3]; |
|
|
3514
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
924 |
|
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
925 |
(*By default only o_apply is built-in. But in the presence of eta-expansion |
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
926 |
this means that some terms displayed as (f o g) will be rewritten, and others |
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
927 |
will not!*) |
|
eb16b8e8d872
Moved some declarations to Message from Public and Shared
paulson
parents:
3476
diff
changeset
|
928 |
Addsimps [o_def]; |