1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory Fv
|
1806
|
2 |
imports "../Nominal-General/Nominal2_Atoms"
|
|
3 |
"Abs" "Perm" "Rsp" "Nominal2_FSet"
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
begin
|
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
|
1505
|
6 |
(* The bindings data structure:
|
|
7 |
|
|
8 |
Bindings are a list of lists of lists of triples.
|
1172
|
9 |
|
1358
|
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.
|
1172
|
14 |
|
1358
|
15 |
Every triple consists of a function, the binding and
|
|
16 |
the body.
|
1169
|
17 |
|
|
18 |
Eg:
|
|
19 |
nominal_datatype
|
|
20 |
|
|
21 |
C1
|
|
22 |
| C2 x y z bind x in z
|
1172
|
23 |
| C3 x y z bind f x in z bind g y in z
|
1169
|
24 |
|
|
25 |
yields:
|
1172
|
26 |
[
|
|
27 |
[],
|
1358
|
28 |
[(NONE, 0, 2)],
|
|
29 |
[(SOME (Const f), 0, 2), (Some (Const g), 1, 2)]]
|
1185
|
30 |
|
1358
|
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.
|
1505
|
34 |
*)
|
1185
|
35 |
|
1505
|
36 |
(*
|
1510
|
37 |
An overview of the generation of free variables:
|
1505
|
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:
|
1191
|
42 |
|
1505
|
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 |
|
1514
|
54 |
2) fv_ty functions generated for all types being defined:
|
1191
|
55 |
|
1514
|
56 |
fv_ty for a constructor is a union of values for the arguments.
|
1505
|
57 |
|
1508
|
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
|
1505
|
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.
|
1358
|
65 |
|
1505
|
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
|
1169
|
78 |
*)
|
|
79 |
|
1510
|
80 |
(*
|
|
81 |
An overview of the generation of alpha-equivalence:
|
1513
|
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 |
|
1514
|
96 |
2) alpha_ty relations are generated for all the types being defined:
|
|
97 |
|
1516
|
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.
|
1514
|
102 |
|
1516
|
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.
|
1514
|
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
|
1516
|
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 |
|
1514
|
121 |
- Argument is bound in some deep non-recursive bindings.
|
|
122 |
We return: alpha_bn argl argr
|
1516
|
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 |
|
1514
|
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 |
|
1510
|
137 |
*)
|
|
138 |
|
1362
|
139 |
ML {*
|
1911
|
140 |
datatype alpha_mode = AlphaGen | AlphaRes | AlphaLst;
|
1669
|
141 |
*}
|
|
142 |
|
|
143 |
ML {*
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
144 |
fun atyp_const AlphaGen = @{const_name alpha_gen}
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
145 |
| atyp_const AlphaRes = @{const_name alpha_res}
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
146 |
| atyp_const AlphaLst = @{const_name alpha_lst}
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
147 |
*}
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
148 |
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
149 |
(* TODO: make sure that parser checks that bindings are compatible *)
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
150 |
ML {*
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
151 |
fun alpha_const_for_binds [] = atyp_const AlphaGen
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
152 |
| alpha_const_for_binds ((NONE, _, _, at) :: t) = atyp_const at
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
153 |
| alpha_const_for_binds ((SOME (_, _), _, _, at) :: _) = atyp_const at
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
154 |
*}
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
155 |
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
156 |
ML {*
|
1362
|
157 |
fun is_atom thy typ =
|
|
158 |
Sign.of_sort thy (typ, @{sort at})
|
1366
|
159 |
|
|
160 |
fun is_atom_set thy (Type ("fun", [t, @{typ bool}])) = is_atom thy t
|
1669
|
161 |
| is_atom_set _ _ = false;
|
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
162 |
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
163 |
fun is_atom_fset thy (Type ("FSet.fset", [t])) = is_atom thy t
|
1669
|
164 |
| is_atom_fset _ _ = false;
|
1362
|
165 |
*}
|
|
166 |
|
|
167 |
|
1358
|
168 |
(* Like map2, only if the second list is empty passes empty lists insted of error *)
|
1302
|
169 |
ML {*
|
|
170 |
fun map2i _ [] [] = []
|
|
171 |
| map2i f (x :: xs) (y :: ys) = f x y :: map2i f xs ys
|
|
172 |
| map2i f (x :: xs) [] = f x [] :: map2i f xs []
|
|
173 |
| map2i _ _ _ = raise UnequalLengths;
|
|
174 |
*}
|
|
175 |
|
1358
|
176 |
(* Finds bindings with the same function and binding, and gathers all
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
177 |
bodys for such pairs
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
178 |
*)
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
179 |
ML {*
|
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
180 |
fun gather_binds binds =
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
181 |
let
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
182 |
fun gather_binds_cons binds =
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
183 |
let
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
184 |
val common = map (fn (f, bi, _, aty) => (f, bi, aty)) binds
|
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
185 |
val nodups = distinct (op =) common
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
186 |
fun find_bodys (sf, sbi, sty) =
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
187 |
filter (fn (f, bi, _, aty) => f = sf andalso bi = sbi andalso aty = sty) binds
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
188 |
val bodys = map ((map (fn (_, _, bo, _) => bo)) o find_bodys) nodups
|
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
189 |
in
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
190 |
nodups ~~ bodys
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
191 |
end
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
192 |
in
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
193 |
map (map gather_binds_cons) binds
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
194 |
end
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
195 |
*}
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
196 |
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
197 |
ML {*
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
198 |
fun un_gather_binds_cons binds =
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
199 |
flat (map (fn (((f, bi, aty), bos), pi) => map (fn bo => ((f, bi, bo, aty), pi)) bos) binds)
|
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
200 |
*}
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
201 |
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
202 |
ML {*
|
1175
|
203 |
open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *);
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
204 |
*}
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
205 |
ML {*
|
1178
|
206 |
(* TODO: It is the same as one in 'nominal_atoms' *)
|
1175
|
207 |
fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom});
|
|
208 |
val noatoms = @{term "{} :: atom set"};
|
|
209 |
fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x];
|
|
210 |
fun mk_union sets =
|
|
211 |
fold (fn a => fn b =>
|
|
212 |
if a = noatoms then b else
|
|
213 |
if b = noatoms then a else
|
1323
|
214 |
if a = b then a else
|
1325
|
215 |
HOLogic.mk_binop @{const_name sup} (a, b)) (rev sets) noatoms;
|
|
216 |
val mk_inter = foldr1 (HOLogic.mk_binop @{const_name inf})
|
1175
|
217 |
fun mk_diff a b =
|
|
218 |
if b = noatoms then a else
|
|
219 |
if b = a then noatoms else
|
|
220 |
HOLogic.mk_binop @{const_name minus} (a, b);
|
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
221 |
fun mk_atom_set t =
|
1185
|
222 |
let
|
|
223 |
val ty = fastype_of t;
|
|
224 |
val atom_ty = HOLogic.dest_setT ty --> @{typ atom};
|
|
225 |
val img_ty = atom_ty --> ty --> @{typ "atom set"};
|
|
226 |
in
|
|
227 |
(Const (@{const_name image}, img_ty) $ Const (@{const_name atom}, atom_ty) $ t)
|
|
228 |
end;
|
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
229 |
fun mk_atom_fset t =
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
230 |
let
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
231 |
val ty = fastype_of t;
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
232 |
val atom_ty = dest_fsetT ty --> @{typ atom};
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
233 |
val fmap_ty = atom_ty --> ty --> @{typ "atom fset"};
|
1656
|
234 |
val fset_to_set = @{term "fset_to_set :: atom fset \<Rightarrow> atom set"}
|
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
235 |
in
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
236 |
fset_to_set $ ((Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t))
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
237 |
end;
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
238 |
(* Similar to one in USyntax *)
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
239 |
fun mk_pair (fst, snd) =
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
240 |
let val ty1 = fastype_of fst
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
241 |
val ty2 = fastype_of snd
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
242 |
val c = HOLogic.pair_const ty1 ty2
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
243 |
in c $ fst $ snd
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
244 |
end;
|
1468
|
245 |
*}
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
246 |
|
1468
|
247 |
(* Given [fv1, fv2, fv3] creates %(x, y, z). fv1 x u fv2 y u fv3 z *)
|
|
248 |
ML {*
|
|
249 |
fun mk_compound_fv fvs =
|
|
250 |
let
|
|
251 |
val nos = (length fvs - 1) downto 0;
|
|
252 |
val fvs_applied = map (fn (fv, no) => fv $ Bound no) (fvs ~~ nos);
|
|
253 |
val fvs_union = mk_union fvs_applied;
|
|
254 |
val (tyh :: tys) = rev (map (domain_type o fastype_of) fvs);
|
|
255 |
fun fold_fun ty t = HOLogic.mk_split (Abs ("", ty, t))
|
|
256 |
in
|
|
257 |
fold fold_fun tys (Abs ("", tyh, fvs_union))
|
|
258 |
end;
|
1175
|
259 |
*}
|
|
260 |
|
1468
|
261 |
(* Given [R1, R2, R3] creates %(x,x'). %(y,y'). %(z,z'). R x x' \<and> R y y' \<and> R z z' *)
|
|
262 |
ML {*
|
|
263 |
fun mk_compound_alpha Rs =
|
|
264 |
let
|
|
265 |
val nos = (length Rs - 1) downto 0;
|
|
266 |
val nos2 = (2 * length Rs - 1) downto length Rs;
|
|
267 |
val Rs_applied = map (fn (R, (no2, no)) => R $ Bound no2 $ Bound no) (Rs ~~ (nos2 ~~ nos));
|
|
268 |
val Rs_conj = mk_conjl Rs_applied;
|
|
269 |
val (tyh :: tys) = rev (map (domain_type o fastype_of) Rs);
|
|
270 |
fun fold_fun ty t = HOLogic.mk_split (Abs ("", ty, t))
|
|
271 |
val abs_rhs = fold fold_fun tys (Abs ("", tyh, Rs_conj))
|
|
272 |
in
|
|
273 |
fold fold_fun tys (Abs ("", tyh, abs_rhs))
|
|
274 |
end;
|
|
275 |
*}
|
|
276 |
|
1288
|
277 |
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
278 |
ML {*
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
279 |
fun non_rec_binds l =
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
280 |
let
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
281 |
fun is_non_rec (SOME (f, false), _, _, _) = SOME f
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
282 |
| is_non_rec _ = NONE
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
283 |
in
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
284 |
distinct (op =) (map_filter is_non_rec (flat (flat l)))
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
285 |
end
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
286 |
*}
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
287 |
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
288 |
(* We assume no bindings in the type on which bn is defined *)
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
289 |
ML {*
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
290 |
fun fv_bn thy (dt_info : Datatype_Aux.info) fv_frees bn_fvbn (fvbn, (bn, ith_dtyp, args_in_bns)) =
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
291 |
let
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
292 |
val {descr, sorts, ...} = dt_info;
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
293 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
|
1379
|
294 |
fun fv_bn_constr (cname, dts) args_in_bn =
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
295 |
let
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
296 |
val Ts = map (typ_of_dtyp descr sorts) dts;
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
297 |
val names = Datatype_Prop.make_tnames Ts;
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
298 |
val args = map Free (names ~~ Ts);
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
299 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp));
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
300 |
fun fv_arg ((dt, x), arg_no) =
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
301 |
let
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
302 |
val ty = fastype_of x
|
1622
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
303 |
(* val _ = tracing ("B 1" ^ PolyML.makestring args_in_bn);*)
|
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
304 |
(* val _ = tracing ("B 2" ^ PolyML.makestring bn_fvbn);*)
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
305 |
in
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
306 |
case AList.lookup (op=) args_in_bn arg_no of
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
307 |
SOME NONE => @{term "{} :: atom set"}
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
308 |
| SOME (SOME (f : term)) => (the (AList.lookup (op=) bn_fvbn f)) $ x
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
309 |
| NONE =>
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
310 |
if is_atom thy ty then mk_single_atom x else
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
311 |
if is_atom_set thy ty then mk_atom_set x else
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
312 |
if is_atom_fset thy ty then mk_atom_fset x else
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
313 |
if is_rec_type dt then nth fv_frees (body_index dt) $ x else
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
314 |
@{term "{} :: atom set"}
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
315 |
end;
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
316 |
val arg_nos = 0 upto (length dts - 1)
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
317 |
in
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
318 |
HOLogic.mk_Trueprop (HOLogic.mk_eq
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
319 |
(fvbn $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos))))
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
320 |
end;
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
321 |
val (_, (_, _, constrs)) = nth descr ith_dtyp;
|
1379
|
322 |
val eqs = map2i fv_bn_constr constrs args_in_bns
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
323 |
in
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
324 |
((bn, fvbn), eqs)
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
325 |
end
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
326 |
*}
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
327 |
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
328 |
ML {* print_depth 100 *}
|
1385
|
329 |
ML {*
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
330 |
fun fv_bns thy dt_info fv_frees rel_bns =
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
331 |
let
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
332 |
fun mk_fvbn_free (bn, ith, _) =
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
333 |
let
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
334 |
val fvbn_name = "fv_" ^ (Long_Name.base_name (fst (dest_Const bn)));
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
335 |
in
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
336 |
(fvbn_name, Free (fvbn_name, fastype_of (nth fv_frees ith)))
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
337 |
end;
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
338 |
val (fvbn_names, fvbn_frees) = split_list (map mk_fvbn_free rel_bns);
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
339 |
val bn_fvbn = (map (fn (bn, _, _) => bn) rel_bns) ~~ fvbn_frees
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
340 |
val (l1, l2) = split_list (map (fv_bn thy dt_info fv_frees bn_fvbn) (fvbn_frees ~~ rel_bns));
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
341 |
in
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
342 |
(l1, (fvbn_names ~~ l2))
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
343 |
end
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
344 |
*}
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
345 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
346 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
347 |
ML {*
|
1669
|
348 |
fun alpha_bn (dt_info : Datatype_Aux.info) alpha_frees bn_alphabn ((bn, ith_dtyp, args_in_bns), (alpha_bn_free, _ (*is_rec*) )) =
|
1385
|
349 |
let
|
|
350 |
val {descr, sorts, ...} = dt_info;
|
|
351 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
|
1386
|
352 |
fun alpha_bn_constr (cname, dts) args_in_bn =
|
|
353 |
let
|
|
354 |
val Ts = map (typ_of_dtyp descr sorts) dts;
|
|
355 |
val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
|
|
356 |
val names2 = Name.variant_list ("pi" :: names) (Datatype_Prop.make_tnames Ts);
|
|
357 |
val args = map Free (names ~~ Ts);
|
|
358 |
val args2 = map Free (names2 ~~ Ts);
|
|
359 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp));
|
1387
|
360 |
val rhs = HOLogic.mk_Trueprop
|
1462
|
361 |
(alpha_bn_free $ (list_comb (c, args)) $ (list_comb (c, args2)));
|
1387
|
362 |
fun lhs_arg ((dt, arg_no), (arg, arg2)) =
|
1669
|
363 |
case AList.lookup (op=) args_in_bn arg_no of
|
|
364 |
SOME NONE => @{term True}
|
|
365 |
| SOME (SOME f) => (the (AList.lookup (op=) bn_alphabn f)) $ arg $ arg2
|
|
366 |
| NONE =>
|
|
367 |
if is_rec_type dt then (nth alpha_frees (body_index dt)) $ arg $ arg2
|
|
368 |
else HOLogic.mk_eq (arg, arg2)
|
1387
|
369 |
val arg_nos = 0 upto (length dts - 1)
|
|
370 |
val lhss = mk_conjl (map lhs_arg (dts ~~ arg_nos ~~ (args ~~ args2)))
|
|
371 |
val eq = Logic.mk_implies (HOLogic.mk_Trueprop lhss, rhs)
|
1386
|
372 |
in
|
1387
|
373 |
eq
|
1386
|
374 |
end
|
1385
|
375 |
val (_, (_, _, constrs)) = nth descr ith_dtyp;
|
1386
|
376 |
val eqs = map2i alpha_bn_constr constrs args_in_bns
|
1385
|
377 |
in
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
378 |
((bn, alpha_bn_free), eqs)
|
1385
|
379 |
end
|
|
380 |
*}
|
|
381 |
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
382 |
ML {*
|
1669
|
383 |
fun alpha_bns dt_info alpha_frees rel_bns bns_rec =
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
384 |
let
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
385 |
val {descr, sorts, ...} = dt_info;
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
386 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
387 |
fun mk_alphabn_free (bn, ith, _) =
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
388 |
let
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
389 |
val alphabn_name = "alpha_" ^ (Long_Name.base_name (fst (dest_Const bn)));
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
390 |
val alphabn_type = nth_dtyp ith --> nth_dtyp ith --> @{typ bool};
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
391 |
val alphabn_free = Free(alphabn_name, alphabn_type);
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
392 |
in
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
393 |
(alphabn_name, alphabn_free)
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
394 |
end;
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
395 |
val (alphabn_names, alphabn_frees) = split_list (map mk_alphabn_free rel_bns);
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
396 |
val bn_alphabn = (map (fn (bn, _, _) => bn) rel_bns) ~~ alphabn_frees;
|
1669
|
397 |
val pair = split_list (map (alpha_bn dt_info alpha_frees bn_alphabn)
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
398 |
(rel_bns ~~ (alphabn_frees ~~ bns_rec)))
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
399 |
in
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
400 |
(alphabn_names, pair)
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
401 |
end
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
402 |
*}
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
403 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
404 |
|
1397
|
405 |
(* Checks that a list of bindings contains only compatible ones *)
|
|
406 |
ML {*
|
|
407 |
fun bns_same l =
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
408 |
length (distinct (op =) (map (fn ((b, _, _, atyp), _) => (b, atyp)) l)) = 1
|
1397
|
409 |
*}
|
|
410 |
|
1679
|
411 |
ML {*
|
|
412 |
fun setify x =
|
|
413 |
if fastype_of x = @{typ "atom list"} then
|
|
414 |
Const (@{const_name set}, @{typ "atom list \<Rightarrow> atom set"}) $ x else x
|
|
415 |
*}
|
|
416 |
|
1175
|
417 |
ML {*
|
1836
|
418 |
fun define_fv (dt_info : Datatype_Aux.info) bindsall bns lthy =
|
|
419 |
let
|
|
420 |
val thy = ProofContext.theory_of lthy;
|
|
421 |
val {descr, sorts, ...} = dt_info;
|
|
422 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
|
|
423 |
val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
|
|
424 |
"fv_" ^ name_of_typ (nth_dtyp i)) descr);
|
|
425 |
val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr;
|
|
426 |
val fv_frees = map Free (fv_names ~~ fv_types);
|
|
427 |
(* TODO: We need a transitive closure, but instead we do this hack considering
|
|
428 |
all binding functions as recursive or not *)
|
|
429 |
val nr_bns =
|
|
430 |
if (non_rec_binds bindsall) = [] then []
|
|
431 |
else map (fn (bn, _, _) => bn) bns;
|
|
432 |
val rel_bns = filter (fn (bn, _, _) => bn mem nr_bns) bns;
|
|
433 |
val (bn_fv_bns, fv_bn_names_eqs) = fv_bns thy dt_info fv_frees rel_bns;
|
|
434 |
val fvbns = map snd bn_fv_bns;
|
|
435 |
val (fv_bn_names, fv_bn_eqs) = split_list fv_bn_names_eqs;
|
|
436 |
|
|
437 |
fun fv_constr ith_dtyp (cname, dts) bindcs =
|
|
438 |
let
|
|
439 |
val Ts = map (typ_of_dtyp descr sorts) dts;
|
|
440 |
val bindslen = length bindcs
|
|
441 |
val pi_strs_same = replicate bindslen "pi"
|
|
442 |
val pi_strs = Name.variant_list [] pi_strs_same;
|
|
443 |
val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs;
|
|
444 |
val bind_pis_gath = bindcs ~~ pis;
|
|
445 |
val bind_pis = un_gather_binds_cons bind_pis_gath;
|
|
446 |
val bindcs = map fst bind_pis;
|
|
447 |
val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts);
|
|
448 |
val args = map Free (names ~~ Ts);
|
|
449 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp));
|
|
450 |
val fv_c = nth fv_frees ith_dtyp;
|
|
451 |
val arg_nos = 0 upto (length dts - 1)
|
|
452 |
fun fv_bind args (NONE, i, _, _) =
|
|
453 |
if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else
|
|
454 |
if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else
|
|
455 |
if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else
|
|
456 |
if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else
|
|
457 |
(* TODO goes the code for preiously defined nominal datatypes *)
|
|
458 |
@{term "{} :: atom set"}
|
|
459 |
| fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i)
|
|
460 |
fun fv_binds_as_set args relevant = mk_union (map (setify o fv_bind args) relevant)
|
|
461 |
fun find_nonrec_binder j (SOME (f, false), i, _, _) = if i = j then SOME f else NONE
|
|
462 |
| find_nonrec_binder _ _ = NONE
|
|
463 |
fun fv_arg ((dt, x), arg_no) =
|
|
464 |
case get_first (find_nonrec_binder arg_no) bindcs of
|
|
465 |
SOME f =>
|
|
466 |
(case get_first (fn (x, y) => if x = f then SOME y else NONE) bn_fv_bns of
|
|
467 |
SOME fv_bn => fv_bn $ x
|
|
468 |
| NONE => error "bn specified in a non-rec binding but not in bn list")
|
|
469 |
| NONE =>
|
|
470 |
let
|
|
471 |
val arg =
|
|
472 |
if is_rec_type dt then nth fv_frees (body_index dt) $ x else
|
|
473 |
if ((is_atom thy) o fastype_of) x then mk_single_atom x else
|
|
474 |
if ((is_atom_set thy) o fastype_of) x then mk_atom_set x else
|
|
475 |
if ((is_atom_fset thy) o fastype_of) x then mk_atom_fset x else
|
|
476 |
(* TODO goes the code for preiously defined nominal datatypes *)
|
|
477 |
@{term "{} :: atom set"};
|
|
478 |
(* If i = j then we generate it only once *)
|
|
479 |
val relevant = filter (fn (_, i, j, _) => ((i = arg_no) orelse (j = arg_no))) bindcs;
|
|
480 |
val sub = fv_binds_as_set args relevant
|
|
481 |
in
|
|
482 |
mk_diff arg sub
|
|
483 |
end;
|
|
484 |
val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq
|
|
485 |
(fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos))))
|
|
486 |
in
|
|
487 |
fv_eq
|
|
488 |
end;
|
|
489 |
fun fv_eq (i, (_, _, constrs)) binds = map2i (fv_constr i) constrs binds;
|
|
490 |
val fveqs = map2i fv_eq descr (gather_binds bindsall)
|
|
491 |
val fv_eqs_perfv = fveqs
|
|
492 |
val rel_bns_nos = map (fn (_, i, _) => i) rel_bns;
|
|
493 |
fun filter_fun (_, b) = b mem rel_bns_nos;
|
|
494 |
val all_fvs = (fv_names ~~ fv_eqs_perfv) ~~ (0 upto (length fv_names - 1))
|
|
495 |
val (fv_names_fst, fv_eqs_fst) = apsnd flat (split_list (map fst (filter_out filter_fun all_fvs)))
|
|
496 |
val (fv_names_snd, fv_eqs_snd) = apsnd flat (split_list (map fst (filter filter_fun all_fvs)))
|
|
497 |
val fv_eqs_all = fv_eqs_fst @ (flat fv_bn_eqs);
|
|
498 |
val fv_names_all = fv_names_fst @ fv_bn_names;
|
|
499 |
val add_binds = map (fn x => (Attrib.empty_binding, x))
|
|
500 |
(* Function_Fun.add_fun Function_Common.default_config ... true *)
|
|
501 |
val (fvs, lthy') = (Primrec.add_primrec
|
|
502 |
(map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_all) (add_binds fv_eqs_all) lthy)
|
|
503 |
val (fvs2, lthy'') =
|
|
504 |
if fv_eqs_snd = [] then (([], []), lthy') else
|
|
505 |
(Primrec.add_primrec
|
|
506 |
(map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_snd) (add_binds fv_eqs_snd) lthy')
|
|
507 |
val ordered_fvs = fv_frees @ fvbns;
|
|
508 |
val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2)
|
|
509 |
in
|
|
510 |
((all_fvs, ordered_fvs), lthy'')
|
|
511 |
end
|
|
512 |
*}
|
|
513 |
|
|
514 |
ML {*
|
1838
|
515 |
fun define_alpha (dt_info : Datatype_Aux.info) bindsall bns fv_frees lthy =
|
1178
|
516 |
let
|
1366
|
517 |
val thy = ProofContext.theory_of lthy;
|
1277
|
518 |
val {descr, sorts, ...} = dt_info;
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
519 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
|
1622
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
520 |
(* TODO: We need a transitive closure, but instead we do this hack considering
|
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
521 |
all binding functions as recursive or not *)
|
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
522 |
val nr_bns =
|
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
523 |
if (non_rec_binds bindsall) = [] then []
|
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
524 |
else map (fn (bn, _, _) => bn) bns;
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
525 |
val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
526 |
"alpha_" ^ name_of_typ (nth_dtyp i)) descr);
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
527 |
val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr;
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
528 |
val alpha_frees = map Free (alpha_names ~~ alpha_types);
|
1385
|
529 |
(* We assume that a bn is either recursive or not *)
|
|
530 |
val bns_rec = map (fn (bn, _, _) => not (bn mem nr_bns)) bns;
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
531 |
val (alpha_bn_names, (bn_alpha_bns, alpha_bn_eqs)) =
|
1669
|
532 |
alpha_bns dt_info alpha_frees bns bns_rec
|
1389
|
533 |
val alpha_bn_frees = map snd bn_alpha_bns;
|
|
534 |
val alpha_bn_types = map fastype_of alpha_bn_frees;
|
1679
|
535 |
|
1838
|
536 |
fun alpha_constr ith_dtyp (cname, dts) bindcs =
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
537 |
let
|
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
538 |
val Ts = map (typ_of_dtyp descr sorts) dts;
|
1288
|
539 |
val bindslen = length bindcs
|
1323
|
540 |
val pi_strs_same = replicate bindslen "pi"
|
|
541 |
val pi_strs = Name.variant_list [] pi_strs_same;
|
1288
|
542 |
val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs;
|
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
543 |
val bind_pis_gath = bindcs ~~ pis;
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
544 |
val bind_pis = un_gather_binds_cons bind_pis_gath;
|
1288
|
545 |
val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts);
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
546 |
val args = map Free (names ~~ Ts);
|
1288
|
547 |
val names2 = Name.variant_list (pi_strs @ names) (Datatype_Prop.make_tnames Ts);
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
548 |
val args2 = map Free (names2 ~~ Ts);
|
1288
|
549 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp));
|
|
550 |
val alpha = nth alpha_frees ith_dtyp;
|
|
551 |
val arg_nos = 0 upto (length dts - 1)
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
552 |
fun fv_bind args (NONE, i, _, _) =
|
1177
|
553 |
if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else
|
1366
|
554 |
if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else
|
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
555 |
if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
556 |
if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else
|
1830
|
557 |
(* TODO goes the code for preiously defined nominal datatypes *)
|
1366
|
558 |
@{term "{} :: atom set"}
|
1679
|
559 |
| fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i)
|
1288
|
560 |
fun fv_binds args relevant = mk_union (map (fv_bind args) relevant)
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
561 |
val alpha_rhs =
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
562 |
HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2)));
|
1288
|
563 |
fun alpha_arg ((dt, arg_no), (arg, arg2)) =
|
|
564 |
let
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
565 |
val rel_in_simp_binds = filter (fn ((NONE, i, _, _), _) => i = arg_no | _ => false) bind_pis;
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
566 |
val rel_in_comp_binds = filter (fn ((SOME _, i, _, _), _) => i = arg_no | _ => false) bind_pis;
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
567 |
val rel_has_binds = filter (fn ((NONE, _, j, _), _) => j = arg_no
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
568 |
| ((SOME (_, false), _, j, _), _) => j = arg_no
|
1468
|
569 |
| _ => false) bind_pis;
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
570 |
val rel_has_rec_binds = filter
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
571 |
(fn ((SOME (_, true), _, j, _), _) => j = arg_no | _ => false) bind_pis;
|
1288
|
572 |
in
|
1472
|
573 |
case (rel_in_simp_binds, rel_in_comp_binds, rel_has_binds, rel_has_rec_binds) of
|
|
574 |
([], [], [], []) =>
|
1383
|
575 |
if is_rec_type dt then (nth alpha_frees (body_index dt) $ arg $ arg2)
|
|
576 |
else (HOLogic.mk_eq (arg, arg2))
|
1472
|
577 |
| (_, [], [], []) => @{term True}
|
|
578 |
| ([], [], [], _) => @{term True}
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
579 |
| ([], ((((SOME (bn, is_rec)), _, _, atyp), _) :: _), [], []) =>
|
1462
|
580 |
if not (bns_same rel_in_comp_binds) then error "incompatible bindings for an argument" else
|
|
581 |
if is_rec then
|
|
582 |
let
|
|
583 |
val (rbinds, rpis) = split_list rel_in_comp_binds
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
584 |
val bound_in_nos = map (fn (_, _, i, _) => i) rbinds
|
1468
|
585 |
val bound_in_ty_nos = map (fn i => body_index (nth dts i)) bound_in_nos;
|
|
586 |
val bound_args = arg :: map (nth args) bound_in_nos;
|
|
587 |
val bound_args2 = arg2 :: map (nth args2) bound_in_nos;
|
1462
|
588 |
val lhs_binds = fv_binds args rbinds
|
1468
|
589 |
val lhs_arg = foldr1 HOLogic.mk_prod bound_args
|
|
590 |
val lhs = mk_pair (lhs_binds, lhs_arg);
|
1462
|
591 |
val rhs_binds = fv_binds args2 rbinds;
|
1468
|
592 |
val rhs_arg = foldr1 HOLogic.mk_prod bound_args2;
|
|
593 |
val rhs = mk_pair (rhs_binds, rhs_arg);
|
|
594 |
val fvs = map (nth fv_frees) ((body_index dt) :: bound_in_ty_nos);
|
|
595 |
val fv = mk_compound_fv fvs;
|
|
596 |
val alphas = map (nth alpha_frees) ((body_index dt) :: bound_in_ty_nos);
|
|
597 |
val alpha = mk_compound_alpha alphas;
|
1896
996d4411e95e
tuned; fleshed out some library functions about permutations; closed Datatype_Aux structure (increases readability)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
598 |
val pi = foldr1 (uncurry mk_plus) (distinct (op =) rpis);
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
599 |
val alpha_gen_pre = Const (atyp_const atyp, dummyT) $ lhs $ alpha $ fv $ pi $ rhs;
|
1462
|
600 |
val alpha_gen = Syntax.check_term lthy alpha_gen_pre
|
|
601 |
in
|
|
602 |
alpha_gen
|
|
603 |
end
|
|
604 |
else
|
|
605 |
let
|
|
606 |
val alpha_bn_const =
|
|
607 |
nth alpha_bn_frees (find_index (fn (b, _, _) => b = bn) bns)
|
|
608 |
in
|
|
609 |
alpha_bn_const $ arg $ arg2
|
|
610 |
end
|
1472
|
611 |
| ([], [], relevant, []) =>
|
1383
|
612 |
let
|
1288
|
613 |
val (rbinds, rpis) = split_list relevant
|
|
614 |
val lhs_binds = fv_binds args rbinds
|
|
615 |
val lhs = mk_pair (lhs_binds, arg);
|
|
616 |
val rhs_binds = fv_binds args2 rbinds;
|
|
617 |
val rhs = mk_pair (rhs_binds, arg2);
|
|
618 |
val alpha = nth alpha_frees (body_index dt);
|
|
619 |
val fv = nth fv_frees (body_index dt);
|
1896
996d4411e95e
tuned; fleshed out some library functions about permutations; closed Datatype_Aux structure (increases readability)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
620 |
val pi = foldr1 (uncurry mk_plus) (distinct (op =) rpis);
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
621 |
val alpha_const = alpha_const_for_binds rbinds;
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
622 |
val alpha_gen_pre = Const (alpha_const, dummyT) $ lhs $ alpha $ fv $ pi $ rhs;
|
1325
|
623 |
val alpha_gen = Syntax.check_term lthy alpha_gen_pre
|
1288
|
624 |
in
|
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
625 |
alpha_gen
|
1383
|
626 |
end
|
1385
|
627 |
| _ => error "Fv.alpha: not supported binding structure"
|
1288
|
628 |
end
|
|
629 |
val alphas = map alpha_arg (dts ~~ arg_nos ~~ (args ~~ args2))
|
|
630 |
val alpha_lhss = mk_conjl alphas
|
|
631 |
val alpha_lhss_ex =
|
|
632 |
fold (fn pi_str => fn t => HOLogic.mk_exists (pi_str, @{typ perm}, t)) pi_strs alpha_lhss
|
|
633 |
val alpha_eq = Logic.mk_implies (HOLogic.mk_Trueprop alpha_lhss_ex, alpha_rhs)
|
1173
|
634 |
in
|
1838
|
635 |
alpha_eq
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
636 |
end;
|
1838
|
637 |
fun alpha_eq (i, (_, _, constrs)) binds = map2i (alpha_constr i) constrs binds;
|
|
638 |
val alphaeqs = map2i alpha_eq descr (gather_binds bindsall)
|
|
639 |
val alpha_eqs = flat alphaeqs
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
640 |
val add_binds = map (fn x => (Attrib.empty_binding, x))
|
1838
|
641 |
val (alphas, lthy') = (Inductive.add_inductive_i
|
1325
|
642 |
{quiet_mode = true, verbose = false, alt_name = Binding.empty,
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
643 |
coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false}
|
1389
|
644 |
(map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) (alpha_names @ alpha_bn_names)
|
|
645 |
(alpha_types @ alpha_bn_types)) []
|
1838
|
646 |
(add_binds (alpha_eqs @ flat alpha_bn_eqs)) [] lthy)
|
1178
|
647 |
in
|
1838
|
648 |
(alphas, lthy')
|
1178
|
649 |
end
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
650 |
*}
|
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
651 |
|
1196
|
652 |
end
|