| author | wenzelm |
| Fri, 24 Sep 2010 14:14:21 +0200 | |
| changeset 39637 | cc3452317b5f |
| parent 39541 | 6605c1e87c7f |
| child 39756 | 6c8e83d94536 |
| 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:
34948
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
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:
39253
diff
changeset
|
18 |
Proof.context -> Proof.context; |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
19 |
|
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
20 |
val tracing : bool Unsynchronized.ref; |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
21 |
val quiet : bool Unsynchronized.ref; |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
22 |
val quickcheck_compile_term : Predicate_Compile_Aux.compilation -> bool -> bool -> int -> |
|
39253
0c47d615a69b
removing report from the arguments of the quickcheck functions and refering to it by picking it from the context
bulwahn
parents:
38797
diff
changeset
|
23 |
Proof.context -> term -> int -> term list option * (bool list * bool); |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
24 |
(* val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *) |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
25 |
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:
34948
diff
changeset
|
26 |
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:
34948
diff
changeset
|
27 |
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:
34948
diff
changeset
|
28 |
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
|
29 |
end; |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
30 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
31 |
structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK = |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
32 |
struct |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
33 |
|
| 33257 | 34 |
open Predicate_Compile_Aux; |
35 |
||
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
36 |
structure Pred_Result = Proof_Data ( |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
37 |
type T = 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:
39253
diff
changeset
|
38 |
fun init _ () = error "Pred_Result" |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
39 |
); |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
40 |
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:
34028
diff
changeset
|
41 |
|
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
42 |
structure Dseq_Result = Proof_Data ( |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
43 |
type T = 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:
39253
diff
changeset
|
44 |
fun init _ () = error "Dseq_Result" |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
45 |
); |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
46 |
val put_dseq_result = Dseq_Result.put; |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
47 |
|
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
48 |
structure Lseq_Result = Proof_Data ( |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
49 |
type T = 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:
39253
diff
changeset
|
50 |
fun init _ () = error "Lseq_Result" |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
51 |
); |
|
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
52 |
val put_lseq_result = Lseq_Result.put; |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
53 |
|
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
54 |
val tracing = Unsynchronized.ref false; |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
55 |
|
|
35537
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents:
35380
diff
changeset
|
56 |
val quiet = Unsynchronized.ref true; |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
57 |
|
|
35537
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents:
35380
diff
changeset
|
58 |
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:
34948
diff
changeset
|
59 |
|
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
60 |
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:
34948
diff
changeset
|
61 |
|
|
35537
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents:
35380
diff
changeset
|
62 |
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:
34948
diff
changeset
|
63 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
64 |
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:
34948
diff
changeset
|
65 |
|
| 35538 | 66 |
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:
34948
diff
changeset
|
67 |
|
| 33257 | 68 |
val options = Options {
|
69 |
expected_modes = NONE, |
|
|
39382
c797f3ab2ae1
proposed modes for code_pred now supports modes for mutual predicates
bulwahn
parents:
39253
diff
changeset
|
70 |
proposed_modes = [], |
|
33651
e4aad90618ad
adding the predicate compiler quickcheck to the ex/ROOT.ML; adopting this quickcheck to the latest changes
bulwahn
parents:
33486
diff
changeset
|
71 |
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:
34948
diff
changeset
|
72 |
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:
34948
diff
changeset
|
73 |
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:
34948
diff
changeset
|
74 |
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:
34948
diff
changeset
|
75 |
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:
34948
diff
changeset
|
76 |
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:
34948
diff
changeset
|
77 |
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:
34948
diff
changeset
|
78 |
show_caught_failures = false, |
|
39383
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents:
39382
diff
changeset
|
79 |
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:
34948
diff
changeset
|
80 |
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:
34948
diff
changeset
|
81 |
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:
34948
diff
changeset
|
82 |
inductify = true, |
| 36256 | 83 |
specialise = true, |
84 |
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:
34948
diff
changeset
|
85 |
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:
34948
diff
changeset
|
86 |
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:
34948
diff
changeset
|
87 |
no_higher_order_predicate = [], |
| 36250 | 88 |
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:
34948
diff
changeset
|
89 |
} |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
90 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
91 |
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:
34948
diff
changeset
|
92 |
expected_modes = NONE, |
|
39383
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents:
39382
diff
changeset
|
93 |
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:
34948
diff
changeset
|
94 |
proposed_names = [], |
| 33257 | 95 |
show_steps = true, |
96 |
show_intermediate_results = true, |
|
97 |
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:
34948
diff
changeset
|
98 |
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:
34948
diff
changeset
|
99 |
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:
34028
diff
changeset
|
100 |
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:
34948
diff
changeset
|
101 |
show_caught_failures = true, |
|
39383
ddfafa97da2f
adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents:
39382
diff
changeset
|
102 |
show_invalid_clauses = false, |
| 33257 | 103 |
skip_proof = false, |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
104 |
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:
34948
diff
changeset
|
105 |
inductify = true, |
| 36256 | 106 |
specialise = true, |
107 |
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:
34948
diff
changeset
|
108 |
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:
34948
diff
changeset
|
109 |
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:
34948
diff
changeset
|
110 |
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:
34948
diff
changeset
|
111 |
no_topmost_reordering = true |
| 33257 | 112 |
} |
113 |
||
|
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:
34948
diff
changeset
|
114 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
115 |
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:
34948
diff
changeset
|
116 |
(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:
34948
diff
changeset
|
117 |
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:
39382
diff
changeset
|
118 |
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:
39382
diff
changeset
|
119 |
show_invalid_clauses = s_ic, skip_proof = s_p, |
| 36256 | 120 |
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
121 |
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
122 |
no_topmost_reordering = re}) = |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
123 |
(Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
124 |
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:
39382
diff
changeset
|
125 |
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:
39382
diff
changeset
|
126 |
show_invalid_clauses = s_ic, skip_proof = s_p, |
| 36256 | 127 |
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:
34948
diff
changeset
|
128 |
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
129 |
no_topmost_reordering = re}) |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
130 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
131 |
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:
34948
diff
changeset
|
132 |
(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:
34948
diff
changeset
|
133 |
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:
39382
diff
changeset
|
134 |
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:
39382
diff
changeset
|
135 |
show_invalid_clauses = s_ic, skip_proof = s_p, |
| 36256 | 136 |
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:
34948
diff
changeset
|
137 |
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
138 |
no_topmost_reordering = re}) = |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
139 |
(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:
34948
diff
changeset
|
140 |
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:
39382
diff
changeset
|
141 |
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:
39382
diff
changeset
|
142 |
show_invalid_clauses = s_ic, skip_proof = s_p, |
| 36256 | 143 |
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:
34948
diff
changeset
|
144 |
fail_safe_function_flattening = b, no_higher_order_predicate = no_ho, |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
145 |
no_topmost_reordering = re}) |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
146 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
147 |
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:
34948
diff
changeset
|
148 |
(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:
34948
diff
changeset
|
149 |
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:
39382
diff
changeset
|
150 |
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:
39382
diff
changeset
|
151 |
show_invalid_clauses = s_ic, skip_proof = s_p, |
| 36256 | 152 |
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:
34948
diff
changeset
|
153 |
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho, |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
154 |
no_topmost_reordering = re}) = |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
155 |
(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:
34948
diff
changeset
|
156 |
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:
39382
diff
changeset
|
157 |
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:
39382
diff
changeset
|
158 |
show_invalid_clauses = s_ic, skip_proof = s_p, |
| 36256 | 159 |
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:
34948
diff
changeset
|
160 |
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss, no_topmost_reordering = re}) |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
161 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
162 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
163 |
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:
34948
diff
changeset
|
164 |
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:
34948
diff
changeset
|
165 |
(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:
34948
diff
changeset
|
166 |
(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:
34948
diff
changeset
|
167 |
|
|
36046
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
168 |
val mk_predT = Predicate_Compile_Aux.mk_predT PredicateCompFuns.compfuns |
|
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
169 |
val mk_return' = Predicate_Compile_Aux.mk_single PredicateCompFuns.compfuns |
|
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
170 |
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:
35886
diff
changeset
|
171 |
|
|
36046
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
172 |
val mk_randompredT = Predicate_Compile_Aux.mk_predT RandomPredCompFuns.compfuns |
|
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
173 |
val mk_return = Predicate_Compile_Aux.mk_single RandomPredCompFuns.compfuns |
|
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
174 |
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
|
175 |
|
|
36046
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
176 |
val mk_new_randompredT = Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.compfuns |
|
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
177 |
val mk_new_return = Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.compfuns |
|
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
178 |
val mk_new_bind = Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.compfuns |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
179 |
|
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
180 |
fun mk_split_lambda [] t = Abs ("u", HOLogic.unitT, t)
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
181 |
| mk_split_lambda [x] t = lambda x t |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
182 |
| mk_split_lambda xs t = |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
183 |
let |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
184 |
fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t)) |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
185 |
| mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t)) |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
186 |
in |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
187 |
mk_split_lambda' xs t |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
188 |
end; |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
189 |
|
|
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:
34948
diff
changeset
|
190 |
fun cpu_time description f = |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
191 |
let |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
192 |
val start = start_timing () |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
193 |
val result = Exn.capture f () |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
194 |
val time = Time.toMilliseconds (#cpu (end_timing start)) |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
195 |
in (Exn.release result, (description, time)) end |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
196 |
|
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
197 |
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:
34948
diff
changeset
|
198 |
let |
|
38755
a37d39fe32f8
standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
wenzelm
parents:
38549
diff
changeset
|
199 |
val thy = Theory.copy (ProofContext.theory_of ctxt) |
|
39541
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents:
39471
diff
changeset
|
200 |
val ((((full_constname, constT), vs'), intro), thy1) = |
|
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents:
39471
diff
changeset
|
201 |
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:
34948
diff
changeset
|
202 |
val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 |
| 36256 | 203 |
val (thy3, preproc_time) = cpu_time "predicate preprocessing" |
|
39541
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents:
39471
diff
changeset
|
204 |
(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:
34948
diff
changeset
|
205 |
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:
35886
diff
changeset
|
206 |
(fn () => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
207 |
case compilation of |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
208 |
Pos_Random_DSeq => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
209 |
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:
35886
diff
changeset
|
210 |
| New_Pos_Random_DSeq => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
211 |
Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3 |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
212 |
(*| Depth_Limited_Random => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
213 |
Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*)) |
| 36256 | 214 |
(*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*) |
215 |
val _ = Output.tracing ("Preprocessing time: " ^ string_of_int (snd preproc_time))
|
|
216 |
val _ = Output.tracing ("Core compilation time: " ^ string_of_int (snd core_comp_time))
|
|
|
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
36610
diff
changeset
|
217 |
val ctxt4 = ProofContext.init_global thy4 |
|
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
36610
diff
changeset
|
218 |
val modes = Predicate_Compile_Core.modes_of compilation ctxt4 full_constname |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
219 |
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
|
220 |
val prog = |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
221 |
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
|
222 |
let |
|
37005
842a73dc6d0e
changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents:
36610
diff
changeset
|
223 |
val name = Predicate_Compile_Core.function_name_of compilation ctxt4 |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
224 |
full_constname output_mode |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
225 |
val T = |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
226 |
case compilation of |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
227 |
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:
35886
diff
changeset
|
228 |
| New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs')) |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
229 |
| Depth_Limited_Random => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
230 |
[@{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:
35886
diff
changeset
|
231 |
@{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:
34028
diff
changeset
|
232 |
in |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
233 |
Const (name, T) |
|
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
234 |
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:
34948
diff
changeset
|
235 |
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:
35886
diff
changeset
|
236 |
val qc_term = |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
237 |
case compilation of |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
238 |
Pos_Random_DSeq => mk_bind (prog, |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
239 |
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:
35886
diff
changeset
|
240 |
(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:
35886
diff
changeset
|
241 |
| New_Pos_Random_DSeq => mk_new_bind (prog, |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
242 |
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:
35886
diff
changeset
|
243 |
(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:
35886
diff
changeset
|
244 |
| Depth_Limited_Random => fold_rev (curry absdummy) |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
245 |
[@{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:
35886
diff
changeset
|
246 |
@{typ "code_numeral * code_numeral"}]
|
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
247 |
(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:
35886
diff
changeset
|
248 |
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:
35886
diff
changeset
|
249 |
(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:
35886
diff
changeset
|
250 |
val prog = |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
251 |
case compilation of |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
252 |
Pos_Random_DSeq => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
253 |
let |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
254 |
val compiled_term = |
| 39471 | 255 |
Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result") |
256 |
thy4 (SOME target) |
|
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
257 |
(fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc) |
| 39471 | 258 |
qc_term [] |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
259 |
in |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
260 |
(fn size => fn nrandom => fn depth => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
261 |
Option.map fst (DSequence.yield |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
262 |
(compiled_term nrandom size |> Random_Engine.run) depth true)) |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
263 |
end |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
264 |
| New_Pos_Random_DSeq => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
265 |
let |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
266 |
val compiled_term = |
| 39471 | 267 |
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:
39253
diff
changeset
|
268 |
(Lseq_Result.get, put_lseq_result, "Predicate_Compile_Quickcheck.put_lseq_result") |
| 39471 | 269 |
thy4 (SOME target) |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
270 |
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth => |
| 36533 | 271 |
g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc) |
| 39471 | 272 |
qc_term [] |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
273 |
in |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
274 |
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:
35886
diff
changeset
|
275 |
( |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
276 |
let |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
277 |
val seed = Random_Engine.next_seed () |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
278 |
in compiled_term nrandom size seed depth end)) |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
279 |
end |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
280 |
| Depth_Limited_Random => |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
281 |
let |
| 39471 | 282 |
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:
39253
diff
changeset
|
283 |
(Pred_Result.get, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result") |
| 39471 | 284 |
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:
35886
diff
changeset
|
285 |
g depth nrandom size seed |> (Predicate.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:
35886
diff
changeset
|
287 |
in |
|
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
288 |
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:
35886
diff
changeset
|
289 |
(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:
35886
diff
changeset
|
290 |
end |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
291 |
in |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
292 |
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:
34948
diff
changeset
|
293 |
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:
34948
diff
changeset
|
294 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
295 |
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:
34948
diff
changeset
|
296 |
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:
34948
diff
changeset
|
297 |
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:
34948
diff
changeset
|
298 |
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:
34948
diff
changeset
|
299 |
let |
|
35380
6ac5b81a763d
adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents:
35378
diff
changeset
|
300 |
val _ = if quiet then () else priority ("Executing depth " ^ string_of_int j)
|
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
301 |
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:
34948
diff
changeset
|
302 |
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:
34948
diff
changeset
|
303 |
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:
34948
diff
changeset
|
304 |
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:
34948
diff
changeset
|
305 |
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:
34948
diff
changeset
|
306 |
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:
34948
diff
changeset
|
307 |
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:
34948
diff
changeset
|
308 |
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:
34948
diff
changeset
|
309 |
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:
34948
diff
changeset
|
310 |
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:
34948
diff
changeset
|
311 |
|
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
312 |
(* 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:
34948
diff
changeset
|
313 |
|
|
39253
0c47d615a69b
removing report from the arguments of the quickcheck functions and refering to it by picking it from the context
bulwahn
parents:
38797
diff
changeset
|
314 |
fun compile_term' compilation options depth 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:
34948
diff
changeset
|
315 |
let |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
316 |
val c = compile_term compilation options ctxt t |
|
35378
95d0e3adf38e
added basic reporting of test cases to quickcheck
bulwahn
parents:
35324
diff
changeset
|
317 |
val dummy_report = ([], false) |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
318 |
in |
|
35537
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents:
35380
diff
changeset
|
319 |
fn size => (try_upto (!quiet) (c size (!nrandom)) depth, dummy_report) |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
320 |
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:
34948
diff
changeset
|
321 |
|
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
322 |
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:
34948
diff
changeset
|
323 |
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:
34948
diff
changeset
|
324 |
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:
34948
diff
changeset
|
325 |
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:
34948
diff
changeset
|
326 |
(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:
34948
diff
changeset
|
327 |
in |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
328 |
compile_term' compilation options depth |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
329 |
end |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
330 |
|
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
331 |
end; |