author | wenzelm |
Sat, 10 Jan 2009 21:32:30 +0100 | |
changeset 29435 | a5f84ac14609 |
parent 29105 | 8f38bf68d42e |
child 30510 | 4120fc59dd85 |
permissions | -rw-r--r-- |
24285 | 1 |
(* Title: HOL/Code_Setup.thy |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann |
|
4 |
*) |
|
5 |
||
28400 | 6 |
header {* Setup of code generators and related tools *} |
24285 | 7 |
|
8 |
theory Code_Setup |
|
9 |
imports HOL |
|
28400 | 10 |
begin |
11 |
||
12 |
subsection {* Generic code generator foundation *} |
|
13 |
||
14 |
text {* Datatypes *} |
|
15 |
||
16 |
code_datatype True False |
|
17 |
||
18 |
code_datatype "TYPE('a\<Colon>{})" |
|
19 |
||
20 |
code_datatype Trueprop "prop" |
|
21 |
||
22 |
text {* Code equations *} |
|
23 |
||
28562 | 24 |
lemma [code]: |
28740
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
25 |
shows "(True \<Longrightarrow> PROP P) \<equiv> PROP P" |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
26 |
and "(False \<Longrightarrow> Q) \<equiv> Trueprop True" |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
27 |
and "(PROP P \<Longrightarrow> True) \<equiv> Trueprop True" |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
28 |
and "(Q \<Longrightarrow> False) \<equiv> Trueprop (\<not> Q)" by (auto intro!: equal_intr_rule) |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
29 |
|
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
30 |
lemma [code]: |
28400 | 31 |
shows "False \<and> x \<longleftrightarrow> False" |
32 |
and "True \<and> x \<longleftrightarrow> x" |
|
33 |
and "x \<and> False \<longleftrightarrow> False" |
|
34 |
and "x \<and> True \<longleftrightarrow> x" by simp_all |
|
35 |
||
28562 | 36 |
lemma [code]: |
28400 | 37 |
shows "False \<or> x \<longleftrightarrow> x" |
38 |
and "True \<or> x \<longleftrightarrow> True" |
|
39 |
and "x \<or> False \<longleftrightarrow> x" |
|
40 |
and "x \<or> True \<longleftrightarrow> True" by simp_all |
|
41 |
||
28562 | 42 |
lemma [code]: |
28400 | 43 |
shows "\<not> True \<longleftrightarrow> False" |
44 |
and "\<not> False \<longleftrightarrow> True" by (rule HOL.simp_thms)+ |
|
45 |
||
28562 | 46 |
lemmas [code] = Let_def if_True if_False |
28400 | 47 |
|
28562 | 48 |
lemmas [code, code unfold, symmetric, code post] = imp_conv_disj |
28400 | 49 |
|
50 |
text {* Equality *} |
|
51 |
||
52 |
context eq |
|
24285 | 53 |
begin |
54 |
||
28562 | 55 |
lemma equals_eq [code inline, code]: "op = \<equiv> eq" |
28400 | 56 |
by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals) |
57 |
||
28512 | 58 |
declare eq [code unfold, code inline del] |
59 |
||
28400 | 60 |
declare equals_eq [symmetric, code post] |
61 |
||
62 |
end |
|
63 |
||
64 |
declare simp_thms(6) [code nbe] |
|
65 |
||
66 |
hide (open) const eq |
|
67 |
hide const eq |
|
68 |
||
69 |
setup {* |
|
70 |
Code_Unit.add_const_alias @{thm equals_eq} |
|
71 |
*} |
|
72 |
||
73 |
text {* Cases *} |
|
74 |
||
75 |
lemma Let_case_cert: |
|
76 |
assumes "CASE \<equiv> (\<lambda>x. Let x f)" |
|
77 |
shows "CASE x \<equiv> f x" |
|
78 |
using assms by simp_all |
|
79 |
||
80 |
lemma If_case_cert: |
|
81 |
assumes "CASE \<equiv> (\<lambda>b. If b f g)" |
|
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28740
diff
changeset
|
82 |
shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)" |
28400 | 83 |
using assms by simp_all |
84 |
||
85 |
setup {* |
|
86 |
Code.add_case @{thm Let_case_cert} |
|
87 |
#> Code.add_case @{thm If_case_cert} |
|
88 |
#> Code.add_undefined @{const_name undefined} |
|
89 |
*} |
|
90 |
||
91 |
code_abort undefined |
|
92 |
||
93 |
||
94 |
subsection {* Generic code generator preprocessor *} |
|
95 |
||
96 |
setup {* |
|
97 |
Code.map_pre (K HOL_basic_ss) |
|
98 |
#> Code.map_post (K HOL_basic_ss) |
|
99 |
*} |
|
100 |
||
24285 | 101 |
|
28400 | 102 |
subsection {* Generic code generator target languages *} |
103 |
||
104 |
text {* type bool *} |
|
105 |
||
106 |
code_type bool |
|
107 |
(SML "bool") |
|
108 |
(OCaml "bool") |
|
109 |
(Haskell "Bool") |
|
110 |
||
111 |
code_const True and False and Not and "op &" and "op |" and If |
|
112 |
(SML "true" and "false" and "not" |
|
113 |
and infixl 1 "andalso" and infixl 0 "orelse" |
|
114 |
and "!(if (_)/ then (_)/ else (_))") |
|
115 |
(OCaml "true" and "false" and "not" |
|
116 |
and infixl 4 "&&" and infixl 2 "||" |
|
117 |
and "!(if (_)/ then (_)/ else (_))") |
|
118 |
(Haskell "True" and "False" and "not" |
|
119 |
and infixl 3 "&&" and infixl 2 "||" |
|
120 |
and "!(if (_)/ then (_)/ else (_))") |
|
121 |
||
122 |
code_reserved SML |
|
123 |
bool true false not |
|
124 |
||
125 |
code_reserved OCaml |
|
126 |
bool not |
|
127 |
||
128 |
text {* using built-in Haskell equality *} |
|
129 |
||
130 |
code_class eq |
|
28687 | 131 |
(Haskell "Eq") |
28400 | 132 |
|
133 |
code_const "eq_class.eq" |
|
134 |
(Haskell infixl 4 "==") |
|
135 |
||
136 |
code_const "op =" |
|
137 |
(Haskell infixl 4 "==") |
|
138 |
||
139 |
text {* undefined *} |
|
140 |
||
141 |
code_const undefined |
|
28512 | 142 |
(SML "!(raise/ Fail/ \"undefined\")") |
28400 | 143 |
(OCaml "failwith/ \"undefined\"") |
144 |
(Haskell "error/ \"undefined\"") |
|
145 |
||
146 |
||
147 |
subsection {* SML code generator setup *} |
|
24285 | 148 |
|
149 |
types_code |
|
150 |
"bool" ("bool") |
|
151 |
attach (term_of) {* |
|
152 |
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; |
|
153 |
*} |
|
154 |
attach (test) {* |
|
25885 | 155 |
fun gen_bool i = |
156 |
let val b = one_of [false, true] |
|
157 |
in (b, fn () => term_of_bool b) end; |
|
24285 | 158 |
*} |
159 |
"prop" ("bool") |
|
160 |
attach (term_of) {* |
|
161 |
fun term_of_prop b = |
|
162 |
HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); |
|
163 |
*} |
|
164 |
||
165 |
consts_code |
|
166 |
"Trueprop" ("(_)") |
|
167 |
"True" ("true") |
|
168 |
"False" ("false") |
|
169 |
"Not" ("Bool.not") |
|
170 |
"op |" ("(_ orelse/ _)") |
|
171 |
"op &" ("(_ andalso/ _)") |
|
172 |
"If" ("(if _/ then _/ else _)") |
|
173 |
||
174 |
setup {* |
|
175 |
let |
|
176 |
||
28537
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
177 |
fun eq_codegen thy defs dep thyname b t gr = |
24285 | 178 |
(case strip_comb t of |
179 |
(Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE |
|
180 |
| (Const ("op =", _), [t, u]) => |
|
181 |
let |
|
28537
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
182 |
val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr; |
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
183 |
val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr'; |
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
184 |
val (_, gr''') = Codegen.invoke_tycodegen thy defs dep thyname false HOLogic.boolT gr''; |
24285 | 185 |
in |
28537
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
186 |
SOME (Codegen.parens |
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
187 |
(Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''') |
24285 | 188 |
end |
189 |
| (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen |
|
28537
1e84256d1a8a
established canonical argument order in SML code generators
haftmann
parents:
28512
diff
changeset
|
190 |
thy defs dep thyname b (Codegen.eta_expand t ts 2) gr) |
24285 | 191 |
| _ => NONE); |
192 |
||
193 |
in |
|
194 |
Codegen.add_codegen "eq_codegen" eq_codegen |
|
195 |
end |
|
196 |
*} |
|
197 |
||
24463
841c2e24761f
Smaller size and fewer iterations for quickcheck.
berghofe
parents:
24293
diff
changeset
|
198 |
|
28400 | 199 |
subsection {* Evaluation and normalization by evaluation *} |
24285 | 200 |
|
29105 | 201 |
setup {* |
202 |
Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of) |
|
203 |
*} |
|
204 |
||
27103 | 205 |
ML {* |
206 |
structure Eval_Method = |
|
207 |
struct |
|
208 |
||
209 |
val eval_ref : (unit -> bool) option ref = ref NONE; |
|
210 |
||
211 |
end; |
|
212 |
*} |
|
213 |
||
28290 | 214 |
oracle eval_oracle = {* fn ct => |
215 |
let |
|
216 |
val thy = Thm.theory_of_cterm ct; |
|
217 |
val t = Thm.term_of ct; |
|
28740
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
218 |
val dummy = @{cprop True}; |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
219 |
in case try HOLogic.dest_Trueprop t |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
220 |
of SOME t' => if Code_ML.eval_term |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
221 |
("Eval_Method.eval_ref", Eval_Method.eval_ref) thy t' [] |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
222 |
then Thm.capply (Thm.capply @{cterm "op \<equiv> \<Colon> prop \<Rightarrow> prop \<Rightarrow> prop"} ct) dummy |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
223 |
else dummy |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
224 |
| NONE => dummy |
28290 | 225 |
end |
24285 | 226 |
*} |
227 |
||
28740
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
228 |
ML {* |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
229 |
fun gen_eval_method conv = Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD' |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
230 |
(CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt) |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
231 |
THEN' rtac TrueI)) |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
232 |
*} |
24285 | 233 |
|
28740
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
234 |
method_setup eval = {* gen_eval_method eval_oracle *} |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
235 |
"solve goal by evaluation" |
24285 | 236 |
|
28740
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
237 |
method_setup evaluation = {* gen_eval_method Codegen.evaluation_conv *} |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
238 |
"solve goal by evaluation" |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
239 |
|
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
240 |
method_setup normalization = {* (Method.no_args o Method.SIMPLE_METHOD') |
22a8125d66fa
improved handling of !!/==> for eval and normalization
haftmann
parents:
28699
diff
changeset
|
241 |
(CONVERSION Nbe.norm_conv THEN' (fn k => TRY (rtac TrueI k))) |
24285 | 242 |
*} "solve goal by normalization" |
243 |
||
28400 | 244 |
|
245 |
subsection {* Quickcheck *} |
|
246 |
||
29105 | 247 |
setup {* |
248 |
Quickcheck.add_generator ("SML", Codegen.test_term) |
|
249 |
*} |
|
250 |
||
28400 | 251 |
quickcheck_params [size = 5, iterations = 50] |
252 |
||
24285 | 253 |
end |