| author | hoelzl | 
| Tue, 04 Dec 2012 18:00:37 +0100 | |
| changeset 50347 | 77e3effa50b6 | 
| parent 50055 | 94041d602ecb | 
| child 51126 | df86080de4cb | 
| permissions | -rw-r--r-- | 
| 33265 | 1 | (* Title: HOL/Predicate_Compile.thy | 
| 2 | Author: Stefan Berghofer, Lukas Bulwahn, Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 4 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 5 | header {* A compiler for predicates defined by introduction rules *}
 | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 6 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 7 | theory Predicate_Compile | 
| 50055 | 8 | imports New_Random_Sequence Quickcheck_Exhaustive | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46664diff
changeset | 9 | keywords "code_pred" :: thy_goal and "values" :: diag | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 10 | begin | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 11 | |
| 48891 | 12 | ML_file "Tools/Predicate_Compile/predicate_compile_aux.ML" | 
| 13 | ML_file "Tools/Predicate_Compile/predicate_compile_compilations.ML" | |
| 14 | ML_file "Tools/Predicate_Compile/core_data.ML" | |
| 15 | ML_file "Tools/Predicate_Compile/mode_inference.ML" | |
| 16 | ML_file "Tools/Predicate_Compile/predicate_compile_proof.ML" | |
| 17 | ML_file "Tools/Predicate_Compile/predicate_compile_core.ML" | |
| 18 | ML_file "Tools/Predicate_Compile/predicate_compile_data.ML" | |
| 19 | ML_file "Tools/Predicate_Compile/predicate_compile_fun.ML" | |
| 20 | ML_file "Tools/Predicate_Compile/predicate_compile_pred.ML" | |
| 21 | ML_file "Tools/Predicate_Compile/predicate_compile_specialisation.ML" | |
| 22 | ML_file "Tools/Predicate_Compile/predicate_compile.ML" | |
| 23 | ||
| 33265 | 24 | setup Predicate_Compile.setup | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 25 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
33265diff
changeset | 26 | end |