| author | wenzelm | 
| Sun, 26 Feb 2023 14:27:21 +0100 | |
| changeset 77377 | 82fdc7cf9d44 | 
| parent 70009 | 435fb018e8ee | 
| child 82774 | 2865a6618cba | 
| 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: 
33561diff
changeset | 8 | imports Pure | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
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: 
52138diff
changeset | 12 | "code_monad" "code_reflect" :: thy_decl and | 
| 63435 | 13 | "checking" and | 
| 70009 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 wenzelm parents: 
69605diff
changeset | 14 | "datatypes" "functions" "module_name" "file" "file_prefix" | 
| 63435 | 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 | |
| 69605 | 19 | ML_file \<open>~~/src/Tools/cache_io.ML\<close> | 
| 20 | ML_file \<open>~~/src/Tools/Code/code_preproc.ML\<close> | |
| 21 | ML_file \<open>~~/src/Tools/Code/code_symbol.ML\<close> | |
| 22 | ML_file \<open>~~/src/Tools/Code/code_thingol.ML\<close> | |
| 23 | ML_file \<open>~~/src/Tools/Code/code_simp.ML\<close> | |
| 24 | ML_file \<open>~~/src/Tools/Code/code_printer.ML\<close> | |
| 25 | ML_file \<open>~~/src/Tools/Code/code_target.ML\<close> | |
| 26 | ML_file \<open>~~/src/Tools/Code/code_namespace.ML\<close> | |
| 27 | ML_file \<open>~~/src/Tools/Code/code_ml.ML\<close> | |
| 28 | ML_file \<open>~~/src/Tools/Code/code_haskell.ML\<close> | |
| 29 | ML_file \<open>~~/src/Tools/Code/code_scala.ML\<close> | |
| 48891 | 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: 
39401diff
changeset | 32 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
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: 
39401diff
changeset | 35 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
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: 
39401diff
changeset | 38 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
changeset | 40 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
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: 
39401diff
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: 
39401diff
changeset | 44 | proof - | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
changeset | 46 | proof | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
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: 
39401diff
changeset | 49 | next | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
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: 
39401diff
changeset | 52 | qed | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 53 | next | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
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: 
39401diff
changeset | 56 | qed | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 57 | |
| 69605 | 58 | ML_file \<open>~~/src/Tools/Code/code_runtime.ML\<close> | 
| 59 | ML_file \<open>~~/src/Tools/nbe.ML\<close> | |
| 39474 | 60 | |
| 39421 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
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: 
39401diff
changeset | 62 | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
31775diff
changeset | 63 | end |