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