| author | kuncar | 
| Fri, 09 Oct 2015 01:37:57 +0200 | |
| changeset 61367 | 46af4f577c7e | 
| parent 60801 | 7664e0916eec | 
| child 61667 | 4b53042d7a40 | 
| permissions | -rw-r--r-- | 
| 39564 | 1 | (* Title: HOL/Tools/code_evaluation.ML | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | ||
| 4 | Evaluation and reconstruction of terms in ML. | |
| 5 | *) | |
| 6 | ||
| 7 | signature CODE_EVALUATION = | |
| 8 | sig | |
| 55757 | 9 | val dynamic_value: Proof.context -> term -> term option | 
| 10 | val dynamic_value_strict: Proof.context -> term -> term | |
| 11 | val dynamic_value_exn: Proof.context -> term -> term Exn.result | |
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 12 |   val static_value: { ctxt: Proof.context, consts: string list, Ts: typ list }
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 13 | -> Proof.context -> term -> term option | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 14 |   val static_value_strict: { ctxt: Proof.context, consts: string list, Ts: typ list }
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 15 | -> Proof.context -> term -> term | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 16 |   val static_value_exn: { ctxt: Proof.context, consts: string list, Ts: typ list }
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 17 | -> Proof.context -> term -> term Exn.result | 
| 55757 | 18 | val dynamic_conv: Proof.context -> conv | 
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 19 |   val static_conv: { ctxt: Proof.context, consts: string list, Ts: typ list }
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 20 | -> Proof.context -> conv | 
| 39564 | 21 | val put_term: (unit -> term) -> Proof.context -> Proof.context | 
| 22 | val tracing: string -> 'a -> 'a | |
| 23 | end; | |
| 24 | ||
| 25 | structure Code_Evaluation : CODE_EVALUATION = | |
| 26 | struct | |
| 27 | ||
| 28 | (** term_of instances **) | |
| 29 | ||
| 30 | (* formal definition *) | |
| 31 | ||
| 32 | fun add_term_of tyco raw_vs thy = | |
| 33 | let | |
| 34 |     val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
 | |
| 35 | val ty = Type (tyco, map TFree vs); | |
| 36 |     val lhs = Const (@{const_name term_of}, ty --> @{typ term})
 | |
| 37 |       $ Free ("x", ty);
 | |
| 38 |     val rhs = @{term "undefined :: term"};
 | |
| 39 | val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); | |
| 40 | fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst | |
| 41 | o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv"; | |
| 42 | in | |
| 43 | thy | |
| 44 |     |> Class.instantiation ([tyco], vs, @{sort term_of})
 | |
