author | paulson |
Tue, 10 Jan 2023 11:06:20 +0000 | |
changeset 76942 | c732fa27b60f |
parent 74561 | 8e6c973003c8 |
child 80634 | a90ab1ea6458 |
permissions | -rw-r--r-- |
68155 | 1 |
(* Author: Pascal Stoop, ETH Zurich |
2 |
Author: Andreas Lochbihler, Digital Asset *) |
|
3 |
||
4 |
signature CODE_LAZY = |
|
5 |
sig |
|
6 |
type lazy_info = |
|
7 |
{eagerT: typ, |
|
8 |
lazyT: typ, |
|
9 |
ctr: term, |
|
10 |
destr: term, |
|
11 |
lazy_ctrs: term list, |
|
12 |
case_lazy: term, |
|
13 |
active: bool, |
|
14 |
activate: theory -> theory, |
|
15 |
deactivate: theory -> theory}; |
|
16 |
val code_lazy_type: string -> theory -> theory |
|
17 |
val activate_lazy_type: string -> theory -> theory |
|
18 |
val deactivate_lazy_type: string -> theory -> theory |
|
19 |
val activate_lazy_types: theory -> theory |
|
20 |
val deactivate_lazy_types: theory -> theory |
|
21 |
||
22 |
val get_lazy_types: theory -> (string * lazy_info) list |
|
23 |
||
24 |
val print_lazy_types: theory -> unit |
|
25 |
||
26 |
val transform_code_eqs: Proof.context -> (thm * bool) list -> (thm * bool) list option |
|
27 |
end; |
|
28 |
||
29 |
structure Code_Lazy : CODE_LAZY = |
|
30 |
struct |
|
31 |
||
32 |
type lazy_info = |
|
33 |
{eagerT: typ, |
|
34 |
lazyT: typ, |
|
35 |
ctr: term, |
|
36 |
destr: term, |
|
37 |
lazy_ctrs: term list, |
|
38 |
case_lazy: term, |
|
39 |
active: bool, |
|
40 |
activate: theory -> theory, |
|
41 |
deactivate: theory -> theory}; |
|
42 |
||
43 |
fun map_active f {eagerT, lazyT, ctr, destr, lazy_ctrs, case_lazy, active, activate, deactivate} = |
|
44 |
{ eagerT = eagerT, |
|
45 |
lazyT = lazyT, |
|
46 |
ctr = ctr, |
|
47 |
destr = destr, |
|
48 |
lazy_ctrs = lazy_ctrs, |
|
49 |
case_lazy = case_lazy, |
|
50 |
active = f active, |
|
51 |
activate = activate, |
|
52 |
deactivate = deactivate |
|
53 |
} |
|
54 |
||
55 |
structure Laziness_Data = Theory_Data( |
|
56 |
type T = lazy_info Symtab.table; |
|
57 |
val empty = Symtab.empty; |
|
58 |
val merge = Symtab.join (fn _ => fn (_, record) => record); |
|
59 |
); |
|
60 |
||
61 |
fun fold_type type' tfree tvar typ = |
|
62 |
let |
|
63 |
fun go (Type (s, Ts)) = type' (s, map go Ts) |
|
64 |
| go (TFree T) = tfree T |
|
65 |
| go (TVar T) = tvar T |
|
66 |
in |
|
67 |
go typ |
|
68 |
end; |
|
69 |
||
70 |
fun read_typ lthy name = |
|
71 |
let |
|
72 |
val (s, Ts) = Proof_Context.read_type_name {proper = true, strict = true} lthy name |> dest_Type |
|
73 |
val (Ts', _) = Ctr_Sugar_Util.mk_TFrees (length Ts) lthy |
|
74 |
in |
|
75 |
Type (s, Ts') |
|
76 |
end |
|
77 |
||
78 |
fun mk_name prefix suffix name ctxt = |
|
79 |
Ctr_Sugar_Util.mk_fresh_names ctxt 1 (prefix ^ name ^ suffix) |>> hd; |
|
80 |
fun generate_typedef_name name ctxt = mk_name "" "_lazy" name ctxt; |
|
81 |
||
82 |
fun add_syntax_definition short_type_name eagerT lazyT lazy_ctr lthy = |
|
83 |
let |
|
84 |
val (name, _) = mk_name "lazy_" "" short_type_name lthy |
|
85 |
val freeT = HOLogic.unitT --> lazyT |
|
69593 | 86 |
val lazyT' = Type (\<^type_name>\<open>lazy\<close>, [lazyT]) |
68155 | 87 |
val def = Logic.all_const freeT $ absdummy freeT (Logic.mk_equals ( |
88 |
Free (name, (freeT --> eagerT)) $ Bound 0, |
|
69593 | 89 |
lazy_ctr $ (Const (\<^const_name>\<open>delay\<close>, (freeT --> lazyT')) $ Bound 0))) |
72450 | 90 |
val lthy' = (snd o Local_Theory.begin_nested) lthy |
68155 | 91 |
val ((t, (_, thm)), lthy') = Specification.definition NONE [] [] |
92 |
((Thm.def_binding (Binding.name name), []), def) lthy' |
|
74338 | 93 |
val lthy' = Local_Theory.notation true ("", false) [(t, Mixfix.mixfix "_")] lthy' |
72450 | 94 |
val lthy = Local_Theory.end_nested lthy' |
68155 | 95 |
val def_thm = singleton (Proof_Context.export lthy' lthy) thm |
96 |
in |
|
97 |
(def_thm, lthy) |
|
98 |
end; |
|
99 |
||
100 |
fun add_ctr_code raw_ctrs case_thms thy = |
|
101 |
let |
|
102 |
fun mk_case_certificate ctxt raw_thms = |
|
103 |
let |
|
104 |
val thms = raw_thms |
|
105 |
|> Conjunction.intr_balanced |
|
106 |
|> Thm.unvarify_global (Proof_Context.theory_of ctxt) |
|
107 |
|> Conjunction.elim_balanced (length raw_thms) |
|
108 |
|> map Simpdata.mk_meta_eq |
|
109 |
|> map Drule.zero_var_indexes |
|
110 |
val thm1 = case thms of |
|
111 |
thm :: _ => thm |
|
112 |
| _ => raise Empty |
|
113 |
val params = Term.add_free_names (Thm.prop_of thm1) []; |
|
114 |
val rhs = thm1 |
|
115 |
|> Thm.prop_of |> Logic.dest_equals |> fst |> strip_comb |
|
116 |
||> fst o split_last |> list_comb |
|
117 |
val lhs = Free (singleton (Name.variant_list params) "case", fastype_of rhs); |
|
118 |
val assum = Thm.cterm_of ctxt (Logic.mk_equals (lhs, rhs)) |
|
119 |
in |
|
120 |
thms |
|
121 |
|> Conjunction.intr_balanced |
|
122 |
|> rewrite_rule ctxt [Thm.symmetric (Thm.assume assum)] |
|
123 |
|> Thm.implies_intr assum |
|
74266 | 124 |
|> Thm.generalize (Names.empty, Names.make_set params) 0 |
68155 | 125 |
|> Axclass.unoverload ctxt |
126 |
|> Thm.varifyT_global |
|
127 |
end |
|
128 |
val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs |
|
129 |
val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs |
|
130 |
in |
|
131 |
if can (Code.constrset_of_consts thy) unover_ctrs then |
|
132 |
thy |
|
133 |
|> Code.declare_datatype_global ctrs |
|
134 |
|> fold_rev (Code.add_eqn_global o rpair true) case_thms |
|
135 |
|> Code.declare_case_global (mk_case_certificate (Proof_Context.init_global thy) case_thms) |
|
136 |
else |
|
137 |
thy |
|
138 |
end; |
|
139 |
||
140 |
fun not_found s = error (s ^ " has not been added as lazy type"); |
|
141 |
||
142 |
fun validate_type_name thy type_name = |
|
143 |
let |
|
144 |
val lthy = Named_Target.theory_init thy |
|
145 |
val eager_type = read_typ lthy type_name |
|
146 |
val type_name = case eager_type of |
|
147 |
Type (s, _) => s |
|
148 |
| _ => raise Match |
|
149 |
in |
|
150 |
type_name |
|
151 |
end; |
|
152 |
||
153 |
fun set_active_lazy_type value eager_type_string thy = |
|
154 |
let |
|
155 |
val type_name = validate_type_name thy eager_type_string |
|
156 |
val x = |
|
157 |
case Symtab.lookup (Laziness_Data.get thy) type_name of |
|
158 |
NONE => not_found type_name |
|
159 |
| SOME x => x |
|
160 |
val new_x = map_active (K value) x |
|
161 |
val thy1 = if value = #active x |
|
162 |
then thy |
|
163 |
else if value |
|
164 |
then #activate x thy |
|
165 |
else #deactivate x thy |
|
166 |
in |
|
167 |
Laziness_Data.map (Symtab.update (type_name, new_x)) thy1 |
|
168 |
end; |
|
169 |
||
170 |
fun set_active_lazy_types value thy = |
|
171 |
let |
|
172 |
val lazy_type_names = Symtab.keys (Laziness_Data.get thy) |
|
173 |
fun fold_fun value type_name thy = |
|
174 |
let |
|
175 |
val x = |
|
176 |
case Symtab.lookup (Laziness_Data.get thy) type_name of |
|
177 |
SOME x => x |
|
178 |
| NONE => raise Match |
|
179 |
val new_x = map_active (K value) x |
|
180 |
val thy1 = if value = #active x |
|
181 |
then thy |
|
182 |
else if value |
|
183 |
then #activate x thy |
|
184 |
else #deactivate x thy |
|
185 |
in |
|
186 |
Laziness_Data.map (Symtab.update (type_name, new_x)) thy1 |
|
187 |
end |
|
188 |
in |
|
189 |
fold (fold_fun value) lazy_type_names thy |
|
190 |
end; |
|
191 |
||
192 |
(* code_lazy_type : string -> theory -> theory *) |
|
193 |
fun code_lazy_type eager_type_string thy = |
|
194 |
let |
|
195 |
val lthy = Named_Target.theory_init thy |
|
196 |
val eagerT = read_typ lthy eager_type_string |
|
197 |
val (type_name, type_args) = dest_Type eagerT |
|
198 |
val short_type_name = Long_Name.base_name type_name |
|
199 |
val _ = if Symtab.defined (Laziness_Data.get thy) type_name |
|
200 |
then error (type_name ^ " has already been added as lazy type.") |
|
201 |
else () |
|
202 |
val {case_thms, casex, ctrs, ...} = case Ctr_Sugar.ctr_sugar_of lthy type_name of |
|
203 |
SOME x => x |
|
204 |
| _ => error (type_name ^ " is not registered with free constructors.") |
|
205 |
||
206 |
fun substitute_ctr (old_T, new_T) ctr_T lthy = |
|
207 |
let |
|
208 |
val old_ctr_vars = map TVar (Term.add_tvarsT ctr_T []) |
|
209 |
val old_ctr_Ts = map TFree (Term.add_tfreesT ctr_T []) @ old_ctr_vars |
|
210 |
val (new_ctr_Ts, lthy') = Ctr_Sugar_Util.mk_TFrees (length old_ctr_Ts) lthy |
|
211 |
||
212 |
fun double_type_fold Ts = case Ts of |
|
213 |
(Type (_, Ts1), Type (_, Ts2)) => flat (map double_type_fold (Ts1 ~~ Ts2)) |
|
214 |
| (Type _, _) => raise Match |
|
215 |
| (_, Type _) => raise Match |
|
216 |
| Ts => [Ts] |
|
217 |
fun map_fun1 f = List.foldr |
|
218 |
(fn ((T1, T2), f) => fn T => if T = T1 then T2 else f T) |
|
219 |
f |
|
220 |
(double_type_fold (old_T, new_T)) |
|
221 |
val map_fun2 = AList.lookup (op =) (old_ctr_Ts ~~ new_ctr_Ts) #> the |
|
222 |
val map_fun = map_fun1 map_fun2 |
|
223 |
||
224 |
val new_ctr_type = fold_type Type (map_fun o TFree) (map_fun o TVar) ctr_T |
|
225 |
in |
|
226 |
(new_ctr_type, lthy') |
|
227 |
end |
|
228 |
||
229 |
val (short_lazy_type_name, lthy1) = generate_typedef_name short_type_name lthy |
|
230 |
||
231 |
fun mk_lazy_typedef (name, eager_type) lthy = |
|
232 |
let |
|
233 |
val binding = Binding.name name |
|
234 |
val (result, lthy1) = (Typedef.add_typedef |
|
235 |
{ overloaded=false } |
|
236 |
(binding, rev (Term.add_tfreesT eager_type []), Mixfix.NoSyn) |
|
69593 | 237 |
(Const (\<^const_name>\<open>top\<close>, Type (\<^type_name>\<open>set\<close>, [eager_type]))) |
68155 | 238 |
NONE |
239 |
(fn ctxt => resolve_tac ctxt [@{thm UNIV_witness}] 1) |
|
72450 | 240 |
o (snd o Local_Theory.begin_nested)) lthy |
68155 | 241 |
in |
72450 | 242 |
(binding, result, Local_Theory.end_nested lthy1) |
68155 | 243 |
end; |
244 |
||
245 |
val (typedef_binding, (_, info), lthy2) = mk_lazy_typedef (short_lazy_type_name, eagerT) lthy1 |
|
246 |
||
247 |
val lazy_type = Type (Local_Theory.full_name lthy2 typedef_binding, type_args) |
|
248 |
val (Abs_lazy, Rep_lazy) = |
|
249 |
let |
|
250 |
val info = fst info |
|
251 |
val Abs_name = #Abs_name info |
|
252 |
val Rep_name = #Rep_name info |
|
253 |
val Abs_type = eagerT --> lazy_type |
|
254 |
val Rep_type = lazy_type --> eagerT |
|
255 |
in |
|
256 |
(Const (Abs_name, Abs_type), Const (Rep_name, Rep_type)) |
|
257 |
end |
|
258 |
val Abs_inverse = #Abs_inverse (snd info) |
|
259 |
val Rep_inverse = #Rep_inverse (snd info) |
|
260 |
||
261 |
val (ctrs', lthy3) = List.foldr |
|
262 |
(fn (Const (s, T), (ctrs, lthy)) => let |
|
263 |
val (T', lthy') = substitute_ctr (body_type T, eagerT) T lthy |
|
264 |
in |
|
265 |
((Const (s, T')) :: ctrs, lthy') |
|
266 |
end |
|
267 |
) |
|
268 |
([], lthy2) |
|
269 |
ctrs |
|
270 |
||
271 |
fun to_destr_eagerT typ = case typ of |
|
69593 | 272 |
Type (\<^type_name>\<open>fun\<close>, [_, Type (\<^type_name>\<open>fun\<close>, Ts)]) => |
273 |
to_destr_eagerT (Type (\<^type_name>\<open>fun\<close>, Ts)) |
|
274 |
| Type (\<^type_name>\<open>fun\<close>, [T, _]) => T |
|
68155 | 275 |
| _ => raise Match |
276 |
val (case', lthy4) = |
|
277 |
let |
|
278 |
val (eager_case, caseT) = dest_Const casex |
|
279 |
val (caseT', lthy') = substitute_ctr (to_destr_eagerT caseT, eagerT) caseT lthy3 |
|
280 |
in (Const (eager_case, caseT'), lthy') end |
|
281 |
||
282 |
val ctr_names = map (Long_Name.base_name o fst o dest_Const) ctrs |
|
283 |
val ((((lazy_ctr_name, lazy_sel_name), lazy_ctrs_name), lazy_case_name), ctxt) = lthy4 |
|
284 |
|> mk_name "Lazy_" "" short_type_name |
|
285 |
||>> mk_name "unlazy_" "" short_type_name |
|
286 |
||>> fold_map (mk_name "" "_Lazy") ctr_names |
|
287 |
||>> mk_name "case_" "_lazy" short_type_name |
|
288 |
||
289 |
fun mk_def (name, T, rhs) lthy = |
|
290 |
let |
|
291 |
val binding = Binding.name name |
|
292 |
val ((_, (_, thm)), lthy1) = |
|
72450 | 293 |
(snd o Local_Theory.begin_nested) lthy |
68155 | 294 |
|> Specification.definition NONE [] [] ((Thm.def_binding binding, []), rhs) |
72450 | 295 |
val lthy2 = Local_Theory.end_nested lthy1 |
68155 | 296 |
val def_thm = hd (Proof_Context.export lthy1 lthy2 [thm]) |
297 |
in |
|
298 |
({binding = binding, const = Const (Local_Theory.full_name lthy2 binding, T), thm = def_thm}, lthy2) |
|
299 |
end; |
|
300 |
||
69593 | 301 |
val lazy_datatype = Type (\<^type_name>\<open>lazy\<close>, [lazy_type]) |
68155 | 302 |
val Lazy_type = lazy_datatype --> eagerT |
303 |
val unstr_type = eagerT --> lazy_datatype |
|
304 |
||
305 |
fun apply_bounds i n term = |
|
306 |
if n > i then apply_bounds i (n-1) (term $ Bound (n-1)) |
|
307 |
else term |
|
308 |
fun all_abs Ts t = Logic.list_all (map (pair Name.uu) Ts, t) |
|
69593 | 309 |
fun mk_force t = Const (\<^const_name>\<open>force\<close>, lazy_datatype --> lazy_type) $ t |
310 |
fun mk_delay t = Const (\<^const_name>\<open>delay\<close>, (\<^typ>\<open>unit\<close> --> lazy_type) --> lazy_datatype) $ t |
|
68155 | 311 |
|
312 |
val lazy_ctr = all_abs [lazy_datatype] |
|
313 |
(Logic.mk_equals (Free (lazy_ctr_name, Lazy_type) $ Bound 0, Rep_lazy $ mk_force (Bound 0))) |
|
314 |
val (lazy_ctr_def, lthy5) = mk_def (lazy_ctr_name, Lazy_type, lazy_ctr) lthy4 |
|
315 |
||
316 |
val lazy_sel = all_abs [eagerT] |
|
317 |
(Logic.mk_equals (Free (lazy_sel_name, unstr_type) $ Bound 0, |
|
69593 | 318 |
mk_delay (Abs (Name.uu, \<^typ>\<open>unit\<close>, Abs_lazy $ Bound 1)))) |
68155 | 319 |
val (lazy_sel_def, lthy6) = mk_def (lazy_sel_name, unstr_type, lazy_sel) lthy5 |
320 |
||
321 |
fun mk_lazy_ctr (name, eager_ctr) = |
|
322 |
let |
|
323 |
val (_, ctrT) = dest_Const eager_ctr |
|
69593 | 324 |
fun to_lazy_ctrT (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) = T1 --> to_lazy_ctrT T2 |
68155 | 325 |
| to_lazy_ctrT T = if T = eagerT then lazy_type else raise Match |
326 |
val lazy_ctrT = to_lazy_ctrT ctrT |
|
327 |
val argsT = binder_types ctrT |
|
328 |
val lhs = apply_bounds 0 (length argsT) (Free (name, lazy_ctrT)) |
|
329 |
val rhs = Abs_lazy $ apply_bounds 0 (length argsT) eager_ctr |
|
330 |
in |
|
331 |
mk_def (name, lazy_ctrT, all_abs argsT (Logic.mk_equals (lhs, rhs))) |
|
332 |
end |
|
333 |
val (lazy_ctrs_def, lthy7) = fold_map mk_lazy_ctr (lazy_ctrs_name ~~ ctrs') lthy6 |
|
334 |
||
335 |
val (lazy_case_def, lthy8) = |
|
336 |
let |
|
337 |
val (_, caseT) = dest_Const case' |
|
69593 | 338 |
fun to_lazy_caseT (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) = |
68155 | 339 |
if T1 = eagerT then lazy_type --> T2 else T1 --> to_lazy_caseT T2 |
340 |
val lazy_caseT = to_lazy_caseT caseT |
|
341 |
val argsT = binder_types lazy_caseT |
|
342 |
val n = length argsT |
|
343 |
val lhs = apply_bounds 0 n (Free (lazy_case_name, lazy_caseT)) |
|
344 |
val rhs = apply_bounds 1 n case' $ (Rep_lazy $ Bound 0) |
|
345 |
in |
|
346 |
mk_def (lazy_case_name, lazy_caseT, all_abs argsT (Logic.mk_equals (lhs, rhs))) lthy7 |
|
347 |
end |
|
348 |
||
349 |
fun mk_thm ((name, term), thms) lthy = |
|
350 |
let |
|
351 |
val binding = Binding.name name |
|
352 |
fun tac {context, ...} = Simplifier.simp_tac |
|
353 |
(put_simpset HOL_basic_ss context addsimps thms) |
|
354 |
1 |
|
355 |
val thm = Goal.prove lthy [] [] term tac |
|
356 |
val (_, lthy1) = lthy |
|
72450 | 357 |
|> (snd o Local_Theory.begin_nested) |
68155 | 358 |
|> Local_Theory.note ((binding, []), [thm]) |
359 |
in |
|
72450 | 360 |
(thm, Local_Theory.end_nested lthy1) |
68155 | 361 |
end |
362 |
fun mk_thms exec_list lthy = fold_map mk_thm exec_list lthy |
|
363 |
||
364 |
val mk_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq |
|
365 |
||
366 |
val lazy_ctrs = map #const lazy_ctrs_def |
|
367 |
val eager_lazy_ctrs = ctrs' ~~ lazy_ctrs |
|
368 |
||
369 |
val (((((((Lazy_delay_eq_name, Rep_ctr_names), ctrs_lazy_names), force_sel_name), case_lazy_name), |
|
370 |
sel_lazy_name), case_ctrs_name), _) = lthy5 |
|
371 |
|> mk_name "Lazy_" "_delay" short_type_name |
|
372 |
||>> fold_map (mk_name "Rep_" "_Lazy") ctr_names |
|
373 |
||>> fold_map (mk_name "" "_conv_lazy") ctr_names |
|
374 |
||>> mk_name "force_unlazy_" "" short_type_name |
|
375 |
||>> mk_name "case_" "_conv_lazy" short_type_name |
|
376 |
||>> mk_name "Lazy_" "_inverse" short_type_name |
|
377 |
||>> fold_map (mk_name ("case_" ^ short_type_name ^ "_lazy_") "") ctr_names |
|
378 |
||
379 |
val mk_Lazy_delay_eq = |
|
74383 | 380 |
(#const lazy_ctr_def $ mk_delay (Bound 0), Rep_lazy $ (Bound 0 $ \<^Const>\<open>Unity\<close>)) |
381 |
|> mk_eq |> all_abs [\<^Type>\<open>unit\<close> --> lazy_type] |
|
68155 | 382 |
val (Lazy_delay_thm, lthy8a) = mk_thm |
383 |
((Lazy_delay_eq_name, mk_Lazy_delay_eq), [#thm lazy_ctr_def, @{thm force_delay}]) |
|
384 |
lthy8 |
|
385 |
||
386 |
fun mk_lazy_ctr_eq (eager_ctr, lazy_ctr) = |
|
387 |
let |
|
388 |
val (_, ctrT) = dest_Const eager_ctr |
|
389 |
val argsT = binder_types ctrT |
|
390 |
val args = length argsT |
|
391 |
in |
|
392 |
(Rep_lazy $ apply_bounds 0 args lazy_ctr, apply_bounds 0 args eager_ctr) |
|
393 |
|> mk_eq |> all_abs argsT |
|
394 |
end |
|
395 |
val Rep_ctr_eqs = map mk_lazy_ctr_eq eager_lazy_ctrs |
|
396 |
val (Rep_ctr_thms, lthy8b) = mk_thms |
|
397 |
((Rep_ctr_names ~~ Rep_ctr_eqs) ~~ |
|
398 |
(map (fn def => [#thm def, Abs_inverse, @{thm UNIV_I}]) lazy_ctrs_def) |
|
399 |
) |
|
400 |
lthy8a |
|
401 |
||
402 |
fun mk_ctrs_lazy_eq (eager_ctr, lazy_ctr) = |
|
403 |
let |
|
404 |
val argsT = dest_Const eager_ctr |> snd |> binder_types |
|
405 |
val n = length argsT |
|
406 |
val lhs = apply_bounds 0 n eager_ctr |
|
407 |
val rhs = #const lazy_ctr_def $ |
|
69593 | 408 |
(mk_delay (Abs (Name.uu, \<^typ>\<open>unit\<close>, apply_bounds 1 (n + 1) lazy_ctr))) |
68155 | 409 |
in |
410 |
(lhs, rhs) |> mk_eq |> all_abs argsT |
|
411 |
end |
|
412 |
val ctrs_lazy_eq = map mk_ctrs_lazy_eq eager_lazy_ctrs |
|
413 |
val (ctrs_lazy_thms, lthy8c) = mk_thms |
|
414 |
((ctrs_lazy_names ~~ ctrs_lazy_eq) ~~ map (fn thm => [Lazy_delay_thm, thm]) Rep_ctr_thms) |
|
415 |
lthy8b |
|
416 |
||
417 |
val force_sel_eq = |
|
418 |
(mk_force (#const lazy_sel_def $ Bound 0), Abs_lazy $ Bound 0) |
|
419 |
|> mk_eq |> all_abs [eagerT] |
|
420 |
val (force_sel_thm, lthy8d) = mk_thm |
|
421 |
((force_sel_name, force_sel_eq), [#thm lazy_sel_def, @{thm force_delay}]) |
|
422 |
lthy8c |
|
423 |
||
424 |
val case_lazy_eq = |
|
425 |
let |
|
426 |
val (_, caseT) = case' |> dest_Const |
|
427 |
val argsT = binder_types caseT |
|
428 |
val n = length argsT |
|
429 |
val lhs = apply_bounds 0 n case' |
|
430 |
val rhs = apply_bounds 1 n (#const lazy_case_def) $ (mk_force (#const lazy_sel_def $ Bound 0)) |
|
431 |
in |
|
432 |
(lhs, rhs) |> mk_eq |> all_abs argsT |
|
433 |
end |
|
434 |
val (case_lazy_thm, lthy8e) = mk_thm |
|
435 |
((case_lazy_name, case_lazy_eq), |
|
436 |
[#thm lazy_case_def, force_sel_thm, Abs_inverse, @{thm UNIV_I}]) |
|
437 |
lthy8d |
|
438 |
||
439 |
val sel_lazy_eq = |
|
440 |
(#const lazy_sel_def $ (#const lazy_ctr_def $ Bound 0), Bound 0) |
|
441 |
|> mk_eq |> all_abs [lazy_datatype] |
|
442 |
val (sel_lazy_thm, lthy8f) = mk_thm |
|
443 |
((sel_lazy_name, sel_lazy_eq), |
|
444 |
[#thm lazy_sel_def, #thm lazy_ctr_def, Rep_inverse, @{thm delay_force}]) |
|
445 |
lthy8e |
|
446 |
||
447 |
fun mk_case_ctrs_eq (i, lazy_ctr) = |
|
448 |
let |
|
449 |
val lazy_case = #const lazy_case_def |
|
450 |
val (_, ctrT) = dest_Const lazy_ctr |
|
451 |
val ctr_argsT = binder_types ctrT |
|
452 |
val ctr_args_n = length ctr_argsT |
|
453 |
val (_, caseT) = dest_Const lazy_case |
|
454 |
val case_argsT = binder_types caseT |
|
455 |
||
456 |
fun n_bounds_from m n t = |
|
457 |
if n > 0 then n_bounds_from (m - 1) (n - 1) (t $ Bound (m - 1)) else t |
|
458 |
||
459 |
val case_argsT' = take (length case_argsT - 1) case_argsT |
|
460 |
val Ts = case_argsT' @ ctr_argsT |
|
461 |
val num_abs_types = length Ts |
|
462 |
val lhs = n_bounds_from num_abs_types (length case_argsT') lazy_case $ |
|
463 |
apply_bounds 0 ctr_args_n lazy_ctr |
|
464 |
val rhs = apply_bounds 0 ctr_args_n (Bound (num_abs_types - i - 1)) |
|
465 |
in |
|
466 |
(lhs, rhs) |> mk_eq |> all_abs Ts |
|
467 |
end |
|
468 |
val case_ctrs_eq = map_index mk_case_ctrs_eq lazy_ctrs |
|
469 |
val (case_ctrs_thms, lthy9) = mk_thms |
|
470 |
((case_ctrs_name ~~ case_ctrs_eq) ~~ |
|
471 |
map2 (fn thm1 => fn thm2 => [#thm lazy_case_def, thm1, thm2]) Rep_ctr_thms case_thms |
|
472 |
) |
|
473 |
lthy8f |
|
474 |
||
475 |
val (pat_def_thm, lthy10) = |
|
476 |
add_syntax_definition short_type_name eagerT lazy_type (#const lazy_ctr_def) lthy9 |
|
477 |
||
478 |
val add_lazy_ctrs = |
|
479 |
Code.declare_datatype_global [dest_Const (#const lazy_ctr_def)] |
|
480 |
val eager_ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global)) o dest_Const) ctrs |
|
481 |
val add_eager_ctrs = |
|
482 |
fold Code.del_eqn_global ctrs_lazy_thms |
|
483 |
#> Code.declare_datatype_global eager_ctrs |
|
484 |
val add_code_eqs = fold (Code.add_eqn_global o rpair true) |
|
485 |
([case_lazy_thm, sel_lazy_thm]) |
|
486 |
val add_lazy_ctr_thms = fold (Code.add_eqn_global o rpair true) ctrs_lazy_thms |
|
487 |
val add_lazy_case_thms = |
|
488 |
fold Code.del_eqn_global case_thms |
|
68549
bbc742358156
declare case theorems as proper code equations
Andreas Lochbihler
parents:
68155
diff
changeset
|
489 |
#> Code.add_eqn_global (case_lazy_thm, true) |
68155 | 490 |
val add_eager_case_thms = Code.del_eqn_global case_lazy_thm |
68549
bbc742358156
declare case theorems as proper code equations
Andreas Lochbihler
parents:
68155
diff
changeset
|
491 |
#> fold (Code.add_eqn_global o rpair true) case_thms |
68155 | 492 |
|
493 |
val delay_dummy_thm = (pat_def_thm RS @{thm symmetric}) |
|
494 |
|> Drule.infer_instantiate' lthy10 |
|
69593 | 495 |
[SOME (Thm.cterm_of lthy10 (Const (\<^const_name>\<open>Pure.dummy_pattern\<close>, HOLogic.unitT --> lazy_type)))] |
74200
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
wenzelm
parents:
72450
diff
changeset
|
496 |
|> Thm.generalize |
74266 | 497 |
(Names.make_set (map (fst o dest_TFree) type_args), Names.empty) |
74200
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
wenzelm
parents:
72450
diff
changeset
|
498 |
(Variable.maxidx_of lthy10 + 1); |
68155 | 499 |
|
500 |
val ctr_post = delay_dummy_thm :: map (fn thm => thm RS @{thm sym}) ctrs_lazy_thms |
|
501 |
val ctr_thms_abs = map (fn thm => Drule.abs_def (thm RS @{thm eq_reflection})) ctrs_lazy_thms |
|
502 |
val case_thm_abs = Drule.abs_def (case_lazy_thm RS @{thm eq_reflection}) |
|
503 |
val add_simps = Code_Preproc.map_pre |
|
504 |
(fn ctxt => ctxt addsimps (case_thm_abs :: ctr_thms_abs)) |
|
505 |
val del_simps = Code_Preproc.map_pre |
|
506 |
(fn ctxt => ctxt delsimps (case_thm_abs :: ctr_thms_abs)) |
|
507 |
val add_post = Code_Preproc.map_post |
|
508 |
(fn ctxt => ctxt addsimps ctr_post) |
|
509 |
val del_post = Code_Preproc.map_post |
|
510 |
(fn ctxt => ctxt delsimps ctr_post) |
|
511 |
in |
|
512 |
Local_Theory.exit_global lthy10 |
|
513 |
|> Laziness_Data.map (Symtab.update (type_name, |
|
514 |
{eagerT = eagerT, |
|
515 |
lazyT = lazy_type, |
|
516 |
ctr = #const lazy_ctr_def, |
|
517 |
destr = #const lazy_sel_def, |
|
518 |
lazy_ctrs = map #const lazy_ctrs_def, |
|
519 |
case_lazy = #const lazy_case_def, |
|
520 |
active = true, |
|
521 |
activate = add_lazy_ctrs #> add_lazy_ctr_thms #> add_lazy_case_thms #> add_simps #> add_post, |
|
522 |
deactivate = add_eager_ctrs #> add_eager_case_thms #> del_simps #> del_post})) |
|
523 |
|> add_lazy_ctrs |
|
524 |
|> add_ctr_code (map (dest_Const o #const) lazy_ctrs_def) case_ctrs_thms |
|
525 |
|> add_code_eqs |
|
526 |
|> add_lazy_ctr_thms |
|
527 |
|> add_simps |
|
528 |
|> add_post |
|
529 |
end; |
|
530 |
||
531 |
fun transform_code_eqs _ [] = NONE |
|
532 |
| transform_code_eqs ctxt eqs = |
|
533 |
let |
|
69568
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
534 |
fun replace_ctr ctxt = |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
535 |
let |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
536 |
val thy = Proof_Context.theory_of ctxt |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
537 |
val table = Laziness_Data.get thy |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
538 |
in fn (s1, s2) => case Symtab.lookup table s1 of |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
539 |
NONE => false |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
540 |
| SOME x => #active x andalso s2 <> (#ctr x |> dest_Const |> fst) |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
541 |
end |
68155 | 542 |
val thy = Proof_Context.theory_of ctxt |
543 |
val table = Laziness_Data.get thy |
|
544 |
fun num_consts_fun (_, T) = |
|
545 |
let |
|
546 |
val s = body_type T |> dest_Type |> fst |
|
547 |
in |
|
548 |
if Symtab.defined table s |
|
69568
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
549 |
then Ctr_Sugar.ctr_sugar_of ctxt s |> the |> #ctrs |> length |
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
550 |
else Code.get_type thy s |> fst |> snd |> length |
68155 | 551 |
end |
552 |
val eqs = map (apfst (Thm.transfer thy)) eqs; |
|
553 |
||
554 |
val ((code_eqs, nbe_eqs), pure) = |
|
555 |
((case hd eqs |> fst |> Thm.prop_of of |
|
69593 | 556 |
Const (\<^const_name>\<open>Pure.eq\<close>, _) $ _ $ _ => |
68155 | 557 |
(map (apfst (fn x => x RS @{thm meta_eq_to_obj_eq})) eqs, true) |
558 |
| _ => (eqs, false)) |
|
559 |
|> apfst (List.partition snd)) |
|
560 |
handle THM _ => (([], eqs), false) |
|
561 |
val to_original_eq = if pure then map (apfst (fn x => x RS @{thm eq_reflection})) else I |
|
562 |
in |
|
69568
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
563 |
case Case_Converter.to_case ctxt (Case_Converter.replace_by_type replace_ctr) num_consts_fun (map fst code_eqs) of |
68155 | 564 |
NONE => NONE |
565 |
| SOME thms => SOME (nbe_eqs @ map (rpair true) thms |> to_original_eq) |
|
69568
de09a7261120
new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents:
68549
diff
changeset
|
566 |
end |
68155 | 567 |
|
568 |
val activate_lazy_type = set_active_lazy_type true; |
|
569 |
val deactivate_lazy_type = set_active_lazy_type false; |
|
570 |
val activate_lazy_types = set_active_lazy_types true; |
|
571 |
val deactivate_lazy_types = set_active_lazy_types false; |
|
572 |
||
573 |
fun get_lazy_types thy = Symtab.dest (Laziness_Data.get thy) |
|
574 |
||
575 |
fun print_lazy_type thy (name, info : lazy_info) = |
|
576 |
let |
|
577 |
val ctxt = Proof_Context.init_global thy |
|
578 |
fun pretty_ctr ctr = |
|
579 |
let |
|
580 |
val argsT = dest_Const ctr |> snd |> binder_types |
|
581 |
in |
|
582 |
Pretty.block [ |
|
583 |
Syntax.pretty_term ctxt ctr, |
|
584 |
Pretty.brk 1, |
|
585 |
Pretty.block (Pretty.separate "" (map (Pretty.quote o Syntax.pretty_typ ctxt) argsT)) |
|
586 |
] |
|
587 |
end |
|
588 |
in |
|
589 |
Pretty.block [ |
|
590 |
Pretty.str (name ^ (if #active info then "" else " (inactive)") ^ ":"), |
|
591 |
Pretty.brk 1, |
|
592 |
Pretty.block [ |
|
593 |
Syntax.pretty_typ ctxt (#eagerT info), |
|
594 |
Pretty.brk 1, |
|
595 |
Pretty.str "=", |
|
596 |
Pretty.brk 1, |
|
597 |
Syntax.pretty_term ctxt (#ctr info), |
|
598 |
Pretty.brk 1, |
|
599 |
Pretty.block [ |
|
600 |
Pretty.str "(", |
|
601 |
Syntax.pretty_term ctxt (#destr info), |
|
602 |
Pretty.str ":", |
|
603 |
Pretty.brk 1, |
|
69593 | 604 |
Syntax.pretty_typ ctxt (Type (\<^type_name>\<open>lazy\<close>, [#lazyT info])), |
68155 | 605 |
Pretty.str ")" |
606 |
] |
|
607 |
], |
|
608 |
Pretty.fbrk, |
|
609 |
Pretty.keyword2 "and", |
|
610 |
Pretty.brk 1, |
|
611 |
Pretty.block ([ |
|
612 |
Syntax.pretty_typ ctxt (#lazyT info), |
|
613 |
Pretty.brk 1, |
|
614 |
Pretty.str "=", |
|
615 |
Pretty.brk 1] @ |
|
616 |
Pretty.separate " |" (map pretty_ctr (#lazy_ctrs info)) @ [ |
|
617 |
Pretty.fbrk, |
|
618 |
Pretty.keyword2 "for", |
|
619 |
Pretty.brk 1, |
|
620 |
Pretty.str "case:", |
|
621 |
Pretty.brk 1, |
|
622 |
Syntax.pretty_term ctxt (#case_lazy info) |
|
623 |
]) |
|
624 |
] |
|
625 |
end; |
|
626 |
||
627 |
fun print_lazy_types thy = |
|
628 |
let |
|
629 |
fun cmp ((name1, _), (name2, _)) = string_ord (name1, name2) |
|
630 |
val infos = Laziness_Data.get thy |> Symtab.dest |> map (apfst Long_Name.base_name) |> sort cmp |
|
631 |
in |
|
632 |
Pretty.writeln_chunks (map (print_lazy_type thy) infos) |
|
633 |
end; |
|
634 |
||
635 |
||
636 |
val _ = |
|
69593 | 637 |
Outer_Syntax.command \<^command_keyword>\<open>code_lazy_type\<close> |
68155 | 638 |
"make a lazy copy of the datatype and activate substitution" |
639 |
(Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> code_lazy_type))); |
|
640 |
val _ = |
|
69593 | 641 |
Outer_Syntax.command \<^command_keyword>\<open>activate_lazy_type\<close> |
68155 | 642 |
"activate substitution on a specific (lazy) type" |
643 |
(Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> activate_lazy_type))); |
|
644 |
val _ = |
|
69593 | 645 |
Outer_Syntax.command \<^command_keyword>\<open>deactivate_lazy_type\<close> |
68155 | 646 |
"deactivate substitution on a specific (lazy) type" |
647 |
(Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> deactivate_lazy_type))); |
|
648 |
val _ = |
|
69593 | 649 |
Outer_Syntax.command \<^command_keyword>\<open>activate_lazy_types\<close> |
68155 | 650 |
"activate substitution on all (lazy) types" |
651 |
(pair (Toplevel.theory activate_lazy_types)); |
|
652 |
val _ = |
|
69593 | 653 |
Outer_Syntax.command \<^command_keyword>\<open>deactivate_lazy_types\<close> |
68155 | 654 |
"deactivate substitution on all (lazy) type" |
655 |
(pair (Toplevel.theory deactivate_lazy_types)); |
|
656 |
val _ = |
|
69593 | 657 |
Outer_Syntax.command \<^command_keyword>\<open>print_lazy_types\<close> |
68155 | 658 |
"print the types that have been declared as lazy and their substitution state" |
659 |
(pair (Toplevel.theory (tap print_lazy_types))); |
|
660 |
||
661 |
end |