| author | haftmann | 
| Wed, 18 Feb 2009 13:39:05 +0100 | |
| changeset 29966 | 27e29256e9f1 | 
| parent 29575 | 41d604e59e93 | 
| child 30427 | dfd31c1db060 | 
| permissions | -rw-r--r-- | 
| 28228 | 1  | 
(* Title: HOL/Code_Eval.thy  | 
2  | 
Author: Florian Haftmann, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Term evaluation using the generic code generator *}
 | 
|
6  | 
||
7  | 
theory Code_Eval  | 
|
| 
28952
 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 
haftmann 
parents: 
28661 
diff
changeset
 | 
8  | 
imports Plain Typerep  | 
| 28228 | 9  | 
begin  | 
10  | 
||
11  | 
subsection {* Term representation *}
 | 
|
12  | 
||
13  | 
subsubsection {* Terms and class @{text term_of} *}
 | 
|
14  | 
||
15  | 
datatype "term" = dummy_term  | 
|
16  | 
||
| 28661 | 17  | 
definition Const :: "message_string \<Rightarrow> typerep \<Rightarrow> term" where  | 
| 28228 | 18  | 
"Const _ _ = dummy_term"  | 
19  | 
||
| 28661 | 20  | 
definition App :: "term \<Rightarrow> term \<Rightarrow> term" where  | 
| 28228 | 21  | 
"App _ _ = dummy_term"  | 
22  | 
||
23  | 
code_datatype Const App  | 
|
24  | 
||
| 28335 | 25  | 
class term_of = typerep +  | 
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
28965 
diff
changeset
 | 
26  | 
  fixes term_of :: "'a::{} \<Rightarrow> term"
 | 
| 28228 | 27  | 
|
28  | 
lemma term_of_anything: "term_of x \<equiv> t"  | 
|
29  | 
by (rule eq_reflection) (cases "term_of x", cases t, simp)  | 
|
30  | 
||
31  | 
ML {*
 | 
|
32  | 
structure Eval =  | 
|
33  | 
struct  | 
|
34  | 
||
35  | 
fun mk_term f g (Const (c, ty)) =  | 
|
36  | 
      @{term Const} $ Message_String.mk c $ g ty
 | 
|
37  | 
| mk_term f g (t1 $ t2) =  | 
|
38  | 
      @{term App} $ mk_term f g t1 $ mk_term f g t2
 | 
|
39  | 
| mk_term f g (Free v) = f v  | 
|
40  | 
| mk_term f g (Bound i) = Bound i  | 
|
41  | 
  | mk_term f g (Abs (v, _, t)) = Abs (v, @{typ term}, mk_term f g t);
 | 
|
42  | 
||
43  | 
fun mk_term_of ty t = Const (@{const_name term_of}, ty --> @{typ term}) $ t;
 | 
|
44  | 
||
45  | 
end;  | 
|
46  | 
*}  | 
|
47  | 
||
48  | 
||
49  | 
subsubsection {* @{text term_of} instances *}
 | 
|
50  | 
||
51  | 
setup {*
 | 
|
52  | 
let  | 
|
53  | 
fun add_term_of_def ty vs tyco thy =  | 
|
54  | 
let  | 
|
55  | 
      val lhs = Const (@{const_name term_of}, ty --> @{typ term})
 | 
|
56  | 
        $ Free ("x", ty);
 | 
|
57  | 
      val rhs = @{term "undefined \<Colon> term"};
 | 
|
58  | 
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));  | 
|
| 28243 | 59  | 
fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst  | 
60  | 
o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";  | 
|
| 28228 | 61  | 
in  | 
62  | 
thy  | 
|
63  | 
      |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
 | 