| 45 | |> `(fn lthy => Syntax.check_term lthy eq) | |
| 46 | |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq))) | |
| 47 | |> snd | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59323diff
changeset | 48 | |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) | 
| 39564 | 49 | end; | 
| 50 | ||
| 51 | fun ensure_term_of (tyco, (raw_vs, _)) thy = | |
| 52 | let | |
| 48272 | 53 |     val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of})
 | 
| 54 |       andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep};
 | |
| 39564 | 55 | in if need_inst then add_term_of tyco raw_vs thy else thy end; | 
| 56 | ||
| 57 | ||
| 58 | (* code equations for datatypes *) | |
| 59 | ||
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
39567diff
changeset | 60 | fun mk_term_of_eq thy ty (c, (_, tys)) = | 
| 39564 | 61 | let | 
| 62 | val t = list_comb (Const (c, tys ---> ty), | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
42402diff
changeset | 63 | map Free (Name.invent_names Name.context "a" tys)); | 
| 39564 | 64 | val (arg, rhs) = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59617diff
changeset | 65 | 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: 
43329diff
changeset | 66 | (t, | 
| 
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
 wenzelm parents: 
43329diff
changeset | 67 | 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: 
43329diff
changeset | 68 | (HOLogic.reflect_term t)); | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59617diff
changeset | 69 | val cty = Thm.global_ctyp_of thy ty; | 
| 39564 | 70 | in | 
| 71 |     @{thm term_of_anything}
 | |
| 60801 | 72 | |> Thm.instantiate' [SOME cty] [SOME arg, SOME rhs] | 
| 39564 | 73 | |> Thm.varifyT_global | 
| 74 | end; | |
| 75 | ||
| 76 | fun add_term_of_code tyco raw_vs raw_cs thy = | |
| 77 | let | |
| 78 | val algebra = Sign.classes_of thy; | |
| 79 | val vs = map (fn (v, sort) => | |
| 80 |       (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
 | |
| 81 | val ty = Type (tyco, map TFree vs); | |
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
39567diff
changeset | 82 | val cs = (map o apsnd o apsnd o map o map_atyps) | 
| 39564 | 83 | (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs; | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
48272diff
changeset | 84 |     val const = Axclass.param_of_inst thy (@{const_name term_of}, tyco);
 | 
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 85 | val eqs = map (mk_term_of_eq thy ty) cs; | 
| 39564 | 86 | in | 
| 87 | thy | |
| 88 | |> Code.del_eqns const | |
| 89 | |> fold Code.add_eqn eqs | |
| 90 | end; | |
| 91 | ||
| 92 | fun ensure_term_of_code (tyco, (raw_vs, cs)) thy = | |
| 93 | let | |
| 48272 | 94 |     val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of};
 | 
| 39564 | 95 | in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end; | 
| 96 | ||
| 97 | ||
| 98 | (* code equations for abstypes *) | |
| 99 | ||
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 100 | fun mk_abs_term_of_eq thy ty abs ty_rep proj = | 
| 39564 | 101 | let | 
| 102 |     val arg = Var (("x", 0), ty);
 | |
| 103 |     val rhs = Abs ("y", @{typ term}, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $
 | |
| 104 | (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: 
59617diff
changeset | 105 | |> Thm.global_cterm_of thy; | 
| 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59617diff
changeset | 106 | val cty = Thm.global_ctyp_of thy ty; | 
| 39564 | 107 | in | 
| 108 |     @{thm term_of_anything}
 | |
| 60801 | 109 | |> Thm.instantiate' [SOME cty] [SOME (Thm.global_cterm_of thy arg), SOME rhs] | 
| 39564 | 110 | |> Thm.varifyT_global | 
| 111 | end; | |
| 112 | ||
| 113 | fun add_abs_term_of_code tyco raw_vs abs raw_ty_rep proj thy = | |
| 114 | let | |
| 115 | val algebra = Sign.classes_of thy; | |
| 116 | val vs = map (fn (v, sort) => | |
| 117 |       (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
 | |
| 118 | val ty = Type (tyco, map TFree vs); | |
| 119 | val ty_rep = map_atyps | |
| 120 | (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_ty_rep; | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
48272diff
changeset | 121 |     val const = Axclass.param_of_inst thy (@{const_name term_of}, tyco);
 | 
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 122 | val eq = mk_abs_term_of_eq thy ty abs ty_rep proj; | 
| 39564 | 123 | in | 
| 124 | thy | |
| 125 | |> Code.del_eqns const | |
| 126 | |> Code.add_eqn eq | |
| 127 | end; | |
| 128 | ||
| 40726 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 haftmann parents: 
39567diff
changeset | 129 | fun ensure_abs_term_of_code (tyco, (raw_vs, ((abs, (_, ty)), (proj, _)))) thy = | 
| 39564 | 130 | let | 
| 48272 | 131 |     val has_inst = Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of};
 | 
| 39564 | 132 | in if has_inst then add_abs_term_of_code tyco raw_vs abs ty proj thy else thy end; | 
| 133 | ||
| 134 | ||
| 56926 | 135 | (* setup *) | 
| 136 | ||
| 59323 | 137 | val _ = Theory.setup | 
| 56926 | 138 | (Code.datatype_interpretation ensure_term_of | 
| 139 | #> Code.abstype_interpretation ensure_term_of | |
| 140 | #> Code.datatype_interpretation ensure_term_of_code | |
| 59323 | 141 | #> Code.abstype_interpretation ensure_abs_term_of_code); | 
| 56926 | 142 | |
| 143 | ||
| 39564 | 144 | (** termifying syntax **) | 
| 145 | ||
| 146 | fun map_default f xs = | |
| 147 | let val ys = map f xs | |
| 148 | in if exists is_some ys | |
| 149 | then SOME (map2 the_default xs ys) | |
| 150 | else NONE | |
| 151 | end; | |
| 152 | ||
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 153 | fun subst_termify_app (Const (@{const_name termify}, _), [t]) =
 | 
| 39564 | 154 | if not (Term.has_abs t) | 
| 155 | then if fold_aterms (fn Const _ => I | _ => K false) t true | |
| 156 | then SOME (HOLogic.reflect_term t) | |
| 51714 | 157 | else error "Cannot termify expression containing variable" | 
| 39564 | 158 | else error "Cannot termify expression containing abstraction" | 
| 159 | | subst_termify_app (t, ts) = case map_default subst_termify ts | |
| 160 | of SOME ts' => SOME (list_comb (t, ts')) | |
| 161 | | NONE => NONE | |
| 162 | and subst_termify (Abs (v, T, t)) = (case subst_termify t | |
| 163 | of SOME t' => SOME (Abs (v, T, t')) | |
| 164 | | NONE => NONE) | |
| 165 | | subst_termify t = subst_termify_app (strip_comb t) | |
| 166 | ||
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 167 | fun check_termify _ ts = the_default ts (map_default subst_termify ts); | 
| 39564 | 168 | |
| 56926 | 169 | val _ = Context.>> (Syntax_Phases.term_check 0 "termify" check_termify); | 
| 170 | ||
| 39564 | 171 | |
| 172 | (** evaluation **) | |
| 173 | ||
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41247diff
changeset | 174 | structure Evaluation = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41247diff
changeset | 175 | ( | 
| 39564 | 176 | type T = unit -> term | 
| 59153 | 177 | val empty: T = fn () => raise Fail "Evaluation" | 
| 178 | fun init _ = empty | |
| 39564 | 179 | ); | 
| 180 | val put_term = Evaluation.put; | |
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 181 | val cookie = (Evaluation.get, put_term, "Code_Evaluation.put_term"); | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 182 | |
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 183 | fun mk_term_of t = HOLogic.mk_term_of (fastype_of t) t; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 184 | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
48272diff
changeset | 185 | fun term_of_const_for thy = Axclass.unoverload_const thy o dest_Const o HOLogic.term_of_const; | 
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 186 | |
| 55757 | 187 | fun gen_dynamic_value dynamic_value ctxt t = | 
| 188 | dynamic_value cookie ctxt NONE I (mk_term_of t) []; | |
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 189 | |
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 190 | val dynamic_value = gen_dynamic_value Code_Runtime.dynamic_value; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 191 | val dynamic_value_strict = gen_dynamic_value Code_Runtime.dynamic_value_strict; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 192 | val dynamic_value_exn = gen_dynamic_value Code_Runtime.dynamic_value_exn; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 193 | |
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 194 | fun gen_static_value static_value { ctxt, consts, Ts } =
 | 
| 55757 | 195 | let | 
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 196 | val static_value' = static_value cookie | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 197 |       { ctxt = ctxt, target = NONE, lift_postproc = I, consts =
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 198 | union (op =) (map (term_of_const_for (Proof_Context.theory_of ctxt)) Ts) consts } | 
| 55757 | 199 | in fn ctxt' => fn t => static_value' ctxt' (mk_term_of t) end; | 
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 200 | |
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 201 | val static_value = gen_static_value Code_Runtime.static_value; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 202 | val static_value_strict = gen_static_value Code_Runtime.static_value_strict; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 203 | val static_value_exn = gen_static_value Code_Runtime.static_value_exn; | 
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 204 | |
| 55757 | 205 | fun certify_eval ctxt value conv ct = | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 206 | let | 
| 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 207 | val t = Thm.term_of ct; | 
| 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 208 | val T = fastype_of t; | 
| 59617 | 209 | val mk_eq = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59617diff
changeset | 210 |       Thm.mk_binop (Thm.cterm_of ctxt (Const (@{const_name Pure.eq}, T --> T --> propT)));
 | 
| 55757 | 211 | in case value ctxt t | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 212 | of NONE => Thm.reflexive ct | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59617diff
changeset | 213 |     | SOME t' => conv ctxt (mk_eq ct (Thm.cterm_of ctxt t')) RS @{thm eq_eq_TrueD}
 | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 214 | handle THM _ => | 
| 55757 | 215 |           error ("Failed to certify evaluation result of " ^ Syntax.string_of_term ctxt t)
 | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 216 | end; | 
| 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 217 | |
| 55757 | 218 | fun dynamic_conv ctxt = certify_eval ctxt dynamic_value | 
| 219 | Code_Runtime.dynamic_holds_conv; | |
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 220 | |
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 221 | fun static_conv { ctxt, consts, Ts }  =
 | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 222 | let | 
| 56245 | 223 |     val eqs = @{const_name Pure.eq} :: @{const_name HOL.eq} ::
 | 
| 55757 | 224 | map (fn T => Axclass.unoverload_const (Proof_Context.theory_of ctxt) | 
| 56245 | 225 |         (@{const_name HOL.equal}, T)) Ts; (*assumes particular code equations for Pure.eq etc.*)
 | 
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 226 |     val value = static_value { ctxt = ctxt, consts = consts, Ts = Ts };
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56926diff
changeset | 227 |     val holds = Code_Runtime.static_holds_conv { ctxt = ctxt, consts = union (op =) eqs consts };
 | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 228 | in | 
| 55757 | 229 | fn ctxt' => certify_eval ctxt' value holds | 
| 39567 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 230 | end; | 
| 
5ee997fbe5cc
dynamic_eval_conv static_eval_conv: certification of previously unreliably reconstructed evaluated term
 haftmann parents: 
39565diff
changeset | 231 | |
| 39565 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 232 | |
| 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 haftmann parents: 
39564diff
changeset | 233 | (** diagnostic **) | 
| 39564 | 234 | |
| 235 | fun tracing s x = (Output.tracing s; x); | |
| 236 | ||
| 237 | end; |