author | krauss |
Thu, 25 Aug 2011 14:06:34 +0200 | |
changeset 44490 | e3e8d20a6ebc |
parent 44121 | 44adaa6db327 |
child 45190 | 58e33a125f32 |
permissions | -rw-r--r-- |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
1 |
(* Title: Tools/Code_Generator.thy |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
2 |
Author: Florian Haftmann, TU Muenchen |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
3 |
*) |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
4 |
|
39911 | 5 |
header {* Loading the code generator and related modules *} |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
6 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
7 |
theory Code_Generator |
33820
082d9bc6992d
load ML directly into theory Code_Generator (quickcheck also requires this);
wenzelm
parents:
33561
diff
changeset
|
8 |
imports Pure |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
9 |
uses |
44121 | 10 |
"~~/src/Tools/misc_legacy.ML" |
44120 | 11 |
"~~/src/Tools/codegen.ML" |
37818
dd65033fed78
load cache_io before code generator; moved adhoc-overloading to generic tools
haftmann
parents:
37442
diff
changeset
|
12 |
"~~/src/Tools/cache_io.ML" |
43018 | 13 |
"~~/src/Tools/try.ML" |
40116
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents:
39911
diff
changeset
|
14 |
"~~/src/Tools/solve_direct.ML" |
33889
4328de748fb2
some rearangement of load order to keep preferences adjacent -- slightly fragile;
wenzelm
parents:
33820
diff
changeset
|
15 |
"~~/src/Tools/quickcheck.ML" |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
16 |
"~~/src/Tools/value.ML" |
31775 | 17 |
"~~/src/Tools/Code/code_preproc.ML" |
18 |
"~~/src/Tools/Code/code_thingol.ML" |
|
37442 | 19 |
"~~/src/Tools/Code/code_simp.ML" |
31775 | 20 |
"~~/src/Tools/Code/code_printer.ML" |
21 |
"~~/src/Tools/Code/code_target.ML" |
|
38970
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
37818
diff
changeset
|
22 |
"~~/src/Tools/Code/code_namespace.ML" |
31775 | 23 |
"~~/src/Tools/Code/code_ml.ML" |
24 |
"~~/src/Tools/Code/code_haskell.ML" |
|
34294 | 25 |
"~~/src/Tools/Code/code_scala.ML" |
39474 | 26 |
("~~/src/Tools/Code/code_runtime.ML") |
39911 | 27 |
("~~/src/Tools/nbe.ML") |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
28 |
begin |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
29 |
|
39423 | 30 |
setup {* |
40116
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
blanchet
parents:
39911
diff
changeset
|
31 |
Solve_Direct.setup |
39423 | 32 |
#> Code_Preproc.setup |
33 |
#> Code_Simp.setup |
|
43564
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
43018
diff
changeset
|
34 |
#> Code_Target.setup |
39423 | 35 |
#> Code_ML.setup |
36 |
#> Code_Haskell.setup |
|
37 |
#> Code_Scala.setup |
|
38 |
#> Quickcheck.setup |
|
43612 | 39 |
#> Value.setup |
39423 | 40 |
*} |
41 |
||
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
42 |
code_datatype "TYPE('a\<Colon>{})" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
43 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
44 |
definition holds :: "prop" where |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
45 |
"holds \<equiv> ((\<lambda>x::prop. x) \<equiv> (\<lambda>x. x))" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
46 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
47 |
lemma holds: "PROP holds" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
48 |
by (unfold holds_def) (rule reflexive) |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
49 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
50 |
code_datatype holds |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
51 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
52 |
lemma implies_code [code]: |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
53 |
"(PROP holds \<Longrightarrow> PROP P) \<equiv> PROP P" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
54 |
"(PROP P \<Longrightarrow> PROP holds) \<equiv> PROP holds" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
55 |
proof - |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
56 |
show "(PROP holds \<Longrightarrow> PROP P) \<equiv> PROP P" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
57 |
proof |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
58 |
assume "PROP holds \<Longrightarrow> PROP P" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
59 |
then show "PROP P" using holds . |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
60 |
next |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
61 |
assume "PROP P" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
62 |
then show "PROP P" . |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
63 |
qed |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
64 |
next |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
65 |
show "(PROP P \<Longrightarrow> PROP holds) \<equiv> PROP holds" |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
66 |
by rule (rule holds)+ |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
67 |
qed |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
68 |
|
39474 | 69 |
use "~~/src/Tools/Code/code_runtime.ML" |
39911 | 70 |
use "~~/src/Tools/nbe.ML" |
71 |
||
72 |
setup Nbe.setup |
|
39474 | 73 |
|
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
74 |
hide_const (open) holds |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
75 |
|
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
31775
diff
changeset
|
76 |
end |