| author | huffman | 
| Wed, 24 Aug 2011 11:56:57 -0700 | |
| changeset 44514 | d02b01e5ab8f | 
| parent 44241 | 7943b69f0188 | 
| child 45214 | 66ba67adafab | 
| 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*) | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 10 | val put_pred_result : | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 11 | (unit -> int -> int -> int -> int * int -> term list Predicate.pred) -> | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 12 | Proof.context -> Proof.context; | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 13 | val put_dseq_result : | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 14 | (unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) -> | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 15 | Proof.context -> Proof.context; | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 16 | val put_lseq_result : | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 17 | (unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence) -> | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 18 | Proof.context -> Proof.context; | 
| 40103 | 19 | val put_new_dseq_result : (unit -> int -> term list Lazy_Sequence.lazy_sequence) -> | 
| 20 | Proof.context -> Proof.context | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 21 | |
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 22 | val tracing : bool Unsynchronized.ref; | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 23 | val quiet : bool Unsynchronized.ref; | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 24 | val test_goals : (Predicate_Compile_Aux.compilation * bool * bool * int) -> | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 25 | Proof.context -> bool * bool -> (string * typ) list -> (term * term list) list | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 26 | -> Quickcheck.result 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 | 27 | 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 | 28 | 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 | 29 | 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 | 30 | val no_higher_order_predicate : string list Unsynchronized.ref; | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 31 | val setup : theory -> theory | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 32 | end; | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 33 | |
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 34 | structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK = | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 35 | struct | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 36 | |
| 33257 | 37 | open Predicate_Compile_Aux; | 
| 38 | ||
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 39 | (* FIXME just one data slot (record) per program unit *) | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 40 | |
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 41 | structure Pred_Result = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 42 | ( | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 43 | type T = unit -> int -> int -> int -> int * int -> term list Predicate.pred | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 44 | (* FIXME avoid user error with non-user text *) | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 45 | fun init _ () = error "Pred_Result" | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 46 | ); | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 47 | val put_pred_result = Pred_Result.put; | 
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 48 | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 49 | structure Dseq_Result = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 50 | ( | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 51 | type T = unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int) | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 52 | (* FIXME avoid user error with non-user text *) | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 53 | fun init _ () = error "Dseq_Result" | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 54 | ); | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 55 | val put_dseq_result = Dseq_Result.put; | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 56 | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 57 | structure Lseq_Result = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 58 | ( | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 59 | type T = unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 60 | (* FIXME avoid user error with non-user text *) | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 61 | fun init _ () = error "Lseq_Result" | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 62 | ); | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 63 | val put_lseq_result = Lseq_Result.put; | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 64 | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 65 | structure New_Dseq_Result = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 66 | ( | 
| 40103 | 67 | type T = unit -> int -> term list Lazy_Sequence.lazy_sequence | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
40924diff
changeset | 68 | (* FIXME avoid user error with non-user text *) | 
| 40103 | 69 | fun init _ () = error "New_Dseq_Random_Result" | 
| 70 | ); | |
| 71 | val put_new_dseq_result = New_Dseq_Result.put; | |
| 72 | ||
| 34948 
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
 bulwahn parents: 
34028diff
changeset | 73 | val tracing = Unsynchronized.ref false; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 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 quiet = Unsynchronized.ref true; | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 76 | |
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 77 | 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 | 78 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 79 | 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 | 80 | |
| 35537 
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
 bulwahn parents: 
