| author | wenzelm | 
| Fri, 05 Aug 2016 20:26:13 +0200 | |
| changeset 63615 | d786d54efc70 | 
| parent 61140 | 78ece168f5b5 | 
| child 69593 | 3dda49e08b9d | 
| 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 | 
| 50057 | 9 | type seed = Random_Engine.seed | 
| 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 | 10 | (*val quickcheck : Proof.context -> term -> int -> term list option*) | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 11 | val put_pred_result : | 
| 55437 | 12 | (unit -> Code_Numeral.natural -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> | 
| 13 | term list Predicate.pred) -> | |
| 14 | Proof.context -> Proof.context | |
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 15 | val put_dseq_result : | 
| 55437 | 16 | (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> | 
| 17 | term list Limited_Sequence.dseq * seed) -> | |
| 18 | Proof.context -> Proof.context | |
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 19 | val put_lseq_result : | 
| 55437 | 20 | (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> | 
| 21 | term list Lazy_Sequence.lazy_sequence) -> | |
| 22 | Proof.context -> Proof.context | |
| 23 | val put_new_dseq_result : | |
| 24 | (unit -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence) -> | |
| 40103 | 25 | Proof.context -> Proof.context | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 26 | val put_cps_result : (unit -> Code_Numeral.natural -> (bool * term list) option) -> | 
| 45450 | 27 | Proof.context -> Proof.context | 
| 28 | val test_goals : (Predicate_Compile_Aux.compilation * bool) -> | |
| 61140 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 Andreas Lochbihler parents: 
59154diff
changeset | 29 | Proof.context -> bool -> (string * typ) list -> (term * term list) list -> | 
| 55437 | 30 | Quickcheck.result list | 
| 31 | val nrandom : int Unsynchronized.ref | |
| 32 | val debug : bool Unsynchronized.ref | |
| 33 | 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 | 34 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 35 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 36 | structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 37 | struct | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 38 | |
| 33257 | 39 | open Predicate_Compile_Aux; | 
| 40 | ||
| 50093 
a2886be4d615
repaired slip accidentally introduced in 57209cfbf16b
 haftmann parents: 
50057diff
changeset | 41 | type seed = Random_Engine.seed; | 
| 50057 | 42 | |
| 59154 | 43 | structure Data = Proof_Data | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 44 | ( | 
| 59154 | 45 | type T = | 
| 46 | (unit -> Code_Numeral.natural -> Code_Numeral.natural -> | |
| 47 | Code_Numeral.natural -> seed -> term list Predicate.pred) * | |
| 48 | (unit -> Code_Numeral.natural -> Code_Numeral.natural -> | |
| 49 | seed -> term list Limited_Sequence.dseq * seed) * | |
| 50 | (unit -> Code_Numeral.natural -> Code_Numeral.natural -> | |
| 51 | seed -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence) * | |
| 52 | (unit -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence) * | |
| 53 | (unit -> Code_Numeral.natural -> (bool * term list) option); | |
| 54 | val empty: T = | |
| 55 | (fn () => raise Fail "pred_result", | |
| 56 | fn () => raise Fail "dseq_result", | |
| 57 | fn () => raise Fail "lseq_result", | |
| 58 | fn () => raise Fail "new_dseq_result", | |
| 59 | fn () => raise Fail "cps_result"); | |
| 60 | fun init _ = empty; | |
| 61 | ); | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 62 | |
| 59154 | 63 | val get_pred_result = #1 o Data.get; | 
| 64 | val get_dseq_result = #2 o Data.get; | |
| 65 | val get_lseq_result = #3 o Data.get; | |
| 66 | val get_new_dseq_result = #4 o Data.get; | |
| 67 | val get_cps_result = #5 o Data.get; | |
| 40103 | 68 | |
| 59154 | 69 | val put_pred_result = Data.map o @{apply 5(1)} o K;
 | 
| 70 | val put_dseq_result = Data.map o @{apply 5(2)} o K;
 | |
