| author | wenzelm | 
| Sun, 05 Sep 2010 19:47:40 +0200 | |
| changeset 39133 | 70d3915c92f0 | 
| parent 38797 | abe92b33ac9f | 
| child 39253 | 0c47d615a69b | 
| permissions | -rw-r--r-- | 
| 33265 | 1 | (* Title: HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML | 
| 2 | Author: Lukas Bulwahn, TU Muenchen | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 3 | |
| 33265 | 4 | A quickcheck generator based on the predicate compiler. | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 5 | *) | 
| 
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 | signature PREDICATE_COMPILE_QUICKCHECK = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 8 | sig | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 9 | (*val quickcheck : Proof.context -> term -> int -> term list option*) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 10 | val test_ref : | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 11 | ((unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option) Unsynchronized.ref | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 12 | val new_test_ref : | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 13 | ((unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence) option) Unsynchronized.ref; | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 14 | val eval_random_ref : | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 15 | ((unit -> int -> int -> int -> int * int -> term list Predicate.pred) option) Unsynchronized.ref; | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 16 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 17 | val tracing : bool Unsynchronized.ref; | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 18 | val quiet : bool Unsynchronized.ref; | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 19 | val quickcheck_compile_term : Predicate_Compile_Aux.compilation -> bool -> bool -> int -> | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 20 | Proof.context -> bool -> term -> int -> term list option * (bool list * bool); | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 21 | (* val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 22 | val nrandom : int Unsynchronized.ref; | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 23 | val debug : bool Unsynchronized.ref; | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 24 | val function_flattening : bool Unsynchronized.ref; | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 25 | val no_higher_order_predicate : string list Unsynchronized.ref; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 26 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 27 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 28 | structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 29 | struct | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 30 | |
| 33257 | 31 | open Predicate_Compile_Aux; | 
| 32 | ||
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 33 | val test_ref = | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 34 | Unsynchronized.ref (NONE : (unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option); | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 35 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 36 | val new_test_ref = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 37 | Unsynchronized.ref (NONE : (unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence) option) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 38 | |
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 39 | val eval_random_ref = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 40 | Unsynchronized.ref (NONE : (unit -> int -> int -> int -> int * int -> term list Predicate.pred) option); | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 41 | |
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 42 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 43 | val tracing = Unsynchronized.ref false; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 44 | |
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 45 | val quiet = Unsynchronized.ref true; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 46 | |
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 47 | val target = "Quickcheck" | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 48 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 49 | val nrandom = Unsynchronized.ref 3; | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 50 | |
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 51 | val debug = Unsynchronized.ref false; | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 52 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 53 | val function_flattening = Unsynchronized.ref true; | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 54 | |
| 35538 | 55 | val no_higher_order_predicate = Unsynchronized.ref ([] : string list); | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 56 | |
| 33257 | 57 | val options = Options {
 | 
| 58 | expected_modes = NONE, | |
| 33752 
9aa8e961f850
adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
 bulwahn parents: 
33651diff
changeset | 59 | proposed_modes = NONE, | 
| 33651 
e4aad90618ad
adding the predicate compiler quickcheck to the ex/ROOT.ML; adopting this quickcheck to the latest changes
 bulwahn parents: 
33486diff
changeset | 60 | proposed_names = [], | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 61 | show_steps = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 62 | show_intermediate_results = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 63 | show_proof_trace = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 64 | show_modes = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 65 | show_mode_inference = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 66 | show_compilation = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 67 | show_caught_failures = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 68 | skip_proof = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 69 | compilation = Random, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 70 | inductify = true, | 
| 36256 | 71 | specialise = true, | 
| 72 | detect_switches = false, | |
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 73 | function_flattening = true, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 74 | fail_safe_function_flattening = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 75 | no_higher_order_predicate = [], | 
| 36250 | 76 | no_topmost_reordering = false | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 77 | } | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 78 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 79 | val debug_options = Options {
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 80 | expected_modes = NONE, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 81 | proposed_modes = NONE, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 82 | proposed_names = [], | 
| 33257 | 83 | show_steps = true, | 
| 84 | show_intermediate_results = true, | |
| 85 | show_proof_trace = false, | |
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 86 | show_modes = true, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 87 | show_mode_inference = true, | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 88 | show_compilation = false, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 89 | show_caught_failures = true, | 
| 33257 | 90 | skip_proof = false, | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 91 | compilation = Random, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 92 | inductify = true, | 
| 36256 | 93 | specialise = true, | 
| 94 | detect_switches = false, | |
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 95 | function_flattening = true, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 96 | fail_safe_function_flattening = false, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 97 | no_higher_order_predicate = [], | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 98 | no_topmost_reordering = true | 
| 33257 | 99 | } | 
| 100 | ||
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 101 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 102 | fun set_function_flattening b | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 103 |   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 104 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 105 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p, | 
| 36256 | 106 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 107 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 108 | no_topmost_reordering = re}) = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 109 |   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 110 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 111 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p, | 
| 36256 | 112 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = b, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 113 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 114 | no_topmost_reordering = re}) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 115 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 116 | fun set_fail_safe_function_flattening b | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 117 |   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 118 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 119 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p, | 
| 36256 | 120 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 121 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 122 | no_topmost_reordering = re}) = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 123 |   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 124 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 125 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p, | 
| 36256 | 126 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 127 | fail_safe_function_flattening = b, no_higher_order_predicate = no_ho, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 128 | no_topmost_reordering = re}) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 129 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 130 | fun set_no_higher_order_predicate ss | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 131 |   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 132 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 133 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p, | 
| 36256 | 134 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 135 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 136 | no_topmost_reordering = re}) = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 137 |   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 138 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 139 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p, | 
| 36256 | 140 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 141 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss, no_topmost_reordering = re}) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 142 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 143 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 144 | fun get_options () = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 145 | set_no_higher_order_predicate (!no_higher_order_predicate) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 146 | (set_function_flattening (!function_flattening) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 147 | (if !debug then debug_options else options)) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 148 | |
| 36046 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 149 | val mk_predT = Predicate_Compile_Aux.mk_predT PredicateCompFuns.compfuns | 
| 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 150 | val mk_return' = Predicate_Compile_Aux.mk_single PredicateCompFuns.compfuns | 
| 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 151 | val mk_bind' = Predicate_Compile_Aux.mk_bind PredicateCompFuns.compfuns | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 152 | |
| 36046 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 153 | val mk_randompredT = Predicate_Compile_Aux.mk_predT RandomPredCompFuns.compfuns | 
| 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 154 | val mk_return = Predicate_Compile_Aux.mk_single RandomPredCompFuns.compfuns | 
| 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 155 | val mk_bind = Predicate_Compile_Aux.mk_bind RandomPredCompFuns.compfuns | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 156 | |
| 36046 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 157 | val mk_new_randompredT = Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.compfuns | 
| 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 158 | val mk_new_return = Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.compfuns | 
| 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 159 | val mk_new_bind = Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.compfuns | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 160 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 161 | fun mk_split_lambda [] t = Abs ("u", HOLogic.unitT, t)
 | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 162 | | mk_split_lambda [x] t = lambda x t | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 163 | | mk_split_lambda xs t = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 164 | let | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 165 | fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t)) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 166 | | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t)) | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 167 | in | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 168 | mk_split_lambda' xs t | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 169 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 170 | |
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38549diff
changeset | 171 | fun strip_imp_prems (Const(@{const_name HOL.implies}, _) $ A $ B) = A :: strip_imp_prems B
 | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 172 | | strip_imp_prems _ = []; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 173 | |
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38549diff
changeset | 174 | fun strip_imp_concl (Const(@{const_name HOL.implies}, _) $ A $ B) = strip_imp_concl B
 | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 175 | | strip_imp_concl A = A : term; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 176 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 177 | fun strip_horn A = (strip_imp_prems A, strip_imp_concl A); | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 178 | |
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 179 | fun cpu_time description f = | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 180 | let | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 181 | val start = start_timing () | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 182 | val result = Exn.capture f () | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 183 | val time = Time.toMilliseconds (#cpu (end_timing start)) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 184 | in (Exn.release result, (description, time)) end | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 185 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 186 | fun compile_term compilation options ctxt t = | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 187 | let | 
| 38755 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38549diff
changeset | 188 | val thy = Theory.copy (ProofContext.theory_of ctxt) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 189 | val (vs, t') = strip_abs t | 
| 38755 
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
 wenzelm parents: 
38549diff
changeset | 190 | val vs' = Variable.variant_frees ctxt [] vs | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 191 | val t'' = subst_bounds (map Free (rev vs'), t') | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 192 | val (prems, concl) = strip_horn t'' | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 193 | val constname = "pred_compile_quickcheck" | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 194 | val full_constname = Sign.full_bname thy constname | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 195 |     val constT = map snd vs' ---> @{typ bool}
 | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 196 | val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 197 | val const = Const (full_constname, constT) | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 198 | val t = Logic.list_implies | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 199 | (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]), | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 200 | HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs'))) | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 201 | val tac = fn _ => Skip_Proof.cheat_tac thy1 | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36533diff