35380diff
changeset | 81 | 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 | 82 | |
| 
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 | 83 | 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 | 84 | |
| 35538 | 85 | 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 | 86 | |
| 33257 | 87 | val options = Options {
 | 
| 88 | expected_modes = NONE, | |
| 39382 
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
 bulwahn parents: 
39253diff
changeset | 89 | 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 | 90 | 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 | 91 | 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 | 92 | 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 | 93 | 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 | 94 | 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 | 95 | 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 | 96 | 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 | 97 | show_caught_failures = false, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 98 | 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 | 99 | 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 | 100 | 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 | 101 | inductify = true, | 
| 36256 | 102 | specialise = true, | 
| 103 | 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 | 104 | 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 | 105 | 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 | 106 | no_higher_order_predicate = [], | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 107 | smart_depth_limiting = true, | 
| 36250 | 108 | 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 | 109 | } | 
| 
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 | |
| 
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 | 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 | 112 | expected_modes = NONE, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 113 | 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 | 114 | proposed_names = [], | 
| 33257 | 115 | show_steps = true, | 
| 116 | show_intermediate_results = true, | |
| 117 | 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 | 118 | 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 | 119 | 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 | 120 | 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 | 121 | 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 | 122 | show_invalid_clauses = false, | 
| 33257 | 123 | 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 | 124 | 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 | 125 | inductify = true, | 
| 36256 | 126 | specialise = true, | 
| 127 | 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 | 128 | 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 | 129 | 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 | 130 | no_higher_order_predicate = [], | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 131 | 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 | 132 | no_topmost_reordering = true | 
| 33257 | 133 | } | 
| 134 | ||
| 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 | |
| 
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 | 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 | 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, | 
| 39383 
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
 bulwahn parents: 
