src/HOL/Tools/Quickcheck/quickcheck_common.ML
author wenzelm
Fri, 11 Apr 2014 23:26:31 +0200
changeset 56547 e9bb73d7b6cf
parent 56375 32e0da92c786
child 57996 ca917ea6969c
permissions -rw-r--r--
added spell-checker based on jortho-1.0;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Quickcheck/quickcheck_common.ML
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     2
    Author:     Florian Haftmann, Lukas Bulwahn, TU Muenchen
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     3
41938
645cca858c69 tuned headers;
wenzelm
parents: 41935
diff changeset
     4
Common functions for quickcheck's generators.
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     5
*)
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     6
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     7
signature QUICKCHECK_COMMON =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
     8
sig
42195
1e7b62c93f5d adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents: 42159
diff changeset
     9
  val strip_imp : term -> (term list * term)
45721
d1fb55c2ed65 quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
bulwahn
parents: 45719
diff changeset
    10
  val reflect_bool : bool -> term
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    11
  val define_functions : ((term list -> term list) * (Proof.context -> tactic) option)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    12
    -> string -> string list -> string list -> typ list -> Proof.context -> Proof.context 
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    13
  val perhaps_constrain: theory -> (typ * sort) list -> (string * sort) list
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    14
    -> (string * sort -> string * sort) option
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    15
  val instantiate_goals: Proof.context -> (string * typ) list -> (term * term list) list
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    16
    -> (typ option * (term * term list)) list list
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
    17
  val register_predicate : term * string -> Context.generic -> Context.generic
