| author | wenzelm | 
| Sat, 04 Mar 2017 20:26:32 +0100 | |
| changeset 65105 | 1f47b92021de | 
| parent 65062 | dc746d43f40e | 
| child 66251 | cd935b7cb3fb | 
| 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  | 
|
| 39564 | 12  | 
val put_term: (unit -> term) -> Proof.context -> Proof.context  | 
13  | 
val tracing: string -> 'a -> 'a  | 
|
14  | 
end;  | 
|
15  | 
||
16  | 
structure Code_Evaluation : CODE_EVALUATION =  | 
|
17  | 
struct  | 
|
18  | 
||
19  | 
(** term_of instances **)  | 
|
20  | 
||
21  | 
(* formal definition *)  | 
|
22  | 
||
23  | 
fun add_term_of tyco raw_vs thy =  | 
|
24  | 
let  | 
|
25  | 
    val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
 | 
|
26  | 
val ty = Type (tyco, map TFree vs);  | 
|
27  | 
    val lhs = Const (@{const_name term_of}, ty --> @{typ term})
 | 
|
28  | 
      $ Free ("x", ty);
 | 
|
29  | 
    val rhs = @{term "undefined :: term"};
 | 
|
30  | 
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));  | 
|
31  | 
fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst  | 
|
32  | 
o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";  | 
|
33  | 
in  | 
|
34  | 
thy  | 
|
35  | 
    |> Class.instantiation ([tyco], vs, @{sort term_of})
 | 
