src/HOL/Tools/code_evaluation.ML
author wenzelm
Thu, 26 Jul 2012 16:54:44 +0200
changeset 48518 0c86acc069ad
parent 48272 db75b4005d9a
child 51685 385ef6706252
permissions -rw-r--r--
proper arguments for old usedir;
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
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
     9
  val dynamic_value: theory -> term -> term option
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    10
  val dynamic_value_strict: theory -> term -> term
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    11
  val dynamic_value_exn: theory -> term -> term Exn.result
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    12
  val static_value: theory -> string list -> typ list -> term -> term option
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    13
  val static_value_strict: theory -> string list -> typ list -> term -> term
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    14
  val static_value_exn: theory -> string list -> typ list -> term -> term Exn.result
41247
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41184
diff changeset
    15
  val dynamic_conv: theory -> conv
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40726
diff changeset
    16
  val static_conv: theory -> string list -> typ list -> conv
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    17
  val put_term: (unit -> term) -> Proof.context -> Proof.context
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    18
  val tracing: string -> 'a -> 'a
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    19
  val setup: theory -> theory
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    20
end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    21
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    22
structure Code_Evaluation : CODE_EVALUATION =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    23
struct
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    24
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    25
(** term_of instances **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    26
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    27
(* formal definition *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    28
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    29
fun add_term_of tyco raw_vs thy =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    30
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    31
    val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    32
    val ty = Type (tyco, map TFree vs);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    33
    val lhs = Const (@{const_name term_of}, ty --> @{typ term})
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    34
      $ Free ("x", ty);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    35
    val rhs = @{term "undefined :: term"};
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    36
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    37
    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
    38
      o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    39
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    40
    thy
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    41
    |> Class.instantiation ([tyco], vs, @{sort term_of})
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    42
    |> `(fn lthy => Syntax.check_term lthy eq)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    43
    |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    44
    |> snd
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    45
    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    46
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    47
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    48
fun ensure_term_of (tyco, (raw_vs, _)) thy =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    49
  let
48272
db75b4005d9a more direct Sorts.has_instance;
wenzelm
parents: 45429
diff changeset
    50
    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
    51
      andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep};
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    52
  in if need_inst then add_term_of tyco raw_vs thy else thy end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    53
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    54
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    55
(* code equations for datatypes *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    56
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
    57
fun mk_term_of_eq thy ty (c, (_, tys)) =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    58
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    59
    val t = list_comb (Const (c, tys ---> ty),
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 42402
diff changeset
    60
      map Free (Name.invent_names Name.context "a" tys));
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    61
    val (arg, rhs) =
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    62
      pairself (Thm.cterm_of thy o Logic.unvarify_types_global o Logic.varify_global)
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    63
        (t,
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 43329
diff changeset
    64
          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
    65
            (HOLogic.reflect_term t));
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    66
    val cty = Thm.ctyp_of thy ty;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    67
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    68
    @{thm term_of_anything}
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    69
    |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    70
    |> Thm.varifyT_global
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    71
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    72
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    73
fun add_term_of_code tyco raw_vs raw_cs thy =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    74
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    75
    val algebra = Sign.classes_of thy;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    76
    val vs = map (fn (v, sort) =>
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    77
      (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    78
    val ty = Type (tyco, map TFree vs);
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
    79
    val cs = (map o apsnd o apsnd o map o map_atyps)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    80
      (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    81
    val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
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
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    85
    |> Code.del_eqns const
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    86
    |> fold Code.add_eqn eqs
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
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    89
fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    90
  let
48272
db75b4005d9a more direct Sorts.has_instance;
wenzelm
parents: 45429
diff changeset
    91
    val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of};
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    92
  in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    93
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    94
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    95
(* code equations for abstypes *)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    96
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
    97
fun mk_abs_term_of_eq thy ty abs ty_rep proj =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    98
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
    99
    val arg = Var (("x", 0), ty);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   100
    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
   101
      (HOLogic.mk_term_of ty_rep (Const (proj, ty --> ty_rep) $ arg))
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   102
      |> Thm.cterm_of thy;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   103
    val cty = Thm.ctyp_of thy ty;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   104
  in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   105
    @{thm term_of_anything}
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   106
    |> Drule.instantiate' [SOME cty] [SOME (Thm.cterm_of thy arg), SOME rhs]
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   107
    |> Thm.varifyT_global
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   108
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   109
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   110
fun add_abs_term_of_code tyco raw_vs abs raw_ty_rep proj thy =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   111
  let
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   112
    val algebra = Sign.classes_of thy;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   113
    val vs = map (fn (v, sort) =>
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   114
      (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   115
    val ty = Type (tyco, map TFree vs);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   116
    val ty_rep = map_atyps
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   117
      (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_ty_rep;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   118
    val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   119
    val eq = mk_abs_term_of_eq thy ty abs ty_rep proj;
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   120
 in
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   121
    thy
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   122
    |> Code.del_eqns const
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   123
    |> Code.add_eqn eq
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   124
  end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   125
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 39567
diff changeset
   126
fun ensure_abs_term_of_code (tyco, (raw_vs, ((abs, (_, ty)), (proj, _)))) thy =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   127
  let
48272
db75b4005d9a more direct Sorts.has_instance;
wenzelm
parents: 45429
diff changeset
   128
    val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of};
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   129
  in if has_inst then add_abs_term_of_code tyco raw_vs abs ty proj thy else thy end;
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   130
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   131
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
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   141
fun subst_termify_app (Const (@{const_name termify}, _), [t]) =
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   142
      if not (Term.has_abs t)
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)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   145
        else error "Cannot termify expression containing variables"
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
42402
c7139609b67d simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents: 42361
diff changeset
   155
fun check_termify ctxt ts =
c7139609b67d simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents: 42361
diff changeset
   156
  the_default ts (map_default subst_termify ts);
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   157
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   158
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   159
(** evaluation **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   160
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41247
diff changeset
   161
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
   162
(
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   163
  type T = unit -> term
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41247
diff changeset
   164
  (* FIXME avoid user error with non-user text *)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   165
  fun init _ () = error "Evaluation"
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   166
);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   167
val put_term = Evaluation.put;
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   168
val cookie = (Evaluation.get, put_term, "Code_Evaluation.put_term");
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   169
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   170
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
   171
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   172
fun term_of_const_for thy = AxClass.unoverload_const thy o dest_Const o HOLogic.term_of_const;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   173
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   174
fun gen_dynamic_value dynamic_value thy t =
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   175
  dynamic_value cookie thy NONE I (mk_term_of t) [];
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
fun gen_static_value static_value thy consts Ts =
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   182
  static_value cookie thy NONE I (union (op =) (map (term_of_const_for thy) Ts) consts)
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   183
  o mk_term_of;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   184
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   185
val static_value = gen_static_value Code_Runtime.static_value;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   186
val static_value_strict = gen_static_value Code_Runtime.static_value_strict;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   187
val static_value_exn = gen_static_value Code_Runtime.static_value_exn;
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   188
39567
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   189
fun certify_eval thy value conv ct =
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   190
  let
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   191
    val t = Thm.term_of ct;
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   192
    val T = fastype_of t;
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   193
    val mk_eq = Thm.mk_binop (Thm.cterm_of thy (Const ("==", T --> T --> propT)));
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   194
  in case value t
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   195
   of NONE => Thm.reflexive ct
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   196
    | SOME t' => conv (mk_eq ct (Thm.cterm_of thy t')) RS @{thm eq_eq_TrueD}
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   197
        handle THM _ =>
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   198
          error ("Failed to certify evaluation result of " ^ Syntax.string_of_term_global thy t)
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   199
  end;
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   200
41247
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41184
diff changeset
   201
fun dynamic_conv thy = certify_eval thy (dynamic_value thy)
c5cb19ecbd41 avoid slightly odd Conv.tap_thy
haftmann
parents: 41184
diff changeset
   202
  (Code_Runtime.dynamic_holds_conv thy);
39567
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   203
41184
5c6f44d22f51 simplified evaluation function names
haftmann
parents: 40726
diff changeset
   204
fun static_conv thy consts Ts =
39567
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   205
  let
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   206
    val eqs = "==" :: @{const_name HOL.eq} ::
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   207
      map (fn T => AxClass.unoverload_const thy (@{const_name HOL.equal}, T)) Ts;
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   208
        (*assumes particular code equations for "==" etc.*)
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   209
  in
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   210
    certify_eval thy (static_value thy consts Ts)
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   211
      (Code_Runtime.static_holds_conv thy (union (op =) eqs consts))
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   212
  end;
5ee997fbe5cc dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
haftmann
parents: 39565
diff changeset
   213
39565
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   214
f4f87c6e2fad full palette of dynamic/static value(_strict/exn)
haftmann
parents: 39564
diff changeset
   215
(** diagnostic **)
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   216
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   217
fun tracing s x = (Output.tracing s; x);
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   218
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   219
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   220
(** setup **)
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   221
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   222
val setup =
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   223
  Code.datatype_interpretation ensure_term_of
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   224
  #> Code.abstype_interpretation ensure_term_of
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   225
  #> Code.datatype_interpretation ensure_term_of_code
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   226
  #> Code.abstype_interpretation ensure_abs_term_of_code
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45344
diff changeset
   227
  #> Context.theory_map (Syntax_Phases.term_check 0 "termify" check_termify)
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41472
diff changeset
   228
  #> Value.add_evaluator ("code", dynamic_value_strict o Proof_Context.theory_of);
39564
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   229
acfd10e38e80 Factored out ML into separate file
haftmann
parents:
diff changeset
   230
end;