berghofe@5177
|
1 |
(* Title: HOL/Tools/datatype_aux.ML
|
berghofe@5177
|
2 |
ID: $Id$
|
wenzelm@11539
|
3 |
Author: Stefan Berghofer, TU Muenchen
|
berghofe@5177
|
4 |
|
wenzelm@11539
|
5 |
Auxiliary functions for defining datatypes.
|
berghofe@5177
|
6 |
*)
|
berghofe@5177
|
7 |
|
berghofe@5177
|
8 |
signature DATATYPE_AUX =
|
berghofe@5177
|
9 |
sig
|
berghofe@5661
|
10 |
val quiet_mode : bool ref
|
berghofe@5661
|
11 |
val message : string -> unit
|
berghofe@5661
|
12 |
|
berghofe@5177
|
13 |
val foldl1 : ('a * 'a -> 'a) -> 'a list -> 'a
|
berghofe@5177
|
14 |
|
berghofe@5661
|
15 |
val add_path : bool -> string -> theory -> theory
|
berghofe@5661
|
16 |
val parent_path : bool -> theory -> theory
|
berghofe@5661
|
17 |
|
wenzelm@18728
|
18 |
val store_thmss_atts : string -> string list -> attribute list list -> thm list list
|
berghofe@18349
|
19 |
-> theory -> thm list list * theory
|
haftmann@18314
|
20 |
val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory
|
wenzelm@18728
|
21 |
val store_thms_atts : string -> string list -> attribute list list -> thm list
|
haftmann@18314
|
22 |
-> theory -> thm list * theory
|
haftmann@18314
|
23 |
val store_thms : string -> string list -> thm list -> theory -> thm list * theory
|
berghofe@5177
|
24 |
|
berghofe@5177
|
25 |
val split_conj_thm : thm -> thm list
|
berghofe@5177
|
26 |
val mk_conj : term list -> term
|
berghofe@5177
|
27 |
val mk_disj : term list -> term
|
berghofe@5177
|
28 |
|
berghofe@13641
|
29 |
val app_bnds : term -> int -> term
|
berghofe@13641
|
30 |
|
berghofe@13641
|
31 |
val cong_tac : int -> tactic
|
berghofe@5177
|
32 |
val indtac : thm -> int -> tactic
|
berghofe@5177
|
33 |
val exh_tac : (string -> thm) -> int -> tactic
|
berghofe@5177
|
34 |
|
berghofe@7015
|
35 |
datatype simproc_dist = QuickAndDirty
|
berghofe@7015
|
36 |
| FewConstrs of thm list
|
berghofe@7015
|
37 |
| ManyConstrs of thm * simpset;
|
berghofe@7015
|
38 |
|
berghofe@5177
|
39 |
datatype dtyp =
|
berghofe@5177
|
40 |
DtTFree of string
|
berghofe@5177
|
41 |
| DtType of string * (dtyp list)
|
berghofe@5177
|
42 |
| DtRec of int;
|
wenzelm@8404
|
43 |
type descr
|
berghofe@5177
|
44 |
type datatype_info
|
berghofe@5177
|
45 |
|
berghofe@7015
|
46 |
exception Datatype
|
berghofe@14887
|
47 |
exception Datatype_Empty of string
|
berghofe@9740
|
48 |
val name_of_typ : typ -> string
|
berghofe@5177
|
49 |
val dtyp_of_typ : (string * string list) list -> typ -> dtyp
|
berghofe@5177
|
50 |
val mk_Free : string -> typ -> int -> term
|
berghofe@5177
|
51 |
val is_rec_type : dtyp -> bool
|
berghofe@13641
|
52 |
val typ_of_dtyp : descr -> (string * sort) list -> dtyp -> typ
|
berghofe@5177
|
53 |
val dest_DtTFree : dtyp -> string
|
berghofe@5177
|
54 |
val dest_DtRec : dtyp -> int
|
berghofe@13641
|
55 |
val strip_dtyp : dtyp -> dtyp list * dtyp
|
berghofe@13641
|
56 |
val body_index : dtyp -> int
|
berghofe@13641
|
57 |
val mk_fun_dtyp : dtyp list -> dtyp -> dtyp
|
berghofe@5177
|
58 |
val dest_TFree : typ -> string
|
berghofe@13641
|
59 |
val get_nonrec_types : descr -> (string * sort) list -> typ list
|
berghofe@13641
|
60 |
val get_branching_types : descr -> (string * sort) list -> typ list
|
berghofe@13641
|
61 |
val get_arities : descr -> int list
|
berghofe@13641
|
62 |
val get_rec_types : descr -> (string * sort) list -> typ list
|
berghofe@13641
|
63 |
val check_nonempty : descr list -> unit
|
berghofe@5177
|
64 |
val unfold_datatypes :
|
berghofe@13641
|
65 |
Sign.sg -> descr -> (string * sort) list -> datatype_info Symtab.table ->
|
berghofe@13641
|
66 |
descr -> int -> descr list * int
|
berghofe@5177
|
67 |
end;
|
berghofe@5177
|
68 |
|
berghofe@5177
|
69 |
structure DatatypeAux : DATATYPE_AUX =
|
berghofe@5177
|
70 |
struct
|
berghofe@5177
|
71 |
|
berghofe@5661
|
72 |
val quiet_mode = ref false;
|
berghofe@5661
|
73 |
fun message s = if !quiet_mode then () else writeln s;
|
berghofe@5661
|
74 |
|
berghofe@5177
|
75 |
(* FIXME: move to library ? *)
|
skalberg@15570
|
76 |
fun foldl1 f (x::xs) = Library.foldl f (x, xs);
|
berghofe@5177
|
77 |
|
berghofe@5661
|
78 |
fun add_path flat_names s = if flat_names then I else Theory.add_path s;
|
berghofe@5661
|
79 |
fun parent_path flat_names = if flat_names then I else Theory.parent_path;
|
berghofe@5661
|
80 |
|
wenzelm@8435
|
81 |
|
berghofe@5177
|
82 |
(* store theorems in theory *)
|
berghofe@5177
|
83 |
|
haftmann@18377
|
84 |
fun store_thmss_atts label tnames attss thmss =
|
haftmann@18377
|
85 |
fold_map (fn ((tname, atts), thms) =>
|
haftmann@18377
|
86 |
Theory.add_path tname
|
haftmann@18377
|
87 |
#> PureThy.add_thmss [((label, thms), atts)]
|
haftmann@18377
|
88 |
#-> (fn thm::_ => Theory.parent_path #> pair thm)
|
haftmann@18377
|
89 |
) (tnames ~~ attss ~~ thmss);
|
berghofe@5177
|
90 |
|
berghofe@18101
|
91 |
fun store_thmss label tnames = store_thmss_atts label tnames (replicate (length tnames) []);
|
berghofe@18101
|
92 |
|
haftmann@18377
|
93 |
fun store_thms_atts label tnames attss thmss =
|
haftmann@18377
|
94 |
fold_map (fn ((tname, atts), thms) =>
|
haftmann@18377
|
95 |
Theory.add_path tname
|
haftmann@18377
|
96 |
#> PureThy.add_thms [((label, thms), atts)]
|
haftmann@18377
|
97 |
#-> (fn thm::_ => Theory.parent_path #> pair thm)
|
haftmann@18377
|
98 |
) (tnames ~~ attss ~~ thmss);
|
wenzelm@8435
|
99 |
|
wenzelm@8435
|
100 |
fun store_thms label tnames = store_thms_atts label tnames (replicate (length tnames) []);
|
wenzelm@8435
|
101 |
|
berghofe@5177
|
102 |
|
berghofe@5177
|
103 |
(* split theorem thm_1 & ... & thm_n into n theorems *)
|
berghofe@5177
|
104 |
|
berghofe@5177
|
105 |
fun split_conj_thm th =
|
berghofe@7015
|
106 |
((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
|
berghofe@5177
|
107 |
|
berghofe@5177
|
108 |
val mk_conj = foldr1 (HOLogic.mk_binop "op &");
|
berghofe@5177
|
109 |
val mk_disj = foldr1 (HOLogic.mk_binop "op |");
|
berghofe@5177
|
110 |
|
berghofe@13641
|
111 |
fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
|
berghofe@13641
|
112 |
|
berghofe@13641
|
113 |
|
berghofe@13641
|
114 |
fun cong_tac i st = (case Logic.strip_assums_concl
|
skalberg@15570
|
115 |
(List.nth (prems_of st, i - 1)) of
|
berghofe@13641
|
116 |
_ $ (_ $ (f $ x) $ (g $ y)) =>
|
berghofe@13641
|
117 |
let
|
wenzelm@18145
|
118 |
val cong' = Thm.lift_rule (Thm.cprem_of st i) cong;
|
berghofe@13641
|
119 |
val _ $ (_ $ (f' $ x') $ (g' $ y')) =
|
berghofe@13641
|
120 |
Logic.strip_assums_concl (prop_of cong');
|
berghofe@13641
|
121 |
val insts = map (pairself (cterm_of (#sign (rep_thm st))) o
|
berghofe@13641
|
122 |
apsnd (curry list_abs (Logic.strip_params (concl_of cong'))) o
|
berghofe@13641
|
123 |
apfst head_of) [(f', f), (g', g), (x', x), (y', y)]
|
berghofe@13641
|
124 |
in compose_tac (false, cterm_instantiate insts cong', 2) i st
|
berghofe@13641
|
125 |
handle THM _ => no_tac st
|
berghofe@13641
|
126 |
end
|
berghofe@13641
|
127 |
| _ => no_tac st);
|
berghofe@5177
|
128 |
|
berghofe@5177
|
129 |
(* instantiate induction rule *)
|
berghofe@5177
|
130 |
|
berghofe@5177
|
131 |
fun indtac indrule i st =
|
berghofe@5177
|
132 |
let
|
wenzelm@8305
|
133 |
val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
|
wenzelm@8305
|
134 |
val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
|
skalberg@15570
|
135 |
(Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
|
berghofe@5177
|
136 |
val getP = if can HOLogic.dest_imp (hd ts) then
|
skalberg@15531
|
137 |
(apfst SOME) o HOLogic.dest_imp else pair NONE;
|
berghofe@5177
|
138 |
fun abstr (t1, t2) = (case t1 of
|
skalberg@15531
|
139 |
NONE => let val [Free (s, T)] = add_term_frees (t2, [])
|
berghofe@5177
|
140 |
in absfree (s, T, t2) end
|
skalberg@15531
|
141 |
| SOME (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2)))
|
wenzelm@6394
|
142 |
val cert = cterm_of (Thm.sign_of_thm st);
|
berghofe@5177
|
143 |
val Ps = map (cert o head_of o snd o getP) ts;
|
berghofe@5177
|
144 |
val indrule' = cterm_instantiate (Ps ~~
|
berghofe@5177
|
145 |
(map (cert o abstr o getP) ts')) indrule
|
berghofe@5177
|
146 |
in
|
berghofe@5177
|
147 |
rtac indrule' i st
|
berghofe@5177
|
148 |
end;
|
berghofe@5177
|
149 |
|
berghofe@5177
|
150 |
(* perform exhaustive case analysis on last parameter of subgoal i *)
|
berghofe@5177
|
151 |
|
berghofe@5177
|
152 |
fun exh_tac exh_thm_of i state =
|
berghofe@5177
|
153 |
let
|
wenzelm@6394
|
154 |
val sg = Thm.sign_of_thm state;
|
skalberg@15570
|
155 |
val prem = List.nth (prems_of state, i - 1);
|
berghofe@5177
|
156 |
val params = Logic.strip_params prem;
|
berghofe@5177
|
157 |
val (_, Type (tname, _)) = hd (rev params);
|
wenzelm@18145
|
158 |
val exhaustion = Thm.lift_rule (Thm.cprem_of state i) (exh_thm_of tname);
|
berghofe@5177
|
159 |
val prem' = hd (prems_of exhaustion);
|
berghofe@5177
|
160 |
val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
|
berghofe@5177
|
161 |
val exhaustion' = cterm_instantiate [(cterm_of sg (head_of lhs),
|
skalberg@15574
|
162 |
cterm_of sg (foldr (fn ((_, T), t) => Abs ("z", T, t))
|
skalberg@15574
|
163 |
(Bound 0) params))] exhaustion
|
berghofe@5177
|
164 |
in compose_tac (false, exhaustion', nprems_of exhaustion) i state
|
berghofe@5177
|
165 |
end;
|
berghofe@5177
|
166 |
|
berghofe@7015
|
167 |
(* handling of distinctness theorems *)
|
berghofe@7015
|
168 |
|
berghofe@7015
|
169 |
datatype simproc_dist = QuickAndDirty
|
berghofe@7015
|
170 |
| FewConstrs of thm list
|
berghofe@7015
|
171 |
| ManyConstrs of thm * simpset;
|
berghofe@7015
|
172 |
|
berghofe@5177
|
173 |
(********************** Internal description of datatypes *********************)
|
berghofe@5177
|
174 |
|
berghofe@5177
|
175 |
datatype dtyp =
|
berghofe@5177
|
176 |
DtTFree of string
|
berghofe@5177
|
177 |
| DtType of string * (dtyp list)
|
berghofe@5177
|
178 |
| DtRec of int;
|
berghofe@5177
|
179 |
|
berghofe@5177
|
180 |
(* information about datatypes *)
|
berghofe@5177
|
181 |
|
berghofe@16901
|
182 |
(* index, datatype name, type arguments, constructor name, types of constructor's arguments *)
|
wenzelm@8404
|
183 |
type descr = (int * (string * dtyp list * (string * dtyp list) list)) list;
|
wenzelm@8404
|
184 |
|
berghofe@5177
|
185 |
type datatype_info =
|
berghofe@5177
|
186 |
{index : int,
|
wenzelm@8404
|
187 |
descr : descr,
|
berghofe@18319
|
188 |
sorts : (string * sort) list,
|
berghofe@5177
|
189 |
rec_names : string list,
|
berghofe@5177
|
190 |
rec_rewrites : thm list,
|
berghofe@5177
|
191 |
case_name : string,
|
berghofe@5177
|
192 |
case_rewrites : thm list,
|
berghofe@5177
|
193 |
induction : thm,
|
berghofe@5177
|
194 |
exhaustion : thm,
|
berghofe@7015
|
195 |
distinct : simproc_dist,
|
berghofe@5177
|
196 |
inject : thm list,
|
berghofe@5177
|
197 |
nchotomy : thm,
|
wenzelm@10120
|
198 |
case_cong : thm,
|
wenzelm@10120
|
199 |
weak_case_cong : thm};
|
berghofe@5177
|
200 |
|
berghofe@5177
|
201 |
fun mk_Free s T i = Free (s ^ (string_of_int i), T);
|
berghofe@5177
|
202 |
|
berghofe@5177
|
203 |
fun subst_DtTFree _ substs (T as (DtTFree name)) =
|
haftmann@17485
|
204 |
AList.lookup (op =) substs name |> the_default T
|
berghofe@5177
|
205 |
| subst_DtTFree i substs (DtType (name, ts)) =
|
berghofe@5177
|
206 |
DtType (name, map (subst_DtTFree i substs) ts)
|
berghofe@5177
|
207 |
| subst_DtTFree i _ (DtRec j) = DtRec (i + j);
|
berghofe@5177
|
208 |
|
berghofe@7015
|
209 |
exception Datatype;
|
berghofe@14887
|
210 |
exception Datatype_Empty of string;
|
berghofe@7015
|
211 |
|
berghofe@7015
|
212 |
fun dest_DtTFree (DtTFree a) = a
|
berghofe@7015
|
213 |
| dest_DtTFree _ = raise Datatype;
|
berghofe@7015
|
214 |
|
berghofe@7015
|
215 |
fun dest_DtRec (DtRec i) = i
|
berghofe@7015
|
216 |
| dest_DtRec _ = raise Datatype;
|
berghofe@5177
|
217 |
|
berghofe@5177
|
218 |
fun is_rec_type (DtType (_, dts)) = exists is_rec_type dts
|
berghofe@5177
|
219 |
| is_rec_type (DtRec _) = true
|
berghofe@5177
|
220 |
| is_rec_type _ = false;
|
berghofe@5177
|
221 |
|
berghofe@13641
|
222 |
fun strip_dtyp (DtType ("fun", [T, U])) = apfst (cons T) (strip_dtyp U)
|
berghofe@13641
|
223 |
| strip_dtyp T = ([], T);
|
berghofe@13641
|
224 |
|
berghofe@13641
|
225 |
val body_index = dest_DtRec o snd o strip_dtyp;
|
berghofe@13641
|
226 |
|
berghofe@13641
|
227 |
fun mk_fun_dtyp [] U = U
|
berghofe@13641
|
228 |
| mk_fun_dtyp (T :: Ts) U = DtType ("fun", [T, mk_fun_dtyp Ts U]);
|
berghofe@13641
|
229 |
|
berghofe@5177
|
230 |
fun dest_TFree (TFree (n, _)) = n;
|
berghofe@5177
|
231 |
|
berghofe@13707
|
232 |
fun name_of_typ (Type (s, Ts)) =
|
berghofe@13707
|
233 |
let val s' = Sign.base_name s
|
skalberg@15570
|
234 |
in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
|
wenzelm@14673
|
235 |
[if Syntax.is_identifier s' then s' else "x"])
|
berghofe@13707
|
236 |
end
|
berghofe@9740
|
237 |
| name_of_typ _ = "";
|
berghofe@9740
|
238 |
|
berghofe@5177
|
239 |
fun dtyp_of_typ _ (TFree (n, _)) = DtTFree n
|
berghofe@5177
|
240 |
| dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
|
berghofe@5177
|
241 |
| dtyp_of_typ new_dts (Type (tname, Ts)) =
|
haftmann@17485
|
242 |
(case AList.lookup (op =) new_dts tname of
|
skalberg@15531
|
243 |
NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
|
skalberg@15531
|
244 |
| SOME vs => if map (try dest_TFree) Ts = map SOME vs then
|
berghofe@5177
|
245 |
DtRec (find_index (curry op = tname o fst) new_dts)
|
huffman@18069
|
246 |
else error ("Illegal occurrence of recursive type " ^ tname));
|
berghofe@5177
|
247 |
|
haftmann@17485
|
248 |
fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, (the o AList.lookup (op =) sorts) a)
|
berghofe@5177
|
249 |
| typ_of_dtyp descr sorts (DtRec i) =
|
haftmann@17485
|
250 |
let val (s, ds, _) = (the o AList.lookup (op =) descr) i
|
berghofe@5177
|
251 |
in Type (s, map (typ_of_dtyp descr sorts) ds) end
|
berghofe@5177
|
252 |
| typ_of_dtyp descr sorts (DtType (s, ds)) =
|
berghofe@5177
|
253 |
Type (s, map (typ_of_dtyp descr sorts) ds);
|
berghofe@5177
|
254 |
|
berghofe@5177
|
255 |
(* find all non-recursive types in datatype description *)
|
berghofe@5177
|
256 |
|
berghofe@5177
|
257 |
fun get_nonrec_types descr sorts =
|
skalberg@15570
|
258 |
map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
|
skalberg@15570
|
259 |
Library.foldl (fn (Ts', (_, cargs)) =>
|
berghofe@13641
|
260 |
filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr));
|
berghofe@5177
|
261 |
|
berghofe@5177
|
262 |
(* get all recursive types in datatype description *)
|
berghofe@5177
|
263 |
|
berghofe@5177
|
264 |
fun get_rec_types descr sorts = map (fn (_ , (s, ds, _)) =>
|
berghofe@5177
|
265 |
Type (s, map (typ_of_dtyp descr sorts) ds)) descr;
|
berghofe@5177
|
266 |
|
berghofe@7015
|
267 |
(* get all branching types *)
|
berghofe@7015
|
268 |
|
berghofe@13641
|
269 |
fun get_branching_types descr sorts =
|
skalberg@15570
|
270 |
map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
|
skalberg@15574
|
271 |
Library.foldl (fn (Ts', (_, cargs)) => foldr op union Ts' (map (fst o strip_dtyp)
|
skalberg@15574
|
272 |
cargs)) (Ts, constrs)) ([], descr));
|
berghofe@13641
|
273 |
|
skalberg@15570
|
274 |
fun get_arities descr = Library.foldl (fn (is, (_, (_, _, constrs))) =>
|
skalberg@15570
|
275 |
Library.foldl (fn (is', (_, cargs)) => map (length o fst o strip_dtyp)
|
skalberg@15570
|
276 |
(List.filter is_rec_type cargs) union is') (is, constrs)) ([], descr);
|
berghofe@7015
|
277 |
|
berghofe@5177
|
278 |
(* nonemptiness check for datatypes *)
|
berghofe@5177
|
279 |
|
berghofe@5177
|
280 |
fun check_nonempty descr =
|
berghofe@5177
|
281 |
let
|
skalberg@15570
|
282 |
val descr' = List.concat descr;
|
berghofe@5177
|
283 |
fun is_nonempty_dt is i =
|
berghofe@5177
|
284 |
let
|
haftmann@17485
|
285 |
val (_, _, constrs) = (the o AList.lookup (op =) descr') i;
|
berghofe@13641
|
286 |
fun arg_nonempty (_, DtRec i) = if i mem is then false
|
berghofe@5177
|
287 |
else is_nonempty_dt (i::is) i
|
berghofe@5177
|
288 |
| arg_nonempty _ = true;
|
berghofe@13641
|
289 |
in exists ((forall (arg_nonempty o strip_dtyp)) o snd) constrs
|
berghofe@5177
|
290 |
end
|
berghofe@5177
|
291 |
in assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
|
berghofe@14887
|
292 |
(fn (_, (s, _, _)) => raise Datatype_Empty s)
|
berghofe@5177
|
293 |
end;
|
berghofe@5177
|
294 |
|
berghofe@5177
|
295 |
(* unfold a list of mutually recursive datatype specifications *)
|
berghofe@5177
|
296 |
(* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
|
berghofe@5177
|
297 |
(* need to be unfolded *)
|
berghofe@5177
|
298 |
|
berghofe@7015
|
299 |
fun unfold_datatypes sign orig_descr sorts (dt_info : datatype_info Symtab.table) descr i =
|
berghofe@5177
|
300 |
let
|
berghofe@7015
|
301 |
fun typ_error T msg = error ("Non-admissible type expression\n" ^
|
berghofe@7015
|
302 |
Sign.string_of_typ sign (typ_of_dtyp (orig_descr @ descr) sorts T) ^ "\n" ^ msg);
|
berghofe@7015
|
303 |
|
berghofe@7015
|
304 |
fun get_dt_descr T i tname dts =
|
wenzelm@17412
|
305 |
(case Symtab.lookup dt_info tname of
|
skalberg@15531
|
306 |
NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
|
berghofe@7015
|
307 |
\ nested recursion")
|
skalberg@15531
|
308 |
| (SOME {index, descr, ...}) =>
|
haftmann@17485
|
309 |
let val (_, vars, _) = (the o AList.lookup (op =) descr) index;
|
skalberg@15570
|
310 |
val subst = ((map dest_DtTFree vars) ~~ dts) handle UnequalLengths =>
|
berghofe@7015
|
311 |
typ_error T ("Type constructor " ^ tname ^ " used with wrong\
|
berghofe@5177
|
312 |
\ number of arguments")
|
berghofe@5177
|
313 |
in (i + index, map (fn (j, (tn, args, cs)) => (i + j,
|
berghofe@5177
|
314 |
(tn, map (subst_DtTFree i subst) args,
|
berghofe@5177
|
315 |
map (apsnd (map (subst_DtTFree i subst))) cs))) descr)
|
berghofe@5177
|
316 |
end);
|
berghofe@5177
|
317 |
|
berghofe@5177
|
318 |
(* unfold a single constructor argument *)
|
berghofe@5177
|
319 |
|
berghofe@13641
|
320 |
fun unfold_arg ((i, Ts, descrs), T) =
|
berghofe@13641
|
321 |
if is_rec_type T then
|
berghofe@13641
|
322 |
let val (Us, U) = strip_dtyp T
|
berghofe@13641
|
323 |
in if exists is_rec_type Us then
|
berghofe@13641
|
324 |
typ_error T "Non-strictly positive recursive occurrence of type"
|
berghofe@13641
|
325 |
else (case U of
|
berghofe@13641
|
326 |
DtType (tname, dts) =>
|
berghofe@13641
|
327 |
let
|
berghofe@13641
|
328 |
val (index, descr) = get_dt_descr T i tname dts;
|
berghofe@13641
|
329 |
val (descr', i') = unfold_datatypes sign orig_descr sorts
|
berghofe@13641
|
330 |
dt_info descr (i + length descr)
|
berghofe@13641
|
331 |
in (i', Ts @ [mk_fun_dtyp Us (DtRec index)], descrs @ descr') end
|
berghofe@13641
|
332 |
| _ => (i, Ts @ [T], descrs))
|
berghofe@13641
|
333 |
end
|
berghofe@13641
|
334 |
else (i, Ts @ [T], descrs);
|
berghofe@5177
|
335 |
|
berghofe@5177
|
336 |
(* unfold a constructor *)
|
berghofe@5177
|
337 |
|
berghofe@5177
|
338 |
fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
|
skalberg@15570
|
339 |
let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
|
berghofe@5177
|
340 |
in (i', constrs @ [(cname, cargs')], descrs') end;
|
berghofe@5177
|
341 |
|
berghofe@5177
|
342 |
(* unfold a single datatype *)
|
berghofe@5177
|
343 |
|
berghofe@5177
|
344 |
fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
|
berghofe@5177
|
345 |
let val (i', constrs', descrs') =
|
skalberg@15570
|
346 |
Library.foldl unfold_constr ((i, [], descrs), constrs)
|
berghofe@5177
|
347 |
in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
|
berghofe@5177
|
348 |
end;
|
berghofe@5177
|
349 |
|
skalberg@15570
|
350 |
val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
|
berghofe@5177
|
351 |
|
berghofe@5177
|
352 |
in (descr' :: descrs, i') end;
|
berghofe@5177
|
353 |
|
berghofe@5177
|
354 |
end;
|