| author | wenzelm | 
| Thu, 26 Jul 2012 16:54:44 +0200 | |
| changeset 48518 | 0c86acc069ad | 
| parent 48272 | db75b4005d9a | 
| child 50046 | 0051dc4f301f | 
| permissions | -rw-r--r-- | 
| 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 | 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: 
42159diff
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: 
45719diff
changeset | 10 | val reflect_bool : bool -> term | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 11 | val define_functions : ((term list -> term list) * (Proof.context -> tactic) option) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
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: 
45039diff
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: 
45039diff
changeset | 16 | -> (typ option * (term * term list)) list list | 
| 46565 | 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: 
45761diff
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: 
45039diff
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: 
46327diff
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: 
46327diff
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: 
46327diff
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: 
45419diff
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: 
46327diff
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: 
46327diff
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: 
45921diff
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: 
45921diff
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: 
45921diff
changeset | 28 | val ensure_sort : | 
| 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
changeset | 29 | (((sort * sort) * sort) * | 
| 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
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: 
45921diff
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: 
45921diff
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: 
45921diff
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: 
45921diff
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: 
42110diff
changeset | 36 | val gen_mk_parametric_generator_expr : | 
| 42229 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 37 | (((Proof.context -> term * term list -> term) * term) * typ) | 
| 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 38 | -> Proof.context -> (term * term list) list -> term | 
| 45039 
632a033616e9
adding post-processing of terms to narrowing-based Quickcheck
 bulwahn parents: 
44241diff
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: 
41927diff
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: 
46327diff
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: 
42195diff
changeset | 47 | (* static options *) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 48 | |
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 49 | val define_foundationally = false | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 50 | |
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42159diff
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: 
42159diff
changeset | 52 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42159diff
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: 
42159diff
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: 
42159diff
changeset | 55 | |
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45719diff
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: 
45719diff
changeset | 57 | |
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
changeset | 61 | |
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
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: 
46327diff
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: 
46327diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
changeset | 77 | |
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
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: 
45039diff
changeset | 79 | let | 
| 45757 
e32dd098f57a
renaming potential flag to genuine_only flag with an inverse semantics
 bulwahn parents: 
45755diff
changeset | 80 | val genuine_only = Config.get ctxt Quickcheck.genuine_only | 
| 46478 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
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: 
45039diff
changeset | 85 | val current_result = Unsynchronized.ref Quickcheck.empty_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: 
45039diff
changeset | 86 | fun excipit () = | 
| 
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: 
45039diff
changeset | 87 | "Quickcheck ran out of time while testing at size " ^ string_of_int (!current_size) | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 88 | 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: 
45753diff
changeset | 89 | val (test_fun, comp_time) = cpu_time "quickcheck compilation" | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 90 | (fn () => act (compile ctxt) [(t, eval_terms)]); | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 91 | val _ = Quickcheck.add_timing comp_time current_result | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 92 | 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: 
45039diff
changeset | 93 | 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: 
45039diff
changeset | 94 | 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: 
45039diff
changeset | 95 | 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: 
45039diff
changeset | 96 | let | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45763diff
changeset | 97 | val _ = Quickcheck.verbose_message ctxt | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 98 |             ("[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: 
45039diff
changeset | 99 | val _ = current_size := k | 
| 46565 | 100 | val ((result, report), time) = | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 101 |             cpu_time ("size " ^ string_of_int k) (fn () => test_fun genuine_only [1, k - 1])
 | 
| 46565 | 102 | val _ = if Config.get ctxt Quickcheck.timing then | 
| 103 | Quickcheck.verbose_message ctxt (fst time ^ ": " ^ string_of_int (snd time) ^ " ms") else () | |
| 104 | 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: 
45039diff
changeset | 105 | 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: 
45039diff
changeset | 106 | in | 
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45728diff
changeset | 107 | case result of | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 108 | NONE => with_size test_fun genuine_only (k + 1) | 
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45728diff
changeset | 109 | | SOME (true, ts) => SOME (true, ts) | 
| 45755 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 bulwahn parents: 
45754diff
changeset | 110 | | SOME (false, ts) => | 
| 46478 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 111 | if abort_potential then SOME (false, ts) | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 112 | else | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 113 | let | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 114 | val (ts1, ts2) = chop (length names) ts | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 115 | val (eval_terms', _) = chop (length ts2) eval_terms | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 116 | val cex = SOME ((false, names ~~ ts1), eval_terms' ~~ ts2) | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 117 | in | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 118 | (Quickcheck.message ctxt (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex)); | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 119 | Quickcheck.message ctxt "Quickcheck continues to find a genuine counterexample..."; | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 120 | with_size test_fun true k) | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 121 | 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: 
45039diff
changeset | 122 | 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: 
45039diff
changeset | 123 | 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: 
45416diff
changeset | 124 | case test_fun of | 
| 45420 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 125 |       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: 
45419diff
changeset | 126 | !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: 
45416diff
changeset | 127 | | 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: 
45039diff
changeset | 128 | let | 
| 45819 | 129 |         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: 
45039diff
changeset | 130 | val (response, exec_time) = | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 131 | 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: 
45039diff
changeset | 132 | 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: 
45039diff
changeset | 133 | 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: 
45416diff
changeset | 134 | 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: 
45039diff
changeset | 135 | 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: 
45039diff
changeset | 136 | |
| 
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: 
45039diff
changeset | 137 | fun validate_terms ctxt 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: 
45039diff
changeset | 138 | 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: 
45039diff
changeset | 139 | 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: 
45039diff
changeset | 140 | val size = 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: 
45039diff
changeset | 141 | val (test_funs, comp_time) = cpu_time "quickcheck batch compilation" | 
| 
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: 
45039diff
changeset | 142 | (fn () => Quickcheck.mk_batch_validator ctxt 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: 
45039diff
changeset | 143 | fun with_size tester k = | 
| 
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: 
45039diff
changeset | 144 | if k > size then true | 
| 
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: 
45039diff
changeset | 145 | else if tester k then with_size tester (k + 1) else false | 
| 
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: 
45039diff
changeset | 146 | val (results, exec_time) = cpu_time "quickcheck batch execution" (fn () => | 
| 45420 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 147 | Option.map (map (fn test_fun => | 
| 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 148 | TimeLimit.timeLimit (seconds (Config.get ctxt Quickcheck.timeout)) | 
| 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: 
45039diff
changeset | 149 | (fn () => with_size test_fun 1) () | 
| 
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: 
45039diff
changeset | 150 | handle TimeLimit.TimeOut => true)) test_funs) | 
| 
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: 
45039diff
changeset | 151 | 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: 
45039diff
changeset | 152 | (results, [comp_time, exec_time]) | 
| 
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: 
45039diff
changeset | 153 | end | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 154 | |
| 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: 
45039diff
changeset | 155 | fun test_terms ctxt 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: 
45039diff
changeset | 156 | 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: 
45039diff
changeset | 157 | 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: 
45039diff
changeset | 158 | val size = 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: 
45039diff
changeset | 159 | val namess = map (fn t => Term.add_free_names t []) 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: 
45039diff
changeset | 160 | val (test_funs, comp_time) = | 
| 
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: 
45039diff
changeset | 161 | cpu_time "quickcheck batch compilation" (fn () => Quickcheck.mk_batch_tester ctxt 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: 
45039diff
changeset | 162 | fun with_size tester k = | 
| 
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: 
45039diff
changeset | 163 | if k > size then 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: 
45039diff
changeset | 164 | else case tester k of SOME ts => SOME ts | NONE => with_size tester (k + 1) | 
| 
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: 
45039diff
changeset | 165 | val (results, exec_time) = cpu_time "quickcheck batch execution" (fn () => | 
| 45420 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 166 | Option.map (map (fn test_fun => | 
| 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 167 | TimeLimit.timeLimit (seconds (Config.get ctxt Quickcheck.timeout)) | 
| 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: 
45039diff
changeset | 168 | (fn () => with_size test_fun 1) () | 
| 
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: 
45039diff
changeset | 169 | handle TimeLimit.TimeOut => NONE)) test_funs) | 
| 
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: 
45039diff
changeset | 170 | 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: 
45039diff
changeset | 171 | (Option.map (map2 (fn names => Option.map (fn ts => names ~~ ts)) namess) 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: 
45039diff
changeset | 172 | [comp_time, exec_time]) | 
| 
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: 
45039diff
changeset | 173 | 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: 
45039diff
changeset | 174 | |
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
changeset | 175 | 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: 
45039diff
changeset | 176 | let | 
| 45757 
e32dd098f57a
renaming potential flag to genuine_only flag with an inverse semantics
 bulwahn parents: 
45755diff
changeset | 177 | val genuine_only = Config.get ctxt Quickcheck.genuine_only | 
| 46478 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 178 | 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: 
45039diff
changeset | 179 | 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: 
45039diff
changeset | 180 | 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: 
45039diff
changeset | 181 | 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: 
45039diff
changeset | 182 | 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: 
45039diff
changeset | 183 | 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: 
45039diff
changeset | 184 | val current_result = Unsynchronized.ref Quickcheck.empty_result | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 185 | 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: 
45039diff
changeset | 186 | (* 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: 
45039diff
changeset | 187 | let | 
| 45686 
cf22004ad55d
adding more verbose messages to exhaustive quickcheck
 bulwahn parents: 
45639diff
changeset | 188 | val _ = | 
| 45765 
cb6ddee6a463
making the default behaviour of quickcheck a little bit less verbose;
 bulwahn parents: 
45763diff
changeset | 189 |           Quickcheck.verbose_message ctxt ("[Quickcheck-" ^ name ^ "] Test " ^
 | 
| 46674 | 190 | (if size = 0 then "" else "data size: " ^ string_of_int size ^ " and ") ^ | 
| 45686 
cf22004ad55d
adding more verbose messages to exhaustive quickcheck
 bulwahn parents: 
45639diff
changeset | 191 | "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: 
45039diff
changeset | 192 | 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: 
45039diff
changeset | 193 |           cpu_time ("size " ^ string_of_int size ^ " and card " ^ string_of_int card)
 | 
| 46674 | 194 | (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: 
45039diff
changeset | 195 | 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: 
45039diff
changeset | 196 | in | 
| 45755 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 bulwahn parents: 
45754diff
changeset | 197 | 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: 
45039diff
changeset | 198 | 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: 
45039diff
changeset | 199 | val enumeration_card_size = | 
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
changeset | 200 | 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: 
45039diff
changeset | 201 | 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: 
45039diff
changeset | 202 | |> 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: 
46327diff
changeset | 203 | else | 
| 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
changeset | 204 | map (rpair 0) (1 upto (length ts)) | 
| 45419 
10ba32c347b0
quickcheck fails with code generator errors only if one tester is invoked
 bulwahn parents: 
45418diff
changeset | 205 | 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: 
45418diff
changeset | 206 | 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: 
45416diff
changeset | 207 | 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: 
45039diff
changeset | 208 | 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: 
45416diff
changeset | 209 | case test_fun of | 
| 45420 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 210 |       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: 
45419diff
changeset | 211 | !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: 
45416diff
changeset | 212 | | 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: 
45039diff
changeset | 213 | let | 
| 45921 | 214 |         val _ = Quickcheck.message ctxt ("Testing conjecture with Quickcheck-" ^ name ^ "...");
 | 
| 45755 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 bulwahn parents: 
45754diff
changeset | 215 | 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: 
45754diff
changeset | 216 | SOME ((card, _), (true, ts)) => | 
| 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 bulwahn parents: 
45754diff
changeset | 217 | 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: 
45754diff
changeset | 218 | | SOME ((card, size), (false, ts)) => | 
| 46478 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 219 | if abort_potential then | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 220 | Quickcheck.add_response names (nth eval_terms (card - 1)) (SOME (false, ts)) current_result | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 221 | else | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 222 | let | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 223 | val (ts1, ts2) = chop (length names) ts | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 224 | val (eval_terms', _) = chop (length ts2) (nth eval_terms (card - 1)) | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 225 | val cex = SOME ((false, names ~~ ts1), eval_terms' ~~ ts2) | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 226 | in | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 227 | (Quickcheck.message ctxt (Pretty.string_of (Quickcheck.pretty_counterex ctxt false cex)); | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 228 | Quickcheck.message ctxt "Quickcheck continues to find a genuine counterexample..."; | 
| 
cf1bcfb34c82
adding abort_potential functionality in quickcheck
 bulwahn parents: 
46331diff
changeset | 229 | test true (snd (take_prefix (fn x => not (x = (card, size))) enum))) | 
| 45755 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 bulwahn parents: 
45754diff
changeset | 230 | 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: 
45039diff
changeset | 231 | | NONE => () | 
| 45755 
b27a06dfb2ef
outputing the potentially spurious counterexample and continue search
 bulwahn parents: 
45754diff
changeset | 232 | 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: 
45039diff
changeset | 233 | 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: 
45039diff
changeset | 234 | |
| 
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: 
45039diff
changeset | 235 | 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: 
45039diff
changeset | 236 | fst (chop (Config.get ctxt Quickcheck.finite_type_size) | 
| 45416 | 237 |     [@{typ "Enum.finite_1"}, @{typ "Enum.finite_2"}, @{typ "Enum.finite_3"},
 | 
| 238 |      @{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: 
45039diff
changeset | 239 | |
| 
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: 
45039diff
changeset | 240 | 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: 
45039diff
changeset | 241 | |
| 
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: 
45039diff
changeset | 242 | 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: 
45039diff
changeset | 243 | 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: 
45039diff
changeset | 244 | 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: 
45039diff
changeset | 245 | 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: 
45039diff
changeset | 246 | 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: 
45039diff
changeset | 247 | |> 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: 
45039diff
changeset | 248 | 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: 
45039diff
changeset | 249 |         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: 
45039diff
changeset | 250 | 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: 
45039diff
changeset | 251 | ":\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: 
45039diff
changeset | 252 | " 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: 
45039diff
changeset | 253 | 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: 
45039diff
changeset | 254 | 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: 
45039diff
changeset | 255 | 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: 
45039diff
changeset | 256 | | 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: 
45039diff
changeset | 257 | 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: 
45039diff
changeset | 258 | |
| 
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: 
45039diff
changeset | 259 | 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: 
45039diff
changeset | 260 | |
| 45440 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 261 | (* minimalistic preprocessing *) | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 262 | |
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 263 | fun strip_all (Const (@{const_name HOL.All}, _) $ Abs (a, T, t)) = 
 | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 264 | let | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 265 | val (a', t') = strip_all t | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 266 | in ((a, T) :: a', t') end | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 267 | | strip_all t = ([], t); | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 268 | |
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 269 | fun preprocess ctxt t = | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 270 | let | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 271 | val thy = Proof_Context.theory_of ctxt | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 272 | 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: 
45420diff
changeset | 273 |     val rewrs = map (swap o dest) @{thms all_simps} @
 | 
| 46327 | 274 |       (map dest [@{thm not_ex}, @{thm not_all}, @{thm imp_conjL}, @{thm fun_eq_iff},
 | 
| 275 |         @{thm bot_fun_def}, @{thm less_bool_def}])
 | |
| 45440 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 276 | 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: 
45420diff
changeset | 277 | val (vs, body) = strip_all t' | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 278 | val vs' = Variable.variant_frees ctxt [t'] vs | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 279 | in | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 280 | subst_bounds (map Free (rev vs'), body) | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 281 | end | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 282 | |
| 46565 | 283 | |
| 284 | structure Subtype_Predicates = Generic_Data | |
| 285 | ( | |
| 286 | type T = (term * string) list | |
| 287 | val empty = [] | |
| 288 | val extend = I | |
| 46580 | 289 | fun merge data : T = AList.merge (op =) (K true) data | 
| 46565 | 290 | ) | 
| 291 | ||
| 292 | val register_predicate = Subtype_Predicates.map o AList.update (op =) | |
| 293 | ||
| 294 | fun subtype_preprocess ctxt (T, (t, ts)) = | |
| 295 | let | |
| 296 | val preds = Subtype_Predicates.get (Context.Proof ctxt) | |
| 297 | fun matches (p $ x) = AList.defined Term.could_unify preds p | |
| 298 | fun get_match (p $ x) = Option.map (rpair x) (AList.lookup Term.could_unify preds p) | |
| 299 | fun subst_of (tyco, v as Free (x, repT)) = | |
| 300 | let | |
| 301 | val [(info, _)] = Typedef.get_info ctxt tyco | |
| 302 | val repT' = Logic.varifyT_global (#rep_type info) | |
| 303 | val substT = Sign.typ_match (Proof_Context.theory_of ctxt) (repT', repT) Vartab.empty | |
| 304 | val absT = Envir.subst_type substT (Logic.varifyT_global (#abs_type info)) | |
| 305 | in (v, Const (#Rep_name info, absT --> repT) $ Free (x, absT)) end | |
| 306 | val (prems, concl) = strip_imp t | |
| 307 | val subst = map subst_of (map_filter get_match prems) | |
| 308 | val t' = Term.subst_free subst | |
| 309 | (fold_rev (curry HOLogic.mk_imp) (filter_out matches prems) concl) | |
| 310 | in | |
| 311 | (T, (t', ts)) | |
| 312 | end | |
| 313 | ||
| 45440 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 314 | (* instantiation of type variables with concrete types *) | 
| 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 315 | |
| 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: 
45039diff
changeset | 316 | 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: 
45039diff
changeset | 317 | 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: 
45039diff
changeset | 318 | 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: 
45039diff
changeset | 319 | 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: 
45039diff
changeset | 320 | val default_insts = | 
| 45420 
d17556b9a89b
more precise messages with the tester's name in quickcheck; tuned
 bulwahn parents: 
45419diff
changeset | 321 | 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: 
45039diff
changeset | 322 | 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: 
45039diff
changeset | 323 | 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: 
45039diff
changeset | 324 | 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: 
45039diff
changeset | 325 | map (fn T => | 
| 45440 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 326 | (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: 
45039diff
changeset | 327 | (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: 
45419diff
changeset | 328 | 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: 
45039diff
changeset | 329 | else | 
| 45440 
9f4d3e68ae98
adding a minimalistic preprocessing rewriting common boolean operators; tuned
 bulwahn parents: 
45420diff
changeset | 330 | [(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: 
45039diff
changeset | 331 | 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: 
45039diff
changeset | 332 | cat_lines | 
| 
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: 
45039diff
changeset | 333 | (maps (map_filter (fn (_, Term t) => NONE | (_, Wellsorted_Error s) => SOME s)) 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: 
45039diff
changeset | 334 | fun is_wellsorted_term (T, Term t) = SOME (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: 
45039diff
changeset | 335 | | is_wellsorted_term (_, Wellsorted_Error s) = 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: 
45039diff
changeset | 336 | 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: 
45039diff
changeset | 337 | 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: 
45039diff
changeset | 338 | [[]] => 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: 
45039diff
changeset | 339 | | 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: 
45039diff
changeset | 340 | 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: 
45039diff
changeset | 341 | 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: 
45039diff
changeset | 342 | 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: 
45039diff
changeset | 343 | 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: 
45039diff
changeset | 344 | |
| 45718 
8979b2463fc8
quickcheck random can also find potential counterexamples;
 bulwahn parents: 
45687diff
changeset | 345 | (* compilation of testing functions *) | 
| 
8979b2463fc8
quickcheck random can also find potential counterexamples;
 bulwahn parents: 
45687diff
changeset | 346 | |
| 45763 
3bb2bdf654f7
random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
 bulwahn parents: 
45761diff
changeset | 347 | 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: 
45728diff
changeset | 348 | let | 
| 45763 
3bb2bdf654f7
random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
 bulwahn parents: 
45761diff
changeset | 349 | val T = fastype_of then_t | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 350 |     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: 
45728diff
changeset | 351 | in | 
| 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45728diff
changeset | 352 |     Const (@{const_name "Quickcheck.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: 
45757diff
changeset | 353 | (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: 
45761diff
changeset | 354 | (if_t $ genuine_only $ none $ else_t false) | 
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45728diff
changeset | 355 | end | 
| 45718 
8979b2463fc8
quickcheck random can also find potential counterexamples;
 bulwahn parents: 
45687diff
changeset | 356 | |
| 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: 
45039diff
changeset | 357 | 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: 
45039diff
changeset | 358 | | 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: 
45039diff
changeset | 359 | 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: 
45039diff
changeset | 360 | 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: 
45039diff
changeset | 361 | 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: 
45039diff
changeset | 362 | 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: 
45039diff
changeset | 363 | (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: 
45039diff
changeset | 364 | 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: 
45039diff
changeset | 365 | 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: 
45039diff
changeset | 366 | 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: 
45039diff
changeset | 367 | |
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
changeset | 368 | 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: 
45039diff
changeset | 369 | let | 
| 46565 | 370 | 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: 
45686diff
changeset | 371 | 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: 
45686diff
changeset | 372 | 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: 
45686diff
changeset | 373 | 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: 
45686diff
changeset | 374 | 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: 
45686diff
changeset | 375 | | 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: 
45039diff
changeset | 376 | 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: 
45039diff
changeset | 377 | case goal of | 
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
changeset | 378 | [(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: 
46327diff
changeset | 379 | | 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: 
45039diff
changeset | 380 | val goals' = instantiate_goals ctxt insts goals | 
| 46565 | 381 | |> (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: 
45039diff
changeset | 382 | |> 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: 
45039diff
changeset | 383 | 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: 
45039diff
changeset | 384 | 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: 
45039diff
changeset | 385 | 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: 
45039diff
changeset | 386 | else | 
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46327diff
changeset | 387 | 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: 
45039diff
changeset | 388 | (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: 
45039diff
changeset | 389 | 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: 
45039diff
changeset | 390 | |
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 391 | (* defining functions *) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 392 | |
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 393 | fun pat_completeness_auto ctxt = | 
| 42793 | 394 | Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 395 | |
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 396 | fun define_functions (mk_equations, termination_tac) prfx argnames names Ts = | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 397 | if define_foundationally andalso is_some termination_tac then | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 398 | let | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 399 | 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: 
42195diff
changeset | 400 | in | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 401 | Function.add_function | 
| 42287 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 wenzelm parents: 
42229diff
changeset | 402 | (map (fn (name, T) => (Binding.conceal (Binding.name name), SOME T, NoSyn)) | 
| 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 wenzelm parents: 
42229diff
changeset | 403 | (names ~~ Ts)) | 
| 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 wenzelm parents: 
42229diff
changeset | 404 | (map (pair (apfst Binding.conceal Attrib.empty_binding)) eqs_t) | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 405 | Function_Common.default_config pat_completeness_auto | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 406 | #> snd | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 407 | #> (fn lthy => Function.prove_termination NONE (the termination_tac lthy) lthy) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 408 | #> snd | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 409 | end | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 410 | else | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 411 | fold_map (fn (name, T) => Local_Theory.define | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 412 | ((Binding.conceal (Binding.name name), NoSyn), | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 413 | (apfst Binding.conceal Attrib.empty_binding, mk_undefined T)) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 414 | #> apfst fst) (names ~~ Ts) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 415 | #> (fn (consts, lthy) => | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 416 | let | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 417 | val eqs_t = mk_equations consts | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 418 | val eqs = map (fn eq => Goal.prove lthy argnames [] eq | 
| 42361 | 419 | (fn _ => Skip_Proof.cheat_tac (Proof_Context.theory_of lthy))) eqs_t | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 420 | in | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 421 | 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: 
46674diff
changeset | 422 | ((Binding.qualify true prfx | 
| 
6212bcc94bb0
refine bindings in quickcheck_common: do not conceal and do not declare as simps
 bulwahn parents: 
46674diff
changeset | 423 | (Binding.qualify true name (Binding.name "simps")), | 
| 
6212bcc94bb0
refine bindings in quickcheck_common: do not conceal and do not declare as simps
 bulwahn parents: 
46674diff
changeset | 424 | [Code.add_default_eqn_attrib]), [eq]) #> snd) | 
| 45592 | 425 | (names ~~ eqs) lthy | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 426 | end) | 
| 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 427 | |
| 41935 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 428 | (** ensuring sort constraints **) | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 429 | |
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
changeset | 430 | 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: 
45921diff
changeset | 431 | -> 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: 
45921diff
changeset | 432 | |
| 41927 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 433 | fun perhaps_constrain thy insts raw_vs = | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 434 | let | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 435 | 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 | 436 | (Logic.varifyT_global T, sort); | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 437 | val vtab = Vartab.empty | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 438 | |> 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 | 439 | |> fold meet insts; | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 440 | 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 | 441 | end handle Sorts.CLASS_ERROR _ => NONE; | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 442 | |
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
changeset | 443 | 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 | 444 | let | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 445 | 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: 
45921diff
changeset | 446 | 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: 
42214diff
changeset | 447 | val vs = (map o apsnd) (curry (Sorts.inter_sort algebra) sort_vs) raw_vs | 
| 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 448 | fun insts_of sort constr = (map (rpair sort) o flat o maps snd o maps snd) | 
| 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 449 | (Datatype_Aux.interpret_construction descr vs constr) | 
| 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 450 |     val insts = insts_of sort  { atyp = single, dtyp = (K o K o K) [] }
 | 
| 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 451 |       @ insts_of aux_sort { atyp = K [], dtyp = K o K }
 | 
| 48272 | 452 | 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 | 453 | in if has_inst then thy | 
| 42229 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 454 | 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: 
45921diff
changeset | 455 | of SOME constrain => instantiate config descr | 
| 42229 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 456 | (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 | 457 | ((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 | 458 | | NONE => thy | 
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 459 | end; | 
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
changeset | 460 | |
| 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
changeset | 461 | 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: 
45921diff
changeset | 462 |   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: 
45921diff
changeset | 463 | |
| 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45921diff
changeset | 464 | 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: 
41927diff
changeset | 465 | |
| 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: 
42110diff
changeset | 466 | (** 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: 
42110diff
changeset | 467 | |
| 
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: 
42110diff
changeset | 468 | 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: 
42110diff
changeset | 469 | 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: 
42110diff
changeset | 470 |     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: 
45719diff
changeset | 471 | fun mk_if (index, (t, eval_terms)) else_t = if_t $ | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45719diff
changeset | 472 |         (HOLogic.eq_const @{typ code_numeral} $ Bound 0 $ HOLogic.mk_number @{typ code_numeral} 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: 
42110diff
changeset | 473 | (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: 
42110diff
changeset | 474 | in | 
| 44241 | 475 |     absdummy @{typ "code_numeral"} (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: 
42110diff
changeset | 476 | 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: 
42110diff
changeset | 477 | |
| 41935 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 478 | (** post-processing of function terms **) | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 479 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 480 | 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: 
41927diff
changeset | 481 |   | dest_fun_upd t = raise TERM ("dest_fun_upd", [t])
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 482 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 483 | fun mk_fun_upd T1 T2 (t1, t2) t = | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 484 |   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: 
41927diff
changeset | 485 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 486 | fun dest_fun_upds t = | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 487 | case try dest_fun_upd t of | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 488 | NONE => | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 489 | (case t of | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 490 | Abs (_, _, _) => ([], t) | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 491 |       | _ => raise TERM ("dest_fun_upds", [t]))
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 492 | | 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: 
41927diff
changeset | 493 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 494 | 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: 
41927diff
changeset | 495 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 496 | 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: 
41927diff
changeset | 497 |   | make_set T1 ((_, @{const False}) :: tps) = make_set T1 tps
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 498 |   | make_set T1 ((t1, @{const True}) :: tps) =
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 499 |     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: 
41927diff
changeset | 500 | $ t1 $ (make_set T1 tps) | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 501 |   | make_set T1 ((_, t) :: tps) = raise TERM ("make_set", [t])
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 502 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 503 | 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: 
41927diff
changeset | 504 | | make_coset T tps = | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 505 | let | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 506 |       val U = T --> @{typ bool}
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 507 |       fun invert @{const False} = @{const True}
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 508 |         | invert @{const True} = @{const False}
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 509 | in | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 510 |       Const (@{const_name "Groups.minus_class.minus"}, U --> U --> U)
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 511 |         $ 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: 
41927diff
changeset | 512 | end | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 513 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 514 | 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: 
41927diff
changeset | 515 |   | 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: 
41927diff
changeset | 516 | | 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: 
41927diff
changeset | 517 | |
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 518 | fun post_process_term t = | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 519 | let | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 520 | fun map_Abs f t = | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 521 |       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: 
41927diff
changeset | 522 | fun process_args t = case strip_comb t of | 
| 42110 
17e0cd9bc259
fixed postprocessing for term presentation in quickcheck; tuned spacing
 bulwahn parents: 
41938diff
changeset | 523 | (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: 
41927diff
changeset | 524 | in | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 525 | case fastype_of t of | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 526 |       Type (@{type_name fun}, [T1, T2]) =>
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 527 | (case try dest_fun_upds t of | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 528 | SOME (tps, t) => | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 529 | (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: 
41927diff
changeset | 530 | |> (case T2 of | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 531 |               @{typ bool} => 
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 532 | (case t of | 
| 42110 
17e0cd9bc259
fixed postprocessing for term presentation in quickcheck; tuned spacing
 bulwahn parents: 
41938diff
changeset | 533 |                    Abs(_, _, @{const False}) => fst #> rev #> make_set T1
 | 
| 
17e0cd9bc259
fixed postprocessing for term presentation in quickcheck; tuned spacing
 bulwahn parents: 
41938diff
changeset | 534 |                  | Abs(_, _, @{const True}) => fst #> rev #> make_coset T1
 | 
| 41935 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 535 |                  | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> rev #> make_set T1
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 536 |                  | _ => raise TERM ("post_process_term", [t]))
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 537 |             | Type (@{type_name option}, _) =>
 | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 538 | (case t of | 
| 42110 
17e0cd9bc259
fixed postprocessing for term presentation in quickcheck; tuned spacing
 bulwahn parents: 
41938diff
changeset | 539 |                   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: 
41927diff
changeset | 540 |                 | Abs(_, _, Const (@{const_name undefined}, _)) => fst #> make_map T1 T2
 | 
| 42110 
17e0cd9bc259
fixed postprocessing for term presentation in quickcheck; tuned spacing
 bulwahn parents: 
41938diff
changeset | 541 | | _ => make_fun_upds T1 T2) | 
| 41935 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 542 | | _ => make_fun_upds T1 T2) | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 543 | | NONE => process_args t) | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 544 | | _ => process_args t | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41927diff
changeset | 545 | end | 
| 41927 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 546 | |
| 
8759e9d043f9
adding file quickcheck_common to carry common functions of all quickcheck generators
 bulwahn parents: diff
changeset | 547 | end; |