|
1 theory Fv |
|
2 imports "../Nominal-General/Nominal2_Atoms" |
|
3 "Abs" "Perm" "Rsp" "Nominal2_FSet" |
|
4 begin |
|
5 |
|
6 (* The bindings data structure: |
|
7 |
|
8 Bindings are a list of lists of lists of triples. |
|
9 |
|
10 The first list represents the datatypes defined. |
|
11 The second list represents the constructors. |
|
12 The internal list is a list of all the bndings that |
|
13 concern the constructor. |
|
14 |
|
15 Every triple consists of a function, the binding and |
|
16 the body. |
|
17 |
|
18 Eg: |
|
19 nominal_datatype |
|
20 |
|
21 C1 |
|
22 | C2 x y z bind x in z |
|
23 | C3 x y z bind f x in z bind g y in z |
|
24 |
|
25 yields: |
|
26 [ |
|
27 [], |
|
28 [(NONE, 0, 2)], |
|
29 [(SOME (Const f), 0, 2), (Some (Const g), 1, 2)]] |
|
30 |
|
31 A SOME binding has to have a function which takes an appropriate |
|
32 argument and returns an atom set. A NONE binding has to be on an |
|
33 argument that is an atom or an atom set. |
|
34 *) |
|
35 |
|
36 (* |
|
37 An overview of the generation of free variables: |
|
38 |
|
39 1) fv_bn functions are generated only for the non-recursive binds. |
|
40 |
|
41 An fv_bn for a constructor is a union of values for the arguments: |
|
42 |
|
43 For an argument x that is in the bn function |
|
44 - if it is a recursive argument bn' we return: fv_bn' x |
|
45 - otherwise empty |
|
46 |
|
47 For an argument x that is not in the bn function |
|
48 - for atom we return: {atom x} |
|
49 - for atom set we return: atom ` x |
|
50 - for a recursive call to type ty' we return: fv_ty' x |
|
51 with fv of the appropriate type |
|
52 - otherwise empty |
|
53 |
|
54 2) fv_ty functions generated for all types being defined: |
|
55 |
|
56 fv_ty for a constructor is a union of values for the arguments. |
|
57 |
|
58 For an argument that is bound in a shallow binding we return empty. |
|
59 |
|
60 For an argument x that bound in a non-recursive deep binding |
|
61 we return: fv_bn x. |
|
62 |
|
63 Otherwise we return the free variables of the argument minus the |
|
64 bound variables of the argument. |
|
65 |
|
66 The free variables for an argument x are: |
|
67 - for an atom: {atom x} |
|
68 - for atom set: atom ` x |
|
69 - for recursive call to type ty' return: fv_ty' x |
|
70 - for nominal datatype ty' return: fv_ty' x |
|
71 |
|
72 The bound variables are a union of results of all bindings that |
|
73 involve the given argument. For a paricular binding: |
|
74 |
|
75 - for a binding function bn: bn x |
|
76 - for a recursive argument of type ty': fv_fy' x |
|
77 - for nominal datatype ty' return: fv_ty' x |
|
78 *) |
|
79 |
|
80 (* |
|
81 An overview of the generation of alpha-equivalence: |
|
82 |
|
83 1) alpha_bn relations are generated for binding functions. |
|
84 |
|
85 An alpha_bn for a constructor is true if a conjunction of |
|
86 propositions for each argument holds. |
|
87 |
|
88 For an argument a proposition is build as follows from |
|
89 th: |
|
90 |
|
91 - for a recursive argument in the bn function, we return: alpha_bn argl argr |
|
92 - for a recursive argument for type ty not in bn, we return: alpha_ty argl argr |
|
93 - for other arguments in the bn function we return: True |
|
94 - for other arguments not in the bn function we return: argl = argr |
|
95 |
|
96 2) alpha_ty relations are generated for all the types being defined: |
|
97 |
|
98 For each constructor we gather all the arguments that are bound, |
|
99 and for each of those we add a permutation. We associate those |
|
100 permutations with the bindings. Note that two bindings can have |
|
101 the same permutation if the arguments being bound are the same. |
|
102 |
|
103 An alpha_ty for a constructor is true if there exist permutations |
|
104 as above such that a conjunction of propositions for all arguments holds. |
|
105 |
|
106 For an argument we allow bindings where only one of the following |
|
107 holds: |
|
108 |
|
109 - Argument is bound in some shallow bindings: We return true |
|
110 - Argument of type ty is bound recursively in some other |
|
111 arguments [i1, .. in] with one binding function bn. |
|
112 We return: |
|
113 |
|
114 (bn argl, (argl, argl_i1, ..., argl_in)) \<approx>gen |
|
115 \<lambda>(argl,argl1,..,argln) (argr,argr1,..,argrn). |
|
116 (alpha_ty argl argr) \<and> (alpha_i1 argl1 argr1) \<and> .. \<and> (alpha_in argln argrn) |
|
117 \<lambda>(arg,arg1,..,argn). (fv_ty arg) \<union> (fv_i1 arg1) \<union> .. \<union> (fv_in argn) |
|
118 pi |
|
119 (bn argr, (argr, argr_i1, ..., argr_in)) |
|
120 |
|
121 - Argument is bound in some deep non-recursive bindings. |
|
122 We return: alpha_bn argl argr |
|
123 - Argument of type ty has some shallow bindings [b1..bn] and/or |
|
124 non-recursive bindings [f1 a1, .., fm am], where the bindings |
|
125 have the permutations p1..pl. We return: |
|
126 |
|
127 (b1l \<union>..\<union> bnl \<union> f1 a1l \<union>..\<union> fn anl, argl) \<approx>gen |
|
128 alpha_ty fv_ty (p1 +..+ pl) |
|
129 (b1r \<union>..\<union> bnr \<union> f1 a1r \<union>..\<union> fn anr, argr) |
|
130 |
|
131 - Argument has some recursive bindings. The bindings were |
|
132 already treated in 2nd case so we return: True |
|
133 - Argument has no bindings and is not bound. |
|
134 If it is recursive for type ty, we return: alpha_ty argl argr |
|
135 Otherwise we return: argl = argr |
|
136 |
|
137 *) |
|
138 |
|
139 |
|
140 ML {* |
|
141 datatype alpha_mode = AlphaGen | AlphaRes | AlphaLst; |
|
142 *} |
|
143 |
|
144 ML {* |
|
145 fun atyp_const AlphaGen = @{const_name alpha_gen} |
|
146 | atyp_const AlphaRes = @{const_name alpha_res} |
|
147 | atyp_const AlphaLst = @{const_name alpha_lst} |
|
148 *} |
|
149 |
|
150 (* TODO: make sure that parser checks that bindings are compatible *) |
|
151 ML {* |
|
152 fun alpha_const_for_binds [] = atyp_const AlphaGen |
|
153 | alpha_const_for_binds ((NONE, _, _, at) :: t) = atyp_const at |
|
154 | alpha_const_for_binds ((SOME (_, _), _, _, at) :: _) = atyp_const at |
|
155 *} |
|
156 |
|
157 ML {* |
|
158 fun is_atom thy typ = |
|
159 Sign.of_sort thy (typ, @{sort at}) |
|
160 |
|
161 fun is_atom_set thy (Type ("fun", [t, @{typ bool}])) = is_atom thy t |
|
162 | is_atom_set _ _ = false; |
|
163 |
|
164 fun is_atom_fset thy (Type ("FSet.fset", [t])) = is_atom thy t |
|
165 | is_atom_fset _ _ = false; |
|
166 *} |
|
167 |
|
168 |
|
169 (* Like map2, only if the second list is empty passes empty lists insted of error *) |
|
170 ML {* |
|
171 fun map2i _ [] [] = [] |
|
172 | map2i f (x :: xs) (y :: ys) = f x y :: map2i f xs ys |
|
173 | map2i f (x :: xs) [] = f x [] :: map2i f xs [] |
|
174 | map2i _ _ _ = raise UnequalLengths; |
|
175 *} |
|
176 |
|
177 (* Finds bindings with the same function and binding, and gathers all |
|
178 bodys for such pairs |
|
179 *) |
|
180 ML {* |
|
181 fun gather_binds binds = |
|
182 let |
|
183 fun gather_binds_cons binds = |
|
184 let |
|
185 val common = map (fn (f, bi, _, aty) => (f, bi, aty)) binds |
|
186 val nodups = distinct (op =) common |
|
187 fun find_bodys (sf, sbi, sty) = |
|
188 filter (fn (f, bi, _, aty) => f = sf andalso bi = sbi andalso aty = sty) binds |
|
189 val bodys = map ((map (fn (_, _, bo, _) => bo)) o find_bodys) nodups |
|
190 in |
|
191 nodups ~~ bodys |
|
192 end |
|
193 in |
|
194 map (map gather_binds_cons) binds |
|
195 end |
|
196 *} |
|
197 |
|
198 ML {* |
|
199 fun un_gather_binds_cons binds = |
|
200 flat (map (fn (((f, bi, aty), bos), pi) => map (fn bo => ((f, bi, bo, aty), pi)) bos) binds) |
|
201 *} |
|
202 |
|
203 ML {* |
|
204 open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *); |
|
205 *} |
|
206 ML {* |
|
207 (* TODO: It is the same as one in 'nominal_atoms' *) |
|
208 fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom}); |
|
209 val noatoms = @{term "{} :: atom set"}; |
|
210 fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]; |
|
211 fun mk_union sets = |
|
212 fold (fn a => fn b => |
|
213 if a = noatoms then b else |
|
214 if b = noatoms then a else |
|
215 if a = b then a else |
|
216 HOLogic.mk_binop @{const_name sup} (a, b)) (rev sets) noatoms; |
|
217 val mk_inter = foldr1 (HOLogic.mk_binop @{const_name inf}) |
|
218 fun mk_diff a b = |
|
219 if b = noatoms then a else |
|
220 if b = a then noatoms else |
|
221 HOLogic.mk_binop @{const_name minus} (a, b); |
|
222 fun mk_atom_set t = |
|
223 let |
|
224 val ty = fastype_of t; |
|
225 val atom_ty = HOLogic.dest_setT ty --> @{typ atom}; |
|
226 val img_ty = atom_ty --> ty --> @{typ "atom set"}; |
|
227 in |
|
228 (Const (@{const_name image}, img_ty) $ Const (@{const_name atom}, atom_ty) $ t) |
|
229 end; |
|
230 fun mk_atom_fset t = |
|
231 let |
|
232 val ty = fastype_of t; |
|
233 val atom_ty = dest_fsetT ty --> @{typ atom}; |
|
234 val fmap_ty = atom_ty --> ty --> @{typ "atom fset"}; |
|
235 val fset_to_set = @{term "fset_to_set :: atom fset \<Rightarrow> atom set"} |
|
236 in |
|
237 fset_to_set $ ((Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)) |
|
238 end; |
|
239 (* Similar to one in USyntax *) |
|
240 fun mk_pair (fst, snd) = |
|
241 let val ty1 = fastype_of fst |
|
242 val ty2 = fastype_of snd |
|
243 val c = HOLogic.pair_const ty1 ty2 |
|
244 in c $ fst $ snd |
|
245 end; |
|
246 *} |
|
247 |
|
248 (* Given [fv1, fv2, fv3] creates %(x, y, z). fv1 x u fv2 y u fv3 z *) |
|
249 ML {* |
|
250 fun mk_compound_fv fvs = |
|
251 let |
|
252 val nos = (length fvs - 1) downto 0; |
|
253 val fvs_applied = map (fn (fv, no) => fv $ Bound no) (fvs ~~ nos); |
|
254 val fvs_union = mk_union fvs_applied; |
|
255 val (tyh :: tys) = rev (map (domain_type o fastype_of) fvs); |
|
256 fun fold_fun ty t = HOLogic.mk_split (Abs ("", ty, t)) |
|
257 in |
|
258 fold fold_fun tys (Abs ("", tyh, fvs_union)) |
|
259 end; |
|
260 *} |
|
261 |
|
262 (* Given [R1, R2, R3] creates %(x,x'). %(y,y'). %(z,z'). R x x' \<and> R y y' \<and> R z z' *) |
|
263 ML {* |
|
264 fun mk_compound_alpha Rs = |
|
265 let |
|
266 val nos = (length Rs - 1) downto 0; |
|
267 val nos2 = (2 * length Rs - 1) downto length Rs; |
|
268 val Rs_applied = map (fn (R, (no2, no)) => R $ Bound no2 $ Bound no) (Rs ~~ (nos2 ~~ nos)); |
|
269 val Rs_conj = mk_conjl Rs_applied; |
|
270 val (tyh :: tys) = rev (map (domain_type o fastype_of) Rs); |
|
271 fun fold_fun ty t = HOLogic.mk_split (Abs ("", ty, t)) |
|
272 val abs_rhs = fold fold_fun tys (Abs ("", tyh, Rs_conj)) |
|
273 in |
|
274 fold fold_fun tys (Abs ("", tyh, abs_rhs)) |
|
275 end; |
|
276 *} |
|
277 |
|
278 |
|
279 ML {* |
|
280 fun non_rec_binds l = |
|
281 let |
|
282 fun is_non_rec (SOME (f, false), _, _, _) = SOME f |
|
283 | is_non_rec _ = NONE |
|
284 in |
|
285 distinct (op =) (map_filter is_non_rec (flat (flat l))) |
|
286 end |
|
287 *} |
|
288 |
|
289 (* We assume no bindings in the type on which bn is defined *) |
|
290 ML {* |
|
291 fun fv_bn thy (dt_info : Datatype_Aux.info) fv_frees bn_fvbn (fvbn, (bn, ith_dtyp, args_in_bns)) = |
|
292 let |
|
293 val {descr, sorts, ...} = dt_info; |
|
294 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
295 fun fv_bn_constr (cname, dts) args_in_bn = |
|
296 let |
|
297 val Ts = map (typ_of_dtyp descr sorts) dts; |
|
298 val names = Datatype_Prop.make_tnames Ts; |
|
299 val args = map Free (names ~~ Ts); |
|
300 val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
|
301 fun fv_arg ((dt, x), arg_no) = |
|
302 let |
|
303 val ty = fastype_of x |
|
304 (* val _ = tracing ("B 1" ^ PolyML.makestring args_in_bn);*) |
|
305 (* val _ = tracing ("B 2" ^ PolyML.makestring bn_fvbn);*) |
|
306 in |
|
307 case AList.lookup (op=) args_in_bn arg_no of |
|
308 SOME NONE => @{term "{} :: atom set"} |
|
309 | SOME (SOME (f : term)) => (the (AList.lookup (op=) bn_fvbn f)) $ x |
|
310 | NONE => |
|
311 if is_atom thy ty then mk_single_atom x else |
|
312 if is_atom_set thy ty then mk_atom_set x else |
|
313 if is_atom_fset thy ty then mk_atom_fset x else |
|
314 if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
|
315 @{term "{} :: atom set"} |
|
316 end; |
|
317 val arg_nos = 0 upto (length dts - 1) |
|
318 in |
|
319 HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
320 (fvbn $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos)))) |
|
321 end; |
|
322 val (_, (_, _, constrs)) = nth descr ith_dtyp; |
|
323 val eqs = map2i fv_bn_constr constrs args_in_bns |
|
324 in |
|
325 ((bn, fvbn), eqs) |
|
326 end |
|
327 *} |
|
328 |
|
329 ML {* print_depth 100 *} |
|
330 ML {* |
|
331 fun fv_bns thy dt_info fv_frees rel_bns = |
|
332 let |
|
333 fun mk_fvbn_free (bn, ith, _) = |
|
334 let |
|
335 val fvbn_name = "fv_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
|
336 in |
|
337 (fvbn_name, Free (fvbn_name, fastype_of (nth fv_frees ith))) |
|
338 end; |
|
339 val (fvbn_names, fvbn_frees) = split_list (map mk_fvbn_free rel_bns); |
|
340 val bn_fvbn = (map (fn (bn, _, _) => bn) rel_bns) ~~ fvbn_frees |
|
341 val (l1, l2) = split_list (map (fv_bn thy dt_info fv_frees bn_fvbn) (fvbn_frees ~~ rel_bns)); |
|
342 in |
|
343 (l1, (fvbn_names ~~ l2)) |
|
344 end |
|
345 *} |
|
346 |
|
347 |
|
348 ML {* |
|
349 fun alpha_bn (dt_info : Datatype_Aux.info) alpha_frees bn_alphabn ((bn, ith_dtyp, args_in_bns), (alpha_bn_free, _ (*is_rec*) )) = |
|
350 let |
|
351 val {descr, sorts, ...} = dt_info; |
|
352 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
353 fun alpha_bn_constr (cname, dts) args_in_bn = |
|
354 let |
|
355 val Ts = map (typ_of_dtyp descr sorts) dts; |
|
356 val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); |
|
357 val names2 = Name.variant_list ("pi" :: names) (Datatype_Prop.make_tnames Ts); |
|
358 val args = map Free (names ~~ Ts); |
|
359 val args2 = map Free (names2 ~~ Ts); |
|
360 val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
|
361 val rhs = HOLogic.mk_Trueprop |
|
362 (alpha_bn_free $ (list_comb (c, args)) $ (list_comb (c, args2))); |
|
363 fun lhs_arg ((dt, arg_no), (arg, arg2)) = |
|
364 case AList.lookup (op=) args_in_bn arg_no of |
|
365 SOME NONE => @{term True} |
|
366 | SOME (SOME f) => (the (AList.lookup (op=) bn_alphabn f)) $ arg $ arg2 |
|
367 | NONE => |
|
368 if is_rec_type dt then (nth alpha_frees (body_index dt)) $ arg $ arg2 |
|
369 else HOLogic.mk_eq (arg, arg2) |
|
370 val arg_nos = 0 upto (length dts - 1) |
|
371 val lhss = mk_conjl (map lhs_arg (dts ~~ arg_nos ~~ (args ~~ args2))) |
|
372 val eq = Logic.mk_implies (HOLogic.mk_Trueprop lhss, rhs) |
|
373 in |
|
374 eq |
|
375 end |
|
376 val (_, (_, _, constrs)) = nth descr ith_dtyp; |
|
377 val eqs = map2i alpha_bn_constr constrs args_in_bns |
|
378 in |
|
379 ((bn, alpha_bn_free), eqs) |
|
380 end |
|
381 *} |
|
382 |
|
383 ML {* |
|
384 fun alpha_bns dt_info alpha_frees rel_bns bns_rec = |
|
385 let |
|
386 val {descr, sorts, ...} = dt_info; |
|
387 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
388 fun mk_alphabn_free (bn, ith, _) = |
|
389 let |
|
390 val alphabn_name = "alpha_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
|
391 val alphabn_type = nth_dtyp ith --> nth_dtyp ith --> @{typ bool}; |
|
392 val alphabn_free = Free(alphabn_name, alphabn_type); |
|
393 in |
|
394 (alphabn_name, alphabn_free) |
|
395 end; |
|
396 val (alphabn_names, alphabn_frees) = split_list (map mk_alphabn_free rel_bns); |
|
397 val bn_alphabn = (map (fn (bn, _, _) => bn) rel_bns) ~~ alphabn_frees; |
|
398 val pair = split_list (map (alpha_bn dt_info alpha_frees bn_alphabn) |
|
399 (rel_bns ~~ (alphabn_frees ~~ bns_rec))) |
|
400 in |
|
401 (alphabn_names, pair) |
|
402 end |
|
403 *} |
|
404 |
|
405 |
|
406 (* Checks that a list of bindings contains only compatible ones *) |
|
407 ML {* |
|
408 fun bns_same l = |
|
409 length (distinct (op =) (map (fn ((b, _, _, atyp), _) => (b, atyp)) l)) = 1 |
|
410 *} |
|
411 |
|
412 ML {* |
|
413 fun setify x = |
|
414 if fastype_of x = @{typ "atom list"} then |
|
415 Const (@{const_name set}, @{typ "atom list \<Rightarrow> atom set"}) $ x else x |
|
416 *} |
|
417 |
|
418 ML {* |
|
419 fun define_fv (dt_info : Datatype_Aux.info) bindsall bns lthy = |
|
420 let |
|
421 val thy = ProofContext.theory_of lthy; |
|
422 val {descr, sorts, ...} = dt_info; |
|
423 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
424 val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
|
425 "fv_" ^ name_of_typ (nth_dtyp i)) descr); |
|
426 val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr; |
|
427 val fv_frees = map Free (fv_names ~~ fv_types); |
|
428 (* TODO: We need a transitive closure, but instead we do this hack considering |
|
429 all binding functions as recursive or not *) |
|
430 val nr_bns = |
|
431 if (non_rec_binds bindsall) = [] then [] |
|
432 else map (fn (bn, _, _) => bn) bns; |
|
433 val rel_bns = filter (fn (bn, _, _) => bn mem nr_bns) bns; |
|
434 val (bn_fv_bns, fv_bn_names_eqs) = fv_bns thy dt_info fv_frees rel_bns; |
|
435 val fvbns = map snd bn_fv_bns; |
|
436 val (fv_bn_names, fv_bn_eqs) = split_list fv_bn_names_eqs; |
|
437 |
|
438 fun fv_constr ith_dtyp (cname, dts) bindcs = |
|
439 let |
|
440 val Ts = map (typ_of_dtyp descr sorts) dts; |
|
441 val bindslen = length bindcs |
|
442 val pi_strs_same = replicate bindslen "pi" |
|
443 val pi_strs = Name.variant_list [] pi_strs_same; |
|
444 val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs; |
|
445 val bind_pis_gath = bindcs ~~ pis; |
|
446 val bind_pis = un_gather_binds_cons bind_pis_gath; |
|
447 val bindcs = map fst bind_pis; |
|
448 val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts); |
|
449 val args = map Free (names ~~ Ts); |
|
450 val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
|
451 val fv_c = nth fv_frees ith_dtyp; |
|
452 val arg_nos = 0 upto (length dts - 1) |
|
453 fun fv_bind args (NONE, i, _, _) = |
|
454 if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else |
|
455 if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else |
|
456 if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else |
|
457 if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else |
|
458 (* TODO goes the code for preiously defined nominal datatypes *) |
|
459 @{term "{} :: atom set"} |
|
460 | fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i) |
|
461 fun fv_binds_as_set args relevant = mk_union (map (setify o fv_bind args) relevant) |
|
462 fun find_nonrec_binder j (SOME (f, false), i, _, _) = if i = j then SOME f else NONE |
|
463 | find_nonrec_binder _ _ = NONE |
|
464 fun fv_arg ((dt, x), arg_no) = |
|
465 case get_first (find_nonrec_binder arg_no) bindcs of |
|
466 SOME f => |
|
467 (case get_first (fn (x, y) => if x = f then SOME y else NONE) bn_fv_bns of |
|
468 SOME fv_bn => fv_bn $ x |
|
469 | NONE => error "bn specified in a non-rec binding but not in bn list") |
|
470 | NONE => |
|
471 let |
|
472 val arg = |
|
473 if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
|
474 if ((is_atom thy) o fastype_of) x then mk_single_atom x else |
|
475 if ((is_atom_set thy) o fastype_of) x then mk_atom_set x else |
|
476 if ((is_atom_fset thy) o fastype_of) x then mk_atom_fset x else |
|
477 (* TODO goes the code for preiously defined nominal datatypes *) |
|
478 @{term "{} :: atom set"}; |
|
479 (* If i = j then we generate it only once *) |
|
480 val relevant = filter (fn (_, i, j, _) => ((i = arg_no) orelse (j = arg_no))) bindcs; |
|
481 val sub = fv_binds_as_set args relevant |
|
482 in |
|
483 mk_diff arg sub |
|
484 end; |
|
485 val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
486 (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos)))) |
|
487 in |
|
488 fv_eq |
|
489 end; |
|
490 fun fv_eq (i, (_, _, constrs)) binds = map2i (fv_constr i) constrs binds; |
|
491 val fveqs = map2i fv_eq descr (gather_binds bindsall) |
|
492 val fv_eqs_perfv = fveqs |
|
493 val rel_bns_nos = map (fn (_, i, _) => i) rel_bns; |
|
494 fun filter_fun (_, b) = b mem rel_bns_nos; |
|
495 val all_fvs = (fv_names ~~ fv_eqs_perfv) ~~ (0 upto (length fv_names - 1)) |
|
496 val (fv_names_fst, fv_eqs_fst) = apsnd flat (split_list (map fst (filter_out filter_fun all_fvs))) |
|
497 val (fv_names_snd, fv_eqs_snd) = apsnd flat (split_list (map fst (filter filter_fun all_fvs))) |
|
498 val fv_eqs_all = fv_eqs_fst @ (flat fv_bn_eqs); |
|
499 val fv_names_all = fv_names_fst @ fv_bn_names; |
|
500 val add_binds = map (fn x => (Attrib.empty_binding, x)) |
|
501 (* Function_Fun.add_fun Function_Common.default_config ... true *) |
|
502 val (fvs, lthy') = (Primrec.add_primrec |
|
503 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_all) (add_binds fv_eqs_all) lthy) |
|
504 val (fvs2, lthy'') = |
|
505 if fv_eqs_snd = [] then (([], []), lthy') else |
|
506 (Primrec.add_primrec |
|
507 (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_snd) (add_binds fv_eqs_snd) lthy') |
|
508 val ordered_fvs = fv_frees @ fvbns; |
|
509 val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2) |
|
510 in |
|
511 ((all_fvs, ordered_fvs), lthy'') |
|
512 end |
|
513 *} |
|
514 |
|
515 ML {* |
|
516 fun define_alpha (dt_info : Datatype_Aux.info) bindsall bns fv_frees lthy = |
|
517 let |
|
518 val thy = ProofContext.theory_of lthy; |
|
519 val {descr, sorts, ...} = dt_info; |
|
520 fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
|
521 (* TODO: We need a transitive closure, but instead we do this hack considering |
|
522 all binding functions as recursive or not *) |
|
523 val nr_bns = |
|
524 if (non_rec_binds bindsall) = [] then [] |
|
525 else map (fn (bn, _, _) => bn) bns; |
|
526 val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
|
527 "alpha_" ^ name_of_typ (nth_dtyp i)) descr); |
|
528 val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr; |
|
529 val alpha_frees = map Free (alpha_names ~~ alpha_types); |
|
530 (* We assume that a bn is either recursive or not *) |
|
531 val bns_rec = map (fn (bn, _, _) => not (bn mem nr_bns)) bns; |
|
532 val (alpha_bn_names, (bn_alpha_bns, alpha_bn_eqs)) = |
|
533 alpha_bns dt_info alpha_frees bns bns_rec |
|
534 val alpha_bn_frees = map snd bn_alpha_bns; |
|
535 val alpha_bn_types = map fastype_of alpha_bn_frees; |
|
536 |
|
537 fun alpha_constr ith_dtyp (cname, dts) bindcs = |
|
538 let |
|
539 val Ts = map (typ_of_dtyp descr sorts) dts; |
|
540 val bindslen = length bindcs |
|
541 val pi_strs_same = replicate bindslen "pi" |
|
542 val pi_strs = Name.variant_list [] pi_strs_same; |
|
543 val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs; |
|
544 val bind_pis_gath = bindcs ~~ pis; |
|
545 val bind_pis = un_gather_binds_cons bind_pis_gath; |
|
546 val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts); |
|
547 val args = map Free (names ~~ Ts); |
|
548 val names2 = Name.variant_list (pi_strs @ names) (Datatype_Prop.make_tnames Ts); |
|
549 val args2 = map Free (names2 ~~ Ts); |
|
550 val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
|
551 val alpha = nth alpha_frees ith_dtyp; |
|
552 val arg_nos = 0 upto (length dts - 1) |
|
553 fun fv_bind args (NONE, i, _, _) = |
|
554 if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else |
|
555 if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else |
|
556 if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else |
|
557 if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else |
|
558 (* TODO goes the code for preiously defined nominal datatypes *) |
|
559 @{term "{} :: atom set"} |
|
560 | fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i) |
|
561 fun fv_binds args relevant = mk_union (map (fv_bind args) relevant) |
|
562 val alpha_rhs = |
|
563 HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2))); |
|
564 fun alpha_arg ((dt, arg_no), (arg, arg2)) = |
|
565 let |
|
566 val rel_in_simp_binds = filter (fn ((NONE, i, _, _), _) => i = arg_no | _ => false) bind_pis; |
|
567 val rel_in_comp_binds = filter (fn ((SOME _, i, _, _), _) => i = arg_no | _ => false) bind_pis; |
|
568 val rel_has_binds = filter (fn ((NONE, _, j, _), _) => j = arg_no |
|
569 | ((SOME (_, false), _, j, _), _) => j = arg_no |
|
570 | _ => false) bind_pis; |
|
571 val rel_has_rec_binds = filter |
|
572 (fn ((SOME (_, true), _, j, _), _) => j = arg_no | _ => false) bind_pis; |
|
573 in |
|
574 case (rel_in_simp_binds, rel_in_comp_binds, rel_has_binds, rel_has_rec_binds) of |
|
575 ([], [], [], []) => |
|
576 if is_rec_type dt then (nth alpha_frees (body_index dt) $ arg $ arg2) |
|
577 else (HOLogic.mk_eq (arg, arg2)) |
|
578 | (_, [], [], []) => @{term True} |
|
579 | ([], [], [], _) => @{term True} |
|
580 | ([], ((((SOME (bn, is_rec)), _, _, atyp), _) :: _), [], []) => |
|
581 if not (bns_same rel_in_comp_binds) then error "incompatible bindings for an argument" else |
|
582 if is_rec then |
|
583 let |
|
584 val (rbinds, rpis) = split_list rel_in_comp_binds |
|
585 val bound_in_nos = map (fn (_, _, i, _) => i) rbinds |
|
586 val bound_in_ty_nos = map (fn i => body_index (nth dts i)) bound_in_nos; |
|
587 val bound_args = arg :: map (nth args) bound_in_nos; |
|
588 val bound_args2 = arg2 :: map (nth args2) bound_in_nos; |
|
589 val lhs_binds = fv_binds args rbinds |
|
590 val lhs_arg = foldr1 HOLogic.mk_prod bound_args |
|
591 val lhs = mk_pair (lhs_binds, lhs_arg); |
|
592 val rhs_binds = fv_binds args2 rbinds; |
|
593 val rhs_arg = foldr1 HOLogic.mk_prod bound_args2; |
|
594 val rhs = mk_pair (rhs_binds, rhs_arg); |
|
595 val fvs = map (nth fv_frees) ((body_index dt) :: bound_in_ty_nos); |
|
596 val fv = mk_compound_fv fvs; |
|
597 val alphas = map (nth alpha_frees) ((body_index dt) :: bound_in_ty_nos); |
|
598 val alpha = mk_compound_alpha alphas; |
|
599 val pi = foldr1 (uncurry mk_plus) (distinct (op =) rpis); |
|
600 val alpha_gen_pre = Const (atyp_const atyp, dummyT) $ lhs $ alpha $ fv $ pi $ rhs; |
|
601 val alpha_gen = Syntax.check_term lthy alpha_gen_pre |
|
602 in |
|
603 alpha_gen |
|
604 end |
|
605 else |
|
606 let |
|
607 val alpha_bn_const = |
|
608 nth alpha_bn_frees (find_index (fn (b, _, _) => b = bn) bns) |
|
609 in |
|
610 alpha_bn_const $ arg $ arg2 |
|
611 end |
|
612 | ([], [], relevant, []) => |
|
613 let |
|
614 val (rbinds, rpis) = split_list relevant |
|
615 val lhs_binds = fv_binds args rbinds |
|
616 val lhs = mk_pair (lhs_binds, arg); |
|
617 val rhs_binds = fv_binds args2 rbinds; |
|
618 val rhs = mk_pair (rhs_binds, arg2); |
|
619 val alpha = nth alpha_frees (body_index dt); |
|
620 val fv = nth fv_frees (body_index dt); |
|
621 val pi = foldr1 (uncurry mk_plus) (distinct (op =) rpis); |
|
622 val alpha_const = alpha_const_for_binds rbinds; |
|
623 val alpha_gen_pre = Const (alpha_const, dummyT) $ lhs $ alpha $ fv $ pi $ rhs; |
|
624 val alpha_gen = Syntax.check_term lthy alpha_gen_pre |
|
625 in |
|
626 alpha_gen |
|
627 end |
|
628 | _ => error "Fv.alpha: not supported binding structure" |
|
629 end |
|
630 val alphas = map alpha_arg (dts ~~ arg_nos ~~ (args ~~ args2)) |
|
631 val alpha_lhss = mk_conjl alphas |
|
632 val alpha_lhss_ex = |
|
633 fold (fn pi_str => fn t => HOLogic.mk_exists (pi_str, @{typ perm}, t)) pi_strs alpha_lhss |
|
634 val alpha_eq = Logic.mk_implies (HOLogic.mk_Trueprop alpha_lhss_ex, alpha_rhs) |
|
635 in |
|
636 alpha_eq |
|
637 end; |
|
638 fun alpha_eq (i, (_, _, constrs)) binds = map2i (alpha_constr i) constrs binds; |
|
639 val alphaeqs = map2i alpha_eq descr (gather_binds bindsall) |
|
640 val alpha_eqs = flat alphaeqs |
|
641 val add_binds = map (fn x => (Attrib.empty_binding, x)) |
|
642 val (alphas, lthy') = (Inductive.add_inductive_i |
|
643 {quiet_mode = true, verbose = false, alt_name = Binding.empty, |
|
644 coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false} |
|
645 (map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) (alpha_names @ alpha_bn_names) |
|
646 (alpha_types @ alpha_bn_types)) [] |
|
647 (add_binds (alpha_eqs @ flat alpha_bn_eqs)) [] lthy) |
|
648 in |
|
649 (alphas, lthy') |
|
650 end |
|
651 *} |
|
652 |
|
653 end |