author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 82379 | 3f875966c3e1 |
permissions | -rw-r--r-- |
68155 | 1 |
(* Author: Pascal Stoop, ETH Zurich |
2 |
Author: Andreas Lochbihler, Digital Asset *) |
|
3 |
||
4 |
section \<open>Lazy types in generated code\<close> |
|
5 |
||
6 |
theory Code_Lazy |
|
69567
6b4c41037649
separate case converter into a separate theory
Andreas Lochbihler
parents:
69528
diff
changeset
|
7 |
imports Case_Converter |
68155 | 8 |
keywords |
9 |
"code_lazy_type" |
|
10 |
"activate_lazy_type" |
|
11 |
"deactivate_lazy_type" |
|
12 |
"activate_lazy_types" |
|
13 |
"deactivate_lazy_types" |
|
14 |
"print_lazy_types" :: thy_decl |
|
15 |
begin |
|
16 |
||
17 |
text \<open> |
|
76987 | 18 |
This theory and the CodeLazy tool described in \<^cite>\<open>"LochbihlerStoop2018"\<close>. |
68155 | 19 |
|
68390 | 20 |
It hooks into Isabelle's code generator such that the generated code evaluates a user-specified |
68155 | 21 |
set of type constructors lazily, even in target languages with eager evaluation. |
22 |
The lazy type must be algebraic, i.e., values must be built from constructors and a |
|
23 |
corresponding case operator decomposes them. Every datatype and codatatype is algebraic |
|
24 |
and thus eligible for lazification. |
|
25 |
\<close> |
|
26 |
||
69272 | 27 |
subsection \<open>The type \<open>lazy\<close>\<close> |
68155 | 28 |
|
29 |
typedef 'a lazy = "UNIV :: 'a set" .. |
|
30 |
setup_lifting type_definition_lazy |
|
31 |
lift_definition delay :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a lazy" is "\<lambda>f. f ()" . |
|
32 |
lift_definition force :: "'a lazy \<Rightarrow> 'a" is "\<lambda>x. x" . |
|
33 |
||
34 |
code_datatype delay |
|
69528 | 35 |
lemma force_delay [code]: "force (delay f) = f ()" by transfer (rule refl) |
36 |
lemma delay_force: "delay (\<lambda>_. force s) = s" by transfer (rule refl) |
|
37 |
||
38 |
definition termify_lazy2 :: "'a :: typerep lazy \<Rightarrow> term" |
|
39 |
where "termify_lazy2 x = |
|
40 |
Code_Evaluation.App (Code_Evaluation.Const (STR ''Code_Lazy.delay'') (TYPEREP((unit \<Rightarrow> 'a) \<Rightarrow> 'a lazy))) |
|
41 |
(Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (TYPEREP((unit \<Rightarrow> 'a))))" |
|
42 |
||
43 |
definition termify_lazy :: |
|
44 |
"(String.literal \<Rightarrow> 'typerep \<Rightarrow> 'term) \<Rightarrow> |
|
45 |
('term \<Rightarrow> 'term \<Rightarrow> 'term) \<Rightarrow> |
|
46 |
(String.literal \<Rightarrow> 'typerep \<Rightarrow> 'term \<Rightarrow> 'term) \<Rightarrow> |
|
47 |
'typerep \<Rightarrow> ('typerep \<Rightarrow> 'typerep \<Rightarrow> 'typerep) \<Rightarrow> ('typerep \<Rightarrow> 'typerep) \<Rightarrow> |
|
48 |
('a \<Rightarrow> 'term) \<Rightarrow> 'typerep \<Rightarrow> 'a :: typerep lazy \<Rightarrow> 'term \<Rightarrow> term" |
|
49 |
where "termify_lazy _ _ _ _ _ _ _ _ x _ = termify_lazy2 x" |
|
68155 | 50 |
|
69528 | 51 |
declare [[code drop: "Code_Evaluation.term_of :: _ lazy \<Rightarrow> _"]] |
52 |
||
53 |
lemma term_of_lazy_code [code]: |
|
54 |
"Code_Evaluation.term_of x \<equiv> |
|
55 |
termify_lazy |
|
56 |
Code_Evaluation.Const Code_Evaluation.App Code_Evaluation.Abs |
|
57 |
TYPEREP(unit) (\<lambda>T U. typerep.Typerep (STR ''fun'') [T, U]) (\<lambda>T. typerep.Typerep (STR ''Code_Lazy.lazy'') [T]) |
|
58 |
Code_Evaluation.term_of TYPEREP('a) x (Code_Evaluation.Const (STR '''') (TYPEREP(unit)))" |
|
59 |
for x :: "'a :: {typerep, term_of} lazy" |
|
60 |
by (rule term_of_anything) |
|
61 |
||
62 |
text \<open> |
|
69593 | 63 |
The implementations of \<^typ>\<open>_ lazy\<close> using language primitives cache forced values. |
69528 | 64 |
|
65 |
Term reconstruction for lazy looks into the lazy value and reconstructs it to the depth it has been evaluated. |
|
66 |
This is not done for Haskell as we do not know of any portable way to inspect whether a lazy value |
|
67 |
has been evaluated to or not. |
|
68 |
\<close> |
|
68155 | 69 |
|
82379 | 70 |
code_printing code_module Lazy \<rightharpoonup> (SML) file "~~/src/HOL/Library/Tools/lazy.ML" |
71 |
for type_constructor lazy constant delay force termify_lazy |
|
69528 | 72 |
| type_constructor lazy \<rightharpoonup> (SML) "_ Lazy.lazy" |
68155 | 73 |
| constant delay \<rightharpoonup> (SML) "Lazy.lazy" |
74 |
| constant force \<rightharpoonup> (SML) "Lazy.force" |
|
69528 | 75 |
| constant termify_lazy \<rightharpoonup> (SML) "Lazy.termify'_lazy" |
76 |
||
81706 | 77 |
code_reserved (SML) Lazy |
68155 | 78 |
|
79 |
code_printing \<comment> \<open>For code generation within the Isabelle environment, we reuse the thread-safe |
|
69593 | 80 |
implementation of lazy from \<^file>\<open>~~/src/Pure/Concurrent/lazy.ML\<close>\<close> |
69528 | 81 |
code_module Lazy \<rightharpoonup> (Eval) \<open>\<close> for constant undefined |
68155 | 82 |
| type_constructor lazy \<rightharpoonup> (Eval) "_ Lazy.lazy" |
83 |
| constant delay \<rightharpoonup> (Eval) "Lazy.lazy" |
|
84 |
| constant force \<rightharpoonup> (Eval) "Lazy.force" |
|
82379 | 85 |
| code_module Termify_Lazy \<rightharpoonup> (Eval) file "~~/src/HOL/Library/Tools/termify_lazy.ML" |
86 |
for constant termify_lazy |
|
69528 | 87 |
| constant termify_lazy \<rightharpoonup> (Eval) "Termify'_Lazy.termify'_lazy" |
88 |
||
81706 | 89 |
code_reserved (Eval) Termify_Lazy |
69528 | 90 |
|
91 |
code_printing |
|
92 |
type_constructor lazy \<rightharpoonup> (OCaml) "_ Lazy.t" |
|
93 |
| constant delay \<rightharpoonup> (OCaml) "Lazy.from'_fun" |
|
94 |
| constant force \<rightharpoonup> (OCaml) "Lazy.force" |
|
82379 | 95 |
| code_module Termify_Lazy \<rightharpoonup> (OCaml) file "~~/src/HOL/Library/Tools/termify_lazy.ocaml" |
96 |
for constant termify_lazy |
|
69528 | 97 |
| constant termify_lazy \<rightharpoonup> (OCaml) "Termify'_Lazy.termify'_lazy" |
98 |
||
81706 | 99 |
code_reserved (OCaml) Lazy Termify_Lazy |
69528 | 100 |
|
68155 | 101 |
code_printing |
82379 | 102 |
code_module Lazy \<rightharpoonup> (Haskell) file "~~/src/HOL/Library/Tools/lazy.hs" |
103 |
for type_constructor lazy constant delay force |
|
68155 | 104 |
| type_constructor lazy \<rightharpoonup> (Haskell) "Lazy.Lazy _" |
105 |
| constant delay \<rightharpoonup> (Haskell) "Lazy.delay" |
|
106 |
| constant force \<rightharpoonup> (Haskell) "Lazy.force" |
|
69528 | 107 |
|
81706 | 108 |
code_reserved (Haskell) Lazy |
68155 | 109 |
|
110 |
code_printing |
|
82379 | 111 |
code_module Lazy \<rightharpoonup> (Scala) file "~~/src/HOL/Library/Tools/lazy.scala" |
112 |
for type_constructor lazy constant delay force termify_lazy |
|
68155 | 113 |
| type_constructor lazy \<rightharpoonup> (Scala) "Lazy.Lazy[_]" |
114 |
| constant delay \<rightharpoonup> (Scala) "Lazy.delay" |
|
115 |
| constant force \<rightharpoonup> (Scala) "Lazy.force" |
|
69528 | 116 |
| constant termify_lazy \<rightharpoonup> (Scala) "Lazy.termify'_lazy" |
68155 | 117 |
|
81706 | 118 |
code_reserved (Scala) Lazy |
68155 | 119 |
|
69593 | 120 |
text \<open>Make evaluation with the simplifier respect \<^term>\<open>delay\<close>s.\<close> |
68155 | 121 |
|
122 |
lemma delay_lazy_cong: "delay f = delay f" by simp |
|
123 |
setup \<open>Code_Simp.map_ss (Simplifier.add_cong @{thm delay_lazy_cong})\<close> |
|
124 |
||
125 |
subsection \<open>Implementation\<close> |
|
126 |
||
69605 | 127 |
ML_file \<open>code_lazy.ML\<close> |
68155 | 128 |
|
129 |
setup \<open> |
|
69216
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents:
68390
diff
changeset
|
130 |
Code_Preproc.add_functrans ("lazy_datatype", Code_Lazy.transform_code_eqs) |
68155 | 131 |
\<close> |
132 |
||
133 |
end |