|
64  | 
|> `(fn lthy => Syntax.check_term lthy eq)  | 
|
| 28965 | 65  | 
|-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))  | 
| 28228 | 66  | 
|> snd  | 
67  | 
|> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))  | 
|
| 28394 | 68  | 
|> LocalTheory.exit_global  | 
| 28228 | 69  | 
end;  | 
70  | 
fun interpretator (tyco, (raw_vs, _)) thy =  | 
|
71  | 
let  | 
|
72  | 
      val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
 | 
|
73  | 
val constrain_sort =  | 
|
74  | 
        curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
 | 
|
75  | 
val vs = (map o apsnd) constrain_sort raw_vs;  | 
|
76  | 
val ty = Type (tyco, map TFree vs);  | 
|
77  | 
in  | 
|
78  | 
thy  | 
|
| 28335 | 79  | 
|> Typerep.perhaps_add_def tyco  | 
| 28228 | 80  | 
|> not has_inst ? add_term_of_def ty vs tyco  | 
81  | 
end;  | 
|
82  | 
in  | 
|
83  | 
Code.type_interpretation interpretator  | 
|
84  | 
end  | 
|
85  | 
*}  | 
|
86  | 
||
87  | 
setup {*
 | 
|
88  | 
let  | 
|
89  | 
fun mk_term_of_eq ty vs tyco (c, tys) =  | 
|
90  | 
let  | 
|
91  | 
val t = list_comb (Const (c, tys ---> ty),  | 
|
92  | 
map Free (Name.names Name.context "a" tys));  | 
|
93  | 
in (map_aterms (fn Free (v, ty) => Var ((v, 0), ty) | t => t) t, Eval.mk_term  | 
|
94  | 
(fn (v, ty) => Eval.mk_term_of ty (Var ((v, 0), ty)))  | 
|
| 28335 | 95  | 
(Typerep.mk (fn (v, sort) => Typerep.typerep (TFree (v, sort)))) t)  | 
| 28228 | 96  | 
end;  | 
97  | 
fun prove_term_of_eq ty eq thy =  | 
|
98  | 
let  | 
|
99  | 
val cty = Thm.ctyp_of thy ty;  | 
|
100  | 
val (arg, rhs) = pairself (Thm.cterm_of thy) eq;  | 
|
101  | 
      val thm = @{thm term_of_anything}
 | 
|
102  | 
|> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]  | 
|
103  | 
|> Thm.varifyT;  | 
|
104  | 
in  | 
|
105  | 
thy  | 
|
| 28370 | 106  | 
|> Code.add_eqn thm  | 
| 28228 | 107  | 
end;  | 
108  | 
fun interpretator (tyco, (raw_vs, raw_cs)) thy =  | 
|
109  | 
let  | 
|
110  | 
val constrain_sort =  | 
|
111  | 
        curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
 | 
|
112  | 
val vs = (map o apsnd) constrain_sort raw_vs;  | 
|
113  | 
val cs = (map o apsnd o map o map_atyps)  | 
|
114  | 
(fn TFree (v, sort) => TFree (v, constrain_sort sort)) raw_cs;  | 
|
115  | 
val ty = Type (tyco, map TFree vs);  | 
|
116  | 
val eqs = map (mk_term_of_eq ty vs tyco) cs;  | 
|
117  | 
      val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
 | 
|
118  | 
in  | 
|
119  | 
thy  | 
|
| 28370 | 120  | 
|> Code.del_eqns const  | 
| 28228 | 121  | 
|> fold (prove_term_of_eq ty) eqs  | 
122  | 
end;  | 
|
123  | 
in  | 
|
124  | 
Code.type_interpretation interpretator  | 
|
125  | 
end  | 
|
126  | 
*}  | 
|
127  | 
||
128  | 
||
129  | 
subsubsection {* Code generator setup *}
 | 
|
130  | 
||
| 28562 | 131  | 
lemmas [code del] = term.recs term.cases term.size  | 
132  | 
lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..  | 
|
| 28228 | 133  | 
|
| 28562 | 134  | 
lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..  | 
135  | 
lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..  | 
|
136  | 
lemma [code, code del]: "(term_of \<Colon> message_string \<Rightarrow> term) = term_of" ..  | 
|
| 28228 | 137  | 
|
| 28562 | 138  | 
lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_nibble_def, code]: "Code_Eval.term_of c =  | 
| 28243 | 139  | 
(let (n, m) = nibble_pair_of_char c  | 
| 28335 | 140  | 
in Code_Eval.App (Code_Eval.App (Code_Eval.Const (STR ''Pair'') (TYPEREP(nibble \<Rightarrow> nibble \<Rightarrow> char)))  | 
| 28243 | 141  | 
(Code_Eval.term_of n)) (Code_Eval.term_of m))"  | 
142  | 
by (subst term_of_anything) rule  | 
|
143  | 
||
| 28228 | 144  | 
code_type "term"  | 
145  | 
(SML "Term.term")  | 
|
146  | 
||
147  | 
code_const Const and App  | 
|
148  | 
(SML "Term.Const/ (_, _)" and "Term.$/ (_, _)")  | 
|
149  | 
||
150  | 
code_const "term_of \<Colon> message_string \<Rightarrow> term"  | 
|
151  | 
(SML "Message'_String.mk")  | 
|
152  | 
||
153  | 
||
154  | 
subsection {* Evaluation setup *}
 | 
|
155  | 
||
156  | 
ML {*
 | 
|
157  | 
signature EVAL =  | 
|
158  | 
sig  | 
|
159  | 
val mk_term: ((string * typ) -> term) -> (typ -> term) -> term -> term  | 
|
160  | 
val eval_ref: (unit -> term) option ref  | 
|
161  | 
val eval_term: theory -> term -> term  | 
|
162  | 
end;  | 
|
163  | 
||
164  | 
structure Eval : EVAL =  | 
|
165  | 
struct  | 
|
166  | 
||
167  | 
open Eval;  | 
|
168  | 
||
169  | 
val eval_ref = ref (NONE : (unit -> term) option);  | 
|
170  | 
||
171  | 
fun eval_term thy t =  | 
|
172  | 
t  | 
|
173  | 
|> Eval.mk_term_of (fastype_of t)  | 
|
174  | 
  |> (fn t => Code_ML.eval_term ("Eval.eval_ref", eval_ref) thy t [])
 | 
|
175  | 
|> Code.postprocess_term thy;  | 
|
176  | 
||
177  | 
end;  | 
|
178  | 
*}  | 
|
179  | 
||
180  | 
setup {*
 | 
|
181  | 
  Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
 | 
|
182  | 
*}  | 
|
183  | 
||
| 28243 | 184  | 
|
185  | 
subsubsection {* Syntax *}
 | 
|
186  | 
||
187  | 
print_translation {*
 | 
|
188  | 
let  | 
|
189  | 
  val term = Const ("<TERM>", dummyT);
 | 
|
190  | 
fun tr1' [_, _] = term;  | 
|
191  | 
fun tr2' [] = term;  | 
|
192  | 
in  | 
|
193  | 
  [(@{const_syntax Const}, tr1'),
 | 
|
194  | 
    (@{const_syntax App}, tr1'),
 | 
|
195  | 
    (@{const_syntax dummy_term}, tr2')]
 | 
|
| 28228 | 196  | 
end  | 
| 28243 | 197  | 
*}  | 
198  | 
||
199  | 
hide const dummy_term  | 
|
200  | 
hide (open) const Const App  | 
|
201  | 
hide (open) const term_of  | 
|
202  | 
||
203  | 
end  |