author | wenzelm |
Fri, 13 Mar 2009 19:58:26 +0100 | |
changeset 30510 | 4120fc59dd85 |
parent 24179 | c89d77d97f84 |
child 30549 | d2d7874648bd |
permissions | -rw-r--r-- |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
1 |
(* Title: HOL/Auth/NSP_Bad |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
4 |
Copyright 1996 University of Cambridge |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
5 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
6 |
Original file is ../Auth/NS_Public_Bad |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
7 |
*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
8 |
|
14203 | 9 |
header{*Analyzing the Needham-Schroeder Public-Key Protocol in UNITY*} |
10 |
||
24179 | 11 |
theory NSP_Bad imports "../../Auth/Public" "../UNITY_Main" begin |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
12 |
|
14203 | 13 |
text{*This is the flawed version, vulnerable to Lowe's attack. |
14 |
From page 260 of |
|
15 |
Burrows, Abadi and Needham. A Logic of Authentication. |
|
16 |
Proc. Royal Soc. 426 (1989). |
|
17 |
*} |
|
18 |
||
19 |
types state = "event list" |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
20 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
21 |
constdefs |
14203 | 22 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
23 |
(*The spy MAY say anything he CAN say. We do not expect him to |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
24 |
invent new nonces here, but he can also use NS1. Common to |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
25 |
all similar protocols.*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
26 |
Fake :: "(state*state) set" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
27 |
"Fake == {(s,s'). |
14203 | 28 |
\<exists>B X. s' = Says Spy B X # s |
29 |
& X \<in> synth (analz (spies s))}" |
|
30 |
||
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
31 |
(*The numeric suffixes on A identify the rule*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
32 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
33 |
(*Alice initiates a protocol run, sending a nonce to Bob*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
34 |
NS1 :: "(state*state) set" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
35 |
"NS1 == {(s1,s'). |
14203 | 36 |
\<exists>A1 B NA. |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
37 |
s' = Says A1 B (Crypt (pubK B) {|Nonce NA, Agent A1|}) # s1 |
14203 | 38 |
& Nonce NA \<notin> used s1}" |
39 |
||
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
40 |
(*Bob responds to Alice's message with a further nonce*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
41 |
NS2 :: "(state*state) set" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
42 |
"NS2 == {(s2,s'). |
14203 | 43 |
\<exists>A' A2 B NA NB. |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
44 |
s' = Says B A2 (Crypt (pubK A2) {|Nonce NA, Nonce NB|}) # s2 |
14203 | 45 |
& Says A' B (Crypt (pubK B) {|Nonce NA, Agent A2|}) \<in> set s2 |
46 |
& Nonce NB \<notin> used s2}" |
|
47 |
||
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
48 |
(*Alice proves her existence by sending NB back to Bob.*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
49 |
NS3 :: "(state*state) set" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
50 |
"NS3 == {(s3,s'). |
14203 | 51 |
\<exists>A3 B' B NA NB. |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
52 |
s' = Says A3 B (Crypt (pubK B) (Nonce NB)) # s3 |
14203 | 53 |
& Says A3 B (Crypt (pubK B) {|Nonce NA, Agent A3|}) \<in> set s3 |
54 |
& Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) \<in> set s3}" |
|
55 |
||
56 |
||
57 |
constdefs |
|
58 |
Nprg :: "state program" |
|
59 |
(*Initial trace is empty*) |
|
60 |
"Nprg == mk_total_program({[]}, {Fake, NS1, NS2, NS3}, UNIV)" |
|
61 |
||
62 |
declare spies_partsEs [elim] |
|
63 |
declare analz_into_parts [dest] |
|
64 |
declare Fake_parts_insert_in_Un [dest] |
|
65 |
||
66 |
text{*For other theories, e.g. Mutex and Lift, using [iff] slows proofs down. |
|
67 |
Here, it facilitates re-use of the Auth proofs.*} |
|
68 |
||
69 |
declare Fake_def [THEN def_act_simp, iff] |
|
70 |
declare NS1_def [THEN def_act_simp, iff] |
|
71 |
declare NS2_def [THEN def_act_simp, iff] |
|
72 |
declare NS3_def [THEN def_act_simp, iff] |
|
73 |
||
74 |
declare Nprg_def [THEN def_prg_Init, simp] |
|
75 |
||
76 |
||
77 |
text{*A "possibility property": there are traces that reach the end. |
|
78 |
Replace by LEADSTO proof!*} |
|
79 |
lemma "A \<noteq> B ==> |
|
80 |
\<exists>NB. \<exists>s \<in> reachable Nprg. Says A B (Crypt (pubK B) (Nonce NB)) \<in> set s" |
|
81 |
apply (intro exI bexI) |
|
82 |
apply (rule_tac [2] act = "totalize_act NS3" in reachable.Acts) |
|
83 |
apply (rule_tac [3] act = "totalize_act NS2" in reachable.Acts) |
|
84 |
apply (rule_tac [4] act = "totalize_act NS1" in reachable.Acts) |
|
85 |
apply (rule_tac [5] reachable.Init) |
|
86 |
apply (simp_all (no_asm_simp) add: Nprg_def totalize_act_def) |
|
87 |
apply auto |
|
88 |
done |
|
89 |
||
90 |
||
91 |
subsection{*Inductive Proofs about @{term ns_public}*} |
|
92 |
||
93 |
lemma ns_constrainsI: |
|
94 |
"(!!act s s'. [| act \<in> {Id, Fake, NS1, NS2, NS3}; |
|
95 |
(s,s') \<in> act; s \<in> A |] ==> s' \<in> A') |
|
96 |
==> Nprg \<in> A co A'" |
|
97 |
apply (simp add: Nprg_def mk_total_program_def) |
|
98 |
apply (rule constrainsI, auto) |
|
99 |
done |
|
100 |
||
101 |
||
102 |
text{*This ML code does the inductions directly.*} |
|
103 |
ML{* |
|
104 |
fun ns_constrains_tac(cs,ss) i = |
|
105 |
SELECT_GOAL |
|
24147 | 106 |
(EVERY [REPEAT (etac @{thm Always_ConstrainsI} 1), |
107 |
REPEAT (resolve_tac [@{thm StableI}, @{thm stableI}, |
|
108 |
@{thm constrains_imp_Constrains}] 1), |
|
109 |
rtac @{thm ns_constrainsI} 1, |
|
14203 | 110 |
full_simp_tac ss 1, |
111 |
REPEAT (FIRSTGOAL (etac disjE)), |
|
24147 | 112 |
ALLGOALS (clarify_tac (cs delrules [impI, @{thm impCE}])), |
14203 | 113 |
REPEAT (FIRSTGOAL analz_mono_contra_tac), |
114 |
ALLGOALS (asm_simp_tac ss)]) i; |
|
115 |
||
116 |
(*Tactic for proving secrecy theorems*) |
|
117 |
fun ns_induct_tac(cs,ss) = |
|
118 |
(SELECT_GOAL o EVERY) |
|
24147 | 119 |
[rtac @{thm AlwaysI} 1, |
14203 | 120 |
force_tac (cs,ss) 1, |
121 |
(*"reachable" gets in here*) |
|
24147 | 122 |
rtac (@{thm Always_reachable} RS @{thm Always_ConstrainsI} RS @{thm StableI}) 1, |
14203 | 123 |
ns_constrains_tac(cs,ss) 1]; |
124 |
*} |
|
125 |
||
126 |
method_setup ns_induct = {* |
|
127 |
Method.ctxt_args (fn ctxt => |
|
30510
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents:
24179
diff
changeset
|
128 |
SIMPLE_METHOD' (ns_induct_tac (local_clasimpset_of ctxt))) *} |
14203 | 129 |
"for inductive reasoning about the Needham-Schroeder protocol" |
130 |
||
131 |
text{*Converts invariants into statements about reachable states*} |
|
132 |
lemmas Always_Collect_reachableD = |
|
133 |
Always_includes_reachable [THEN subsetD, THEN CollectD] |
|
134 |
||
135 |
text{*Spy never sees another agent's private key! (unless it's bad at start)*} |
|
136 |
lemma Spy_see_priK: |
|
137 |
"Nprg \<in> Always {s. (Key (priK A) \<in> parts (spies s)) = (A \<in> bad)}" |
|
138 |
apply ns_induct |
|
139 |
apply blast |
|
140 |
done |
|
141 |
declare Spy_see_priK [THEN Always_Collect_reachableD, simp] |
|
142 |
||
143 |
lemma Spy_analz_priK: |
|
144 |
"Nprg \<in> Always {s. (Key (priK A) \<in> analz (spies s)) = (A \<in> bad)}" |
|
145 |
by (rule Always_reachable [THEN Always_weaken], auto) |
|
146 |
declare Spy_analz_priK [THEN Always_Collect_reachableD, simp] |
|
147 |
||
148 |
||
149 |
subsection{*Authenticity properties obtained from NS2*} |
|
150 |
||
151 |
text{*It is impossible to re-use a nonce in both NS1 and NS2 provided the |
|
152 |
nonce is secret. (Honest users generate fresh nonces.)*} |
|
153 |
lemma no_nonce_NS1_NS2: |
|
154 |
"Nprg |
|
155 |
\<in> Always {s. Crypt (pubK C) {|NA', Nonce NA|} \<in> parts (spies s) --> |
|
156 |
Crypt (pubK B) {|Nonce NA, Agent A|} \<in> parts (spies s) --> |
|
157 |
Nonce NA \<in> analz (spies s)}" |
|
158 |
apply ns_induct |
|
159 |
apply (blast intro: analz_insertI)+ |
|
160 |
done |
|
161 |
||
162 |
text{*Adding it to the claset slows down proofs...*} |
|
163 |
lemmas no_nonce_NS1_NS2_reachable = |
|
164 |
no_nonce_NS1_NS2 [THEN Always_Collect_reachableD, rule_format] |
|
165 |
||
166 |
||
167 |
text{*Unicity for NS1: nonce NA identifies agents A and B*} |
|
168 |
lemma unique_NA_lemma: |
|
169 |
"Nprg |
|
170 |
\<in> Always {s. Nonce NA \<notin> analz (spies s) --> |
|
171 |
Crypt(pubK B) {|Nonce NA, Agent A|} \<in> parts(spies s) --> |
|
172 |
Crypt(pubK B') {|Nonce NA, Agent A'|} \<in> parts(spies s) --> |
|
173 |
A=A' & B=B'}" |
|
174 |
apply ns_induct |
|
175 |
apply auto |
|
176 |
txt{*Fake, NS1 are non-trivial*} |
|
177 |
done |
|
178 |
||
179 |
text{*Unicity for NS1: nonce NA identifies agents A and B*} |
|
180 |
lemma unique_NA: |
|
181 |
"[| Crypt(pubK B) {|Nonce NA, Agent A|} \<in> parts(spies s); |
|
182 |
Crypt(pubK B') {|Nonce NA, Agent A'|} \<in> parts(spies s); |
|
183 |
Nonce NA \<notin> analz (spies s); |
|
184 |
s \<in> reachable Nprg |] |
|
185 |
==> A=A' & B=B'" |
|
186 |
by (blast dest: unique_NA_lemma [THEN Always_Collect_reachableD]) |
|
187 |
||
188 |
||
189 |
text{*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*} |
|
190 |
lemma Spy_not_see_NA: |
|
191 |
"[| A \<notin> bad; B \<notin> bad |] |
|
192 |
==> Nprg \<in> Always |
|
193 |
{s. Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) \<in> set s |
|
194 |
--> Nonce NA \<notin> analz (spies s)}" |
|
195 |
apply ns_induct |
|
196 |
txt{*NS3*} |
|
197 |
prefer 4 apply (blast intro: no_nonce_NS1_NS2_reachable) |
|
198 |
txt{*NS2*} |
|
199 |
prefer 3 apply (blast dest: unique_NA) |
|
200 |
txt{*NS1*} |
|
201 |
prefer 2 apply blast |
|
202 |
txt{*Fake*} |
|
203 |
apply spy_analz |
|
204 |
done |
|
205 |
||
206 |
||
207 |
text{*Authentication for A: if she receives message 2 and has used NA |
|
208 |
to start a run, then B has sent message 2.*} |
|
209 |
lemma A_trusts_NS2: |
|
210 |
"[| A \<notin> bad; B \<notin> bad |] |
|
211 |
==> Nprg \<in> Always |
|
212 |
{s. Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) \<in> set s & |
|
213 |
Crypt(pubK A) {|Nonce NA, Nonce NB|} \<in> parts (knows Spy s) |
|
214 |
--> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB|}) \<in> set s}" |
|
215 |
(*insert an invariant for use in some of the subgoals*) |
|
216 |
apply (insert Spy_not_see_NA [of A B NA], simp, ns_induct) |
|
217 |
apply (auto dest: unique_NA) |
|
218 |
done |
|
219 |
||
220 |
||
221 |
text{*If the encrypted message appears then it originated with Alice in NS1*} |
|
222 |
lemma B_trusts_NS1: |
|
223 |
"Nprg \<in> Always |
|
224 |
{s. Nonce NA \<notin> analz (spies s) --> |
|
225 |
Crypt (pubK B) {|Nonce NA, Agent A|} \<in> parts (spies s) |
|
226 |
--> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) \<in> set s}" |
|
227 |
apply ns_induct |
|
228 |
apply blast |
|
229 |
done |
|
230 |
||
231 |
||
232 |
subsection{*Authenticity properties obtained from NS2*} |
|
233 |
||
234 |
text{*Unicity for NS2: nonce NB identifies nonce NA and agent A. |
|
235 |
Proof closely follows that of @{text unique_NA}.*} |
|
236 |
lemma unique_NB_lemma: |
|
237 |
"Nprg |
|
238 |
\<in> Always {s. Nonce NB \<notin> analz (spies s) --> |
|
239 |
Crypt (pubK A) {|Nonce NA, Nonce NB|} \<in> parts (spies s) --> |
|
240 |
Crypt(pubK A'){|Nonce NA', Nonce NB|} \<in> parts(spies s) --> |
|
241 |
A=A' & NA=NA'}" |
|
242 |
apply ns_induct |
|
243 |
apply auto |
|
244 |
txt{*Fake, NS2 are non-trivial*} |
|
245 |
done |
|
246 |
||
247 |
lemma unique_NB: |
|
248 |
"[| Crypt(pubK A) {|Nonce NA, Nonce NB|} \<in> parts(spies s); |
|
249 |
Crypt(pubK A'){|Nonce NA', Nonce NB|} \<in> parts(spies s); |
|
250 |
Nonce NB \<notin> analz (spies s); |
|
251 |
s \<in> reachable Nprg |] |
|
252 |
==> A=A' & NA=NA'" |
|
253 |
apply (blast dest: unique_NB_lemma [THEN Always_Collect_reachableD]) |
|
254 |
done |
|
255 |
||
256 |
||
257 |
text{*NB remains secret PROVIDED Alice never responds with round 3*} |
|
258 |
lemma Spy_not_see_NB: |
|
259 |
"[| A \<notin> bad; B \<notin> bad |] |
|
260 |
==> Nprg \<in> Always |
|
261 |
{s. Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s & |
|
262 |
(ALL C. Says A C (Crypt (pubK C) (Nonce NB)) \<notin> set s) |
|
263 |
--> Nonce NB \<notin> analz (spies s)}" |
|
264 |
apply ns_induct |
|
265 |
apply (simp_all (no_asm_simp) add: all_conj_distrib) |
|
266 |
txt{*NS3: because NB determines A*} |
|
267 |
prefer 4 apply (blast dest: unique_NB) |
|
268 |
txt{*NS2: by freshness and unicity of NB*} |
|
269 |
prefer 3 apply (blast intro: no_nonce_NS1_NS2_reachable) |
|
270 |
txt{*NS1: by freshness*} |
|
271 |
prefer 2 apply blast |
|
272 |
txt{*Fake*} |
|
273 |
apply spy_analz |
|
274 |
done |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
275 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
276 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
277 |
|
14203 | 278 |
text{*Authentication for B: if he receives message 3 and has used NB |
279 |
in message 2, then A has sent message 3--to somebody....*} |
|
280 |
lemma B_trusts_NS3: |
|
281 |
"[| A \<notin> bad; B \<notin> bad |] |
|
282 |
==> Nprg \<in> Always |
|
283 |
{s. Crypt (pubK B) (Nonce NB) \<in> parts (spies s) & |
|
284 |
Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s |
|
285 |
--> (\<exists>C. Says A C (Crypt (pubK C) (Nonce NB)) \<in> set s)}" |
|
286 |
(*insert an invariant for use in some of the subgoals*) |
|
287 |
apply (insert Spy_not_see_NB [of A B NA NB], simp, ns_induct) |
|
288 |
apply (simp_all (no_asm_simp) add: ex_disj_distrib) |
|
289 |
apply auto |
|
290 |
txt{*NS3: because NB determines A. This use of @{text unique_NB} is robust.*} |
|
291 |
apply (blast intro: unique_NB [THEN conjunct1]) |
|
292 |
done |
|
293 |
||
294 |
||
295 |
text{*Can we strengthen the secrecy theorem? NO*} |
|
296 |
lemma "[| A \<notin> bad; B \<notin> bad |] |
|
297 |
==> Nprg \<in> Always |
|
298 |
{s. Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s |
|
299 |
--> Nonce NB \<notin> analz (spies s)}" |
|
300 |
apply ns_induct |
|
301 |
apply auto |
|
302 |
txt{*Fake*} |
|
303 |
apply spy_analz |
|
304 |
txt{*NS2: by freshness and unicity of NB*} |
|
305 |
apply (blast intro: no_nonce_NS1_NS2_reachable) |
|
306 |
txt{*NS3: unicity of NB identifies A and NA, but not B*} |
|
307 |
apply (frule_tac A'=A in Says_imp_spies [THEN parts.Inj, THEN unique_NB]) |
|
308 |
apply (erule Says_imp_spies [THEN parts.Inj], auto) |
|
309 |
apply (rename_tac s B' C) |
|
310 |
txt{*This is the attack! |
|
311 |
@{subgoals[display,indent=0,margin=65]} |
|
312 |
*} |
|
313 |
oops |
|
314 |
||
315 |
||
316 |
(* |
|
317 |
THIS IS THE ATTACK! |
|
318 |
[| A \<notin> bad; B \<notin> bad |] |
|
319 |
==> Nprg |
|
320 |
\<in> Always |
|
321 |
{s. Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s --> |
|
322 |
Nonce NB \<notin> analz (knows Spy s)} |
|
323 |
1. !!s B' C. |
|
324 |
[| A \<notin> bad; B \<notin> bad; s \<in> reachable Nprg |
|
325 |
Says A C (Crypt (pubK C) {|Nonce NA, Agent A|}) \<in> set s; |
|
326 |
Says B' A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s; |
|
327 |
C \<in> bad; Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB|}) \<in> set s; |
|
328 |
Nonce NB \<notin> analz (knows Spy s) |] |
|
329 |
==> False |
|
330 |
*) |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
331 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
332 |
end |