author | blanchet |
Thu, 28 Aug 2014 07:34:23 +0200 | |
changeset 58067 | a7a0af643499 |
parent 57430 | 020cea57eaa4 |
child 58889 | 5b7a9633cfa8 |
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 |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46947
diff
changeset
|
9 |
keywords |
56925 | 10 |
"print_codeproc" "code_thms" "code_deps" :: diag and |
55372 | 11 |
"export_code" "code_identifier" "code_printing" "code_reserved" |
54890
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
52138
diff
changeset
|
12 |
"code_monad" "code_reflect" :: thy_decl and |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46947
diff
changeset
|
13 |
"datatypes" "functions" "module_name" "file" "checking" |
52137
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
haftmann
parents:
52136
diff
changeset
|
14 |
"constant" "type_constructor" "type_class" "class_relation" "class_instance" "code_module" |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
15 |
begin |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
diff
changeset
|
16 |
|
48891 | 17 |
ML_file "~~/src/Tools/cache_io.ML" |
18 |
ML_file "~~/src/Tools/Code/code_preproc.ML" |
|
57430
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
19 |
|
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
20 |
attribute_setup code_preproc_trace = {* |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
21 |
(Scan.lift (Args.$$$ "off" >> K Code_Preproc.trace_none) |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
22 |
|| (Scan.lift (Args.$$$ "only" |-- Args.colon |-- Scan.repeat1 Parse.term)) |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
23 |
>> Code_Preproc.trace_only_ext |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
24 |
|| Scan.succeed Code_Preproc.trace_all) |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
25 |
>> (Thm.declaration_attribute o K) |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
26 |
*} "tracing of the code generator preprocessor" |
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
56970
diff
changeset
|
27 |
|
52136 | 28 |
ML_file "~~/src/Tools/Code/code_symbol.ML" |
48891 | 29 |
ML_file "~~/src/Tools/Code/code_thingol.ML" |
30 |
ML_file "~~/src/Tools/Code/code_simp.ML" |
|
31 |
ML_file "~~/src/Tools/Code/code_printer.ML" |
|
32 |
ML_file "~~/src/Tools/Code/code_target.ML" |
|
33 |
ML_file "~~/src/Tools/Code/code_namespace.ML" |
|
34 |
ML_file "~~/src/Tools/Code/code_ml.ML" |
|
35 |
ML_file "~~/src/Tools/Code/code_haskell.ML" |
|
36 |
ML_file "~~/src/Tools/Code/code_scala.ML" |
|
37 |
||
39423 | 38 |
setup {* |
56970 | 39 |
Code_Simp.setup |
43564
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents:
43018
diff
changeset
|
40 |
#> Code_Target.setup |
39423 | 41 |
#> Code_ML.setup |
42 |
#> Code_Haskell.setup |
|
43 |
#> Code_Scala.setup |
|
44 |
*} |
|
45 |
||
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
46 |
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
|
47 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
48 |
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
|
49 |
"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
|
50 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
51 |
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
|
52 |
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
|
53 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
54 |
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
|
55 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
56 |
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
|
57 |
"(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
|
58 |
"(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
|
59 |
proof - |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
60 |
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
|
61 |
proof |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
62 |
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
|
63 |
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
|
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 |
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
|
66 |
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
|
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 |
next |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
69 |
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
|
70 |
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
|
71 |
qed |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
72 |
|
48891 | 73 |
ML_file "~~/src/Tools/Code/code_runtime.ML" |
74 |
ML_file "~~/src/Tools/nbe.ML" |
|
39474 | 75 |
|
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39401
diff
changeset
|
76 |
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
|
77 |
|
33561
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
blanchet
parents:
31775
diff
changeset
|
78 |
end |
47657
1ba213363d0c
moved modules with only vague relation to the code generator to theory HOL rather than theory Code_Generator
haftmann
parents:
47610
diff
changeset
|
79 |