src/HOL/Tools/code_evaluation.ML
author haftmann
Sun, 08 Oct 2017 22:28:19 +0200
changeset 66798 39bb2462e681
parent 66330 dcb3e6bdc00a
child 67149 e61557884799
permissions -rw-r--r--
fundamental property of division by units
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Tools/code_evaluation.ML
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     3
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     4
Evaluation and reconstruction of terms in ML.
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     5
*)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     6
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     7
signature CODE_EVALUATION =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
     8
sig
55757
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 51714
diff changeset
     9
  val dynamic_value: Proof.context -> term -> term option
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 51714
diff changeset
    10
  val dynamic_value_strict: Proof.context -> term -> term
9fc71814b8c1 prefer proof context over background theory
haftmann
parents: 51714
diff changeset
    11
  val dynamic_value_exn: Proof.context -> term -> term Exn.result
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    12
  val put_term: (unit -> term) -> Proof.context -> Proof.context
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    13
  val tracing: string -> 'a -> 'a
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    14
end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    15
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    16
structure Code_Evaluation : CODE_EVALUATION =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    17
struct
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    18
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    19
(** term_of instances **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    20
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    21
(* formal definition *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    22
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    23
fun add_term_of_inst tyco thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    24
  let
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    25
    val ((raw_vs, _), _) = Code.get_type thy tyco;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    26
    val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    27
    val ty = Type (tyco, map TFree vs);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    28
    val lhs = Const (@{const_name term_of}, ty --> @{typ term})
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    29
      $ Free ("x", ty);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    30
    val rhs = @{term "undefined :: term"};
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    31
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    32
    fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    33
      o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    34
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    35
    thy
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    36
    |> Class.instantiation ([tyco], vs, @{sort term_of})
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    37
    |> `(fn lthy => Syntax.check_term lthy eq)
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63161
diff changeset
    38
    |-> (fn eq => Specification.definition NONE [] [] ((Binding.name (triv_name_of eq), []), eq))
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    39
    |> snd
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59323
diff changeset
    40
    |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    41
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    42
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    43
fun ensure_term_of_inst tyco thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    44
  let
48272
db75b4005d9a more direct Sorts.has_instance;
wenzelm
parents: 45429
diff changeset
    45
    val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of})
db75b4005d9a more direct Sorts.has_instance;
wenzelm
parents: 45429
diff changeset
    46
      andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep};
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    47
  in if need_inst then add_term_of_inst tyco thy else thy end;
61667
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    48
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    49
fun for_term_of_instance tyco vs f thy =
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    50
  let
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    51
    val algebra = Sign.classes_of thy;
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    52
  in
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    53
    case try (Sorts.mg_domain algebra tyco) @{sort term_of} of
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    54
      NONE => thy
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    55
    | SOME sorts => f tyco (map2 (fn (v, sort) => fn sort' =>
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    56
        (v, Sorts.inter_sort algebra (sort, sort'))) vs sorts) thy
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
    57
  end;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    58
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    59
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    60
(* code equations for datatypes *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    61
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
    62
fun mk_term_of_eq thy ty (c, (_, tys)) =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    63
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    64
    val t = list_comb (Const (c, tys ---> ty),
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 42402
diff changeset
    65
      map Free (Name.invent_names Name.context "a" tys));
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    66
    val (arg, rhs) =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
    67
      apply2 (Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global)
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    68
        (t,
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    69
          map_aterms (fn t as Free (_, ty) => HOLogic.mk_term_of ty t | t => t)
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    70
            (HOLogic.reflect_term t));
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
    71
    val cty = Thm.global_ctyp_of thy ty;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    72
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    73
    @{thm term_of_anything}
60801
7664e0916eec tuned signature;
wenzelm
parents: 59621
diff changeset
    74
    |> Thm.instantiate' [SOME cty] [SOME arg, SOME rhs]
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    75
    |> Thm.varifyT_global
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    76
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    77
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    78
fun add_term_of_code_datatype tyco vs raw_cs thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    79
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    80
    val ty = Type (tyco, map TFree vs);
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
    81
    val cs = (map o apsnd o apsnd o map o map_atyps)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    82
      (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    83
    val eqs = map (mk_term_of_eq thy ty) cs;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    84
 in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    85
    thy
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 65062
diff changeset
    86
    |> Code.declare_default_eqns_global (map (rpair true) eqs)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    87
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    88
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    89
fun ensure_term_of_code_datatype (tyco, (vs, cs)) =
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
    90
  for_term_of_instance tyco vs (fn tyco => fn vs => add_term_of_code_datatype tyco vs cs);
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    91
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    92
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    93
(* code equations for abstypes *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    94
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    95
fun mk_abs_term_of_eq thy ty abs ty_rep proj =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    96
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    97
    val arg = Var (("x", 0), ty);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    98
    val rhs = Abs ("y", @{typ term}, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    99
      (HOLogic.mk_term_of ty_rep (Const (proj, ty --> ty_rep) $ arg))
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
   100
      |> Thm.global_cterm_of thy;
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
   101
    val cty = Thm.global_ctyp_of thy ty;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   102
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   103
    @{thm term_of_anything}
60801
7664e0916eec tuned signature;
wenzelm
parents: 59621
diff changeset
   104
    |> Thm.instantiate' [SOME cty] [SOME (Thm.global_cterm_of thy arg), SOME rhs]
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   105
    |> Thm.varifyT_global
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   106
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   107
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   108
fun add_term_of_code_abstype tyco vs abs raw_ty_rep projection thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   109
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   110
    val ty = Type (tyco, map TFree vs);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   111
    val ty_rep = map_atyps
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   112
      (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_ty_rep;
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   113
    val eq = mk_abs_term_of_eq thy ty abs ty_rep projection;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   114
 in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   115
    thy
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 65062
diff changeset
   116
    |> Code.declare_default_eqns_global [(eq, true)]
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   117
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   118
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   119
fun ensure_term_of_code_abstype (tyco, (vs, {abstractor = (abs, (_, ty)),
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66251
diff changeset
   120
    projection, ...})) =
61667
4b53042d7a40 explicit computation of sort arguments for code equations makes less assumption about sort arguments of underlying type class instances
haftmann
parents: 60801
diff changeset
   121
  for_term_of_instance tyco vs
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   122
    (fn tyco => fn vs => add_term_of_code_abstype tyco vs abs ty projection);
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   123
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   124
56926
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   125
(* setup *)
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   126
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 59153
diff changeset
   127
val _ = Theory.setup
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   128
  (Code.type_interpretation ensure_term_of_inst
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   129
  #> Code.datatype_interpretation ensure_term_of_code_datatype
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66310
diff changeset
   130
  #> Code.abstype_interpretation ensure_term_of_code_abstype);
56926
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   131
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   132
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   133
(** termifying syntax **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   134
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   135
fun map_default f xs =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   136
  let val ys = map f xs
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   137
  in if exists is_some ys
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   138
    then SOME (map2 the_default xs ys)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   139
    else NONE
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   140
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   141
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   142
fun subst_termify_app (Const (@{const_name termify}, _), [t]) =
63619
9c870388e87a more tight filtering;
wenzelm
parents: 63239
diff changeset
   143
      if not (Term.exists_subterm (fn Abs _ => true | _ => false) t)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   144
      then if fold_aterms (fn Const _ => I | _ => K false) t true
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   145
        then SOME (HOLogic.reflect_term t)
51714
88f7f38d5cb9 spelling
haftmann
parents: 51685
diff changeset
   146
        else error "Cannot termify expression containing variable"
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   147
      else error "Cannot termify expression containing abstraction"
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   148
  | subst_termify_app (t, ts) = case map_default subst_termify ts
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   149
     of SOME ts' => SOME (list_comb (t, ts'))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   150
      | NONE => NONE
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   151
and subst_termify (Abs (v, T, t)) = (case subst_termify t
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   152
     of SOME t' => SOME (Abs (v, T, t'))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   153
      | NONE => NONE)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   154
  | subst_termify t = subst_termify_app (strip_comb t) 
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   155
62952
wenzelm
parents: 61674
diff changeset
   156
fun check_termify ts = the_default ts (map_default subst_termify ts);
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   157
62952
wenzelm
parents: 61674
diff changeset
   158
val _ = Context.>> (Syntax_Phases.term_check 0 "termify" (K check_termify));
56926
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   159
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   160
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   161
(** evaluation **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   162
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41247
diff changeset
   163
structure Evaluation = Proof_Data
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41247
diff changeset
   164
(
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   165
  type T = unit -> term
59153
wenzelm
parents: 59151
diff changeset
   166
  val empty: T = fn () => raise Fail "Evaluation"
wenzelm
parents: 59151
diff changeset
   167
  fun init _ = empty
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   168
);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   169
val put_term = Evaluation.put;
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   170
val cookie = (Evaluation.get, put_term, "Code_Evaluation.put_term");
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   171
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   172
fun mk_term_of t = HOLogic.mk_term_of (fastype_of t) t;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   173
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63064
diff changeset
   174
fun gen_dynamic_value computation ctxt t =
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63064
diff changeset
   175
  computation cookie ctxt NONE I (mk_term_of t) [];
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   176
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   177
val dynamic_value = gen_dynamic_value Code_Runtime.dynamic_value;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   178
val dynamic_value_strict = gen_dynamic_value Code_Runtime.dynamic_value_strict;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   179
val dynamic_value_exn = gen_dynamic_value Code_Runtime.dynamic_value_exn;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   180
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   181
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   182
(** diagnostic **)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   183
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   184
fun tracing s x = (Output.tracing s; x);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   185
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   186
end;