| 71 | val put_lseq_result = Data.map o @{apply 5(3)} o K;
 | |
| 72 | val put_new_dseq_result = Data.map o @{apply 5(4)} o K;
 | |
| 73 | val put_cps_result = Data.map o @{apply 5(5)} o K;
 | |
| 45450 | 74 | |
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 75 | 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 | 76 | |
| 55437 | 77 | 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 | 78 | |
| 55437 | 79 | 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 | 80 | |
| 55437 | 81 | 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 | 82 | |
| 33257 | 83 | val options = Options {
 | 
| 84 | expected_modes = NONE, | |
| 39382 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 bulwahn parents: 
39253diff
changeset | 85 | proposed_modes = [], | 
| 33651 
e4aad90618ad
adding the predicate compiler quickcheck to the ex/ROOT.ML; adopting this quickcheck to the latest changes
 bulwahn parents: 
33486diff
changeset | 86 | 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 | 87 | 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 | 88 | 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 | 89 | 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 | 90 | 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 | 91 | 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 | 92 | 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 | 93 | show_caught_failures = false, | 
| 55437 | 94 | show_invalid_clauses = 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 | 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 | 96 | 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 | 97 | inductify = true, | 
| 36256 | 98 | specialise = true, | 
| 99 | 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 | 100 | 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 | 101 | 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 | 102 | no_higher_order_predicate = [], | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 103 | smart_depth_limiting = true, | 
| 36250 | 104 | 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 | 105 | } | 
| 
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 | 106 | |
| 
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 | 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 | 108 | expected_modes = NONE, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 109 | proposed_modes = [], | 
| 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 | 110 | proposed_names = [], | 
| 33257 | 111 | show_steps = true, | 
| 112 | show_intermediate_results = true, | |
| 113 | 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 | 114 | 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 | 115 | 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 | 116 | 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 | 117 | show_caught_failures = true, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 118 | show_invalid_clauses = false, | 
| 33257 | 119 | 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 | 120 | 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 | 121 | inductify = true, | 
| 36256 | 122 | specialise = true, | 
| 123 | 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 | 124 | 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 | 125 | 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 | 126 | no_higher_order_predicate = [], | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 127 | smart_depth_limiting = true, | 
| 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 | 128 | no_topmost_reordering = true | 
| 33257 | 129 | } | 
| 130 | ||
| 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 | 131 | |
| 
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 | 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 | 133 |   (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 | 134 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 135 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, | 
| 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 136 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 55437 | 137 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = _, | 
| 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 | 138 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 139 | smart_depth_limiting = sm_dl, no_topmost_reordering = re}) = | 
| 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 | 140 |   (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 | 141 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 142 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, | 
| 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 143 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 144 | 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 | 145 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 146 | smart_depth_limiting = sm_dl, no_topmost_reordering = re}) | 
| 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 | 147 | |
| 
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 | 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 | 149 |   (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 | 150 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 151 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, | 
| 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 152 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 55437 | 153 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 154 | fail_safe_function_flattening = _, no_higher_order_predicate = no_ho, | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 155 | smart_depth_limiting = sm_dl, no_topmost_reordering = re}) = | 
| 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 | 156 |   (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 | 157 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 158 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, | 
| 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 159 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 160 | 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 | 161 | fail_safe_function_flattening = b, no_higher_order_predicate = no_ho, | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 162 | smart_depth_limiting = sm_dl, no_topmost_reordering = re}) | 
| 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 | 163 | |
| 
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 | 164 | 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 | 165 |   (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 | 166 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 167 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, | 
| 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 168 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 55437 | 169 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 170 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = _, | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 171 | smart_depth_limiting = sm_dl, no_topmost_reordering = re}) = | 
| 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 | 172 |   (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 | 173 | show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 174 | show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, | 
| 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 175 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 176 | compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 177 | fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss, | 
| 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 178 | smart_depth_limiting = sm_dl, no_topmost_reordering = re}) | 
| 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 | |
| 
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 | 180 | |
| 55437 | 181 | fun get_options () = | 
| 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 | 182 | set_no_higher_order_predicate (!no_higher_order_predicate) | 
| 45450 | 183 | (if !debug then debug_options else options) | 
| 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 | 184 | |
| 45461 
130c90bb80b4
using more conventional names for monad plus operations
 bulwahn parents: 
