author | haftmann |
Mon, 12 Nov 2012 23:24:40 +0100 | |
changeset 50056 | 72efd6b4038d |
parent 45750 | 17100f4ce0b5 |
child 50057 | 57209cfbf16b |
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; |
40103 | 19 |
val put_new_dseq_result : (unit -> int -> term list Lazy_Sequence.lazy_sequence) -> |
20 |
Proof.context -> Proof.context |
|
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
21 |
val put_cps_result : (unit -> int -> (bool * term list) option) -> |
45450 | 22 |
Proof.context -> Proof.context |
23 |
val test_goals : (Predicate_Compile_Aux.compilation * bool) -> |
|
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
24 |
Proof.context -> bool * bool -> (string * typ) list -> (term * term list) list |
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
25 |
-> 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:
34948
diff
changeset
|
26 |
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
|
27 |
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
|
28 |
val no_higher_order_predicate : string list Unsynchronized.ref; |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
29 |
val setup : theory -> theory |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
30 |
end; |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
31 |
|
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
32 |
structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK = |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
33 |
struct |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
34 |
|
33257 | 35 |
open Predicate_Compile_Aux; |
36 |
||
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
37 |
(* 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:
40924
diff
changeset
|
38 |
|
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
39 |
structure Pred_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
40 |
( |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
41 |
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:
40924
diff
changeset
|
42 |
(* 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:
39253
diff
changeset
|
43 |
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
|
44 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
45 |
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
|
46 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
47 |
structure Dseq_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
48 |
( |
39388
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 -> 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:
40924
diff
changeset
|
50 |
(* 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:
39253
diff
changeset
|
51 |
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
|
52 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
53 |
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
|
54 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
55 |
structure Lseq_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
56 |
( |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
57 |
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:
40924
diff
changeset
|
58 |
(* 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:
39253
diff
changeset
|
59 |
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
|
60 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39253
diff
changeset
|
61 |
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
|
62 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
63 |
structure New_Dseq_Result = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
40924
diff
changeset
|
64 |
( |
40103 | 65 |
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:
40924
diff
changeset
|
66 |
(* FIXME avoid user error with non-user text *) |
40103 | 67 |
fun init _ () = error "New_Dseq_Random_Result" |
68 |
); |
|
69 |
val put_new_dseq_result = New_Dseq_Result.put; |
|
70 |
||
45450 | 71 |
structure CPS_Result = Proof_Data |
72 |
( |
|
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
73 |
type T = unit -> int -> (bool * term list) option |
45450 | 74 |
(* FIXME avoid user error with non-user text *) |
75 |
fun init _ () = error "CPS_Result" |
|
76 |
); |
|
77 |
val put_cps_result = CPS_Result.put; |
|
78 |
||
35537
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents:
35380
diff
changeset
|
79 |
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
|
80 |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
81 |
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
|
82 |
|
35537
59dd6be5834c
adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents:
35380
diff
changeset
|
83 |
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
|
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:
34948
diff
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:
39253
diff
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:
33486
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
changeset
|
106 |
no_higher_order_predicate = [], |
40054
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents:
39756
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
changeset
|
112 |
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
|
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34028
diff
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:
34948
diff
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:
39382
diff
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:
34028
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
changeset
|
130 |
no_higher_order_predicate = [], |
40054
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents:
39756
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
39382
diff
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:
34948
diff
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:
39756
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
39382
diff
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:
34948
diff
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:
39756
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
39382
diff
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:
34948
diff
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:
39756
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
39382
diff
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:
34948
diff
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:
39756
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
39382
diff
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:
34948
diff
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:
39756
diff
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:
34948
diff
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:
34948
diff
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:
39382
diff
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:
39382
diff
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:
39756
diff
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:
39756
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
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:
34948
diff
changeset
|
186 |
set_no_higher_order_predicate (!no_higher_order_predicate) |
45450 | 187 |
(if !debug then debug_options else options) |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
188 |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45450
diff
changeset
|
189 |
val mk_predT = Predicate_Compile_Aux.mk_monadT Predicate_Comp_Funs.compfuns |
45450 | 190 |
val mk_return' = Predicate_Compile_Aux.mk_single Predicate_Comp_Funs.compfuns |
191 |
val mk_bind' = Predicate_Compile_Aux.mk_bind Predicate_Comp_Funs.compfuns |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
192 |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45450
diff
changeset
|
193 |
val mk_randompredT = Predicate_Compile_Aux.mk_monadT RandomPredCompFuns.compfuns |
36046
c3946372f556
putting compilation setup of predicate compiler in a separate file
bulwahn
parents:
36025
diff
changeset
|
194 |
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
|
195 |
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
|
196 |
|
40054
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents:
39756
diff
changeset
|
197 |
val mk_new_randompredT = |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45450
diff
changeset
|
198 |
Predicate_Compile_Aux.mk_monadT New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns |
40054
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents:
39756
diff
changeset
|
199 |
val mk_new_return = |
40103 | 200 |
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:
39756
diff
changeset
|
201 |
val mk_new_bind = |
40103 | 202 |
Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns |
203 |
||
204 |
val mk_new_dseqT = |
|
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45450
diff
changeset
|
205 |
Predicate_Compile_Aux.mk_monadT New_Pos_DSequence_CompFuns.depth_unlimited_compfuns |
40103 | 206 |
val mk_gen_return = |
207 |
Predicate_Compile_Aux.mk_single New_Pos_DSequence_CompFuns.depth_unlimited_compfuns |
|
208 |
val mk_gen_bind = |
|
209 |
Predicate_Compile_Aux.mk_bind New_Pos_DSequence_CompFuns.depth_unlimited_compfuns |
|
210 |
||
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
211 |
|
45450 | 212 |
val mk_cpsT = |
45461
130c90bb80b4
using more conventional names for monad plus operations
bulwahn
parents:
45450
diff
changeset
|
213 |
Predicate_Compile_Aux.mk_monadT Pos_Bounded_CPS_Comp_Funs.compfuns |
45450 | 214 |
val mk_cps_return = |
215 |
Predicate_Compile_Aux.mk_single Pos_Bounded_CPS_Comp_Funs.compfuns |
|
216 |
val mk_cps_bind = |
|
217 |
Predicate_Compile_Aux.mk_bind Pos_Bounded_CPS_Comp_Funs.compfuns |
|
218 |
||
39756
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39541
diff
changeset
|
219 |
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
|
220 |
|
42014
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
wenzelm
parents:
42012
diff
changeset
|
221 |
fun cpu_time description e = |
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
wenzelm
parents:
42012
diff
changeset
|
222 |
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:
42012
diff
changeset
|
223 |
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:
34948
diff
changeset
|
224 |
|
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
225 |
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
|
226 |
let |
44241 | 227 |
val t' = fold_rev absfree (Term.add_frees t []) t |
42361 | 228 |
val thy = Theory.copy (Proof_Context.theory_of ctxt) |
39541
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents:
39471
diff
changeset
|
229 |
val ((((full_constname, constT), vs'), intro), thy1) = |
42031 | 230 |
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
|
231 |
val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1 |
36256 | 232 |
val (thy3, preproc_time) = cpu_time "predicate preprocessing" |
39541
6605c1e87c7f
removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents:
39471
diff
changeset
|
233 |
(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
|
234 |
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
|
235 |
(fn () => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
236 |
case compilation of |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
237 |
Pos_Random_DSeq => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
238 |
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
|
239 |
| New_Pos_Random_DSeq => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
240 |
Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3 |
40103 | 241 |
| Pos_Generator_DSeq => |
242 |
Predicate_Compile_Core.add_generator_dseq_equations options [full_constname] thy3 |
|
45450 | 243 |
| Pos_Generator_CPS => |
244 |
Predicate_Compile_Core.add_generator_cps_equations options [full_constname] thy3 |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
245 |
(*| Depth_Limited_Random => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
246 |
Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*)) |
36256 | 247 |
(*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*) |
42361 | 248 |
val ctxt4 = Proof_Context.init_global thy4 |
40054
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents:
39756
diff
changeset
|
249 |
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:
34028
diff
changeset
|
250 |
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
|
251 |
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
|
252 |
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
|
253 |
let |
40054
cd7b1fa20bce
adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents:
39756
diff
changeset
|
254 |
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:
35886
diff
changeset
|
255 |
val T = |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
256 |
case compilation of |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
257 |
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
|
258 |
| New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs')) |
40103 | 259 |
| 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:
35886
diff
changeset
|
260 |
| Depth_Limited_Random => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
261 |
[@{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
|
262 |
@{typ "code_numeral * code_numeral"}] ---> mk_predT (HOLogic.mk_tupleT (map snd vs')) |
45450 | 263 |
| Pos_Generator_CPS => mk_cpsT (HOLogic.mk_tupleT (map snd vs')) |
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
264 |
in |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
34028
diff
changeset
|
265 |
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
|
266 |
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
|
267 |
else error ("Predicate Compile Quickcheck failed: " ^ commas (map string_of_mode modes)) |
45450 | 268 |
fun mk_Some T = Const (@{const_name "Option.Some"}, T --> Type (@{type_name "Option.option"}, [T])) |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
269 |
val qc_term = |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
270 |
case compilation of |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
271 |
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
|
272 |
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
|
273 |
(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
|
274 |
| 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
|
275 |
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
|
276 |
(map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) |
40103 | 277 |
| Pos_Generator_DSeq => mk_gen_bind (prog, |
278 |
mk_split_lambda (map Free vs') (mk_gen_return (HOLogic.mk_list @{typ term} |
|
279 |
(map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))) |
|
45450 | 280 |
| Pos_Generator_CPS => prog $ |
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
281 |
mk_split_lambda (map Free vs') (mk_Some @{typ "bool * term list"} $ |
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
282 |
HOLogic.mk_prod (@{term "True"}, HOLogic.mk_list @{typ term} |
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
283 |
(map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))) |
44241 | 284 |
| Depth_Limited_Random => fold_rev absdummy |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
285 |
[@{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
|
286 |
@{typ "code_numeral * code_numeral"}] |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
287 |
(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
|
288 |
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
|
289 |
(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
|
290 |
val prog = |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
291 |
case compilation of |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
292 |
Pos_Random_DSeq => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
293 |
let |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
294 |
val compiled_term = |
39471 | 295 |
Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result") |
296 |
thy4 (SOME target) |
|
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
297 |
(fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc) |
39471 | 298 |
qc_term [] |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
299 |
in |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
300 |
(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
|
301 |
Option.map fst (DSequence.yield |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
302 |
(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
|
303 |
end |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
304 |
| New_Pos_Random_DSeq => |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
305 |
let |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
306 |
val compiled_term = |
39471 | 307 |
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
|
308 |
(Lseq_Result.get, put_lseq_result, "Predicate_Compile_Quickcheck.put_lseq_result") |
39471 | 309 |
thy4 (SOME target) |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
310 |
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth => |
36533 | 311 |
g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc) |
39471 | 312 |
qc_term [] |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
313 |
in |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
314 |
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
|
315 |
( |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
316 |
let |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
317 |
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
|
318 |
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
|
319 |
end |
40103 | 320 |
| Pos_Generator_DSeq => |
321 |
let |
|
322 |
val compiled_term = |
|
323 |
Code_Runtime.dynamic_value_strict |
|
324 |
(New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Quickcheck.put_new_dseq_result") |
|
325 |
thy4 (SOME target) |
|
326 |
(fn proc => fn g => fn depth => g depth |> (Lazy_Sequence.mapa o map) proc) |
|
327 |
qc_term [] |
|
328 |
in |
|
329 |
fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield (compiled_term depth)) |
|
330 |
end |
|
45450 | 331 |
| Pos_Generator_CPS => |
332 |
let |
|
333 |
val compiled_term = |
|
334 |
Code_Runtime.dynamic_value_strict |
|
335 |
(CPS_Result.get, put_cps_result, "Predicate_Compile_Quickcheck.put_cps_result") |
|
336 |
thy4 (SOME target) |
|
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
337 |
(fn proc => fn g => fn depth => g depth |> Option.map (apsnd (map proc))) |
45450 | 338 |
qc_term [] |
339 |
in |
|
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
340 |
fn size => fn nrandom => Option.map snd o compiled_term |
45450 | 341 |
end |
40103 | 342 |
| Depth_Limited_Random => |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
343 |
let |
39471 | 344 |
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
|
345 |
(Pred_Result.get, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result") |
39471 | 346 |
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
|
347 |
g depth nrandom size seed |> (Predicate.map o map) proc) |
39471 | 348 |
qc_term [] |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
349 |
in |
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
350 |
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
|
351 |
(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
|
352 |
end |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
353 |
in |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
354 |
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
|
355 |
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
|
356 |
|
45214 | 357 |
fun try_upto_depth ctxt f = |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
358 |
let |
45214 | 359 |
val max_depth = Config.get ctxt Quickcheck.depth |
360 |
fun message s = if Config.get ctxt Quickcheck.quiet then () else Output.urgent_message s |
|
361 |
fun try' i = |
|
362 |
if i <= max_depth then |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
363 |
let |
45214 | 364 |
val _ = message ("Depth: " ^ string_of_int i) |
365 |
val (result, time) = |
|
366 |
cpu_time ("Depth " ^ string_of_int i) (fn () => |
|
367 |
f i handle Match => (if Config.get ctxt Quickcheck.quiet then () |
|
368 |
else warning "Exception Match raised during quickcheck"; NONE)) |
|
369 |
val _ = if Config.get ctxt Quickcheck.timing then |
|
370 |
message (fst time ^ ": " ^ string_of_int (snd time) ^ " ms") else () |
|
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
371 |
in |
45214 | 372 |
case result of NONE => try' (i + 1) | SOME q => SOME q |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
373 |
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
|
374 |
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
|
375 |
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
|
376 |
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
|
377 |
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
|
378 |
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
|
379 |
|
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_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
|
380 |
(* 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
|
381 |
|
45214 | 382 |
fun compile_term' compilation options 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:
34948
diff
changeset
|
383 |
let |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
384 |
val size = Config.get ctxt Quickcheck.size |
36025
d25043e7843f
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents:
35886
diff
changeset
|
385 |
val c = compile_term compilation options ctxt t |
45214 | 386 |
val counterexample = try_upto_depth ctxt (c size (!nrandom)) |
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
|
387 |
in |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
388 |
Quickcheck.Result |
45750
17100f4ce0b5
adding genuine flag to predicate_compile_quickcheck and prolog_quickcheck (cf. 5e46c225370e);
bulwahn
parents:
45484
diff
changeset
|
389 |
{counterexample = Option.map (pair true o (curry (op ~~)) (Term.add_free_names t [])) counterexample, |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
390 |
evaluation_terms = Option.map (K []) counterexample, timings = [], reports = []} |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
391 |
end |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
392 |
|
45214 | 393 |
fun quickcheck_compile_term compilation function_flattening fail_safe_function_flattening = |
35324
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
394 |
let |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
395 |
val options = |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
396 |
set_fail_safe_function_flattening fail_safe_function_flattening |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
397 |
(set_function_flattening function_flattening (get_options ())) |
c9f428269b38
adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents:
34948
diff
changeset
|
398 |
in |
45214 | 399 |
compile_term' compilation options |
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
400 |
end |
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
401 |
|
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
402 |
|
45450 | 403 |
fun test_goals options ctxt catch_code_errors insts goals = |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
404 |
let |
45450 | 405 |
val (compilation, fail_safe_function_flattening) = options |
406 |
val function_flattening = Config.get ctxt (Quickcheck.allow_function_inversion) |
|
45214 | 407 |
val correct_inst_goals = Quickcheck_Common.instantiate_goals ctxt insts goals |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
408 |
val test_term = |
45214 | 409 |
quickcheck_compile_term compilation function_flattening fail_safe_function_flattening |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
410 |
in |
45214 | 411 |
Quickcheck_Common.collect_results (test_term ctxt) |
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
412 |
(maps (map snd) correct_inst_goals) [] |
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
413 |
end |
45450 | 414 |
|
415 |
val smart_exhaustive_active = Attrib.setup_config_bool @{binding quickcheck_smart_exhaustive_active} (K true); |
|
416 |
val smart_slow_exhaustive_active = Attrib.setup_config_bool @{binding quickcheck_slow_smart_exhaustive_active} (K false); |
|
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
417 |
|
45484
23871e17dddb
setting up exhaustive generators which are used for the smart generators
bulwahn
parents:
45461
diff
changeset
|
418 |
val setup = |
23871e17dddb
setting up exhaustive generators which are used for the smart generators
bulwahn
parents:
45461
diff
changeset
|
419 |
Exhaustive_Generators.setup_exhaustive_datatype_interpretation |
23871e17dddb
setting up exhaustive generators which are used for the smart generators
bulwahn
parents:
45461
diff
changeset
|
420 |
#> Context.theory_map (Quickcheck.add_tester ("smart_exhaustive", |
45450 | 421 |
(smart_exhaustive_active, test_goals (Predicate_Compile_Aux.Pos_Generator_CPS, false)))) |
422 |
#> Context.theory_map (Quickcheck.add_tester ("smart_slow_exhaustive", |
|
423 |
(smart_slow_exhaustive_active, test_goals (Predicate_Compile_Aux.Pos_Generator_DSeq, false)))) |
|
43886
bf068e758783
adapting quickcheck based on the analysis of the predicate compiler
bulwahn
parents:
42361
diff
changeset
|
424 |
|
33250
5c2af18a3237
including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff
changeset
|
425 |
end; |