45763
3bb2bdf654f7 random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
bulwahn
parents: 45761
diff changeset
    18
  val mk_safe_if : term -> term -> term * term * (bool -> term) -> bool -> term
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    19
  val collect_results : ('a -> Quickcheck.result) -> 'a list -> Quickcheck.result list -> Quickcheck.result list
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    20
  type result = (bool * term list) option * Quickcheck.report option
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    21
  type generator = string * ((theory -> typ list -> bool) * 
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    22
      (Proof.context -> (term * term list) list -> bool -> int list -> result))
45420
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
    23
  val generator_test_goal_terms :
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    24
    generator -> Proof.context -> bool -> (string * typ) list
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    25
      -> (term * term list) list -> Quickcheck.result list
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    26
  type instantiation = Datatype.config -> Datatype.descr -> (string * sort) list
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    27
     -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    28
  val ensure_sort :
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    29
    (((sort * sort) * sort) *
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    30
      ((theory -> string list -> Datatype_Aux.descr * (string * sort) list * string list
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    31
        * string * (string list * string list) * (typ list * typ list)) * instantiation))
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    32
    -> Datatype.config -> string list -> theory -> theory
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    33
  val ensure_common_sort_datatype :
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    34
    (sort * instantiation) -> Datatype.config -> string list -> theory -> theory
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
    35
  val datatype_interpretation : (sort * instantiation) -> theory -> theory
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
    36
  val gen_mk_parametric_generator_expr :
42229
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
    37
   (((Proof.context -> term * term list -> term) * term) * typ)
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
    38
     -> Proof.context -> (term * term list) list -> term
45039
632a033616e9 adding post-processing of terms to narrowing-based Quickcheck
bulwahn
parents: 44241
diff changeset
    39
  val mk_fun_upd : typ -> typ -> term * term -> term -> term
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
    40
  val post_process_term : term -> term
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    41
  val test_term : generator -> Proof.context -> bool -> term * term list -> Quickcheck.result
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    42
end;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    43
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    44
structure Quickcheck_Common : QUICKCHECK_COMMON =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    45
struct
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
    46
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    47
(* static options *)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    48
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    49
val define_foundationally = false
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    50
42195
1e7b62c93f5d adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents: 42159
diff changeset
    51
(* HOLogic's term functions *)
1e7b62c93f5d adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents: 42159
diff changeset
    52
1e7b62c93f5d adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents: 42159
diff changeset
    53
fun strip_imp (Const(@{const_name HOL.implies},_) $ A $ B) = apfst (cons A) (strip_imp B)
1e7b62c93f5d adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents: 42159
diff changeset
    54
  | strip_imp A = ([], A)
1e7b62c93f5d adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents: 42159
diff changeset
    55
45721
d1fb55c2ed65 quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
bulwahn
parents: 45719
diff changeset
    56
fun reflect_bool b = if b then @{term "True"} else @{term "False"}
d1fb55c2ed65 quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
bulwahn
parents: 45719
diff changeset
    57
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
    58
fun mk_undefined T = Const(@{const_name undefined}, T)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    59
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    60
(* testing functions: testing with increasing sizes (and cardinalities) *)
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    61
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    62
type result = (bool * term list) option * Quickcheck.report option
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    63
type generator = string * ((theory -> typ list -> bool) * 
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    64
      (Proof.context -> (term * term list) list -> bool -> int list -> result))
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    65
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    66
fun check_test_term t =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    67
  let
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    68
    val _ = (null (Term.add_tvars t []) andalso null (Term.add_tfrees t [])) orelse
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    69
      error "Term to be tested contains type variables";
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    70
    val _ = null (Term.add_vars t []) orelse
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    71
      error "Term to be tested contains schematic variables";
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    72
  in () end
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    73
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    74
fun cpu_time description e =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    75
  let val ({cpu, ...}, result) = Timing.timing e ()
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    76
  in (result, (description, Time.toMilliseconds cpu)) end
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    77
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
    78
fun test_term (name, (_, compile)) ctxt catch_code_errors (t, eval_terms) =
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    79
  let
45757
e32dd098f57a renaming potential flag to genuine_only flag with an inverse semantics
bulwahn
parents: 45755
diff changeset
    80
    val genuine_only = Config.get ctxt Quickcheck.genuine_only
46478
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
    81
    val abort_potential = Config.get ctxt Quickcheck.abort_potential
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    82
    val _ = check_test_term t
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    83
    val names = Term.add_free_names t []
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    84
    val current_size = Unsynchronized.ref 0
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    85
    val current_result = Unsynchronized.ref Quickcheck.empty_result 
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    86
    val act = if catch_code_errors then try else (fn f => SOME o f) 
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    87
    val (test_fun, comp_time) = cpu_time "quickcheck compilation"
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    88
        (fn () => act (compile ctxt) [(t, eval_terms)]);
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    89
    val _ = Quickcheck.add_timing comp_time current_result
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    90
    fun with_size test_fun genuine_only k =
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    91
      if k > Config.get ctxt Quickcheck.size then
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    92
        NONE
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    93
      else
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    94
        let
45765
cb6ddee6a463 making the default behaviour of quickcheck a little bit less verbose;
bulwahn
parents: 45763
diff changeset
    95
          val _ = Quickcheck.verbose_message ctxt
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    96
            ("[Quickcheck-" ^ name ^ "] Test data size: " ^ string_of_int k)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
    97
          val _ = current_size := k
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
    98
          val ((result, report), time) =
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
    99
            cpu_time ("size " ^ string_of_int k) (fn () => test_fun genuine_only [1, k - 1])
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   100
          val _ = if Config.get ctxt Quickcheck.timing then
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   101
            Quickcheck.verbose_message ctxt (fst time ^ ": " ^ string_of_int (snd time) ^ " ms") else ()
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   102
          val _ = Quickcheck.add_timing time current_result
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   103
          val _ = Quickcheck.add_report k report current_result
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   104
        in
45753
196697f71488 indicating where the restart should occur; making safe_if dynamic
bulwahn
parents: 45728
diff changeset
   105
          case result of
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
   106
            NONE => with_size test_fun genuine_only (k + 1)
45753
196697f71488 indicating where the restart should occur; making safe_if dynamic
bulwahn
parents: 45728
diff changeset
   107
          | SOME (true, ts) => SOME (true, ts)
45755
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   108
          | SOME (false, ts) =>
46478
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   109
            if abort_potential then SOME (false, ts)
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   110
            else
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   111
              let
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   112
                val (ts1, ts2) = chop (length names) ts
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   113
                val (eval_terms', _) = chop (length ts2) eval_terms
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   114
                val cex = SOME ((false, names ~~ ts1), eval_terms' ~~ ts2)
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   115
              in
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   116
                (Quickcheck.message ctxt (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex));
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   117
                Quickcheck.message ctxt "Quickcheck continues to find a genuine counterexample...";
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   118
                with_size test_fun true k)
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   119
              end
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   120
        end;
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   121
  in
45417
cae3ba9be892 removing deactivated timeout handling; catching compilation errors and only outputing an urgent message to enable parallel sucessful quickcheck compilations and runs to present their result
bulwahn
parents: 45416
diff changeset
   122
    case test_fun of
45420
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
   123
      NONE => (Quickcheck.message ctxt ("Conjecture is not executable with Quickcheck-" ^ name);
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
   124
        !current_result)
45417
cae3ba9be892 removing deactivated timeout handling; catching compilation errors and only outputing an urgent message to enable parallel sucessful quickcheck compilations and runs to present their result
bulwahn
parents: 45416
diff changeset
   125
    | SOME test_fun =>
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   126
      let
45819
b85bb803bcf3 tuned quickcheck's response
bulwahn
parents: 45792
diff changeset
   127
        val _ = Quickcheck.message ctxt ("Testing conjecture with Quickcheck-" ^ name ^ "...");
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   128
        val (response, exec_time) =
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
   129
          cpu_time "quickcheck execution" (fn () => with_size test_fun genuine_only 1)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   130
        val _ = Quickcheck.add_response names eval_terms response current_result
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   131
        val _ = Quickcheck.add_timing exec_time current_result
45417
cae3ba9be892 removing deactivated timeout handling; catching compilation errors and only outputing an urgent message to enable parallel sucessful quickcheck compilations and runs to present their result
bulwahn
parents: 45416
diff changeset
   132
      in !current_result end
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   133
  end;
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   134
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   135
fun test_term_with_cardinality (name, (size_matters_for, compile)) ctxt catch_code_errors ts =
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   136
  let
45757
e32dd098f57a renaming potential flag to genuine_only flag with an inverse semantics
bulwahn
parents: 45755
diff changeset
   137
    val genuine_only = Config.get ctxt Quickcheck.genuine_only
46478
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   138
    val abort_potential = Config.get ctxt Quickcheck.abort_potential
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   139
    val thy = Proof_Context.theory_of ctxt
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   140
    val (ts', eval_terms) = split_list ts
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   141
    val _ = map check_test_term ts'
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   142
    val names = Term.add_free_names (hd ts') []
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   143
    val Ts = map snd (Term.add_frees (hd ts') [])
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   144
    val current_result = Unsynchronized.ref Quickcheck.empty_result
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
   145
    fun test_card_size test_fun genuine_only (card, size) =
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   146
      (* FIXME: why decrement size by one? *)
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   147
      let
45686
cf22004ad55d adding more verbose messages to exhaustive quickcheck
bulwahn
parents: 45639
diff changeset
   148
        val _ =
45765
cb6ddee6a463 making the default behaviour of quickcheck a little bit less verbose;
bulwahn
parents: 45763
diff changeset
   149
          Quickcheck.verbose_message ctxt ("[Quickcheck-" ^ name ^ "] Test " ^
46674
bc03b533b061 slightly changing the enumeration scheme
bulwahn
parents: 46580
diff changeset
   150
            (if size = 0 then "" else "data size: " ^ string_of_int size ^ " and ") ^
45686
cf22004ad55d adding more verbose messages to exhaustive quickcheck
bulwahn
parents: 45639
diff changeset
   151
            "cardinality: " ^ string_of_int card)          
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   152
        val (ts, timing) =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   153
          cpu_time ("size " ^ string_of_int size ^ " and card " ^ string_of_int card)
46674
bc03b533b061 slightly changing the enumeration scheme
bulwahn
parents: 46580
diff changeset
   154
            (fn () => fst (test_fun genuine_only [card, size + 1]))
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   155
        val _ = Quickcheck.add_timing timing current_result
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   156
      in
45755
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   157
        Option.map (pair (card, size)) ts
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   158
      end
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   159
    val enumeration_card_size =
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   160
      if size_matters_for thy Ts then
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   161
        map_product pair (1 upto (length ts)) (1 upto (Config.get ctxt Quickcheck.size))
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   162
        |> sort (fn ((c1, s1), (c2, s2)) => int_ord ((c1 + s1), (c2 + s2)))
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   163
      else
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   164
        map (rpair 0) (1 upto (length ts))
45419
10ba32c347b0 quickcheck fails with code generator errors only if one tester is invoked
bulwahn
parents: 45418
diff changeset
   165
    val act = if catch_code_errors then try else (fn f => SOME o f)
10ba32c347b0 quickcheck fails with code generator errors only if one tester is invoked
bulwahn
parents: 45418
diff changeset
   166
    val (test_fun, comp_time) = cpu_time "quickcheck compilation" (fn () => act (compile ctxt) ts)
45417
cae3ba9be892 removing deactivated timeout handling; catching compilation errors and only outputing an urgent message to enable parallel sucessful quickcheck compilations and runs to present their result
bulwahn
parents: 45416
diff changeset
   167
    val _ = Quickcheck.add_timing comp_time current_result
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   168
  in
45417
cae3ba9be892 removing deactivated timeout handling; catching compilation errors and only outputing an urgent message to enable parallel sucessful quickcheck compilations and runs to present their result
bulwahn
parents: 45416
diff changeset
   169
    case test_fun of
45420
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
   170
      NONE => (Quickcheck.message ctxt ("Conjecture is not executable with Quickcheck-" ^ name);
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
   171
        !current_result)
45417
cae3ba9be892 removing deactivated timeout handling; catching compilation errors and only outputing an urgent message to enable parallel sucessful quickcheck compilations and runs to present their result
bulwahn
parents: 45416
diff changeset
   172
    | SOME test_fun =>
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   173
      let
45921
bulwahn
parents: 45819
diff changeset
   174
        val _ = Quickcheck.message ctxt ("Testing conjecture with Quickcheck-" ^ name ^ "...");
45755
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   175
        fun test genuine_only enum = case get_first (test_card_size test_fun genuine_only) enum of
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   176
          SOME ((card, _), (true, ts)) =>
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   177
            Quickcheck.add_response names (nth eval_terms (card - 1)) (SOME (true, ts)) current_result
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   178
        | SOME ((card, size), (false, ts)) =>
46478
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   179
            if abort_potential then
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   180
              Quickcheck.add_response names (nth eval_terms (card - 1)) (SOME (false, ts)) current_result
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   181
            else
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   182
              let
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   183
                val (ts1, ts2) = chop (length names) ts
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   184
                val (eval_terms', _) = chop (length ts2) (nth eval_terms (card - 1))
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   185
                val cex = SOME ((false, names ~~ ts1), eval_terms' ~~ ts2)
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   186
              in
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   187
                (Quickcheck.message ctxt (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex));
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   188
                Quickcheck.message ctxt "Quickcheck continues to find a genuine counterexample...";
cf1bcfb34c82 adding abort_potential functionality in quickcheck
bulwahn
parents: 46331
diff changeset
   189
                test true (snd (take_prefix (fn x => not (x = (card, size))) enum)))              
45755
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   190
            end
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   191
        | NONE => ()
45755
b27a06dfb2ef outputing the potentially spurious counterexample and continue search
bulwahn
parents: 45754
diff changeset
   192
      in (test genuine_only enumeration_card_size; !current_result) end
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   193
  end
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   194
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   195
fun get_finite_types ctxt =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   196
  fst (chop (Config.get ctxt Quickcheck.finite_type_size)
45416
bulwahn
parents: 45159
diff changeset
   197
    [@{typ "Enum.finite_1"}, @{typ "Enum.finite_2"}, @{typ "Enum.finite_3"},
bulwahn
parents: 45159
diff changeset
   198
     @{typ "Enum.finite_4"}, @{typ "Enum.finite_5"}])
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   199
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   200
exception WELLSORTED of string
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   201
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   202
fun monomorphic_term thy insts default_T =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   203
  let
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   204
    fun subst (T as TFree (v, S)) =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   205
      let
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   206
        val T' = AList.lookup (op =) insts v
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   207
          |> the_default default_T
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   208
      in if Sign.of_sort thy (T', S) then T'
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   209
        else raise (WELLSORTED ("For instantiation with default_type " ^
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   210
          Syntax.string_of_typ_global thy default_T ^
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   211
          ":\n" ^ Syntax.string_of_typ_global thy T' ^
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   212
          " to be substituted for variable " ^
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   213
          Syntax.string_of_typ_global thy T ^ " does not have sort " ^
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   214
          Syntax.string_of_sort_global thy S))
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   215
      end
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   216
      | subst T = T;
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   217
  in (map_types o map_atyps) subst end;
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   218
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   219
datatype wellsorted_error = Wellsorted_Error of string | Term of term * term list
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   220
45440
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   221
(* minimalistic preprocessing *)
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   222
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   223
fun strip_all (Const (@{const_name HOL.All}, _) $ Abs (a, T, t)) = 
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   224
  let
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   225
    val (a', t') = strip_all t
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   226
  in ((a, T) :: a', t') end
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   227
  | strip_all t = ([], t);
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   228
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   229
fun preprocess ctxt t =
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   230
  let
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   231
    val thy = Proof_Context.theory_of ctxt
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   232
    val dest = HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   233
    val rewrs = map (swap o dest) @{thms all_simps} @
46327
ecda23528833 adding some rules to quickcheck's preprocessing
bulwahn
parents: 46312
diff changeset
   234
      (map dest [@{thm not_ex}, @{thm not_all}, @{thm imp_conjL}, @{thm fun_eq_iff},
ecda23528833 adding some rules to quickcheck's preprocessing
bulwahn
parents: 46312
diff changeset
   235
        @{thm bot_fun_def}, @{thm less_bool_def}])
45440
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   236
    val t' = Pattern.rewrite_term thy rewrs [] (Object_Logic.atomize_term thy t)
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   237
    val (vs, body) = strip_all t'
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   238
    val vs' = Variable.variant_frees ctxt [t'] vs
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   239
  in
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   240
    subst_bounds (map Free (rev vs'), body)
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   241
  end
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   242
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   243
  
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   244
structure Subtype_Predicates = Generic_Data
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   245
(
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   246
  type T = (term * string) list
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   247
  val empty = []
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   248
  val extend = I
46580
8d32138811cb made SML/NJ happy;
wenzelm
parents: 46565
diff changeset
   249
  fun merge data : T = AList.merge (op =) (K true) data
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   250
)
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   251
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   252
val register_predicate = Subtype_Predicates.map o AList.update (op =)
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   253
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   254
fun subtype_preprocess ctxt (T, (t, ts)) =
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   255
  let
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   256
    val preds = Subtype_Predicates.get (Context.Proof ctxt)
50046
0051dc4f301f dropped dead code;
haftmann
parents: 48272
diff changeset
   257
    fun matches (p $ _) = AList.defined Term.could_unify preds p  
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   258
    fun get_match (p $ x) = Option.map (rpair x) (AList.lookup Term.could_unify preds p)
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   259
    fun subst_of (tyco, v as Free (x, repT)) =
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   260
      let
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   261
        val [(info, _)] = Typedef.get_info ctxt tyco
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   262
        val repT' = Logic.varifyT_global (#rep_type info)
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   263
        val substT = Sign.typ_match (Proof_Context.theory_of ctxt) (repT', repT) Vartab.empty 
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   264
        val absT = Envir.subst_type substT (Logic.varifyT_global (#abs_type info))
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   265
      in (v, Const (#Rep_name info, absT --> repT) $ Free (x, absT)) end
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   266
    val (prems, concl) = strip_imp t
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   267
    val subst = map subst_of (map_filter get_match prems)
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   268
    val t' = Term.subst_free subst
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   269
     (fold_rev (curry HOLogic.mk_imp) (filter_out matches prems) concl)
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   270
  in
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   271
    (T, (t', ts))
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   272
  end
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   273
45440
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   274
(* instantiation of type variables with concrete types *)
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   275
 
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   276
fun instantiate_goals lthy insts goals =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   277
  let
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   278
    fun map_goal_and_eval_terms f (check_goal, eval_terms) = (f check_goal, map f eval_terms)
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   279
    val thy = Proof_Context.theory_of lthy
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   280
    val default_insts =
45420
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
   281
      if Config.get lthy Quickcheck.finite_types then get_finite_types else Quickcheck.default_type
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   282
    val inst_goals =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   283
      map (fn (check_goal, eval_terms) =>
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   284
        if not (null (Term.add_tfree_names check_goal [])) then
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   285
          map (fn T =>
45440
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   286
            (pair (SOME T) o Term o apfst (preprocess lthy))
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   287
              (map_goal_and_eval_terms (monomorphic_term thy insts T) (check_goal, eval_terms))
45420
d17556b9a89b more precise messages with the tester's name in quickcheck; tuned
bulwahn
parents: 45419
diff changeset
   288
              handle WELLSORTED s => (SOME T, Wellsorted_Error s)) (default_insts lthy)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   289
        else
45440
9f4d3e68ae98 adding a minimalistic preprocessing rewriting common boolean operators; tuned
bulwahn
parents: 45420
diff changeset
   290
          [(NONE, Term (preprocess lthy check_goal, eval_terms))]) goals
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   291
    val error_msg =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   292
      cat_lines
50046
0051dc4f301f dropped dead code;
haftmann
parents: 48272
diff changeset
   293
        (maps (map_filter (fn (_, Term _) => NONE | (_, Wellsorted_Error s) => SOME s)) inst_goals)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   294
    fun is_wellsorted_term (T, Term t) = SOME (T, t)
50046
0051dc4f301f dropped dead code;
haftmann
parents: 48272
diff changeset
   295
      | is_wellsorted_term (_, Wellsorted_Error _) = NONE
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   296
    val correct_inst_goals =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   297
      case map (map_filter is_wellsorted_term) inst_goals of
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   298
        [[]] => error error_msg
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   299
      | xs => xs
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   300
    val _ = if Config.get lthy Quickcheck.quiet then () else warning error_msg
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   301
  in
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   302
    correct_inst_goals
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   303
  end
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   304
45718
8979b2463fc8 quickcheck random can also find potential counterexamples;
bulwahn
parents: 45687
diff changeset
   305
(* compilation of testing functions *)
8979b2463fc8 quickcheck random can also find potential counterexamples;
bulwahn
parents: 45687
diff changeset
   306
45763
3bb2bdf654f7 random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
bulwahn
parents: 45761
diff changeset
   307
fun mk_safe_if genuine_only none (cond, then_t, else_t) genuine =
45753
196697f71488 indicating where the restart should occur; making safe_if dynamic
bulwahn
parents: 45728
diff changeset
   308
  let
45763
3bb2bdf654f7 random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
bulwahn
parents: 45761
diff changeset
   309
    val T = fastype_of then_t
45754
394ecd91434a dynamic genuine_flag in compilation of random and exhaustive
bulwahn
parents: 45753
diff changeset
   310
    val if_t = Const (@{const_name "If"}, @{typ bool} --> T --> T --> T)
45753
196697f71488 indicating where the restart should occur; making safe_if dynamic
bulwahn
parents: 45728
diff changeset
   311
  in
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 50046
diff changeset
   312
    Const (@{const_name "Quickcheck_Random.catch_match"}, T --> T --> T) $ 
45761
90028fd2f1fa exhaustive returns if a counterexample is genuine or potentially spurious in the presence of assumptions more correctly
bulwahn
parents: 45757
diff changeset
   313
      (if_t $ cond $ then_t $ else_t genuine) $
45763
3bb2bdf654f7 random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
bulwahn
parents: 45761
diff changeset
   314
      (if_t $ genuine_only $ none $ else_t false)
45753
196697f71488 indicating where the restart should occur; making safe_if dynamic
bulwahn
parents: 45728
diff changeset
   315
  end
45718
8979b2463fc8 quickcheck random can also find potential counterexamples;
bulwahn
parents: 45687
diff changeset
   316
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   317
fun collect_results f [] results = results
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   318
  | collect_results f (t :: ts) results =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   319
    let
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   320
      val result = f t
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   321
    in
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   322
      if Quickcheck.found_counterexample result then
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   323
        (result :: results)
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   324
      else
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   325
        collect_results f ts (result :: results)
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   326
    end  
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   327
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   328
fun generator_test_goal_terms generator ctxt catch_code_errors insts goals =
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   329
  let
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   330
    val use_subtype = Config.get ctxt Quickcheck.use_subtype
45687
a07654c67f30 quickcheck does not show evaluation terms of equations if they are simply free variables to avoid duplicated output; tuned
bulwahn
parents: 45686
diff changeset
   331
    fun add_eval_term t ts = if is_Free t then ts else ts @ [t]
a07654c67f30 quickcheck does not show evaluation terms of equations if they are simply free variables to avoid duplicated output; tuned
bulwahn
parents: 45686
diff changeset
   332
    fun add_equation_eval_terms (t, eval_terms) =
a07654c67f30 quickcheck does not show evaluation terms of equations if they are simply free variables to avoid duplicated output; tuned
bulwahn
parents: 45686
diff changeset
   333
      case try HOLogic.dest_eq (snd (strip_imp t)) of
a07654c67f30 quickcheck does not show evaluation terms of equations if they are simply free variables to avoid duplicated output; tuned
bulwahn
parents: 45686
diff changeset
   334
        SOME (lhs, rhs) => (t, add_eval_term lhs (add_eval_term rhs eval_terms))
a07654c67f30 quickcheck does not show evaluation terms of equations if they are simply free variables to avoid duplicated output; tuned
bulwahn
parents: 45686
diff changeset
   335
      | NONE => (t, eval_terms)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   336
    fun test_term' goal =
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   337
      case goal of
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   338
        [(NONE, t)] => test_term generator ctxt catch_code_errors t
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   339
      | ts => test_term_with_cardinality generator ctxt catch_code_errors (map snd ts)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   340
    val goals' = instantiate_goals ctxt insts goals
46565
ad21900e0ee9 subtype preprocessing in Quickcheck;
bulwahn
parents: 46478
diff changeset
   341
      |> (if use_subtype then map (map (subtype_preprocess ctxt)) else I)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   342
      |> map (map (apsnd add_equation_eval_terms))
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   343
  in
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   344
    if Config.get ctxt Quickcheck.finite_types then
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   345
      collect_results test_term' goals' []
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   346
    else
46331
f5598b604a54 generalizing check if size matters because it is different for random and exhaustive testing
bulwahn
parents: 46327
diff changeset
   347
      collect_results (test_term generator ctxt catch_code_errors)
45159
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   348
        (maps (map snd) goals') []
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   349
  end;
3f1d1ce024cb moving some common functions from quickcheck to the more HOL-specific quickcheck_common; renamed inductive_SML's configurations to more canonical names; adds automatically left and right hand sides of equations as evaluation terms
bulwahn
parents: 45039
diff changeset
   350
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   351
(* defining functions *)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   352
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   353
fun pat_completeness_auto ctxt =
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42361
diff changeset
   354
  Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   355
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   356
fun define_functions (mk_equations, termination_tac) prfx argnames names Ts =
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   357
  if define_foundationally andalso is_some termination_tac then
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   358
    let
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   359
      val eqs_t = mk_equations (map2 (fn name => fn T => Free (name, T)) names Ts)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   360
    in
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   361
      Function.add_function
42287
d98eb048a2e4 discontinued special treatment of structure Mixfix;
wenzelm
parents: 42229
diff changeset
   362
        (map (fn (name, T) => (Binding.conceal (Binding.name name), SOME T, NoSyn))
d98eb048a2e4 discontinued special treatment of structure Mixfix;
wenzelm
parents: 42229
diff changeset
   363
          (names ~~ Ts))
d98eb048a2e4 discontinued special treatment of structure Mixfix;
wenzelm
parents: 42229
diff changeset
   364
        (map (pair (apfst Binding.conceal Attrib.empty_binding)) eqs_t)
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   365
        Function_Common.default_config pat_completeness_auto
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   366
      #> snd
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   367
      #> (fn lthy => Function.prove_termination NONE (the termination_tac lthy) lthy)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   368
      #> snd
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   369
    end
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   370
  else
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   371
    fold_map (fn (name, T) => Local_Theory.define
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   372
        ((Binding.conceal (Binding.name name), NoSyn),
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   373
          (apfst Binding.conceal Attrib.empty_binding, mk_undefined T))
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   374
      #> apfst fst) (names ~~ Ts)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   375
    #> (fn (consts, lthy) =>
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   376
      let
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   377
        val eqs_t = mk_equations consts
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   378
        val eqs = map (fn eq => Goal.prove lthy argnames [] eq
51552
c713c9505f68 clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents: 51143
diff changeset
   379
          (fn _ => ALLGOALS Skip_Proof.cheat_tac)) eqs_t
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   380
      in
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   381
        fold (fn (name, eq) => Local_Theory.note
47204
6212bcc94bb0 refine bindings in quickcheck_common: do not conceal and do not declare as simps
bulwahn
parents: 46674
diff changeset
   382
          ((Binding.qualify true prfx
6212bcc94bb0 refine bindings in quickcheck_common: do not conceal and do not declare as simps
bulwahn
parents: 46674
diff changeset
   383
              (Binding.qualify true name (Binding.name "simps")),
6212bcc94bb0 refine bindings in quickcheck_common: do not conceal and do not declare as simps
bulwahn
parents: 46674
diff changeset
   384
             [Code.add_default_eqn_attrib]), [eq]) #> snd)
45592
8baa0b7f3f66 added ML antiquotation @{attributes};
wenzelm
parents: 45440
diff changeset
   385
          (names ~~ eqs) lthy
42214
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   386
      end)
9ca13615c619 refactoring generator definition in quickcheck and removing clone
bulwahn
parents: 42195
diff changeset
   387
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   388
(** ensuring sort constraints **)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   389
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   390
type instantiation = Datatype.config -> Datatype.descr -> (string * sort) list
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   391
  -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   392
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   393
fun perhaps_constrain thy insts raw_vs =
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   394
  let
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   395
    fun meet (T, sort) = Sorts.meet_sort (Sign.classes_of thy) 
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   396
      (Logic.varifyT_global T, sort);
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   397
    val vtab = Vartab.empty
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   398
      |> fold (fn (v, sort) => Vartab.update ((v, 0), sort)) raw_vs
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   399
      |> fold meet insts;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   400
  in SOME (fn (v, _) => (v, (the o Vartab.lookup vtab) (v, 0)))
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   401
  end handle Sorts.CLASS_ERROR _ => NONE;
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   402
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   403
fun ensure_sort (((sort_vs, aux_sort), sort), (the_descr, instantiate)) config raw_tycos thy =
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   404
  let
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   405
    val algebra = Sign.classes_of thy;
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   406
    val (descr, raw_vs, tycos, prfx, (names, auxnames), raw_TUs) = the_descr thy raw_tycos
42229
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
   407
    val vs = (map o apsnd) (curry (Sorts.inter_sort algebra) sort_vs) raw_vs
56375
32e0da92c786 use same idiom as used for datatype 'size' function to name constants and theorems emerging from various type interpretations -- reduces the chances of name clashes on theory merges
blanchet
parents: 51552
diff changeset
   408
    fun insts_of sort constr = (map (rpair sort) o flat o maps snd o maps snd)
42229
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
   409
      (Datatype_Aux.interpret_construction descr vs constr)
56375
32e0da92c786 use same idiom as used for datatype 'size' function to name constants and theorems emerging from various type interpretations -- reduces the chances of name clashes on theory merges
blanchet
parents: 51552
diff changeset
   410
    val insts = insts_of sort { atyp = single, dtyp = (K o K o K) [] }
42229
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
   411
      @ insts_of aux_sort { atyp = K [], dtyp = K o K }
48272
db75b4005d9a more direct Sorts.has_instance;
wenzelm
parents: 47204
diff changeset
   412
    val has_inst = exists (fn tyco => Sorts.has_instance algebra tyco sort) tycos;
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   413
  in if has_inst then thy
42229
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
   414
    else case perhaps_constrain thy insts vs
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   415
     of SOME constrain => instantiate config descr
42229
1491b7209e76 generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents: 42214
diff changeset
   416
          (map constrain vs) tycos prfx (names, auxnames)
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   417
            ((pairself o map o map_atyps) (fn TFree v => TFree (constrain v)) raw_TUs) thy
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   418
      | NONE => thy
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   419
  end;
45923
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   420
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   421
fun ensure_common_sort_datatype (sort, instantiate) =
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   422
  ensure_sort (((@{sort typerep}, @{sort term_of}), sort), (Datatype.the_descr, instantiate))
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   423
473b744c23f2 generalize ensure_sort_datatype to ensure_sort in quickcheck_common to allow generators for abstract types;
bulwahn
parents: 45921
diff changeset
   424
val datatype_interpretation = Datatype.interpretation o ensure_common_sort_datatype
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   425
  
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   426
(** generic parametric compilation **)
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   427
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   428
fun gen_mk_parametric_generator_expr ((mk_generator_expr, out_of_bounds), T) ctxt ts =
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   429
  let
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   430
    val if_t = Const (@{const_name "If"}, @{typ bool} --> T --> T --> T)
45721
d1fb55c2ed65 quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
bulwahn
parents: 45719
diff changeset
   431
    fun mk_if (index, (t, eval_terms)) else_t = if_t $
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   432
        (HOLogic.eq_const @{typ natural} $ Bound 0 $ HOLogic.mk_number @{typ natural} index) $
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   433
        (mk_generator_expr ctxt (t, eval_terms)) $ else_t
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   434
  in
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51126
diff changeset
   435
    absdummy @{typ natural} (fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds)
42159
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   436
  end
234ec7011e5d generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents: 42110
diff changeset
   437
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   438
(** post-processing of function terms **)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   439
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   440
fun dest_fun_upd (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = (t0, (t1, t2))
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   441
  | dest_fun_upd t = raise TERM ("dest_fun_upd", [t])
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   442
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   443
fun mk_fun_upd T1 T2 (t1, t2) t = 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   444
  Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2) $ t $ t1 $ t2
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   445
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   446
fun dest_fun_upds t =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   447
  case try dest_fun_upd t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   448
    NONE =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   449
      (case t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   450
        Abs (_, _, _) => ([], t) 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   451
      | _ => raise TERM ("dest_fun_upds", [t]))
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   452
  | SOME (t0, (t1, t2)) => apfst (cons (t1, t2)) (dest_fun_upds t0)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   453
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   454
fun make_fun_upds T1 T2 (tps, t) = fold_rev (mk_fun_upd T1 T2) tps t
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   455
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   456
fun make_set T1 [] = Const (@{const_abbrev Set.empty}, T1 --> @{typ bool})
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   457
  | make_set T1 ((_, @{const False}) :: tps) = make_set T1 tps
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   458
  | make_set T1 ((t1, @{const True}) :: tps) =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   459
    Const (@{const_name insert}, T1 --> (T1 --> @{typ bool}) --> T1 --> @{typ bool})
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   460
      $ t1 $ (make_set T1 tps)
50046
0051dc4f301f dropped dead code;
haftmann
parents: 48272
diff changeset
   461
  | make_set T1 ((_, t) :: _) = raise TERM ("make_set", [t])
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   462
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   463
fun make_coset T [] = Const (@{const_abbrev UNIV}, T --> @{typ bool})
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   464
  | make_coset T tps = 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   465
    let
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   466
      val U = T --> @{typ bool}
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   467
      fun invert @{const False} = @{const True}
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   468
        | invert @{const True} = @{const False}
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   469
    in
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   470
      Const (@{const_name "Groups.minus_class.minus"}, U --> U --> U)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   471
        $ Const (@{const_abbrev UNIV}, U) $ make_set T (map (apsnd invert) tps)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   472
    end
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   473
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   474
fun make_map T1 T2 [] = Const (@{const_abbrev Map.empty}, T1 --> T2)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   475
  | make_map T1 T2 ((_, Const (@{const_name None}, _)) :: tps) = make_map T1 T2 tps
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   476
  | make_map T1 T2 ((t1, t2) :: tps) = mk_fun_upd T1 T2 (t1, t2) (make_map T1 T2 tps)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   477
  
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   478
fun post_process_term t =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   479
  let
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   480
    fun map_Abs f t =
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   481
      case t of Abs (x, T, t') => Abs (x, T, f t') | _ => raise TERM ("map_Abs", [t]) 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   482
    fun process_args t = case strip_comb t of
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   483
      (c as Const (_, _), ts) => list_comb (c, map post_process_term ts)
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   484
  in
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   485
    case fastype_of t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   486
      Type (@{type_name fun}, [T1, T2]) =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   487
        (case try dest_fun_upds t of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   488
          SOME (tps, t) =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   489
            (map (pairself post_process_term) tps, map_Abs post_process_term t)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   490
            |> (case T2 of
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   491
              @{typ bool} => 
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   492
                (case t of
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   493
                   Abs(_, _, @{const False}) => fst #> rev #> make_set T1
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   494
                 | Abs(_, _, @{const True}) => fst #> rev #> make_coset T1
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   495
                 | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> rev #> make_set T1
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   496
                 | _ => raise TERM ("post_process_term", [t]))
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   497
            | Type (@{type_name option}, _) =>
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   498
                (case t of
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   499
                  Abs(_, _, Const (@{const_name None}, _)) => fst #> make_map T1 T2
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   500
                | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> make_map T1 T2
42110
17e0cd9bc259 fixed postprocessing for term presentation in quickcheck; tuned spacing
bulwahn
parents: 41938
diff changeset
   501
                | _ => make_fun_upds T1 T2)
41935
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   502
            | _ => make_fun_upds T1 T2)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   503
        | NONE => process_args t)
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   504
    | _ => process_args t
d786a8a3dc47 minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents: 41927
diff changeset
   505
  end
41927
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   506
8759e9d043f9 adding file quickcheck_common to carry common functions of all quickcheck generators
bulwahn
parents:
diff changeset
   507
end;