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