src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML
author bulwahn
Mon, 29 Mar 2010 17:30:43 +0200
changeset 36025 d25043e7843f
parent 35886 f5422b736c44
child 36046 c3946372f556
permissions -rw-r--r--
made quickcheck generic with respect to which compilation; added random compilation to quickcheck
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33265
01c9c6dbd890 proper headers;
wenzelm
parents: 33257
diff changeset
     1
(*  Title:      HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML
01c9c6dbd890 proper headers;
wenzelm
parents: 33257
diff changeset
     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
01c9c6dbd890 proper headers;
wenzelm
parents: 33257
diff changeset
     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*)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    10
  val test_ref :
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    11
    ((unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option) Unsynchronized.ref
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    12
  val new_test_ref :
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    13
    ((unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence) option) Unsynchronized.ref;
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    14
  val eval_random_ref :
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    15
    ((unit -> int -> int -> int -> int * int -> term list Predicate.pred) option) Unsynchronized.ref;
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    16
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    17
  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
    18
  val quiet : bool Unsynchronized.ref;
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    19
  val quickcheck_compile_term : Predicate_Compile_Aux.compilation -> bool -> bool -> int ->
35378
95d0e3adf38e added basic reporting of test cases to quickcheck
bulwahn
parents: 35324
diff changeset
    20
    Proof.context -> bool -> term -> int -> term list option * (bool list * bool);
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    21
(*  val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    22
  val nrandom : int Unsynchronized.ref;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    23
  val debug : bool Unsynchronized.ref;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    24
  val function_flattening : bool Unsynchronized.ref;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    25
  val no_higher_order_predicate : string list Unsynchronized.ref;
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    26
end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    27
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    28
structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    29
struct
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    30
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    31
open Predicate_Compile_Aux;
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    32
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    33
val test_ref =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    34
  Unsynchronized.ref (NONE : (unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option);
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    35
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    36
val new_test_ref =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    37
  Unsynchronized.ref (NONE : (unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence) option)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    38
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    39
val eval_random_ref =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    40
  Unsynchronized.ref (NONE : (unit -> int -> int -> int -> int * int -> term list Predicate.pred) option);
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    41
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    42
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    43
val tracing = Unsynchronized.ref false;
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    44
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
    45
val quiet = Unsynchronized.ref true;
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    46
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
    47
val target = "Quickcheck"
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    48
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    49
val nrandom = Unsynchronized.ref 3;
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    50
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
    51
val debug = Unsynchronized.ref false;
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    52
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    53
val function_flattening = Unsynchronized.ref true;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    54
35538
94170181a842 made smlnj happy
bulwahn
parents: 35537
diff changeset
    55
val no_higher_order_predicate = Unsynchronized.ref ([] : string list);
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    56
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    57
val options = Options {
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    58
  expected_modes = NONE,
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33651
diff changeset
    59
  proposed_modes = NONE,
33651
e4aad90618ad adding the predicate compiler quickcheck to the ex/ROOT.ML; adopting this quickcheck to the latest changes
bulwahn
parents: 33486
diff changeset
    60
  proposed_names = [],
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    61
  show_steps = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    62
  show_intermediate_results = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    63
  show_proof_trace = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    64
  show_modes = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    65
  show_mode_inference = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    66
  show_compilation = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    67
  show_caught_failures = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    68
  skip_proof = false,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    69
  compilation = Random,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    70
  inductify = 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
    71
  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
    72
  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
    73
  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
    74
  no_topmost_reordering = 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
    75
}
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
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
    78
  expected_modes = NONE,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    79
  proposed_modes = NONE,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    80
  proposed_names = [],
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    81
  show_steps = true,
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    82
  show_intermediate_results = true,
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    83
  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
    84
  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
    85
  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
    86
  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
    87
  show_caught_failures = true,
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    88
  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
    89
  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
    90
  inductify = 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
    91
  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
    92
  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
    93
  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
    94
  no_topmost_reordering = true
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    95
}
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    96
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
    97
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
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
    99
  (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
   100
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   101
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
   102
    compilation = c, inductify = i, function_flattening = f_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
   103
    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
   104
    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
   105
  (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
   106
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   107
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
    compilation = c, inductify = i, 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
   109
    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
   110
    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
   111
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
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
   113
  (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
   114
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   115
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
    compilation = c, inductify = i, function_flattening = f_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
   117
    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
   118
    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
   119
  (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
   120
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   121
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
    compilation = c, inductify = i, function_flattening = f_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
   123
    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
   124
    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
   125
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
   126
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
   127
  (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
   128
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   129
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
    compilation = c, inductify = i, function_flattening = f_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
   131
    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
   132
    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
   133
  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   134
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   135
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
    compilation = c, inductify = i, function_flattening = f_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
   137
    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
   138
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
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
   141
  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
   142
    (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
   143
      (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
   144
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   145
fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   146
val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns)
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   147
val mk_return' = #mk_single (dest_compfuns Predicate_Compile_Core.pred_compfuns)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   148
val mk_bind' = #mk_bind (dest_compfuns Predicate_Compile_Core.pred_compfuns)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   149
33256
b350516cb1f9 adding a prototype of a counter-example generator based on the predicate compiler to HOL/ex
bulwahn
parents: 33250
diff changeset
   150
val mk_randompredT = #mk_predT (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
b350516cb1f9 adding a prototype of a counter-example generator based on the predicate compiler to HOL/ex
bulwahn
parents: 33250
diff changeset
   151
val mk_return = #mk_single (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
b350516cb1f9 adding a prototype of a counter-example generator based on the predicate compiler to HOL/ex
bulwahn
parents: 33250
diff changeset
   152
val mk_bind = #mk_bind (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   153
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   154
val mk_new_randompredT = #mk_predT (dest_compfuns Predicate_Compile_Core.new_randompred_compfuns)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   155
val mk_new_return = #mk_single (dest_compfuns Predicate_Compile_Core.new_randompred_compfuns)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   156
val mk_new_bind = #mk_bind (dest_compfuns Predicate_Compile_Core.new_randompred_compfuns)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   157
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   158
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   159
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
   160
  | 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
   161
  | mk_split_lambda xs t =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   162
  let
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   163
    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
   164
      | 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
   165
  in
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   166
    mk_split_lambda' xs t
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   167
  end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   168
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   169
fun strip_imp_prems (Const("op -->", _) $ A $ B) = A :: strip_imp_prems B
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   170
  | strip_imp_prems _ = [];
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   171
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   172
fun strip_imp_concl (Const("op -->", _) $ A $ B) = strip_imp_concl B
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   173
  | strip_imp_concl A = A : term;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   174
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   175
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   176
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
   177
fun cpu_time description f =
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   178
  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
   179
    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
   180
    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
   181
    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
   182
  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
   183
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   184
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
   185
  let
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   186
    val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   187
    val thy = (ProofContext.theory_of ctxt') 
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   188
    val (vs, t') = strip_abs t
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   189
    val vs' = Variable.variant_frees ctxt' [] vs
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   190
    val t'' = subst_bounds (map Free (rev vs'), t')
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   191
    val (prems, concl) = strip_horn t''
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   192
    val constname = "pred_compile_quickcheck"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   193
    val full_constname = Sign.full_bname thy constname
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   194
    val constT = map snd vs' ---> @{typ bool}
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   195
    val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   196
    val const = Const (full_constname, constT)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   197
    val t = Logic.list_implies
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   198
      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   199
       HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   200
    val tac = fn _ => Skip_Proof.cheat_tac thy1
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   201
    val intro = Goal.prove (ProofContext.init thy1) (map fst vs') [] t tac
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   202
    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
   203
    val (thy3, preproc_time) =  cpu_time "predicate preprocessing"
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
   204
        (fn () => Predicate_Compile.preprocess options const thy2)
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 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*))
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   214
    val _ = Predicate_Compile_Core.print_all_modes compilation thy4
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   215
    val modes = Predicate_Compile_Core.modes_of compilation thy4 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
   216
    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
   217
    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
   218
      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
   219
        let
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   220
          val name = Predicate_Compile_Core.function_name_of compilation thy4
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   221
            full_constname output_mode
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   222
          val T = 
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   223
            case compilation of
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   224
              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
   225
            | 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
   226
            | Depth_Limited_Random =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   227
              [@{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
   228
              @{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
   229
        in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   230
          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
   231
        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
   232
      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
   233
    
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   234
    val qc_term =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   235
      case compilation of
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   236
          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
   237
            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
   238
            (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
   239
        | 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
   240
            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
   241
            (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   242
        | 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
   243
            [@{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
   244
             @{typ "code_numeral * code_numeral"}]
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   245
            (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
   246
            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
   247
            (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
   248
    val prog =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   249
      case compilation of
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   250
        Pos_Random_DSeq =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   251
          let
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   252
            val compiled_term =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   253
              Code_Eval.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   254
                (fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   255
                thy4 qc_term []
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   256
          in
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   257
            (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
   258
              Option.map fst (DSequence.yield
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   259
                (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
   260
          end
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   261
      | New_Pos_Random_DSeq =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   262
          let
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   263
            val compiled_term =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   264
              Code_Eval.eval (SOME target)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   265
                ("Predicate_Compile_Quickcheck.new_test_ref", new_test_ref)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   266
                (fn proc => fn g => fn nrandom => fn size => fn s => fn depth =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   267
                  g nrandom size s depth |> (Lazy_Sequence.map o map) proc)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   268
                  thy4 qc_term []
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   269
          in
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   270
            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
   271
               (
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   272
               let
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   273
                 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
   274
               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
   275
          end
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   276
      | Depth_Limited_Random =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   277
          let
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   278
            val compiled_term = Code_Eval.eval (SOME target)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   279
              ("Predicate_Compile_Quickcheck.eval_random_ref", eval_random_ref)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   280
                (fn proc => fn g => fn depth => fn nrandom => fn size => fn seed =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   281
                  g depth nrandom size seed |> (Predicate.map o map) proc)
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   282
                thy4 qc_term []
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   283
          in
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   284
            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
   285
              (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
   286
          end
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   287
  in
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   288
    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
   289
  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
   290
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
   291
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
   292
  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
   293
    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
   294
      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
   295
        let
35380
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35378
diff changeset
   296
          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
   297
        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
   298
          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
   299
             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
   300
          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
   301
        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
   302
      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
   303
        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
  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
   305
    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
   306
  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
   307
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_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
(* 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
   309
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   310
fun compile_term' compilation options depth ctxt report t =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   311
  let
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   312
    val c = compile_term compilation options ctxt t
35378
95d0e3adf38e added basic reporting of test cases to quickcheck
bulwahn
parents: 35324
diff changeset
   313
    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
   314
  in
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
   315
    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
   316
  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
   317
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   318
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
   319
  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
   320
     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
   321
       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
   322
         (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
   323
  in
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   324
    compile_term' compilation options depth
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   325
  end
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   326
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   327
end;