| author | blanchet | 
| Tue, 08 Nov 2011 08:56:23 +0100 | |
| changeset 45401 | 36478a5f6104 | 
| parent 45190 | 58e33a125f32 | 
| child 46947 | b8c7eb0c2f89 | 
| 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: 
33561diff
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" | 
| 37818 
dd65033fed78
load cache_io before code generator; moved adhoc-overloading to generic tools
 haftmann parents: 
37442diff
changeset | 11 | "~~/src/Tools/cache_io.ML" | 
| 43018 | 12 | "~~/src/Tools/try.ML" | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39911diff
changeset | 13 | "~~/src/Tools/solve_direct.ML" | 
| 33889 
4328de748fb2
some rearangement of load order to keep preferences adjacent -- slightly fragile;
 wenzelm parents: 
33820diff
changeset | 14 | "~~/src/Tools/quickcheck.ML" | 
| 30929 
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
 haftmann parents: diff
changeset | 15 | "~~/src/Tools/value.ML" | 
| 31775 | 16 | "~~/src/Tools/Code/code_preproc.ML" | 
| 17 | "~~/src/Tools/Code/code_thingol.ML" | |
| 37442 | 18 | "~~/src/Tools/Code/code_simp.ML" | 
| 31775 | 19 | "~~/src/Tools/Code/code_printer.ML" | 
| 20 | "~~/src/Tools/Code/code_target.ML" | |
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: 
37818diff
changeset | 21 | "~~/src/Tools/Code/code_namespace.ML" | 
| 31775 | 22 | "~~/src/Tools/Code/code_ml.ML" | 
| 23 | "~~/src/Tools/Code/code_haskell.ML" | |
| 34294 | 24 | "~~/src/Tools/Code/code_scala.ML" | 
| 39474 | 25 |   ("~~/src/Tools/Code/code_runtime.ML")
 | 
| 39911 | 26 |   ("~~/src/Tools/nbe.ML")
 | 
| 30929 
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
 haftmann parents: diff
changeset | 27 | begin | 
| 
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
 haftmann parents: diff
changeset | 28 | |
| 39423 | 29 | setup {*
 | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39911diff
changeset | 30 | Solve_Direct.setup | 
| 39423 | 31 | #> Code_Preproc.setup | 
| 32 | #> Code_Simp.setup | |
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43018diff
changeset | 33 | #> Code_Target.setup | 
| 39423 | 34 | #> Code_ML.setup | 
| 35 | #> Code_Haskell.setup | |
| 36 | #> Code_Scala.setup | |
| 37 | #> Quickcheck.setup | |
| 43612 | 38 | #> Value.setup | 
| 39423 | 39 | *} | 
| 40 | ||
| 39421 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 41 | 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: 
39401diff
changeset | 42 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 43 | 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 | 44 | "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 | 45 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 46 | 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 | 47 | 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 | 48 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 49 | 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 | 50 | |
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 51 | 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 | 52 | "(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 | 53 | "(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 | 54 | proof - | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 55 | 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 | 56 | proof | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 57 | 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 | 58 | 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 | 59 | next | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 60 | 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 | 61 | 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 | 62 | qed | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 63 | next | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 64 | 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 | 65 | 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 | 66 | qed | 
| 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 67 | |
| 39474 | 68 | use "~~/src/Tools/Code/code_runtime.ML" | 
| 39911 | 69 | use "~~/src/Tools/nbe.ML" | 
| 70 | ||
| 71 | setup Nbe.setup | |
| 39474 | 72 | |
| 39421 
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
 haftmann parents: 
39401diff
changeset | 73 | 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 | 74 | |
| 33561 
ab01b72715ef
introduced Auto Nitpick in addition to Auto Quickcheck;
 blanchet parents: 
31775diff
changeset | 75 | end |