| author | wenzelm | 
| Fri, 27 Oct 2017 17:04:41 +0200 | |
| changeset 66928 | 33f9133bed7c | 
| parent 63435 | 7743df69a6b4 | 
| child 69605 | a96320074298 | 
| 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  | 
|
| 62020 | 5  | 
section \<open>Loading the code generator and related modules\<close>  | 
| 
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  | 
| 63435 | 13  | 
"checking" and  | 
14  | 
"datatypes" "functions" "module_name" "file"  | 
|
15  | 
"constant" "type_constructor" "type_class" "class_relation" "class_instance" "code_module"  | 
|
16  | 
:: quasi_command  | 
|
| 
30929
 
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
 
haftmann 
parents:  
diff
changeset
 | 
17  | 
begin  | 
| 
 
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
 
haftmann 
parents:  
diff
changeset
 | 
18  | 
|
| 48891 | 19  | 
ML_file "~~/src/Tools/cache_io.ML"  | 
20  | 
ML_file "~~/src/Tools/Code/code_preproc.ML"  | 
|
| 52136 | 21  | 
ML_file "~~/src/Tools/Code/code_symbol.ML"  | 
| 48891 | 22  | 
ML_file "~~/src/Tools/Code/code_thingol.ML"  | 
23  | 
ML_file "~~/src/Tools/Code/code_simp.ML"  | 
|
24  | 
ML_file "~~/src/Tools/Code/code_printer.ML"  | 
|
25  | 
ML_file "~~/src/Tools/Code/code_target.ML"  | 
|
26  | 
ML_file "~~/src/Tools/Code/code_namespace.ML"  | 
|
27  | 
ML_file "~~/src/Tools/Code/code_ml.ML"  | 
|
28  | 
ML_file "~~/src/Tools/Code/code_haskell.ML"  | 
|
29  | 
ML_file "~~/src/Tools/Code/code_scala.ML"  | 
|
30  | 
||
| 61076 | 31  | 
code_datatype "TYPE('a::{})"
 | 
| 
39421
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
32  | 
|
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
33  | 
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
 | 
34  | 
"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
 | 
35  | 
|
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
36  | 
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
 | 
37  | 
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
 | 
38  | 
|
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
39  | 
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
 | 
40  | 
|
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
41  | 
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
 | 
42  | 
"(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
 | 
43  | 
"(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
 | 
44  | 
proof -  | 
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
45  | 
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
 | 
46  | 
proof  | 
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
47  | 
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
 | 
48  | 
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
 | 
49  | 
next  | 
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
50  | 
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
 | 
51  | 
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
 | 
52  | 
qed  | 
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
53  | 
next  | 
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
54  | 
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
 | 
55  | 
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
 | 
56  | 
qed  | 
| 
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
57  | 
|
| 48891 | 58  | 
ML_file "~~/src/Tools/Code/code_runtime.ML"  | 
59  | 
ML_file "~~/src/Tools/nbe.ML"  | 
|
| 39474 | 60  | 
|
| 
39421
 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 
haftmann 
parents: 
39401 
diff
changeset
 | 
61  | 
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
 | 
62  | 
|
| 
33561
 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 
blanchet 
parents: 
31775 
diff
changeset
 | 
63  | 
end  |