20427
|
1 |
(* ID: $Id$
|
|
2 |
Author: Florian Haftmann, TU Muenchen
|
|
3 |
*)
|
|
4 |
|
|
5 |
header {* A simple embedded term evaluation mechanism *}
|
|
6 |
|
|
7 |
theory CodeEval
|
|
8 |
imports CodeEmbed
|
|
9 |
begin
|
|
10 |
|
|
11 |
section {* A simple embedded term evaluation mechanism *}
|
|
12 |
|
|
13 |
subsection {* The typ_of class *}
|
|
14 |
|
|
15 |
class typ_of =
|
|
16 |
fixes typ_of :: "'a option \<Rightarrow> typ"
|
|
17 |
|
|
18 |
ML {*
|
|
19 |
structure TypOf =
|
|
20 |
struct
|
|
21 |
|
|
22 |
local
|
|
23 |
val thy = the_context ();
|
|
24 |
val const_typ_of = Sign.intern_const thy "typ_of";
|
|
25 |
val const_None = Sign.intern_const thy "None";
|
|
26 |
fun typ_option ty = Type (Sign.intern_type (the_context ()) "option", [ty]);
|
|
27 |
fun term_typ_of ty = Const (const_typ_of, typ_option ty --> Embed.typ_typ);
|
|
28 |
in
|
|
29 |
val class_typ_of = Sign.intern_class thy "typ_of";
|
|
30 |
fun term_typ_of_None ty =
|
|
31 |
term_typ_of ty $ Const (const_None, typ_option ty);
|
|
32 |
fun mk_typ_of_def ty =
|
|
33 |
let
|
|
34 |
val lhs = term_typ_of ty $ Free ("x", typ_option ty)
|
|
35 |
val rhs = Embed.term_typ (fn v => term_typ_of_None (TFree v)) ty
|
|
36 |
in Logic.mk_equals (lhs, rhs) end;
|
|
37 |
end;
|
|
38 |
|
|
39 |
end;
|
|
40 |
*}
|
|
41 |
|
|
42 |
setup {*
|
|
43 |
let
|
|
44 |
fun mk _ arities _ =
|
|
45 |
maps (fn ((tyco, asorts), _) => [(("", []), TypOf.mk_typ_of_def
|
|
46 |
(Type (tyco, map TFree (Name.names Name.context "'a" asorts))))]) arities;
|
|
47 |
fun tac _ = ClassPackage.intro_classes_tac [];
|
|
48 |
fun hook specs =
|
|
49 |
DatatypeCodegen.prove_codetypes_arities tac
|
|
50 |
(map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
|
|
51 |
[TypOf.class_typ_of] mk
|
|
52 |
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
|
|
53 |
*}
|
|
54 |
|
|
55 |
|
|
56 |
subsection {* term_of class *}
|
|
57 |
|
|
58 |
class term_of = typ_of +
|
|
59 |
constrains typ_of :: "'a option \<Rightarrow> typ"
|
|
60 |
fixes term_of :: "'a \<Rightarrow> term"
|
|
61 |
|
|
62 |
ML {*
|
|
63 |
structure TermOf =
|
|
64 |
struct
|
|
65 |
|
|
66 |
local
|
|
67 |
val thy = the_context ();
|
|
68 |
val const_term_of = Sign.intern_const thy "term_of";
|
|
69 |
fun term_term_of ty = Const (const_term_of, ty --> Embed.typ_term);
|
|
70 |
in
|
|
71 |
val class_term_of = Sign.intern_class thy "term_of";
|
|
72 |
fun mk_terms_of_defs vs (tyco, cs) =
|
|
73 |
let
|
|
74 |
val dty = Type (tyco, map TFree vs);
|
|
75 |
fun mk_eq c =
|
|
76 |
let
|
|
77 |
val lhs : term = term_term_of dty $ c;
|
|
78 |
val rhs : term = Embed.term_term
|
|
79 |
(fn (v, ty) => term_term_of ty $ Free (v, ty))
|
|
80 |
(Embed.term_typ (fn (v, sort) => TypOf.term_typ_of_None (TFree (v, sort)))) c
|
|
81 |
in
|
|
82 |
HOLogic.mk_eq (lhs, rhs)
|
|
83 |
end;
|
|
84 |
in map mk_eq cs end;
|
|
85 |
fun mk_term_of t =
|
|
86 |
term_term_of (Term.fastype_of t) $ t;
|
|
87 |
end;
|
|
88 |
|
|
89 |
end;
|
|
90 |
*}
|
|
91 |
|
|
92 |
setup {*
|
|
93 |
let
|
|
94 |
fun mk thy arities css =
|
|
95 |
let
|
|
96 |
val vs = (Name.names Name.context "'a" o snd o fst o hd) arities;
|
|
97 |
val raw_defs = map (TermOf.mk_terms_of_defs vs) css;
|
|
98 |
fun mk' thy' = map (apfst (rpair [])) ((PrimrecPackage.mk_combdefs thy' o flat) raw_defs)
|
|
99 |
in ClassPackage.assume_arities_thy thy arities mk' end;
|
|
100 |
fun tac _ = ClassPackage.intro_classes_tac [];
|
|
101 |
fun hook specs =
|
|
102 |
if (fst o hd) specs = (fst o dest_Type) Embed.typ_typ then I
|
|
103 |
else
|
|
104 |
DatatypeCodegen.prove_codetypes_arities tac
|
|
105 |
(map (fn (tyco, (is_dt, _)) => (tyco, is_dt)) specs)
|
|
106 |
[TermOf.class_term_of] mk
|
|
107 |
in DatatypeCodegen.add_codetypes_hook_bootstrap hook end
|
|
108 |
*}
|
|
109 |
|
|
110 |
|
|
111 |
subsection {* Evaluation infrastructure *}
|
|
112 |
|
|
113 |
lemma lift_eq_Trueprop:
|
|
114 |
"p == q \<Longrightarrow> Trueprop p == Trueprop q" by auto
|
|
115 |
|
|
116 |
ML {*
|
|
117 |
signature EVAL =
|
|
118 |
sig
|
|
119 |
val eval_term: term -> theory -> term * theory
|
|
120 |
val eval_term' : theory -> term -> term
|
|
121 |
val term: string -> unit
|
|
122 |
val eval_ref: term ref
|
|
123 |
val oracle: string * (theory * exn -> term)
|
|
124 |
val method: Method.src -> Proof.context -> Method.method
|
|
125 |
end;
|
|
126 |
|
|
127 |
structure Eval : EVAL =
|
|
128 |
struct
|
|
129 |
|
|
130 |
val eval_ref = ref Logic.protectC;
|
|
131 |
|
|
132 |
fun eval_term t =
|
|
133 |
CodegenPackage.eval_term
|
|
134 |
(("Eval.eval_ref", eval_ref), TermOf.mk_term_of t);
|
|
135 |
|
|
136 |
fun eval_term' thy t =
|
|
137 |
let
|
|
138 |
val thy' = Theory.copy thy;
|
|
139 |
val (t', _) = eval_term t thy';
|
|
140 |
in t' end;
|
|
141 |
|
|
142 |
fun term t =
|
|
143 |
let
|
|
144 |
val thy = the_context ();
|
|
145 |
val t' = eval_term' thy (Sign.read_term thy t);
|
|
146 |
in () end;
|
|
147 |
|
|
148 |
val lift_eq_Trueprop = thm "lift_eq_Trueprop";
|
|
149 |
|
|
150 |
exception Eval of term;
|
|
151 |
|
|
152 |
val oracle = ("Eval", fn (thy, Eval t) =>
|
|
153 |
Logic.mk_equals (t, eval_term' thy t));
|
|
154 |
|
|
155 |
val oracle_name = NameSpace.pack [Context.theory_name (the_context ()), fst oracle];
|
|
156 |
|
|
157 |
fun conv ct =
|
|
158 |
let
|
|
159 |
val {thy, t, ...} = rep_cterm ct;
|
|
160 |
val t' = HOLogic.dest_Trueprop t;
|
|
161 |
val thm' = Thm.invoke_oracle_i thy oracle_name (thy, Eval t');
|
|
162 |
in
|
|
163 |
lift_eq_Trueprop OF [thm']
|
|
164 |
end;
|
|
165 |
|
|
166 |
fun tac i = Tactical.PRIMITIVE (Drule.fconv_rule
|
|
167 |
(Drule.goals_conv (equal i) conv));
|
|
168 |
|
|
169 |
val method =
|
|
170 |
Method.no_args (Method.METHOD (fn _ =>
|
|
171 |
tac 1 THEN rtac TrueI 1));
|
|
172 |
|
|
173 |
end;
|
|
174 |
*}
|
|
175 |
|
|
176 |
setup {*
|
|
177 |
Theory.add_oracle Eval.oracle
|
|
178 |
#> Method.add_method ("eval", Eval.method, "solve goal by evaluation")
|
|
179 |
*}
|
|
180 |
|
|
181 |
|
|
182 |
subsection {* Small examples *}
|
|
183 |
|
|
184 |
ML {* Eval.term "[]::nat list" *}
|
|
185 |
ML {* Eval.term "(Suc 2 + Suc 0) * Suc 3" *}
|
|
186 |
ML {* Eval.term "fst ([]::nat list, Suc 0) = []" *}
|
|
187 |
|
|
188 |
text {* a fancy datatype *}
|
|
189 |
|
|
190 |
datatype ('a, 'b) bair =
|
|
191 |
Bair "'a\<Colon>order" 'b
|
|
192 |
| Shift "('a, 'b) cair"
|
|
193 |
| Dummy unit
|
|
194 |
and ('a, 'b) cair =
|
|
195 |
Cair 'a 'b
|
|
196 |
|
|
197 |
code_generate term_of
|
|
198 |
|
|
199 |
ML {* Eval.term "Shift (Cair (4::nat) [Suc 0])" *}
|
|
200 |
|
|
201 |
lemma
|
|
202 |
"Suc 0 = 1" by eval
|
|
203 |
|
|
204 |
lemma
|
|
205 |
"rev [0, Suc 0, Suc 0] = [Suc 0, Suc 0, 0]" by eval
|
|
206 |
|
|
207 |
lemma
|
20435
|
208 |
"fst (snd (fst ( ((Some (2::nat), (Suc 0, ())), [0::nat]) ))) = Suc (Suc 0) - 1" by eval
|
20427
|
209 |
|
|
210 |
end
|