45450diff
changeset | 185 | val mk_predT = Predicate_Compile_Aux.mk_monadT Predicate_Comp_Funs.compfuns | 
| 45450 | 186 | val mk_return' = Predicate_Compile_Aux.mk_single Predicate_Comp_Funs.compfuns | 
| 187 | val mk_bind' = Predicate_Compile_Aux.mk_bind Predicate_Comp_Funs.compfuns | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 188 | |
| 45461 
130c90bb80b4
using more conventional names for monad plus operations
 bulwahn parents: 
45450diff
changeset | 189 | val mk_randompredT = Predicate_Compile_Aux.mk_monadT RandomPredCompFuns.compfuns | 
| 36046 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 190 | 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 | 191 | 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 | 192 | |
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 193 | val mk_new_randompredT = | 
| 45461 
130c90bb80b4
using more conventional names for monad plus operations
 bulwahn parents: 
45450diff
changeset | 194 | Predicate_Compile_Aux.mk_monadT New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 195 | val mk_new_return = | 
| 40103 | 196 | Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 197 | val mk_new_bind = | 
| 40103 | 198 | Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns | 
| 199 | ||
| 200 | val mk_new_dseqT = | |
| 45461 
130c90bb80b4
using more conventional names for monad plus operations
 bulwahn parents: 
45450diff
changeset | 201 | Predicate_Compile_Aux.mk_monadT New_Pos_DSequence_CompFuns.depth_unlimited_compfuns | 
| 40103 | 202 | val mk_gen_return = | 
| 203 | Predicate_Compile_Aux.mk_single New_Pos_DSequence_CompFuns.depth_unlimited_compfuns | |
| 204 | val mk_gen_bind = | |
| 205 | Predicate_Compile_Aux.mk_bind New_Pos_DSequence_CompFuns.depth_unlimited_compfuns | |
| 55437 | 206 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 207 | |
| 45450 | 208 | val mk_cpsT = | 
| 45461 
130c90bb80b4
using more conventional names for monad plus operations
 bulwahn parents: 
45450diff
changeset | 209 | Predicate_Compile_Aux.mk_monadT Pos_Bounded_CPS_Comp_Funs.compfuns | 
| 45450 | 210 | |
| 39756 
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
 krauss parents: 
39541diff
changeset | 211 | val mk_split_lambda = HOLogic.tupled_lambda o HOLogic.mk_tuple | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 212 | |
| 42014 
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
 wenzelm parents: 
42012diff
changeset | 213 | fun cpu_time description e = | 
| 
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
 wenzelm parents: 
42012diff
changeset | 214 |   let val ({cpu, ...}, result) = Timing.timing e ()
 | 
| 
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
 wenzelm parents: 