39382diff
changeset | 139 | 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 | 140 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 141 | 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 | 142 | 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 | 143 | 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 | 144 |   (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 | 145 | 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 | 146 | 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 | 147 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 148 | 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 | 149 | 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 | 150 | 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 | 151 | |
| 
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 | 152 | 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 | 153 |   (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 | 154 | 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 | 155 | 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 | 156 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 157 | 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 | 158 | 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 | 159 | 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 | 160 |   (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 | 161 | 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 | 162 | 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 | 163 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 164 | 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 | 165 | 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 | 166 | 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 | 167 | |
| 
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 | 168 | 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 | 169 |   (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 | 170 | 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 | 171 | 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 | 172 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 173 | 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 | 174 | 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 | 175 | 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 | 176 |   (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 | 177 | 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 | 178 | 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 | 179 | show_invalid_clauses = s_ic, skip_proof = s_p, | 
| 36256 | 180 | 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 | 181 | 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 | 182 | 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 | 183 | |
| 
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 | |
| 
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 | 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 | 186 | 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 | 187 | (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 | 188 | (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 | 189 | |
| 36046 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 190 | 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 | 191 | 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 | 192 | 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 | 193 | |
| 36046 
c3946372f556
putting compilation setup of predicate compiler in a separate file
 bulwahn parents: 
36025diff
changeset | 194 | 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 | 195 | 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 | 196 | 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 | 197 | |
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 198 | val mk_new_randompredT = | 
| 40103 | 199 | Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 200 | val mk_new_return = | 
| 40103 | 201 | 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 | 202 | val mk_new_bind = | 
| 40103 | 203 | Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns | 
| 204 | ||
| 205 | val mk_new_dseqT = | |
| 206 | Predicate_Compile_Aux.mk_predT New_Pos_DSequence_CompFuns.depth_unlimited_compfuns | |
| 207 | val mk_gen_return = | |
| 208 | Predicate_Compile_Aux.mk_single New_Pos_DSequence_CompFuns.depth_unlimited_compfuns | |
| 209 | val mk_gen_bind = | |
| 210 | Predicate_Compile_Aux.mk_bind New_Pos_DSequence_CompFuns.depth_unlimited_compfuns | |
| 211 | ||
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 212 | |
| 39756 
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
 krauss parents: 
39541diff
changeset | 213 | 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 | 214 | |
| 42014 
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
 wenzelm parents: 
42012diff
changeset | 215 | 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 | 216 |   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 | 217 | 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 | 218 | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 219 | 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 | 220 | let | 
| 44241 | 221 | val t' = fold_rev absfree (Term.add_frees t []) t | 
| 42361 | 222 | val thy = Theory.copy (Proof_Context.theory_of ctxt) | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39471diff
changeset | 223 | val ((((full_constname, constT), vs'), intro), thy1) = | 
| 42031 | 224 | Predicate_Compile_Aux.define_quickcheck_predicate t' thy | 
| 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 | 225 | val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 | 
| 36256 | 226 | val (thy3, preproc_time) = cpu_time "predicate preprocessing" | 
| 39541 
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
 bulwahn parents: 
39471diff
changeset | 227 | (fn () => Predicate_Compile.preprocess options (Const (full_constname, constT)) thy2) | 
| 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 | 228 | 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 | 229 | (fn () => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 230 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 231 | Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 232 | 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 | 233 | | New_Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 234 | Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3 | 
| 40103 | 235 | | Pos_Generator_DSeq => | 
| 236 | Predicate_Compile_Core.add_generator_dseq_equations options [full_constname] thy3 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 237 | (*| Depth_Limited_Random => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 238 | Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*)) | 
| 36256 | 239 | (*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*) | 
| 240 |     val _ = Output.tracing ("Preprocessing time: " ^ string_of_int (snd preproc_time))
 | |
| 241 |     val _ = Output.tracing ("Core compilation time: " ^ string_of_int (snd core_comp_time))
 | |
| 42361 | 242 | val ctxt4 = Proof_Context.init_global thy4 | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 243 | 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 | 244 | 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 | 245 | 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 | 246 | 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 | 247 | let | 
| 40054 
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
 bulwahn parents: 
39756diff
changeset | 248 | val name = Core_Data.function_name_of compilation ctxt4 full_constname output_mode | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 249 | val T = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 250 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 251 | 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 | 252 | | New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs')) | 
| 40103 | 253 | | 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 | 254 | | Depth_Limited_Random => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 255 |               [@{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 | 256 |               @{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 | 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))
 | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 261 | val qc_term = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 262 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 263 | Pos_Random_DSeq => mk_bind (prog, | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 264 |             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 | 265 | (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 | 266 | | 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 | 267 |             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 | 268 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | 
| 40103 | 269 | | Pos_Generator_DSeq => mk_gen_bind (prog, | 
| 270 |             mk_split_lambda (map Free vs') (mk_gen_return (HOLogic.mk_list @{typ term}
 | |
| 271 | (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) | |
| 44241 | 272 | | Depth_Limited_Random => fold_rev absdummy | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 273 |             [@{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 | 274 |              @{typ "code_numeral * code_numeral"}]
 | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 275 | (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 | 276 |             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 | 277 | (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 | 278 | val prog = | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 279 | case compilation of | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 280 | Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 281 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 282 | val compiled_term = | 
| 39471 | 283 | Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result") | 
| 284 | thy4 (SOME target) | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 285 | (fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc) | 
| 39471 | 286 | qc_term [] | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 287 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 288 | (fn size => fn nrandom => fn depth => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 289 | Option.map fst (DSequence.yield | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 290 | (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 | 291 | end | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 292 | | New_Pos_Random_DSeq => | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 293 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 294 | val compiled_term = | 
| 39471 | 295 | Code_Runtime.dynamic_value_strict | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 296 | (Lseq_Result.get, put_lseq_result, "Predicate_Compile_Quickcheck.put_lseq_result") | 
| 39471 | 297 | thy4 (SOME target) | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 298 | (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => | 
| 36533 | 299 | g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc) | 
| 39471 | 300 | qc_term [] | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 301 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 302 | 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 | 303 | ( | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 304 | let | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 305 | val seed = Random_Engine.next_seed () | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 306 | 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 | 307 | end | 
| 40103 | 308 | | Pos_Generator_DSeq => | 
| 309 | let | |
| 310 | val compiled_term = | |
| 311 | Code_Runtime.dynamic_value_strict | |
| 312 | (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Quickcheck.put_new_dseq_result") | |
| 313 | thy4 (SOME target) | |
| 314 | (fn proc => fn g => fn depth => g depth |> (Lazy_Sequence.mapa o map) proc) | |
| 315 | qc_term [] | |
| 316 | in | |
| 317 | fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield (compiled_term depth)) | |
| 318 | end | |
| 319 | | Depth_Limited_Random => | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 320 | let | 
| 39471 | 321 | val compiled_term = Code_Runtime.dynamic_value_strict | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39253diff
changeset | 322 | (Pred_Result.get, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result") | 
| 39471 | 323 | thy4 (SOME target) (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 | 324 | g depth nrandom size seed |> (Predicate.map o map) proc) | 
| 39471 | 325 | qc_term [] | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 326 | in | 
| 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 327 | 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 | 328 | (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 | 329 | end | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 330 | in | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 331 | 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 | 332 | 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 | 333 | |
| 
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 | 334 | 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 | 335 | 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 | 336 | 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 | 337 | 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 | 338 | let | 
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
40103diff
changeset | 339 |           val _ = if quiet then () else Output.urgent_message ("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 | 340 | 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 | 341 | 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 | 342 | 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 | 343 | 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 | 344 | 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 | 345 | 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 | 346 | 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 | 347 | 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 | 348 | 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 | 349 | 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 | 350 | |
| 
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 | 351 | (* 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 | 352 | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 353 | fun compile_term' compilation options depth ctxt (t, eval_terms) = | 
| 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 | 354 | let | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 355 | 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 | 356 | val c = compile_term compilation options ctxt t | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 357 | val counterexample = try_upto (!quiet) (c size (!nrandom)) 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 | 358 | in | 
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 359 | Quickcheck.Result | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 360 |       {counterexample = Option.map ((curry (op ~~)) (Term.add_free_names t [])) counterexample,
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 361 | 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 | 362 | 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 | 363 | |
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 364 | 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 | 365 | 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 | 366 | 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 | 367 | 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 | 368 | (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 | 369 | in | 
| 36025 
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
 bulwahn parents: 
35886diff
changeset | 370 | compile_term' compilation options depth | 
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 371 | end | 
| 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 372 | |
| 43886 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 373 | |
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 374 | fun test_goals options ctxt (limit_time, is_interactive) insts goals = | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 375 | let | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 376 | val (compilation, function_flattening, fail_safe_function_flattening, depth) = options | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 377 | val correct_inst_goals = Quickcheck.instantiate_goals ctxt insts goals | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 378 | val test_term = | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 379 | quickcheck_compile_term compilation function_flattening fail_safe_function_flattening depth | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 380 | in | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 381 | Quickcheck.collect_results (test_term ctxt) | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 382 | (maps (map snd) correct_inst_goals) [] | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 383 | end | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 384 | |
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 385 | val wo_ff_active = Attrib.setup_config_bool @{binding quickcheck_pc_wo_ff_active} (K false);
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 386 | val ff_active = Attrib.setup_config_bool @{binding quickcheck_pc_ff_active} (K false);
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 387 | val ff_nofs_active = Attrib.setup_config_bool @{binding quickcheck_pc_ff_nofs_active} (K false);
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 388 | |
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 389 | val setup = | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 390 |   Context.theory_map (Quickcheck.add_tester ("predicate_compile_wo_ff",
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 391 | (wo_ff_active, test_goals (Predicate_Compile_Aux.New_Pos_Random_DSeq, false, true, 4)))) | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 392 |   #> Context.theory_map (Quickcheck.add_tester ("predicate_compile_ff_fs",
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 393 | (ff_active, test_goals (Predicate_Compile_Aux.New_Pos_Random_DSeq, true, true, 4)))) | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 394 |   #> Context.theory_map (Quickcheck.add_tester ("predicate_compile_ff_nofs",
 | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 395 | (wo_ff_active, test_goals (Predicate_Compile_Aux.New_Pos_Random_DSeq, true, true, 4)))) | 
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 396 | |
| 
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
 bulwahn parents: 
42361diff
changeset | 397 | |
| 33250 
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
 bulwahn parents: diff
changeset | 398 | end; |