author | haftmann |
Fri, 01 Jun 2007 10:44:26 +0200 | |
changeset 23181 | f52b555f8141 |
parent 22922 | 66baa75eae06 |
permissions | -rw-r--r-- |
21046 | 1 |
(* ID: $Id$ |
2 |
Author: Florian Haftmann, TU Muenchen |
|
3 |
*) |
|
4 |
||
5 |
header {* Setup of code generator tools *} |
|
6 |
||
7 |
theory Code_Generator |
|
8 |
imports HOL |
|
22886 | 9 |
uses "~~/src/HOL/Tools/recfun_codegen.ML" |
21046 | 10 |
begin |
11 |
||
22886 | 12 |
ML {* |
13 |
structure HOL = |
|
14 |
struct |
|
15 |
val thy = theory "HOL"; |
|
16 |
end; |
|
17 |
*} -- "belongs to theory HOL" |
|
18 |
||
22385 | 19 |
subsection {* SML code generator setup *} |
21046 | 20 |
|
21 |
types_code |
|
22 |
"bool" ("bool") |
|
23 |
attach (term_of) {* |
|
24 |
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; |
|
25 |
*} |
|
26 |
attach (test) {* |
|
27 |
fun gen_bool i = one_of [false, true]; |
|
28 |
*} |
|
29 |
"prop" ("bool") |
|
30 |
attach (term_of) {* |
|
31 |
fun term_of_prop b = |
|
32 |
HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); |
|
33 |
*} |
|
34 |
||
35 |
consts_code |
|
36 |
"Trueprop" ("(_)") |
|
37 |
"True" ("true") |
|
38 |
"False" ("false") |
|
22922
66baa75eae06
Name of ML function "not" is now qualified in order to avoid
berghofe
parents:
22921
diff
changeset
|
39 |
"Not" ("Bool.not") |
21046 | 40 |
"op |" ("(_ orelse/ _)") |
41 |
"op &" ("(_ andalso/ _)") |
|
22921
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents:
22900
diff
changeset
|
42 |
"If" ("(if _/ then _/ else _)") |
21046 | 43 |
|
44 |
setup {* |
|
45 |
let |
|
46 |
||
47 |
fun eq_codegen thy defs gr dep thyname b t = |
|
48 |
(case strip_comb t of |
|
49 |
(Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE |
|
50 |
| (Const ("op =", _), [t, u]) => |
|
51 |
let |
|
52 |
val (gr', pt) = Codegen.invoke_codegen thy defs dep thyname false (gr, t); |
|
53 |
val (gr'', pu) = Codegen.invoke_codegen thy defs dep thyname false (gr', u); |
|
54 |
val (gr''', _) = Codegen.invoke_tycodegen thy defs dep thyname false (gr'', HOLogic.boolT) |
|
55 |
in |
|
56 |
SOME (gr''', Codegen.parens |
|
57 |
(Pretty.block [pt, Pretty.str " =", Pretty.brk 1, pu])) |
|
58 |
end |
|
59 |
| (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen |
|
60 |
thy defs dep thyname b (gr, Codegen.eta_expand t ts 2)) |
|
61 |
| _ => NONE); |
|
62 |
||
63 |
in |
|
64 |
||
65 |
Codegen.add_codegen "eq_codegen" eq_codegen |
|
22886 | 66 |
#> RecfunCodegen.setup |
21046 | 67 |
|
68 |
end |
|
69 |
*} |
|
70 |
||
71 |
text {* Evaluation *} |
|
72 |
||
22099 | 73 |
method_setup evaluation = {* |
21046 | 74 |
let |
75 |
||
22900 | 76 |
fun evaluation_tac i = Tactical.PRIMITIVE (Conv.fconv_rule |
77 |
(Conv.goals_conv (equal i) Codegen.evaluation_conv)); |
|
21046 | 78 |
|
22099 | 79 |
in Method.no_args (Method.SIMPLE_METHOD' (evaluation_tac THEN' rtac TrueI)) end |
80 |
*} "solve goal by evaluation" |
|
21046 | 81 |
|
82 |
||
83 |
subsection {* Generic code generator setup *} |
|
84 |
||
21904 | 85 |
text {* operational equality for code generation *} |
86 |
||
22473 | 87 |
class eq (attach "op =") = type |
21904 | 88 |
|
89 |
||
22845 | 90 |
text {* using built-in Haskell equality *} |
21904 | 91 |
|
92 |
code_class eq |
|
93 |
(Haskell "Eq" where "op =" \<equiv> "(==)") |
|
94 |
||
95 |
code_const "op =" |
|
96 |
(Haskell infixl 4 "==") |
|
97 |
||
98 |
||
22423 | 99 |
text {* type bool *} |
100 |
||
101 |
code_datatype True False |
|
21904 | 102 |
|
103 |
lemma [code func]: |
|
104 |
shows "(False \<and> x) = False" |
|
105 |
and "(True \<and> x) = x" |
|
106 |
and "(x \<and> False) = False" |
|
107 |
and "(x \<and> True) = x" by simp_all |
|
108 |
||
109 |
lemma [code func]: |
|
110 |
shows "(False \<or> x) = x" |
|
111 |
and "(True \<or> x) = True" |
|
112 |
and "(x \<or> False) = x" |
|
113 |
and "(x \<or> True) = True" by simp_all |
|
114 |
||
115 |
lemma [code func]: |
|
116 |
shows "(\<not> True) = False" |
|
117 |
and "(\<not> False) = True" by (rule HOL.simp_thms)+ |
|
118 |
||
22886 | 119 |
lemmas [code] = imp_conv_disj |
21904 | 120 |
|
121 |
lemmas [code func] = if_True if_False |
|
122 |
||
123 |
instance bool :: eq .. |
|
124 |
||
125 |
lemma [code func]: |
|
22845 | 126 |
shows "True = P \<longleftrightarrow> P" |
127 |
and "False = P \<longleftrightarrow> \<not> P" |
|
128 |
and "P = True \<longleftrightarrow> P" |
|
129 |
and "P = False \<longleftrightarrow> \<not> P" by simp_all |
|
21904 | 130 |
|
131 |
code_type bool |
|
132 |
(SML "bool") |
|
133 |
(OCaml "bool") |
|
134 |
(Haskell "Bool") |
|
135 |
||
136 |
code_instance bool :: eq |
|
137 |
(Haskell -) |
|
138 |
||
139 |
code_const "op = \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool" |
|
140 |
(Haskell infixl 4 "==") |
|
141 |
||
142 |
code_const True and False and Not and "op &" and "op |" and If |
|
143 |
(SML "true" and "false" and "not" |
|
144 |
and infixl 1 "andalso" and infixl 0 "orelse" |
|
145 |
and "!(if (_)/ then (_)/ else (_))") |
|
146 |
(OCaml "true" and "false" and "not" |
|
147 |
and infixl 4 "&&" and infixl 2 "||" |
|
148 |
and "!(if (_)/ then (_)/ else (_))") |
|
149 |
(Haskell "True" and "False" and "not" |
|
150 |
and infixl 3 "&&" and infixl 2 "||" |
|
151 |
and "!(if (_)/ then (_)/ else (_))") |
|
152 |
||
153 |
code_reserved SML |
|
154 |
bool true false not |
|
155 |
||
156 |
code_reserved OCaml |
|
157 |
bool true false not |
|
158 |
||
159 |
||
22423 | 160 |
text {* type prop *} |
161 |
||
162 |
code_datatype Trueprop "prop" |
|
163 |
||
164 |
||
165 |
text {* type itself *} |
|
21046 | 166 |
|
22423 | 167 |
code_datatype "TYPE('a)" |
168 |
||
169 |
||
22480 | 170 |
text {* code generation for undefined as exception *} |
21046 | 171 |
|
22480 | 172 |
code_const undefined |
22499 | 173 |
(SML "raise/ Fail/ \"undefined\"") |
174 |
(OCaml "failwith/ \"undefined\"") |
|
22480 | 175 |
(Haskell "error/ \"undefined\"") |
21046 | 176 |
|
21079 | 177 |
code_reserved SML Fail |
21904 | 178 |
code_reserved OCaml failwith |
21046 | 179 |
|
180 |
||
21378 | 181 |
subsection {* Evaluation oracle *} |
182 |
||
22487 | 183 |
oracle eval_oracle ("term") = {* fn thy => fn t => |
184 |
if CodegenPackage.satisfies thy (HOLogic.dest_Trueprop t) [] |
|
185 |
then t |
|
22886 | 186 |
else HOLogic.Trueprop $ HOLogic.true_const (*dummy*) |
21869 | 187 |
*} |
21378 | 188 |
|
22099 | 189 |
method_setup eval = {* |
190 |
let |
|
22487 | 191 |
fun eval_tac thy = |
192 |
SUBGOAL (fn (t, i) => rtac (eval_oracle thy t) i) |
|
193 |
in |
|
194 |
Method.ctxt_args (fn ctxt => |
|
195 |
Method.SIMPLE_METHOD' (eval_tac (ProofContext.theory_of ctxt))) |
|
196 |
end |
|
22099 | 197 |
*} "solve goal by evaluation" |
21378 | 198 |
|
199 |
||
200 |
subsection {* Normalization by evaluation *} |
|
21046 | 201 |
|
22099 | 202 |
method_setup normalization = {* |
21046 | 203 |
let |
22900 | 204 |
fun normalization_tac i = Tactical.PRIMITIVE (Conv.fconv_rule |
205 |
(Conv.goals_conv (equal i) (HOLogic.Trueprop_conv |
|
21149 | 206 |
NBE.normalization_conv))); |
21046 | 207 |
in |
22099 | 208 |
Method.no_args (Method.SIMPLE_METHOD' (normalization_tac THEN' resolve_tac [TrueI, refl])) |
209 |
end |
|
210 |
*} "solve goal by normalization" |
|
211 |
||
21046 | 212 |
|
21059 | 213 |
text {* lazy @{const If} *} |
214 |
||
215 |
definition |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21378
diff
changeset
|
216 |
if_delayed :: "bool \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> 'a" where |
22845 | 217 |
[code func del]: "if_delayed b f g = (if b then f True else g False)" |
21059 | 218 |
|
219 |
lemma [code func]: |
|
220 |
shows "if_delayed True f g = f True" |
|
221 |
and "if_delayed False f g = g False" |
|
222 |
unfolding if_delayed_def by simp_all |
|
223 |
||
224 |
lemma [normal pre, symmetric, normal post]: |
|
225 |
"(if b then x else y) = if_delayed b (\<lambda>_. x) (\<lambda>_. y)" |
|
226 |
unfolding if_delayed_def .. |
|
227 |
||
21454 | 228 |
hide (open) const if_delayed |
21046 | 229 |
|
22099 | 230 |
end |