changeset | 202 | val intro = Goal.prove (ProofContext.init_global thy1) (map fst vs') [] t tac | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 203 | val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 | 
| 36256 | 204 | val (thy3, preproc_time) = cpu_time "predicate preprocessing" | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 205 | (fn () => Predicate_Compile.preprocess options const thy2) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 206 | val (thy4, core_comp_time) = cpu_time "random_dseq core compilation" | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 207 | (fn () => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 208 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 209 | Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 210 | Predicate_Compile_Core.add_random_dseq_equations options [full_constname] thy3 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 211 | | New_Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 212 | Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 213 | (*| Depth_Limited_Random => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 214 | Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*)) | 
| 36256 | 215 | (*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*) | 
| 216 |     val _ = Output.tracing ("Preprocessing time: " ^ string_of_int (snd preproc_time))
 | |
| 217 |     val _ = Output.tracing ("Core compilation time: " ^ string_of_int (snd core_comp_time))
 | |
| 37005 
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
 bulwahn parents: 
36610diff
changeset | 218 | val ctxt4 = ProofContext.init_global thy4 | 
| 
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
 bulwahn parents: 
36610diff
changeset | 219 | val modes = Predicate_Compile_Core.modes_of compilation ctxt4 full_constname | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 220 | val output_mode = fold_rev (curry Fun) (map (K Output) (binder_types constT)) Bool | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 221 | val prog = | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 222 | if member eq_mode modes output_mode then | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 223 | let | 
| 37005 
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
 bulwahn parents: 
36610diff
changeset | 224 | val name = Predicate_Compile_Core.function_name_of compilation ctxt4 | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 225 | full_constname output_mode | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 226 | val T = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 227 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 228 | Pos_Random_DSeq => mk_randompredT (HOLogic.mk_tupleT (map snd vs')) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 229 | | New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs')) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 230 | | Depth_Limited_Random => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 231 |               [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 232 |               @{typ "code_numeral * code_numeral"}] ---> mk_predT (HOLogic.mk_tupleT (map snd vs'))
 | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 233 | in | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 234 | Const (name, T) | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 235 | end | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 236 |       else error ("Predicate Compile Quickcheck failed: " ^ commas (map string_of_mode modes))
 | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 237 | val qc_term = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 238 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 239 | Pos_Random_DSeq => mk_bind (prog, | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 240 |             mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 241 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 242 | | New_Pos_Random_DSeq => mk_new_bind (prog, | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 243 |             mk_split_lambda (map Free vs') (mk_new_return (HOLogic.mk_list @{typ term}
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 244 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 245 | | Depth_Limited_Random => fold_rev (curry absdummy) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 246 |             [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 247 |              @{typ "code_numeral * code_numeral"}]
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 248 | (mk_bind' (list_comb (prog, map Bound (3 downto 0)), | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 249 |             mk_split_lambda (map Free vs') (mk_return' (HOLogic.mk_list @{typ term}
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 250 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 251 | val prog = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 252 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 253 | Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 254 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 255 | val compiled_term = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 256 |               Code_Eval.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 257 | (fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 258 | thy4 qc_term [] | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 259 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 260 | (fn size => fn nrandom => fn depth => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 261 | Option.map fst (DSequence.yield | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 262 | (compiled_term nrandom size |> Random_Engine.run) depth true)) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 263 | end | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 264 | | New_Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 265 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 266 | val compiled_term = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 267 | Code_Eval.eval (SOME target) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 268 |                 ("Predicate_Compile_Quickcheck.new_test_ref", new_test_ref)
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 269 | (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => | 
| 36533 | 270 | g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc) | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 271 | thy4 qc_term [] | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 272 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 273 | fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 274 | ( | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 275 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 276 | val seed = Random_Engine.next_seed () | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 277 | in compiled_term nrandom size seed depth end)) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 278 | end | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 279 | | Depth_Limited_Random => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 280 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 281 | val compiled_term = Code_Eval.eval (SOME target) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 282 |               ("Predicate_Compile_Quickcheck.eval_random_ref", eval_random_ref)
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 283 | (fn proc => fn g => fn depth => fn nrandom => fn size => fn seed => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 284 | g depth nrandom size seed |> (Predicate.map o map) proc) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 285 | thy4 qc_term [] | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 286 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 287 | fn size => fn nrandom => fn depth => Option.map fst (Predicate.yield | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 288 | (compiled_term depth nrandom size (Random_Engine.run (fn s => (s, s))))) | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 289 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 290 | in | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 291 | prog | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 292 | end | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 293 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 294 | fun try_upto quiet f i = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 295 | let | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 296 | fun try' j = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 297 | if j <= i then | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 298 | let | 
| 35380 
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
 bulwahn parents: 
35378diff
changeset | 299 |           val _ = if quiet then () else priority ("Executing depth " ^ string_of_int j)
 | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 300 | in | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 301 | case f j handle Match => (if quiet then () | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 302 | else warning "Exception Match raised during quickcheck"; NONE) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 303 | of NONE => try' (j + 1) | SOME q => SOME q | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 304 | end | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 305 | else | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 306 | NONE | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 307 | in | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 308 | try' 0 | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 309 | end | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 310 | |
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 311 | (* quickcheck interface functions *) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 312 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 313 | fun compile_term' compilation options depth ctxt report t = | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 314 | let | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 315 | val c = compile_term compilation options ctxt t | 
| 35378 
95d0e3adf38e
added basic reporting of test cases to quickcheck
 bulwahn parents: 
35324diff
changeset | 316 | val dummy_report = ([], false) | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 317 | in | 
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 318 | fn size => (try_upto (!quiet) (c size (!nrandom)) depth, dummy_report) | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 319 | end | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 320 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 321 | fun quickcheck_compile_term compilation function_flattening fail_safe_function_flattening depth = | 
| 35324 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 322 | let | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 323 | val options = | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 324 | set_fail_safe_function_flattening fail_safe_function_flattening | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 325 | (set_function_flattening function_flattening (get_options ())) | 
| 
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
 bulwahn parents: 
34948diff
changeset | 326 | in | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 327 | compile_term' compilation options depth | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 328 | end | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 329 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 330 | end; |