author | wenzelm |
Wed, 25 Jul 2007 17:05:49 +0200 | |
changeset 23981 | 03b71bf91318 |
parent 23934 | 79393cb9c0a6 |
child 24219 | e558fe311376 |
permissions | -rw-r--r-- |
20856 | 1 |
(* Title: Pure/nbe.ML |
2 |
ID: $Id$ |
|
19147 | 3 |
Author: Tobias Nipkow, Florian Haftmann, TU Muenchen |
4 |
||
19177 | 5 |
Toplevel theory interface for "normalization by evaluation" |
19147 | 6 |
*) |
7 |
||
19177 | 8 |
signature NBE = |
19147 | 9 |
sig |
20856 | 10 |
(*preconditions: no Vars/TVars in term*) |
20602 | 11 |
val normalization_conv: cterm -> thm |
19962 | 12 |
val lookup: string -> NBE_Eval.Univ |
19795 | 13 |
val update: string * NBE_Eval.Univ -> unit |
20846 | 14 |
val trace: bool ref |
19147 | 15 |
end; |
16 |
||
19177 | 17 |
structure NBE: NBE = |
19147 | 18 |
struct |
19 |
||
20846 | 20 |
val trace = ref false; |
21124 | 21 |
fun tracing f x = if !trace then (Output.tracing (f x); x) else x; |
19968 | 22 |
|
21068 | 23 |
|
24 |
(** data setup **) |
|
25 |
||
26 |
(* preproc and postproc attributes *) |
|
19968 | 27 |
|
20920 | 28 |
structure NBE_Rewrite = TheoryDataFun |
22846 | 29 |
( |
20920 | 30 |
type T = thm list * thm list |
22846 | 31 |
val empty = ([], []); |
20920 | 32 |
val copy = I; |
33 |
val extend = I; |
|
34 |
fun merge _ ((pres1,posts1), (pres2,posts2)) = |
|
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22213
diff
changeset
|
35 |
(Library.merge Thm.eq_thm (pres1,pres2), Library.merge Thm.eq_thm (posts1,posts2)) |
22846 | 36 |
); |
20920 | 37 |
|
21068 | 38 |
val _ = |
39 |
let |
|
40 |
fun map_data f = Context.mapping (NBE_Rewrite.map f) I; |
|
41 |
fun attr_pre (thy, thm) = |
|
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22213
diff
changeset
|
42 |
((map_data o apfst) (insert Thm.eq_thm thm) thy, thm) |
21068 | 43 |
fun attr_post (thy, thm) = |
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
22213
diff
changeset
|
44 |
((map_data o apsnd) (insert Thm.eq_thm thm) thy, thm) |
21068 | 45 |
val attr = Attrib.syntax (Scan.lift (Args.$$$ "pre" >> K attr_pre |
46 |
|| Args.$$$ "post" >> K attr_post)); |
|
47 |
in |
|
48 |
Context.add_setup ( |
|
22846 | 49 |
Attrib.add_attributes |
21068 | 50 |
[("normal", attr, "declare rewrite theorems for normalization")] |
51 |
) |
|
52 |
end; |
|
20920 | 53 |
|
22213 | 54 |
fun the_pres thy = |
21068 | 55 |
let |
21506 | 56 |
val ctxt = ProofContext.init thy; |
57 |
val pres = (map (LocalDefs.meta_rewrite_rule ctxt) o fst) (NBE_Rewrite.get thy) |
|
22213 | 58 |
in pres end |
20920 | 59 |
|
60 |
fun apply_posts thy = |
|
21068 | 61 |
let |
21506 | 62 |
val ctxt = ProofContext.init thy; |
63 |
val posts = (map (LocalDefs.meta_rewrite_rule ctxt) o snd) (NBE_Rewrite.get thy) |
|
21708 | 64 |
in MetaSimplifier.rewrite false posts end |
20920 | 65 |
|
22213 | 66 |
(* theorem store *) |
67 |
||
23136 | 68 |
structure Funcgr = CodegenFuncgrRetrieval (val rewrites = the_pres); |
20920 | 69 |
|
21068 | 70 |
(* code store *) |
71 |
||
20846 | 72 |
structure NBE_Data = CodeDataFun |
19147 | 73 |
(struct |
22213 | 74 |
type T = NBE_Eval.Univ Symtab.table; |
75 |
val empty = Symtab.empty; |
|
76 |
fun merge _ = Symtab.merge (K true); |
|
77 |
fun purge _ _ _ = Symtab.empty; |
|
19147 | 78 |
end); |
79 |
||
19795 | 80 |
|
21124 | 81 |
(** norm by eval **) |
21068 | 82 |
|
19968 | 83 |
(* sandbox communication *) |
19147 | 84 |
|
19177 | 85 |
val tab : NBE_Eval.Univ Symtab.table ref = ref Symtab.empty; |
86 |
fun lookup s = (the o Symtab.lookup (!tab)) s; |
|
87 |
fun update sx = (tab := Symtab.update sx (!tab)); |
|
88 |
||
20846 | 89 |
local |
90 |
||
21124 | 91 |
(* function generation *) |
20846 | 92 |
|
93 |
fun generate thy funs = |
|
94 |
let |
|
21124 | 95 |
(* FIXME better turn this into a function |
96 |
NBE_Eval.Univ Symtab.table -> NBE_Eval.Univ Symtab.table |
|
97 |
with implicit side effect *) |
|
98 |
fun use_code NONE = () |
|
99 |
| use_code (SOME s) = |
|
21991 | 100 |
(tracing (fn () => "\n---generated code:\n" ^ s) (); |
22144 | 101 |
use_text "" (Output.tracing o enclose "\n---compiler echo:\n" "\n---\n", |
21124 | 102 |
Output.tracing o enclose "\n--- compiler echo (with error!):\n" "\n---\n") |
103 |
(!trace) s); |
|
21991 | 104 |
val _ = tracing (fn () => "new definitions: " ^ (commas o maps (map fst)) funs) (); |
20846 | 105 |
val _ = tab := NBE_Data.get thy;; |
23226 | 106 |
val _ = List.app (use_code o NBE_Codegen.generate thy |
20846 | 107 |
(fn s => Symtab.defined (!tab) s)) funs; |
108 |
in NBE_Data.change thy (K (!tab)) end; |
|
109 |
||
22213 | 110 |
fun ensure_funs thy funcgr t = |
20846 | 111 |
let |
22554
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22360
diff
changeset
|
112 |
fun consts_of thy t = |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22360
diff
changeset
|
113 |
fold_aterms (fn Const c => cons (CodegenConsts.const_of_cexpr thy c) | _ => I) t [] |
d1499fff65d8
simplified constant representation in code generator
haftmann
parents:
22360
diff
changeset
|
114 |
val consts = consts_of thy t; |
20846 | 115 |
val nbe_tab = NBE_Data.get thy; |
22213 | 116 |
in |
117 |
CodegenFuncgr.deps funcgr consts |
|
118 |
|> (map o filter_out) (Symtab.defined nbe_tab o CodegenNames.const thy) |
|
119 |
|> filter_out null |
|
120 |
|> (map o map) (fn c => (CodegenNames.const thy c, CodegenFuncgr.funcs funcgr c)) |
|
121 |
|> generate thy |
|
122 |
end; |
|
21124 | 123 |
|
124 |
(* term evaluation *) |
|
20846 | 125 |
|
22213 | 126 |
fun eval_term thy funcgr t = |
21124 | 127 |
let |
128 |
fun subst_Frees [] = I |
|
129 |
| subst_Frees inst = |
|
23090 | 130 |
Term.map_aterms (fn (t as Free (s, _)) => the_default t (AList.lookup (op =) inst s) |
131 |
| t => t); |
|
21124 | 132 |
val anno_vars = |
133 |
subst_Frees (map (fn (s, T) => (s, Free (s, T))) (Term.add_frees t [])) |
|
23090 | 134 |
#> subst_Vars (map (fn (ixn, T) => (ixn, Var (ixn, T))) (Term.add_vars t [])) |
21124 | 135 |
fun check_tvars t = if null (Term.term_tvars t) then t else |
136 |
error ("Illegal schematic type variables in normalized term: " |
|
137 |
^ setmp show_types true (Sign.string_of_term thy) t); |
|
138 |
val ty = type_of t; |
|
22728 | 139 |
fun constrain t = |
23090 | 140 |
singleton (ProofContext.infer_types_pats (ProofContext.init thy)) (TypeInfer.constrain t ty); |
22213 | 141 |
val _ = ensure_funs thy funcgr t; |
21124 | 142 |
in |
143 |
t |
|
144 |
|> tracing (fn t => "Input:\n" ^ Display.raw_string_of_term t) |
|
145 |
|> NBE_Eval.eval thy (!tab) |
|
146 |
|> tracing (fn nt => "Normalized:\n" ^ NBE_Eval.string_of_nterm nt) |
|
147 |
|> NBE_Codegen.nterm_to_term thy |
|
23090 | 148 |
|> tracing (fn t => "Converted back:\n" ^ setmp show_types true Display.raw_string_of_term t) |
149 |
|> tracing (fn _ => "Term type:\n" ^ Display.raw_string_of_typ ty) |
|
21124 | 150 |
|> anno_vars |
23090 | 151 |
|> tracing (fn t => "Vars typed:\n" ^ setmp show_types true Display.raw_string_of_term t) |
21124 | 152 |
|> constrain |
23090 | 153 |
|> tracing (fn t => "Types inferred:\n" ^ setmp show_types true Display.raw_string_of_term t) |
21124 | 154 |
|> check_tvars |
155 |
end; |
|
19795 | 156 |
|
21124 | 157 |
(* evaluation oracle *) |
158 |
||
22213 | 159 |
exception Normalization of CodegenFuncgr.T * term; |
19795 | 160 |
|
22213 | 161 |
fun normalization_oracle (thy, Normalization (funcgr, t)) = |
162 |
Logic.mk_equals (t, eval_term thy funcgr t); |
|
21124 | 163 |
|
22213 | 164 |
fun normalization_invoke thy funcgr t = |
165 |
Thm.invoke_oracle_i thy "Pure.normalization" (thy, Normalization (funcgr, t)); |
|
19968 | 166 |
|
20846 | 167 |
in |
19968 | 168 |
|
21124 | 169 |
(* interface *) |
170 |
||
171 |
fun normalization_conv ct = |
|
20846 | 172 |
let |
21124 | 173 |
val thy = Thm.theory_of_cterm ct; |
22213 | 174 |
fun mk funcgr drop_classes ct thm1 = |
21129 | 175 |
let |
176 |
val t = Thm.term_of ct; |
|
22213 | 177 |
val thm2 = normalization_invoke thy funcgr t; |
22902
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents:
22846
diff
changeset
|
178 |
val thm3 = apply_posts thy (Thm.rhs_of thm2); |
21129 | 179 |
val thm23 = drop_classes (Thm.transitive thm2 thm3); |
180 |
in |
|
181 |
Thm.transitive thm1 thm23 handle THM _ => |
|
182 |
error ("normalization_conv - could not construct proof:\n" |
|
183 |
^ (cat_lines o map string_of_thm) [thm1, thm2, thm3]) |
|
184 |
end; |
|
22213 | 185 |
in fst (Funcgr.make_term thy mk ct) end; |
20846 | 186 |
|
187 |
fun norm_print_term ctxt modes t = |
|
188 |
let |
|
189 |
val thy = ProofContext.theory_of ctxt; |
|
21124 | 190 |
val ct = Thm.cterm_of thy t; |
191 |
val (_, t') = (Logic.dest_equals o Thm.prop_of o normalization_conv) ct; |
|
20846 | 192 |
val ty = Term.type_of t'; |
23934 | 193 |
val p = PrintMode.with_modes modes (fn () => |
20846 | 194 |
Pretty.block [Pretty.quote (ProofContext.pretty_term ctxt t'), Pretty.fbrk, |
195 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (ProofContext.pretty_typ ctxt ty)]) (); |
|
196 |
in Pretty.writeln p end; |
|
197 |
||
198 |
fun norm_print_term_e (modes, raw_t) state = |
|
21506 | 199 |
let val ctxt = Toplevel.context_of state |
20846 | 200 |
in norm_print_term ctxt modes (ProofContext.read_term ctxt raw_t) end; |
201 |
||
20602 | 202 |
val _ = Context.add_setup |
20677 | 203 |
(Theory.add_oracle ("normalization", normalization_oracle)); |
19177 | 204 |
|
21124 | 205 |
end; (*local*) |
206 |
||
20846 | 207 |
|
19968 | 208 |
(* Isar setup *) |
209 |
||
210 |
local structure P = OuterParse and K = OuterKeyword in |
|
19147 | 211 |
|
20846 | 212 |
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; |
213 |
||
19147 | 214 |
val nbeP = |
20846 | 215 |
OuterSyntax.improper_command "normal_form" "normalize term by evaluation" K.diag |
216 |
(opt_modes -- P.typ >> (Toplevel.keep o norm_print_term_e)); |
|
19968 | 217 |
|
218 |
end; |
|
19147 | 219 |
|
220 |
val _ = OuterSyntax.add_parsers [nbeP]; |
|
221 |
||
19177 | 222 |
end; |