42012diff
changeset | 215 | in (result, (description, Time.toMilliseconds cpu)) 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 | 216 | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 217 | 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 | 218 | let | 
| 44241 | 219 | val t' = fold_rev absfree (Term.add_frees t []) t | 
| 52788 | 220 | val thy = Proof_Context.theory_of ctxt | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39471diff
changeset | 221 | val ((((full_constname, constT), vs'), intro), thy1) = | 
| 42031 | 222 | Predicate_Compile_Aux.define_quickcheck_predicate t' thy | 
| 57962 | 223 | val thy2 = | 
| 224 |       Context.theory_map (Named_Theorems.add_thm @{named_theorems code_pred_def} intro) thy1
 | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 225 | val (thy3, _) = cpu_time "predicate preprocessing" | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39471diff
changeset | 226 | (fn () => Predicate_Compile.preprocess options (Const (full_constname, constT)) thy2) | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 227 | val (thy4, _) = 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 | 228 | (fn () => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 229 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 230 | Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 231 | 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 | 232 | | New_Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 233 | Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3 | 
| 40103 | 234 | | Pos_Generator_DSeq => | 
| 235 | Predicate_Compile_Core.add_generator_dseq_equations options [full_constname] thy3 | |
| 45450 | 236 | | Pos_Generator_CPS => | 
| 237 | Predicate_Compile_Core.add_generator_cps_equations options [full_constname] thy3 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 238 | (*| Depth_Limited_Random => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 239 | Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*)) | 
| 36256 | 240 | (*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*) | 
| 42361 | 241 | val ctxt4 = Proof_Context.init_global thy4 | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 242 | val modes = Core_Data.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 | 243 | 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 | 244 | 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 | 245 | 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 | 246 | let | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 247 | val name = Core_Data.function_name_of compilation ctxt4 full_constname output_mode | 
| 55437 | 248 | val T = | 
| 249 | (case compilation of | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 250 | 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 | 251 | | New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs')) | 
| 40103 | 252 | | Pos_Generator_DSeq => mk_new_dseqT (HOLogic.mk_tupleT (map snd vs')) | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 253 | | Depth_Limited_Random => | 
| 55437 | 254 |                 [@{typ natural}, @{typ natural}, @{typ natural},
 | 
| 255 |                  @{typ Random.seed}] ---> mk_predT (HOLogic.mk_tupleT (map snd vs'))
 | |
| 256 | | Pos_Generator_CPS => mk_cpsT (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 | 257 | in | 
| 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 258 | 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 | 259 | 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 | 260 |       else error ("Predicate Compile Quickcheck failed: " ^ commas (map string_of_mode modes))
 | 
| 45450 | 261 |     fun mk_Some T = Const (@{const_name "Option.Some"}, T --> Type (@{type_name "Option.option"}, [T]))
 | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 262 | val qc_term = | 
| 55437 | 263 | (case compilation of | 
| 264 | Pos_Random_DSeq => mk_bind (prog, | |
| 265 |           mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
 | |
| 266 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | |
| 267 | | New_Pos_Random_DSeq => mk_new_bind (prog, | |
| 268 |           mk_split_lambda (map Free vs') (mk_new_return (HOLogic.mk_list @{typ term}
 | |
| 269 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | |
| 270 | | Pos_Generator_DSeq => mk_gen_bind (prog, | |
| 271 |           mk_split_lambda (map Free vs') (mk_gen_return (HOLogic.mk_list @{typ term}
 | |
| 272 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | |
| 273 | | Pos_Generator_CPS => prog $ | |
| 274 |           mk_split_lambda (map Free vs') (mk_Some @{typ "bool * term list"} $
 | |
| 275 |           HOLogic.mk_prod (@{term "True"}, HOLogic.mk_list @{typ term}
 | |
| 276 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))) | |
| 277 | | Depth_Limited_Random => fold_rev absdummy | |
| 278 |           [@{typ natural}, @{typ natural}, @{typ natural},
 | |
| 279 |            @{typ Random.seed}]
 | |
| 280 | (mk_bind' (list_comb (prog, map Bound (3 downto 0)), | |
| 281 |           mk_split_lambda (map Free vs') (mk_return' (HOLogic.mk_list @{typ term}
 | |
| 282 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))))) | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 283 | val prog = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 284 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 285 | Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 286 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 287 | val compiled_term = | 
| 59154 | 288 | Code_Runtime.dynamic_value_strict | 
| 289 | (get_dseq_result, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result") | |
| 55757 | 290 | (Proof_Context.init_global thy4) (SOME target) | 
| 59154 | 291 | (fn proc => fn g => fn n => fn size => fn s => | 
| 292 | g n size s |>> (Limited_Sequence.map o map) proc) | |
| 39471 | 293 | qc_term [] | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 294 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 295 | (fn size => fn nrandom => fn depth => | 
| 51126 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 haftmann parents: 
50093diff
changeset | 296 | Option.map fst (Limited_Sequence.yield | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 297 | (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 | 298 | end | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 299 | | New_Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 300 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 301 | val compiled_term = | 
| 39471 | 302 | Code_Runtime.dynamic_value_strict | 
| 59154 | 303 | (get_lseq_result, put_lseq_result, "Predicate_Compile_Quickcheck.put_lseq_result") | 
| 55757 | 304 | (Proof_Context.init_global thy4) (SOME target) | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 305 | (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => | 
| 51126 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 haftmann parents: 
50093diff
changeset | 306 | g nrandom size s depth |> (Lazy_Sequence.map o map) proc) | 
| 39471 | 307 | qc_term [] | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 308 | in | 
| 55437 | 309 | fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 310 | ( | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 311 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 312 | val seed = Random_Engine.next_seed () | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 313 | 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 | 314 | end | 
| 40103 | 315 | | Pos_Generator_DSeq => | 
| 316 | let | |
| 317 | val compiled_term = | |
| 318 | Code_Runtime.dynamic_value_strict | |
| 59154 | 319 | (get_new_dseq_result, put_new_dseq_result, | 
| 320 | "Predicate_Compile_Quickcheck.put_new_dseq_result") | |
| 55757 | 321 | (Proof_Context.init_global thy4) (SOME target) | 
| 51126 
df86080de4cb
reform of predicate compiler / quickcheck theories:
 haftmann parents: 
50093diff
changeset | 322 | (fn proc => fn g => fn depth => g depth |> (Lazy_Sequence.map o map) proc) | 
| 40103 | 323 | qc_term [] | 
| 324 | in | |
| 59154 | 325 | fn size => fn nrandom => fn depth => | 
| 326 | Option.map fst (Lazy_Sequence.yield (compiled_term depth)) | |
| 40103 | 327 | end | 
| 45450 | 328 | | Pos_Generator_CPS => | 
| 329 | let | |
| 330 | val compiled_term = | |
| 331 | Code_Runtime.dynamic_value_strict | |
| 59154 | 332 | (get_cps_result, put_cps_result, "Predicate_Compile_Quickcheck.put_cps_result") | 
| 55757 | 333 | (Proof_Context.init_global thy4) (SOME target) | 
| 45750 
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
 bulwahn parents: 
45484diff
changeset | 334 | (fn proc => fn g => fn depth => g depth |> Option.map (apsnd (map proc))) | 
| 45450 | 335 | qc_term [] | 
| 336 | in | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 337 | fn _ => fn _ => Option.map snd o compiled_term | 
| 45450 | 338 | end | 
| 40103 | 339 | | Depth_Limited_Random => | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 340 | let | 
| 39471 | 341 | val compiled_term = Code_Runtime.dynamic_value_strict | 
| 59154 | 342 | (get_pred_result, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result") | 
| 343 | (Proof_Context.init_global thy4) (SOME target) | |
| 344 | (fn proc => fn g => fn depth => fn nrandom => fn size => fn seed => | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 345 | g depth nrandom size seed |> (Predicate.map o map) proc) | 
| 39471 | 346 | qc_term [] | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 347 | in | 
| 55437 | 348 | fn size => fn nrandom => fn depth => Option.map fst (Predicate.yield | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 349 | (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 | 350 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 351 | in | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 352 | 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 | 353 | 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 | 354 | |
| 45214 | 355 | fun try_upto_depth ctxt 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 | 356 | let | 
| 45214 | 357 | val max_depth = Config.get ctxt Quickcheck.depth | 
| 58843 | 358 | fun message s = if Config.get ctxt Quickcheck.quiet then () else writeln s | 
| 45214 | 359 | fun try' i = | 
| 360 | if i <= max_depth then | |
| 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 | 361 | let | 
| 45214 | 362 |           val _ = message ("Depth: " ^ string_of_int i)
 | 
| 363 | val (result, time) = | |
| 364 |             cpu_time ("Depth " ^ string_of_int i) (fn () =>
 | |
| 365 | f i handle Match => (if Config.get ctxt Quickcheck.quiet then () | |
| 366 | else warning "Exception Match raised during quickcheck"; NONE)) | |
| 367 | val _ = if Config.get ctxt Quickcheck.timing then | |
| 368 | message (fst time ^ ": " ^ string_of_int (snd time) ^ " ms") else () | |
| 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 | 369 | in | 
| 55437 | 370 | (case result of NONE => try' (i + 1) | SOME q => SOME q) | 
| 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 | 371 | end | 
| 55437 | 372 | else NONE | 
| 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 | 373 | 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 | 374 | 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 | 375 | 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 | 376 | |
| 55437 | 377 | |
| 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 | 378 | (* 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 | 379 | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 380 | 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 | 381 | let | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 382 | val size = Config.get ctxt Quickcheck.size | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 383 | val c = compile_term compilation options ctxt t | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 384 | val counterexample = try_upto_depth ctxt (c (Code_Numeral.natural_of_integer size) | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 385 | (Code_Numeral.natural_of_integer (!nrandom)) o Code_Numeral.natural_of_integer) | 
| 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 | 386 | in | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 387 | Quickcheck.Result | 
| 55437 | 388 |       {counterexample =
 | 
| 389 | Option.map (pair true o (curry (op ~~)) (Term.add_free_names t [])) counterexample, | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 390 | evaluation_terms = Option.map (K []) counterexample, timings = [], reports = []} | 
| 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 | 391 | 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 | 392 | |
| 45214 | 393 | fun quickcheck_compile_term compilation function_flattening fail_safe_function_flattening = | 
| 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 | 394 | 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 | 395 | 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 | 396 | 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 | 397 | (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 | 398 | in | 
| 45214 | 399 | compile_term' compilation options | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 400 | end | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 401 | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 402 | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51126diff
changeset | 403 | fun test_goals options ctxt _ insts goals = | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 404 | let | 
| 45450 | 405 | val (compilation, fail_safe_function_flattening) = options | 
| 406 | val function_flattening = Config.get ctxt (Quickcheck.allow_function_inversion) | |
| 45214 | 407 | val correct_inst_goals = Quickcheck_Common.instantiate_goals ctxt insts goals | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 408 | val test_term = | 
| 45214 | 409 | quickcheck_compile_term compilation function_flattening fail_safe_function_flattening | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 410 | in | 
| 45214 | 411 | Quickcheck_Common.collect_results (test_term ctxt) | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 412 | (maps (map snd) correct_inst_goals) [] | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 413 | end | 
| 45450 | 414 | |
| 55437 | 415 | val smart_exhaustive_active = | 
| 416 |   Attrib.setup_config_bool @{binding quickcheck_smart_exhaustive_active} (K true)
 | |
| 417 | val smart_slow_exhaustive_active = | |
| 418 |   Attrib.setup_config_bool @{binding quickcheck_slow_smart_exhaustive_active} (K false)
 | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 419 | |
| 58812 | 420 | val _ = | 
| 421 | Theory.setup | |
| 422 | (Exhaustive_Generators.setup_exhaustive_datatype_interpretation #> | |
| 423 |     Context.theory_map (Quickcheck.add_tester ("smart_exhaustive",
 | |
| 424 | (smart_exhaustive_active, test_goals (Predicate_Compile_Aux.Pos_Generator_CPS, false)))) #> | |
| 425 |     Context.theory_map (Quickcheck.add_tester ("smart_slow_exhaustive",
 | |
| 426 | (smart_slow_exhaustive_active, test_goals (Predicate_Compile_Aux.Pos_Generator_DSeq, false))))) | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 427 | |
| 55437 | 428 | end |