src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML
author bulwahn
Fri, 22 Oct 2010 18:38:59 +0200
changeset 40103 ef73a90ab6e6
parent 40054 cd7b1fa20bce
child 40132 7ee65dbffa31
permissions -rw-r--r--
adding generator 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*)
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
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    19
  val put_new_dseq_result : (unit -> int -> term list Lazy_Sequence.lazy_sequence) ->
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    20
    Proof.context -> Proof.context
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    21
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    22
  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
    23
  val quiet : bool Unsynchronized.ref;
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    24
  val quickcheck_compile_term : Predicate_Compile_Aux.compilation -> bool -> bool -> int ->
39253
0c47d615a69b removing report from the arguments of the quickcheck functions and refering to it by picking it from the context
bulwahn
parents: 38797
diff changeset
    25
    Proof.context -> term -> int -> term list option * (bool list * bool);
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    26
(*  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
    27
  val nrandom : int Unsynchronized.ref;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    28
  val debug : bool Unsynchronized.ref;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    29
  val function_flattening : bool Unsynchronized.ref;
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
    30
  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
    31
end;
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    32
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    33
structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    34
struct
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    35
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    36
open Predicate_Compile_Aux;
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    37
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    38
structure Pred_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    39
  type T = unit -> int -> int -> int -> int * int -> term list Predicate.pred
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    40
  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
    41
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    42
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
    43
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    44
structure Dseq_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    45
  type T = unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    46
  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
    47
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    48
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
    49
39388
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    50
structure Lseq_Result = Proof_Data (
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    51
  type T = unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    52
  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
    53
);
fdbb2c55ffc2 replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents: 39253
diff changeset
    54
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
    55
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    56
structure New_Dseq_Result = Proof_Data (
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    57
  type T = unit -> int -> term list Lazy_Sequence.lazy_sequence
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    58
  fun init _ () = error "New_Dseq_Random_Result"
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    59
);
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    60
val put_new_dseq_result = New_Dseq_Result.put;
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
    61
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
    62
val tracing = Unsynchronized.ref false;
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    63
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
    64
val quiet = Unsynchronized.ref true;
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
    65
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
    66
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
    67
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
    68
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
    69
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
    70
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
    71
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype 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
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
    73
35538
94170181a842 made smlnj happy
bulwahn
parents: 35537
diff changeset
    74
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
    75
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    76
val options = Options {
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
    77
  expected_modes = NONE,
39382
c797f3ab2ae1 proposed modes for code_pred now supports modes for mutual predicates
bulwahn
parents: 39253
diff changeset
    78
  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
    79
  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
    80
  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
    81
  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
    82
  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
    83
  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
    84
  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
    85
  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
    86
  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
    87
  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
    88
  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
    89
  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
    90
  inductify = true,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
    91
  specialise = true,
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
    92
  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
    93
  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
    94
  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
    95
  no_higher_order_predicate = [],
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
    96
  smart_depth_limiting = true,
36250
ad558b642a15 switched off no_topmost_reordering
bulwahn
parents: 36046
diff changeset
    97
  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
    98
}
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype 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
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype 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
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
   101
  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
   102
  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
   103
  proposed_names = [],
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
   104
  show_steps = true,
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
   105
  show_intermediate_results = true,
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
   106
  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
   107
  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
   108
  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
   109
  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
   110
  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
   111
  show_invalid_clauses = false,
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
   112
  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
   113
  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
   114
  inductify = true,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   115
  specialise = true,
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   116
  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
   117
  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
   118
  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
   119
  no_higher_order_predicate = [],
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
   120
  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
   121
  no_topmost_reordering = true
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
   122
}
95186fb5653c added examples for quickcheck prototype
bulwahn
parents: 33256
diff changeset
   123
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
   124
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype 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
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
   126
  (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
   127
    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
   128
    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
   129
    show_invalid_clauses = s_ic, skip_proof = s_p,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   130
    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
   131
    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
   132
    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
   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,
39383
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   135
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   136
    show_invalid_clauses = s_ic, skip_proof = s_p,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   137
    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
   138
    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
   139
    smart_depth_limiting = sm_dl, no_topmost_reordering = re})
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   140
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype 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
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
   142
  (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
   143
    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
   144
    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
   145
    show_invalid_clauses = s_ic, skip_proof = s_p,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   146
    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
   147
    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
   148
    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
   149
  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   150
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
39383
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   151
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   152
    show_invalid_clauses = s_ic, skip_proof = s_p,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   153
    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
   154
    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
   155
    smart_depth_limiting = sm_dl, no_topmost_reordering = re})
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   156
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   157
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
   158
  (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
   159
    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
   160
    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
   161
    show_invalid_clauses = s_ic, skip_proof = s_p,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   162
    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
   163
    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
   164
    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
   165
  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   166
    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
39383
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   167
    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
ddfafa97da2f adding option show_invalid_clauses for a more detailed message when modes are not inferred
bulwahn
parents: 39382
diff changeset
   168
    show_invalid_clauses = s_ic, skip_proof = s_p,
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   169
    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
   170
    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
   171
    smart_depth_limiting = sm_dl, no_topmost_reordering = re})
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   172
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   173
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype 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
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
   175
  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
   176
    (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
   177
      (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
   178
36046
c3946372f556 putting compilation setup of predicate compiler in a separate file
bulwahn
parents: 36025
diff changeset
   179
val mk_predT = Predicate_Compile_Aux.mk_predT PredicateCompFuns.compfuns
c3946372f556 putting compilation setup of predicate compiler in a separate file
bulwahn
parents: 36025
diff changeset
   180
val mk_return' = Predicate_Compile_Aux.mk_single PredicateCompFuns.compfuns
c3946372f556 putting compilation setup of predicate compiler in a separate file
bulwahn
parents: 36025
diff changeset
   181
val mk_bind' = Predicate_Compile_Aux.mk_bind PredicateCompFuns.compfuns
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   182
36046
c3946372f556 putting compilation setup of predicate compiler in a separate file
bulwahn
parents: 36025
diff changeset
   183
val mk_randompredT = Predicate_Compile_Aux.mk_predT RandomPredCompFuns.compfuns
c3946372f556 putting compilation setup of predicate compiler in a separate file
bulwahn
parents: 36025
diff changeset
   184
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
   185
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
   186
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
   187
val mk_new_randompredT =
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   188
  Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
   189
val mk_new_return =
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   190
  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
   191
val mk_new_bind =
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   192
  Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   193
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   194
val mk_new_dseqT =
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   195
  Predicate_Compile_Aux.mk_predT New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   196
val mk_gen_return =
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   197
  Predicate_Compile_Aux.mk_single New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   198
val mk_gen_bind =
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   199
  Predicate_Compile_Aux.mk_bind New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   200
  
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   201
39756
6c8e83d94536 consolidated tupled_lambda; moved to structure HOLogic
krauss
parents: 39541
diff changeset
   202
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
   203
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
   204
fun cpu_time description f =
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   205
  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
   206
    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
   207
    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
   208
    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
   209
  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
   210
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   211
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
   212
  let
38755
a37d39fe32f8 standardized Context.copy_thy to Theory.copy alias, with slightly more direct way of using it;
wenzelm
parents: 38549
diff changeset
   213
    val thy = Theory.copy (ProofContext.theory_of ctxt)
39541
6605c1e87c7f removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents: 39471
diff changeset
   214
    val ((((full_constname, constT), vs'), intro), thy1) =
6605c1e87c7f removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents: 39471
diff changeset
   215
      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
   216
    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   217
    val (thy3, preproc_time) = cpu_time "predicate preprocessing"
39541
6605c1e87c7f removing clone in code_prolog and predicate_compile_quickcheck
bulwahn
parents: 39471
diff changeset
   218
        (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
   219
    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
   220
        (fn () =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   221
          case compilation of
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   222
            Pos_Random_DSeq =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   223
              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
   224
          | New_Pos_Random_DSeq =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   225
              Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   226
          | Pos_Generator_DSeq =>
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   227
              Predicate_Compile_Core.add_generator_dseq_equations options [full_constname] thy3
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   228
          (*| Depth_Limited_Random =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   229
              Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*))
36256
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   230
    (*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*)
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   231
    val _ = Output.tracing ("Preprocessing time: " ^ string_of_int (snd preproc_time))
d1d9dee7a4bf adopting quickcheck
bulwahn
parents: 36250
diff changeset
   232
    val _ = Output.tracing ("Core compilation time: " ^ string_of_int (snd core_comp_time))
37005
842a73dc6d0e changing compilation to work only with contexts; adapting quickcheck
bulwahn
parents: 36610
diff changeset
   233
    val ctxt4 = ProofContext.init_global thy4
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
   234
    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
   235
    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
   236
    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
   237
      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
   238
        let
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39756
diff changeset
   239
          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
   240
          val T = 
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   241
            case compilation of
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   242
              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
   243
            | New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs'))
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   244
            | 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
   245
            | Depth_Limited_Random =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   246
              [@{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
   247
              @{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
   248
        in
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 34028
diff changeset
   249
          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
   250
        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
   251
      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
   252
    val qc_term =
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   253
      case compilation of
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   254
          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
   255
            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
   256
            (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
   257
        | 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
   258
            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
   259
            (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   260
        | Pos_Generator_DSeq => mk_gen_bind (prog,
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   261
            mk_split_lambda (map Free vs') (mk_gen_return (HOLogic.mk_list @{typ term}
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   262
            (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   263
        | 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
   264
            [@{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
   265
             @{typ "code_numeral * code_numeral"}]
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   266
            (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
   267
            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
   268
            (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
   269
    val prog =
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 =>
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 compiled_term =
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   274
              Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result")
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   275
                thy4 (SOME target)
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   276
                (fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc)
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   277
                qc_term []
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   278
          in
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   279
            (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
   280
              Option.map fst (DSequence.yield
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   281
                (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
   282
          end
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   283
      | New_Pos_Random_DSeq =>
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   284
          let
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   285
            val compiled_term =
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   286
              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
   287
                (Lseq_Result.get, put_lseq_result, "Predicate_Compile_Quickcheck.put_lseq_result")
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   288
                thy4 (SOME target)
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   289
                (fn proc => fn g => fn nrandom => fn size => fn s => fn depth =>
36533
f8df589ca2a5 dropped unnecessary ML code
haftmann
parents: 36256
diff changeset
   290
                  g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc)
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   291
                  qc_term []
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   292
          in
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   293
            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
   294
               (
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   295
               let
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   296
                 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
   297
               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
   298
          end
40103
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   299
      | Pos_Generator_DSeq =>
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   300
          let
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   301
            val compiled_term =
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   302
              Code_Runtime.dynamic_value_strict
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   303
                (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Quickcheck.put_new_dseq_result")
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   304
                thy4 (SOME target)
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   305
                (fn proc => fn g => fn depth => g depth |> (Lazy_Sequence.mapa o map) proc)
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   306
                qc_term []
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   307
          in
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   308
            fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield (compiled_term depth))
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   309
          end
ef73a90ab6e6 adding generator quickcheck
bulwahn
parents: 40054
diff changeset
   310
       | Depth_Limited_Random =>
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   311
          let
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   312
            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
   313
              (Pred_Result.get, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result")
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   314
                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
   315
                  g depth nrandom size seed |> (Predicate.map o map) proc)
39471
55e0ff582fa4 adjusted to changes in Code_Runtime
haftmann
parents: 39403
diff changeset
   316
                qc_term []
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   317
          in
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   318
            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
   319
              (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
   320
          end
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   321
  in
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   322
    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
   323
  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
   324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   325
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
   326
  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
   327
    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
   328
      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
   329
        let
35380
6ac5b81a763d adopting Mutabelle to quickcheck reporting; improving quickcheck reporting
bulwahn
parents: 35378
diff changeset
   330
          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
   331
        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
   332
          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
   333
             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
   334
          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
   335
        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
   336
      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
   337
        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
   338
  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
   339
    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
   340
  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
   341
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   342
(* 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
   343
39253
0c47d615a69b removing report from the arguments of the quickcheck functions and refering to it by picking it from the context
bulwahn
parents: 38797
diff changeset
   344
fun compile_term' compilation options depth ctxt t =
35324
c9f428269b38 adopting mutabelle and quickcheck to return timing information; exporting make_case_combs in datatype package for predicate compiler; adding Spec_Rules declaration for tail recursive functions; improving the predicate compiler and function flattening
bulwahn
parents: 34948
diff changeset
   345
  let
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   346
    val c = compile_term compilation options ctxt t
35378
95d0e3adf38e added basic reporting of test cases to quickcheck
bulwahn
parents: 35324
diff changeset
   347
    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
   348
  in
35537
59dd6be5834c adding depth to predicate compile quickcheck for mutabelle tests; removing obsolete references in predicate compile quickcheck
bulwahn
parents: 35380
diff changeset
   349
    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
   350
  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
   351
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   352
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
   353
  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
   354
     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
   355
       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
   356
         (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
   357
  in
36025
d25043e7843f made quickcheck generic with respect to which compilation; added random compilation to quickcheck
bulwahn
parents: 35886
diff changeset
   358
    compile_term' compilation options depth
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   359
  end
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   360
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents:
diff changeset
   361
end;