Theory HOL-Library.Code_Lazy
section ‹Lazy types in generated code›
theory Code_Lazy
imports Case_Converter
keywords
"code_lazy_type"
"activate_lazy_type"
"deactivate_lazy_type"
"activate_lazy_types"
"deactivate_lazy_types"
"print_lazy_types" :: thy_decl
begin
text ‹
This theory and the CodeLazy tool described in \<^cite>‹"LochbihlerStoop2018"›.
It hooks into Isabelle's code generator such that the generated code evaluates a user-specified
set of type constructors lazily, even in target languages with eager evaluation.
The lazy type must be algebraic, i.e., values must be built from constructors and a
corresponding case operator decomposes them. Every datatype and codatatype is algebraic
and thus eligible for lazification.
›
subsection ‹The type ‹lazy››
typedef 'a lazy = "UNIV :: 'a set" ..
setup_lifting type_definition_lazy
lift_definition delay :: "(unit ⇒ 'a) ⇒ 'a lazy" is "λf. f ()" .
lift_definition force :: "'a lazy ⇒ 'a" is "λx. x" .
code_datatype delay
lemma force_delay [code]: "force (delay f) = f ()" by transfer (rule refl)
lemma delay_force: "delay (λ_. force s) = s" by transfer (rule refl)
definition termify_lazy2 :: "'a :: typerep lazy ⇒ term"
where "termify_lazy2 x =
Code_Evaluation.App (Code_Evaluation.Const (STR ''Code_Lazy.delay'') (TYPEREP((unit ⇒ 'a) ⇒ 'a lazy)))
(Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (TYPEREP((unit ⇒ 'a))))"
definition termify_lazy ::
"(String.literal ⇒ 'typerep ⇒ 'term) ⇒
('term ⇒ 'term ⇒ 'term) ⇒
(String.literal ⇒ 'typerep ⇒ 'term ⇒ 'term) ⇒
'typerep ⇒ ('typerep ⇒ 'typerep ⇒ 'typerep) ⇒ ('typerep ⇒ 'typerep) ⇒
('a ⇒ 'term) ⇒ 'typerep ⇒ 'a :: typerep lazy ⇒ 'term ⇒ term"
where "termify_lazy _ _ _ _ _ _ _ _ x _ = termify_lazy2 x"
declare [[code drop: "Code_Evaluation.term_of :: _ lazy ⇒ _"]]
lemma term_of_lazy_code [code]:
"Code_Evaluation.term_of x ≡
termify_lazy
Code_Evaluation.Const Code_Evaluation.App Code_Evaluation.Abs
TYPEREP(unit) (λT U. typerep.Typerep (STR ''fun'') [T, U]) (λT. typerep.Typerep (STR ''Code_Lazy.lazy'') [T])
Code_Evaluation.term_of TYPEREP('a) x (Code_Evaluation.Const (STR '''') (TYPEREP(unit)))"
for x :: "'a :: {typerep, term_of} lazy"
by (rule term_of_anything)
text ‹
The implementations of \<^typ>‹_ lazy› using language primitives cache forced values.
Term reconstruction for lazy looks into the lazy value and reconstructs it to the depth it has been evaluated.
This is not done for Haskell as we do not know of any portable way to inspect whether a lazy value
has been evaluated to or not.
›
code_printing code_module Lazy ⇀ (SML)
‹signature LAZY =
sig
type 'a lazy;
val lazy : (unit -> 'a) -> 'a lazy;
val force : 'a lazy -> 'a;
val peek : 'a lazy -> 'a option
val termify_lazy :
(string -> 'typerep -> 'term) ->
('term -> 'term -> 'term) ->
(string -> 'typerep -> 'term -> 'term) ->
'typerep -> ('typerep -> 'typerep -> 'typerep) -> ('typerep -> 'typerep) ->
('a -> 'term) -> 'typerep -> 'a lazy -> 'term -> 'term;
end;
structure Lazy : LAZY =
struct
datatype 'a content =
Delay of unit -> 'a
| Value of 'a
| Exn of exn;
datatype 'a lazy = Lazy of 'a content ref;
fun lazy f = Lazy (ref (Delay f));
fun force (Lazy x) = case !x of
Delay f => (
let val res = f (); val _ = x := Value res; in res end
handle exn => (x := Exn exn; raise exn))
| Value x => x
| Exn exn => raise exn;
fun peek (Lazy x) = case !x of
Value x => SOME x
| _ => NONE;
fun termify_lazy const app abs unitT funT lazyT term_of T x _ =
app (const "Code_Lazy.delay" (funT (funT unitT T) (lazyT T)))
(case peek x of SOME y => abs "_" unitT (term_of y)
| _ => const "Pure.dummy_pattern" (funT unitT T));
end;› for type_constructor lazy constant delay force termify_lazy
| type_constructor lazy ⇀ (SML) "_ Lazy.lazy"
| constant delay ⇀ (SML) "Lazy.lazy"
| constant force ⇀ (SML) "Lazy.force"
| constant termify_lazy ⇀ (SML) "Lazy.termify'_lazy"
code_reserved SML Lazy
code_printing
code_module Lazy ⇀ (Eval) ‹› for constant undefined
| type_constructor lazy ⇀ (Eval) "_ Lazy.lazy"
| constant delay ⇀ (Eval) "Lazy.lazy"
| constant force ⇀ (Eval) "Lazy.force"
| code_module Termify_Lazy ⇀ (Eval)
‹structure Termify_Lazy = struct
fun termify_lazy
(_: string -> typ -> term) (_: term -> term -> term) (_: string -> typ -> term -> term)
(_: typ) (_: typ -> typ -> typ) (_: typ -> typ)
(term_of: 'a -> term) (T: typ) (x: 'a Lazy.lazy) (_: term) =
Const ("Code_Lazy.delay", (HOLogic.unitT --> T) --> Type ("Code_Lazy.lazy", [T])) $
(case Lazy.peek x of
SOME (Exn.Res x) => absdummy HOLogic.unitT (term_of x)
| _ => Const ("Pure.dummy_pattern", HOLogic.unitT --> T));
end;› for constant termify_lazy
| constant termify_lazy ⇀ (Eval) "Termify'_Lazy.termify'_lazy"
code_reserved Eval Termify_Lazy
code_printing
type_constructor lazy ⇀ (OCaml) "_ Lazy.t"
| constant delay ⇀ (OCaml) "Lazy.from'_fun"
| constant force ⇀ (OCaml) "Lazy.force"
| code_module Termify_Lazy ⇀ (OCaml)
‹module Termify_Lazy : sig
val termify_lazy :
(string -> 'typerep -> 'term) ->
('term -> 'term -> 'term) ->
(string -> 'typerep -> 'term -> 'term) ->
'typerep -> ('typerep -> 'typerep -> 'typerep) -> ('typerep -> 'typerep) ->
('a -> 'term) -> 'typerep -> 'a Lazy.t -> 'term -> 'term
end = struct
let termify_lazy const app abs unitT funT lazyT term_of ty x _ =
app (const "Code_Lazy.delay" (funT (funT unitT ty) (lazyT ty)))
(if Lazy.is_val x then abs "_" unitT (term_of (Lazy.force x))
else const "Pure.dummy_pattern" (funT unitT ty));;
end;;› for constant termify_lazy
| constant termify_lazy ⇀ (OCaml) "Termify'_Lazy.termify'_lazy"
code_reserved OCaml Lazy Termify_Lazy
code_printing
code_module Lazy ⇀ (Haskell) ‹
module Lazy(Lazy, delay, force) where
newtype Lazy a = Lazy a
delay f = Lazy (f ())
force (Lazy x) = x› for type_constructor lazy constant delay force
| type_constructor lazy ⇀ (Haskell) "Lazy.Lazy _"
| constant delay ⇀ (Haskell) "Lazy.delay"
| constant force ⇀ (Haskell) "Lazy.force"
code_reserved Haskell Lazy
code_printing
code_module Lazy ⇀ (Scala)
‹object Lazy {
final class Lazy[A] (f: Unit => A) {
var evaluated = false;
lazy val x: A = f(())
def get() : A = {
evaluated = true;
return x
}
}
def force[A] (x: Lazy[A]) : A = {
return x.get()
}
def delay[A] (f: Unit => A) : Lazy[A] = {
return new Lazy[A] (f)
}
def termify_lazy[Typerep, Term, A] (
const: String => Typerep => Term,
app: Term => Term => Term,
abs: String => Typerep => Term => Term,
unitT: Typerep,
funT: Typerep => Typerep => Typerep,
lazyT: Typerep => Typerep,
term_of: A => Term,
ty: Typerep,
x: Lazy[A],
dummy: Term) : Term = {
x.evaluated match {
case true => app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(abs("_")(unitT)(term_of(x.get())))
case false => app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(const("Pure.dummy_pattern")(funT(unitT)(ty)))
}
}
}› for type_constructor lazy constant delay force termify_lazy
| type_constructor lazy ⇀ (Scala) "Lazy.Lazy[_]"
| constant delay ⇀ (Scala) "Lazy.delay"
| constant force ⇀ (Scala) "Lazy.force"
| constant termify_lazy ⇀ (Scala) "Lazy.termify'_lazy"
code_reserved Scala Lazy
text ‹Make evaluation with the simplifier respect \<^term>‹delay›s.›
lemma delay_lazy_cong: "delay f = delay f" by simp
setup ‹Code_Simp.map_ss (Simplifier.add_cong @{thm delay_lazy_cong})›
subsection ‹Implementation›
ML_file ‹code_lazy.ML›
setup ‹
Code_Preproc.add_functrans ("lazy_datatype", Code_Lazy.transform_code_eqs)
›
end