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