src/HOL/Tools/code_evaluation.ML
author wenzelm
Wed, 10 Apr 2019 16:18:12 +0200
changeset 70108 77f978dd8ffb
parent 67149 e61557884799
permissions -rw-r--r--
merged;
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;
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    26
    val vs = map (fn (v, _) => (v, \<^sort>\<open>typerep\<close>)) raw_vs;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    27
    val ty = Type (tyco, map TFree vs);
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    28
    val lhs = Const (\<^const_name>\<open>term_of\<close>, ty --> \<^typ>\<open>term\<close>) $ Free ("x", ty);
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    29
    val rhs = \<^term>\<open>undefined :: term\<close>;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    30
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    31
    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
    32
      o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    33
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    34
    thy
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    35
    |> Class.instantiation ([tyco], vs, \<^sort>\<open>term_of\<close>)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    36
    |> `(fn lthy => Syntax.check_term lthy eq)
63180
ddfd021884b4 clarified check_open_spec / read_open_spec;
wenzelm
parents: 63161
diff changeset
    37
    |-> (fn eq => Specification.definition NONE [] [] ((Binding.name (triv_name_of eq), []), eq))
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    38
    |> snd
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59323
diff changeset
    39
    |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    40
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    41
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
    42
fun ensure_term_of_inst tyco thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    43
  let
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    44
    val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>term_of\<close>)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    45
      andalso Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>typerep\<close>;
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
    46
  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
    47
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
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
    49
  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
    50
    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
    51
  in
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    52
    case try (Sorts.mg_domain algebra tyco) \<^sort>\<open>term_of\<close> of
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
    53
      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
    54
    | 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
    55
        (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
    56
  end;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    57
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    58
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    59
(* code equations for datatypes *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    60
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
    61
fun mk_term_of_eq thy ty (c, (_, tys)) =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    62
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    63
    val t = list_comb (Const (c, tys ---> ty),
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 42402
diff changeset
    64
      map Free (Name.invent_names Name.context "a" tys));
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    65
    val (arg, rhs) =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
    66
      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
    67
        (t,
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    68
          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
    69
            (HOLogic.reflect_term t));
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
    70
    val cty = Thm.global_ctyp_of thy ty;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    71
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    72
    @{thm term_of_anything}
60801
7664e0916eec tuned signature;
wenzelm
parents: 59621
diff changeset
    73
    |> Thm.instantiate' [SOME cty] [SOME arg, SOME rhs]
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    74
    |> Thm.varifyT_global
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    75
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    76
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
    77
fun add_term_of_code_datatype tyco vs raw_cs thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    78
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    79
    val ty = Type (tyco, map TFree vs);
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
    80
    val cs = (map o apsnd o apsnd o map o map_atyps)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    81
      (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
    82
    val eqs = map (mk_term_of_eq thy ty) cs;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    83
 in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    84
    thy
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 65062
diff changeset
    85
    |> Code.declare_default_eqns_global (map (rpair true) eqs)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    86
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    87
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
    88
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
    89
  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
    90
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    91
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    92
(* code equations for abstypes *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    93
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    94
fun mk_abs_term_of_eq thy ty abs ty_rep proj =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    95
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    96
    val arg = Var (("x", 0), ty);
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
    97
    val rhs = Abs ("y", \<^typ>\<open>term\<close>, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    98
      (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
    99
      |> Thm.global_cterm_of thy;
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59617
diff changeset
   100
    val cty = Thm.global_ctyp_of thy ty;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   101
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   102
    @{thm term_of_anything}
60801
7664e0916eec tuned signature;
wenzelm
parents: 59621
diff changeset
   103
    |> 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
   104
    |> Thm.varifyT_global
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   105
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   106
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
   107
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
   108
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   109
    val ty = Type (tyco, map TFree vs);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   110
    val ty_rep = map_atyps
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   111
      (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
   112
    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
   113
 in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   114
    thy
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 65062
diff changeset
   115
    |> Code.declare_default_eqns_global [(eq, true)]
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   116
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   117
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
   118
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
   119
    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
   120
  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
   121
    (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
   122
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   123
56926
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   124
(* setup *)
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   125
59323
468bd3aedfa1 modernized and more uniform style
haftmann
parents: 59153
diff changeset
   126
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
   127
  (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
   128
  #> 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
   129
  #> Code.abstype_interpretation ensure_term_of_code_abstype);
56926
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   130
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   131
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   132
(** termifying syntax **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   133
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   134
fun map_default f xs =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   135
  let val ys = map f xs
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   136
  in if exists is_some ys
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   137
    then SOME (map2 the_default xs ys)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   138
    else NONE
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   139
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   140
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66330
diff changeset
   141
fun subst_termify_app (Const (\<^const_name>\<open>termify\<close>, _), [t]) =
63619
9c870388e87a more tight filtering;
wenzelm
parents: 63239
diff changeset
   142
      if not (Term.exists_subterm (fn Abs _ => true | _ => false) t)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   143
      then if fold_aterms (fn Const _ => I | _ => K false) t true
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   144
        then SOME (HOLogic.reflect_term t)
51714
88f7f38d5cb9 spelling
haftmann
parents: 51685
diff changeset
   145
        else error "Cannot termify expression containing variable"
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   146
      else error "Cannot termify expression containing abstraction"
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   147
  | subst_termify_app (t, ts) = case map_default subst_termify ts
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   148
     of SOME ts' => SOME (list_comb (t, ts'))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   149
      | NONE => NONE
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   150
and subst_termify (Abs (v, T, t)) = (case subst_termify t
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   151
     of SOME t' => SOME (Abs (v, T, t'))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   152
      | NONE => NONE)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   153
  | subst_termify t = subst_termify_app (strip_comb t) 
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   154
62952
wenzelm
parents: 61674
diff changeset
   155
fun check_termify ts = the_default ts (map_default subst_termify ts);
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   156
62952
wenzelm
parents: 61674
diff changeset
   157
val _ = Context.>> (Syntax_Phases.term_check 0 "termify" (K check_termify));
56926
aaea99edc040 modernized setups
haftmann
parents: 56925
diff changeset
   158
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   159
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   160
(** evaluation **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   161
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41247
diff changeset
   162
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
   163
(
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   164
  type T = unit -> term
59153
wenzelm
parents: 59151
diff changeset
   165
  val empty: T = fn () => raise Fail "Evaluation"
wenzelm
parents: 59151
diff changeset
   166
  fun init _ = empty
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   167
);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   168
val put_term = Evaluation.put;
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   169
val cookie = (Evaluation.get, put_term, "Code_Evaluation.put_term");
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   170
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   171
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
   172
63157
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63064
diff changeset
   173
fun gen_dynamic_value computation ctxt t =
65a81a4ef7f8 clarified naming conventions and code for code evaluation sandwiches
haftmann
parents: 63064
diff changeset
   174
  computation cookie ctxt NONE I (mk_term_of t) [];
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   175
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   176
val dynamic_value = gen_dynamic_value Code_Runtime.dynamic_value;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   177
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
   178
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
   179
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
(** diagnostic **)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   182
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   183
fun tracing s x = (Output.tracing s; x);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   184
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   185
end;