|
36  | 
|> `(fn lthy => Syntax.check_term lthy eq)  | 
|
| 63180 | 37  | 
|-> (fn eq => Specification.definition NONE [] [] ((Binding.name (triv_name_of eq), []), eq))  | 
| 39564 | 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 | 40  | 
end;  | 
41  | 
||
| 
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
 | 
42  | 
fun ensure_term_of (tyco, (vs, _)) thy =  | 
| 39564 | 43  | 
let  | 
| 48272 | 44  | 
    val need_inst = not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort term_of})
 | 
45  | 
      andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep};
 | 
|
| 
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
 | 
46  | 
in if need_inst then add_term_of tyco vs thy else thy end;  | 
| 
 
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  | 
| 
 
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  | 
    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
 | 
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 | 57  | 
|
58  | 
||
59  | 
(* code equations for datatypes *)  | 
|
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 | 62  | 
let  | 
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 | 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 | 71  | 
in  | 
72  | 
    @{thm term_of_anything}
 | 
|
| 60801 | 73  | 
|> Thm.instantiate' [SOME cty] [SOME arg, SOME rhs]  | 
| 39564 | 74  | 
|> Thm.varifyT_global  | 
75  | 
end;  | 
|
76  | 
||
| 
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
 | 
77  | 
fun add_term_of_code tyco vs raw_cs thy =  | 
| 39564 | 78  | 
let  | 
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 | 81  | 
(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: 
48272 
diff
changeset
 | 
82  | 
    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
 | 
83  | 
val eqs = map (mk_term_of_eq thy ty) cs;  | 
| 39564 | 84  | 
in  | 
85  | 
thy  | 
|
86  | 
|> Code.del_eqns const  | 
|
| 
63239
 
d562c9948dee
explicit tagging of code equations de-baroquifies interface
 
haftmann 
parents: 
63180 
diff
changeset
 | 
87  | 
|> fold (Code.add_eqn (Code.Equation, false)) eqs  | 
| 39564 | 88  | 
end;  | 
89  | 
||
| 
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
 | 
90  | 
fun ensure_term_of_code (tyco, (vs, cs)) =  | 
| 
 
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
 | 
91  | 
for_term_of_instance tyco vs (fn tyco => fn vs => add_term_of_code tyco vs cs);  | 
| 39564 | 92  | 
|
93  | 
||
94  | 
(* code equations for abstypes *)  | 
|
95  | 
||
| 
39565
 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 
haftmann 
parents: 
39564 
diff
changeset
 | 
96  | 
fun mk_abs_term_of_eq thy ty abs ty_rep proj =  | 
| 39564 | 97  | 
let  | 
98  | 
    val arg = Var (("x", 0), ty);
 | 
|
99  | 
    val rhs = Abs ("y", @{typ term}, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $
 | 
|
100  | 
(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
 | 
101  | 
|> Thm.global_cterm_of thy;  | 
| 
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59617 
diff
changeset
 | 
102  | 
val cty = Thm.global_ctyp_of thy ty;  | 
| 39564 | 103  | 
in  | 
104  | 
    @{thm term_of_anything}
 | 
|
| 60801 | 105  | 
|> Thm.instantiate' [SOME cty] [SOME (Thm.global_cterm_of thy arg), SOME rhs]  | 
| 39564 | 106  | 
|> Thm.varifyT_global  | 
107  | 
end;  | 
|
108  | 
||
| 
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
 | 
109  | 
fun add_abs_term_of_code tyco vs abs raw_ty_rep proj thy =  | 
| 39564 | 110  | 
let  | 
111  | 
val ty = Type (tyco, map TFree vs);  | 
|
112  | 
val ty_rep = map_atyps  | 
|
113  | 
(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: 
48272 
diff
changeset
 | 
114  | 
    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
 | 
115  | 
val eq = mk_abs_term_of_eq thy ty abs ty_rep proj;  | 
| 39564 | 116  | 
in  | 
117  | 
thy  | 
|
118  | 
|> Code.del_eqns const  | 
|
| 
63239
 
d562c9948dee
explicit tagging of code equations de-baroquifies interface
 
haftmann 
parents: 
63180 
diff
changeset
 | 
119  | 
|> Code.add_eqn (Code.Equation, false) eq  | 
| 39564 | 120  | 
end;  | 
121  | 
||
| 
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
 | 
122  | 
fun ensure_abs_term_of_code (tyco, (vs, ((abs, (_, ty)), (proj, _)))) =  | 
| 
 
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
 | 
123  | 
for_term_of_instance tyco vs  | 
| 
 
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
 | 
124  | 
(fn tyco => fn vs => add_abs_term_of_code tyco vs abs ty proj);  | 
| 39564 | 125  | 
|
126  | 
||
| 56926 | 127  | 
(* setup *)  | 
128  | 
||
| 59323 | 129  | 
val _ = Theory.setup  | 
| 56926 | 130  | 
(Code.datatype_interpretation ensure_term_of  | 
131  | 
#> Code.abstype_interpretation ensure_term_of  | 
|
132  | 
#> Code.datatype_interpretation ensure_term_of_code  | 
|
| 59323 | 133  | 
#> Code.abstype_interpretation ensure_abs_term_of_code);  | 
| 56926 | 134  | 
|
135  | 
||
| 39564 | 136  | 
(** termifying syntax **)  | 
137  | 
||
138  | 
fun map_default f xs =  | 
|
139  | 
let val ys = map f xs  | 
|
140  | 
in if exists is_some ys  | 
|
141  | 
then SOME (map2 the_default xs ys)  | 
|
142  | 
else NONE  | 
|
143  | 
end;  | 
|
144  | 
||
| 
39565
 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 
haftmann 
parents: 
39564 
diff
changeset
 | 
145  | 
fun subst_termify_app (Const (@{const_name termify}, _), [t]) =
 | 
| 63619 | 146  | 
if not (Term.exists_subterm (fn Abs _ => true | _ => false) t)  | 
| 39564 | 147  | 
then if fold_aterms (fn Const _ => I | _ => K false) t true  | 
148  | 
then SOME (HOLogic.reflect_term t)  | 
|
| 51714 | 149  | 
else error "Cannot termify expression containing variable"  | 
| 39564 | 150  | 
else error "Cannot termify expression containing abstraction"  | 
151  | 
| subst_termify_app (t, ts) = case map_default subst_termify ts  | 
|
152  | 
of SOME ts' => SOME (list_comb (t, ts'))  | 
|
153  | 
| NONE => NONE  | 
|
154  | 
and subst_termify (Abs (v, T, t)) = (case subst_termify t  | 
|
155  | 
of SOME t' => SOME (Abs (v, T, t'))  | 
|
156  | 
| NONE => NONE)  | 
|
157  | 
| subst_termify t = subst_termify_app (strip_comb t)  | 
|
158  | 
||
| 62952 | 159  | 
fun check_termify ts = the_default ts (map_default subst_termify ts);  | 
| 39564 | 160  | 
|
| 62952 | 161  | 
val _ = Context.>> (Syntax_Phases.term_check 0 "termify" (K check_termify));  | 
| 56926 | 162  | 
|
| 39564 | 163  | 
|
164  | 
(** evaluation **)  | 
|
165  | 
||
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41247 
diff
changeset
 | 
166  | 
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
 | 
167  | 
(  | 
| 39564 | 168  | 
type T = unit -> term  | 
| 59153 | 169  | 
val empty: T = fn () => raise Fail "Evaluation"  | 
170  | 
fun init _ = empty  | 
|
| 39564 | 171  | 
);  | 
172  | 
val put_term = Evaluation.put;  | 
|
| 
39565
 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 
haftmann 
parents: 
39564 
diff
changeset
 | 
173  | 
val cookie = (Evaluation.get, put_term, "Code_Evaluation.put_term");  | 
| 
 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 
haftmann 
parents: 
39564 
diff
changeset
 | 
174  | 
|
| 
 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 
haftmann 
parents: 
39564 
diff
changeset
 | 
175  | 
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
 | 
176  | 
|
| 
63157
 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 
haftmann 
parents: 
63064 
diff
changeset
 | 
177  | 
fun gen_dynamic_value computation ctxt t =  | 
| 
 
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
 
haftmann 
parents: 
63064 
diff
changeset
 | 
178  | 
computation cookie ctxt NONE I (mk_term_of t) [];  | 
| 
39565
 
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  | 
val dynamic_value = gen_dynamic_value Code_Runtime.dynamic_value;  | 
| 
 
f4f87c6e2fad
full palette of dynamic/static value(_strict/exn)
 
haftmann 
parents: 
39564 
diff
changeset
 | 
181  | 
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
 | 
182  | 
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
 | 
183  | 
|
| 
 
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  | 
(** diagnostic **)  | 
| 39564 | 186  | 
|
187  | 
fun tracing s x = (Output.tracing s; x);  | 
|
188  | 
||
189  | 
end;  |