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