| author | wenzelm | 
| Sun, 11 Nov 2012 16:19:55 +0100 | |
| changeset 50080 | 200f749c96db | 
| parent 50046 | 0051dc4f301f | 
| child 51126 | df86080de4cb | 
| permissions | -rw-r--r-- | 
| 41923 
f05fc0711bc7
renaming signatures and structures; correcting header
 bulwahn parents: 
41920diff
changeset | 1 | (* Title: HOL/Tools/Quickcheck/exhaustive_generators.ML | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 3 | |
| 41918 | 4 | Exhaustive generators for various types. | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 5 | *) | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 6 | |
| 41918 | 7 | signature EXHAUSTIVE_GENERATORS = | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 8 | sig | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 9 | val compile_generator_expr: | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 10 | Proof.context -> (term * term list) list -> bool -> int list -> (bool * term list) option * Quickcheck.report option | 
| 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: 
42028diff
changeset | 11 | val compile_generator_exprs: Proof.context -> term list -> (int -> term list option) list | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 12 | val compile_validator_exprs: Proof.context -> term list -> (int -> bool) list | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 13 | val put_counterexample: (unit -> int -> bool -> int -> (bool * term list) option) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 14 | -> Proof.context -> Proof.context | 
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 15 | val put_counterexample_batch: (unit -> (int -> term list option) list) | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 16 | -> Proof.context -> Proof.context | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 17 | val put_validator_batch: (unit -> (int -> bool) list) -> Proof.context -> Proof.context | 
| 42306 | 18 | exception Counterexample of term list | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 19 | val smart_quantifier : bool Config.T | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 20 | val optimise_equality : bool Config.T | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 21 | val quickcheck_pretty : bool Config.T | 
| 45484 
23871e17dddb
setting up exhaustive generators which are used for the smart generators
 bulwahn parents: 
45420diff
changeset | 22 | val setup_exhaustive_datatype_interpretation : theory -> theory | 
| 48178 
0192811f0a96
exporting important function for the "many conjecture refutation" compilation of quickcheck
 bulwahn parents: 
48013diff
changeset | 23 | val setup_bounded_forall_datatype_interpretation : theory -> theory | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 24 | val setup: theory -> theory | 
| 45924 
f03dd48829d8
exporting instantiation functions in quickcheck for their usage in abstract generators
 bulwahn parents: 
45923diff
changeset | 25 | |
| 
f03dd48829d8
exporting instantiation functions in quickcheck for their usage in abstract generators
 bulwahn parents: 
45923diff
changeset | 26 | val instantiate_full_exhaustive_datatype : Datatype_Aux.config -> Datatype_Aux.descr -> | 
| 
f03dd48829d8
exporting instantiation functions in quickcheck for their usage in abstract generators
 bulwahn parents: 
45923diff
changeset | 27 | (string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory | 
| 46565 | 28 | val instantiate_exhaustive_datatype : Datatype_Aux.config -> Datatype_Aux.descr -> | 
| 29 | (string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory | |
| 30 | ||
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 31 | end; | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 32 | |
| 41918 | 33 | structure Exhaustive_Generators : EXHAUSTIVE_GENERATORS = | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 34 | struct | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 35 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 36 | (* basics *) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 37 | |
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 38 | (** dynamic options **) | 
| 40907 | 39 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 40 | val smart_quantifier = Attrib.setup_config_bool @{binding quickcheck_smart_quantifier} (K true)
 | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 41 | val optimise_equality = Attrib.setup_config_bool @{binding quickcheck_optimise_equality} (K true)
 | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 42 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 43 | val fast = Attrib.setup_config_bool @{binding quickcheck_fast} (K false)
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 44 | val bounded_forall = Attrib.setup_config_bool @{binding quickcheck_bounded_forall} (K false)
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 45 | val full_support = Attrib.setup_config_bool @{binding quickcheck_full_support} (K true)
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 46 | val quickcheck_pretty = Attrib.setup_config_bool @{binding quickcheck_pretty} (K true)
 | 
| 41903 
39fd77f0ae59
fixing postprocessing; adding a configuration to enable and disable pretty presentation of quickcheck's counterexample
 bulwahn parents: 
41902diff
changeset | 47 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 48 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 49 | (** abstract syntax **) | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 50 | |
| 40639 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 51 | fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => Code_Evaluation.term"});
 | 
| 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 52 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 53 | val size = @{term "i :: code_numeral"}
 | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 54 | val size_pred = @{term "(i :: code_numeral) - 1"}
 | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 55 | val size_ge_zero = @{term "(i :: code_numeral) > 0"}
 | 
| 42304 | 56 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 57 | fun mk_none_continuation (x, y) = | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 58 | let | 
| 46306 | 59 |     val (T as Type(@{type_name "option"}, _)) = fastype_of x
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 60 | in | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 61 |     Const (@{const_name "Quickcheck_Exhaustive.orelse"}, T --> T --> T) $ x $ y
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 62 | end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 63 | |
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 64 | fun mk_if (b, t, e) = | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 65 | let | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 66 | val T = fastype_of t | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 67 |   in Const (@{const_name "HOL.If"}, @{typ bool} --> T --> T --> T) $ b $ t $ e end
 | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 68 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 69 | (* handling inductive datatypes *) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 70 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 71 | (** constructing generator instances **) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 72 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 73 | exception FUNCTION_TYPE; | 
| 42306 | 74 | |
| 75 | exception Counterexample of term list | |
| 76 | ||
| 45722 | 77 | val resultT =  @{typ "(bool * term list) option"};
 | 
| 78 | ||
| 41918 | 79 | val exhaustiveN = "exhaustive"; | 
| 42304 | 80 | val full_exhaustiveN = "full_exhaustive"; | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 81 | val bounded_forallN = "bounded_forall"; | 
| 42306 | 82 | |
| 83 | fun fast_exhaustiveT T = (T --> @{typ unit})
 | |
| 84 |   --> @{typ code_numeral} --> @{typ unit}
 | |
| 85 | ||
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 86 | fun exhaustiveT T = (T --> resultT) --> @{typ code_numeral} --> resultT
 | 
| 42304 | 87 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 88 | fun bounded_forallT T = (T --> @{typ bool}) --> @{typ code_numeral} --> @{typ bool}
 | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 89 | |
| 45722 | 90 | fun full_exhaustiveT T = (termifyT T --> resultT) --> @{typ code_numeral} --> resultT
 | 
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
40913diff
changeset | 91 | |
| 45722 | 92 | fun check_allT T = (termifyT T --> resultT) --> resultT | 
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
40913diff
changeset | 93 | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 94 | fun mk_equation_terms generics (descr, vs, Ts) = | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 95 | let | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 96 | val (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, exhaustives) = generics | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 97 | val rhss = | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 98 | Datatype_Aux.interpret_construction descr vs | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 99 |         { atyp = mk_call, dtyp = mk_aux_call }
 | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 100 | |> (map o apfst) Type | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 101 | |> map (fn (T, cs) => map (mk_consexpr T) cs) | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 102 | |> map mk_rhs | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 103 | val lhss = map2 (fn t => fn T => t $ test_function T $ size) exhaustives Ts | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 104 | in | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 105 | map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhss ~~ rhss) | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 106 | end | 
| 42306 | 107 | |
| 44241 | 108 | fun gen_mk_call c T = (T, fn t => c T $ absdummy T t $ size_pred) | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 109 | |
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 110 | fun gen_mk_aux_call functerms fTs (k, _) (tyco, Ts) = | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 111 | let | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 112 | val T = Type (tyco, Ts) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 113 | val _ = if not (null fTs) then raise FUNCTION_TYPE else () | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 114 | in | 
| 44241 | 115 | (T, fn t => nth functerms k $ absdummy T t $ size_pred) | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 116 | end | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 117 | |
| 46306 | 118 | fun gen_mk_consexpr test_function simpleT (c, xs) = | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 119 | let | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 120 | val (Ts, fns) = split_list xs | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 121 | val constr = Const (c, Ts ---> simpleT) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 122 | val bounds = map Bound (((length xs) - 1) downto 0) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 123 | val start_term = test_function simpleT $ list_comb (constr, bounds) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 124 | in fold_rev (fn f => fn t => f t) fns start_term end | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 125 | |
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 126 | fun mk_equations functerms = | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 127 | let | 
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 128 |     fun test_function T = Free ("f", T --> resultT)
 | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 129 | val mk_call = gen_mk_call (fn T => | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 130 |       Const (@{const_name "Quickcheck_Exhaustive.exhaustive_class.exhaustive"}, exhaustiveT T))
 | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 131 | val mk_aux_call = gen_mk_aux_call functerms | 
| 46306 | 132 | val mk_consexpr = gen_mk_consexpr test_function | 
| 42304 | 133 | fun mk_rhs exprs = | 
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 134 |       mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (@{const_name "None"}, resultT))
 | 
| 42304 | 135 | in | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 136 | mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) | 
| 42304 | 137 | end | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 138 | |
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 139 | fun mk_bounded_forall_equations functerms = | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 140 | let | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 141 |     fun test_function T = Free ("P", T --> @{typ bool})
 | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 142 | val mk_call = gen_mk_call (fn T => | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 143 |       Const (@{const_name "Quickcheck_Exhaustive.bounded_forall_class.bounded_forall"},
 | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 144 | bounded_forallT T)) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 145 | val mk_aux_call = gen_mk_aux_call functerms | 
| 46306 | 146 | val mk_consexpr = gen_mk_consexpr test_function | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 147 | fun mk_rhs exprs = | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 148 |       mk_if (size_ge_zero, foldr1 HOLogic.mk_conj exprs, @{term "True"})
 | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 149 | in | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 150 | mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 151 | end | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 152 | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 153 | fun mk_full_equations functerms = | 
| 42304 | 154 | let | 
| 45722 | 155 |     fun test_function T = Free ("f", termifyT T --> resultT)
 | 
| 156 |     fun split_const T = HOLogic.split_const (T, @{typ "unit => Code_Evaluation.term"}, resultT)
 | |
| 42304 | 157 | fun mk_call T = | 
| 158 | let | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 159 | val full_exhaustive = | 
| 42310 
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
 bulwahn parents: 
42309diff
changeset | 160 |           Const (@{const_name "Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive"},
 | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 161 | full_exhaustiveT T) | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 162 | in | 
| 44241 | 163 | (T, fn t => full_exhaustive $ | 
| 45722 | 164 |           (split_const T $ absdummy T (absdummy @{typ "unit => Code_Evaluation.term"} t)) $ size_pred)
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 165 | end | 
| 41918 | 166 | fun mk_aux_call fTs (k, _) (tyco, Ts) = | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 167 | let | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 168 | val T = Type (tyco, Ts) | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 169 | val _ = if not (null fTs) then raise FUNCTION_TYPE else () | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 170 | in | 
| 44241 | 171 | (T, fn t => nth functerms k $ | 
| 45722 | 172 |           (split_const T $ absdummy T (absdummy @{typ "unit => Code_Evaluation.term"} t)) $ size_pred)
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 173 | end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 174 | fun mk_consexpr simpleT (c, xs) = | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 175 | let | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 176 | val (Ts, fns) = split_list xs | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 177 | val constr = Const (c, Ts ---> simpleT) | 
| 40639 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 178 | val bounds = map (fn x => Bound (2 * x + 1)) (((length xs) - 1) downto 0) | 
| 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 179 |         val Eval_App = Const ("Code_Evaluation.App", HOLogic.termT --> HOLogic.termT --> HOLogic.termT)
 | 
| 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 180 |         val Eval_Const = Const ("Code_Evaluation.Const", HOLogic.literalT --> @{typ typerep} --> HOLogic.termT)
 | 
| 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 181 |         val term = fold (fn u => fn t => Eval_App $ t $ (u $ @{term "()"}))
 | 
| 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 182 | bounds (Eval_Const $ HOLogic.mk_literal c $ HOLogic.mk_typerep (Ts ---> simpleT)) | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 183 | val start_term = test_function simpleT $ | 
| 40639 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 184 |         (HOLogic.pair_const simpleT @{typ "unit => Code_Evaluation.term"}
 | 
| 44241 | 185 |           $ (list_comb (constr, bounds)) $ absdummy @{typ unit} term)
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 186 | in fold_rev (fn f => fn t => f t) fns start_term end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 187 | fun mk_rhs exprs = | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 188 | mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, | 
| 45722 | 189 |         Const (@{const_name "None"}, resultT))
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 190 | in | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 191 | mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) | 
| 40901 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 192 | end | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 193 | |
| 40901 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 194 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 195 | (** instantiating generator classes **) | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 196 | |
| 42325 | 197 | fun contains_recursive_type_under_function_types xs = | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 198 | exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT => | 
| 42325 | 199 | (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 200 | |
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 201 | fun instantiate_datatype (name, constprfx, sort, mk_equations, mk_T, argnames) | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 202 | config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 203 | if not (contains_recursive_type_under_function_types descr) then | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 204 | let | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 205 |       val _ = Datatype_Aux.message config ("Creating " ^ name ^ "...")
 | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 206 | val fullnames = map (prefix (constprfx ^ "_")) (names @ auxnames) | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 207 | in | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 208 | thy | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 209 | |> Class.instantiation (tycos, vs, sort) | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 210 | |> Quickcheck_Common.define_functions | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 211 | (fn functerms => mk_equations functerms (descr, vs, Ts @ Us), NONE) | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 212 | prfx argnames fullnames (map mk_T (Ts @ Us)) | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 213 | |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 214 | end | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 215 | else | 
| 42230 
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
 bulwahn parents: 
42229diff
changeset | 216 | (Datatype_Aux.message config | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 217 |       ("Creation of " ^ name ^ " failed because the datatype is recursive under a function type");
 | 
| 42230 
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
 bulwahn parents: 
42229diff
changeset | 218 | thy) | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 219 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 220 | val instantiate_bounded_forall_datatype = instantiate_datatype | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 221 |  ("bounded universal quantifiers", bounded_forallN, @{sort bounded_forall},
 | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 222 | mk_bounded_forall_equations, bounded_forallT, ["P", "i"]); | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 223 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 224 | val instantiate_exhaustive_datatype = instantiate_datatype | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 225 |   ("exhaustive generators", exhaustiveN, @{sort exhaustive},
 | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 226 | mk_equations, exhaustiveT, ["f", "i"]) | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 227 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 228 | val instantiate_full_exhaustive_datatype = instantiate_datatype | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 229 |   ("full exhaustive generators", full_exhaustiveN, @{sort full_exhaustive},
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 230 | mk_full_equations, full_exhaustiveT, ["f", "i"]) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 231 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 232 | (* building and compiling generator expressions *) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 233 | |
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 234 | fun mk_let_expr (x, t, e) genuine = | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 235 | let | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 236 | val (T1, T2) = (fastype_of x, fastype_of (e genuine)) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 237 | in | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 238 |     Const (@{const_name Let}, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine)
 | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 239 | end | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 240 | |
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 241 | fun mk_safe_let_expr genuine_only none safe (x, t, e) genuine = | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 242 | let | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 243 | val (T1, T2) = (fastype_of x, fastype_of (e genuine)) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 244 |     val if_t = Const (@{const_name "If"}, @{typ bool} --> T2 --> T2 --> T2)
 | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 245 | in | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 246 |     Const (@{const_name "Quickcheck.catch_match"}, T2 --> T2 --> T2) $ 
 | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 247 |       (Const (@{const_name Let}, T1 --> (T1 --> T2) --> T2) $ t $ lambda x (e genuine)) $
 | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 248 | (if_t $ genuine_only $ none $ safe false) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 249 | end | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 250 | |
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 251 | fun mk_test_term lookup mk_closure mk_if mk_let none_t return ctxt = | 
| 42306 | 252 | let | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 253 | val cnstrs = flat (maps | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 254 | (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 255 | (Symtab.dest (Datatype.get_all (Proof_Context.theory_of ctxt)))) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 256 | fun is_constrt (Const (s, T), ts) = (case (AList.lookup (op =) cnstrs s, body_type T) of | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 257 | (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 258 | | _ => false) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 259 | | is_constrt _ = false | 
| 42306 | 260 | fun mk_naive_test_term t = | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 261 | fold_rev mk_closure (map lookup (Term.add_free_names t [])) | 
| 45761 
90028fd2f1fa
exhaustive returns if a counterexample is genuine or potentially spurious in the presence of assumptions more correctly
 bulwahn parents: 
45754diff
changeset | 262 | (mk_if (t, none_t, return) true) | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 263 | fun mk_test (vars, check) = fold_rev mk_closure (map lookup vars) check | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 264 | fun mk_smart_test_term' concl bound_vars assms genuine = | 
| 42306 | 265 | let | 
| 266 | fun vars_of t = subtract (op =) bound_vars (Term.add_free_names t []) | |
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 267 | fun mk_equality_term (lhs, f as Free (x, _)) c (assm, assms) = | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 268 | if member (op =) (Term.add_free_names lhs bound_vars) x then | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 269 | c (assm, assms) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 270 | else | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 271 | (let | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 272 | val rec_call = mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 273 | fun safe genuine = | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 274 | the_default I (Option.map mk_closure (try lookup x)) (rec_call genuine) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 275 | in | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 276 | mk_test (remove (op =) x (vars_of assm), | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 277 | mk_let safe f (try lookup x) lhs | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 278 | (mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms) genuine) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 279 | |
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 280 | end) | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 281 | | mk_equality_term (lhs, t) c (assm, assms) = | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 282 | if is_constrt (strip_comb t) then | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 283 | let | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 284 | val (constr, args) = strip_comb t | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 285 | val T = fastype_of t | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 286 | val vars = map Free (Variable.variant_frees ctxt (concl :: assms) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 287 |                   (map (fn t => ("x", fastype_of t)) args))
 | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 288 | val varnames = map (fst o dest_Free) vars | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 289 | val dummy_var = Free (singleton | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 290 |                   (Variable.variant_frees ctxt (concl :: assms @ vars)) ("dummy", T))
 | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 291 | val new_assms = map HOLogic.mk_eq (vars ~~ args) | 
| 48273 | 292 | val bound_vars' = union (op =) (vars_of lhs) (union (op =) varnames bound_vars) | 
| 293 | val cont_t = mk_smart_test_term' concl bound_vars' (new_assms @ assms) genuine | |
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 294 | in | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 295 | mk_test (vars_of lhs, Datatype_Case.make_case ctxt Datatype_Case.Quiet [] lhs | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 296 | [(list_comb (constr, vars), cont_t), (dummy_var, none_t)]) | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 297 | end | 
| 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 298 | else c (assm, assms) | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 299 | fun default (assm, assms) = | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 300 | mk_test (vars_of assm, | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 301 | mk_if (HOLogic.mk_not assm, none_t, | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 302 | mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms) genuine) | 
| 42306 | 303 | in | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 304 | case assms of [] => mk_test (vars_of concl, mk_if (concl, none_t, return) genuine) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 305 | | assm :: assms => | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 306 | if Config.get ctxt optimise_equality then | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 307 | (case try HOLogic.dest_eq assm of | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 308 | SOME (lhs, rhs) => | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 309 | mk_equality_term (lhs, rhs) (mk_equality_term (rhs, lhs) default) (assm, assms) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 310 | | NONE => default (assm, assms)) | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 311 | else default (assm, assms) | 
| 42306 | 312 | end | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 313 | val mk_smart_test_term = | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 314 | Quickcheck_Common.strip_imp #> (fn (assms, concl) => mk_smart_test_term' concl [] assms true) | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 315 | in | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 316 | if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 317 | end | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 318 | |
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 319 | fun mk_fast_generator_expr ctxt (t, eval_terms) = | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 320 | let | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 321 | val ctxt' = Variable.auto_fixes t ctxt | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 322 | val names = Term.add_free_names t [] | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 323 | val frees = map Free (Term.add_frees t []) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 324 | fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) | 
| 46306 | 325 | val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt' | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 326 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 327 |     fun return _ = @{term "throw_Counterexample :: term list => unit"} $
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 328 |       (HOLogic.mk_list @{typ "term"}
 | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 329 | (map (fn t => HOLogic.mk_term_of (fastype_of t) t) (frees @ eval_terms))) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 330 | fun mk_exhaustive_closure (free as Free (_, T)) t = | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 331 |       Const (@{const_name "Quickcheck_Exhaustive.fast_exhaustive_class.fast_exhaustive"},
 | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 332 | fast_exhaustiveT T) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 333 | $ lambda free t $ depth | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 334 |     val none_t = @{term "()"}
 | 
| 45761 
90028fd2f1fa
exhaustive returns if a counterexample is genuine or potentially spurious in the presence of assumptions more correctly
 bulwahn parents: 
45754diff
changeset | 335 | fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine) | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 336 | fun mk_let _ def v_opt t e = mk_let_expr (the_default def v_opt, t, e) | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 337 | val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_safe_if mk_let none_t return ctxt | 
| 42306 | 338 |   in lambda depth (@{term "catch_Counterexample :: unit => term list option"} $ mk_test_term t) end
 | 
| 339 | ||
| 45688 
d4ac5e090cd9
also potential counterexamples in the simple exhaustive testing in quickcheck
 bulwahn parents: 
45684diff
changeset | 340 | fun mk_unknown_term T = HOLogic.reflect_term (Const ("Quickcheck_Exhaustive.unknown", T))
 | 
| 
d4ac5e090cd9
also potential counterexamples in the simple exhaustive testing in quickcheck
 bulwahn parents: 
45684diff
changeset | 341 | |
| 45732 
ac5775bbc748
removing catch_match' now that catch_match is polymorphic
 bulwahn parents: 
45728diff
changeset | 342 | fun mk_safe_term t = @{term "Quickcheck.catch_match :: term => term => term"}
 | 
| 45688 
d4ac5e090cd9
also potential counterexamples in the simple exhaustive testing in quickcheck
 bulwahn parents: 
45684diff
changeset | 343 | $ (HOLogic.mk_term_of (fastype_of t) t) $ mk_unknown_term (fastype_of t) | 
| 
d4ac5e090cd9
also potential counterexamples in the simple exhaustive testing in quickcheck
 bulwahn parents: 
45684diff
changeset | 344 | |
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 345 | fun mk_return t genuine = @{term "Some :: bool * term list => (bool * term list) option"} $
 | 
| 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 346 |   (HOLogic.pair_const @{typ bool} @{typ "term list"} $ Quickcheck_Common.reflect_bool genuine $ t)
 | 
| 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 347 | |
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 348 | fun mk_generator_expr ctxt (t, eval_terms) = | 
| 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 349 | let | 
| 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 350 | val ctxt' = Variable.auto_fixes t ctxt | 
| 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 351 | val names = Term.add_free_names t [] | 
| 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 352 | val frees = map Free (Term.add_frees t []) | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 353 | fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) | 
| 46306 | 354 | val ([depth_name, genuine_only_name], _) = | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 355 | Variable.variant_fixes ["depth", "genuine_only"] ctxt' | 
| 42304 | 356 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 357 |     val genuine_only = Free (genuine_only_name, @{typ bool}) 
 | 
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 358 |     val return = mk_return (HOLogic.mk_list @{typ "term"}
 | 
| 45688 
d4ac5e090cd9
also potential counterexamples in the simple exhaustive testing in quickcheck
 bulwahn parents: 
45684diff
changeset | 359 | (map (fn t => HOLogic.mk_term_of (fastype_of t) t) frees @ map mk_safe_term eval_terms)) | 
| 42304 | 360 | fun mk_exhaustive_closure (free as Free (_, T)) t = | 
| 361 |       Const (@{const_name "Quickcheck_Exhaustive.exhaustive_class.exhaustive"}, exhaustiveT T)
 | |
| 362 | $ lambda free t $ depth | |
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 363 |     val none_t = Const (@{const_name "None"}, resultT)
 | 
| 45763 
3bb2bdf654f7
random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
 bulwahn parents: 
45761diff
changeset | 364 | val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 365 | fun mk_let safe def v_opt t e = mk_safe_let_expr genuine_only none_t safe (the_default def v_opt, t, e) | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 366 | val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_if mk_let none_t return ctxt | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 367 | in lambda genuine_only (lambda depth (mk_test_term t)) end | 
| 42304 | 368 | |
| 369 | fun mk_full_generator_expr ctxt (t, eval_terms) = | |
| 370 | let | |
| 42361 | 371 | val thy = Proof_Context.theory_of ctxt | 
| 42304 | 372 | val ctxt' = Variable.auto_fixes t ctxt | 
| 373 | val names = Term.add_free_names t [] | |
| 374 | val frees = map Free (Term.add_frees t []) | |
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 375 | val ([depth_name, genuine_only_name], ctxt'') = | 
| 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 376 | Variable.variant_fixes ["depth", "genuine_only"] ctxt' | 
| 46306 | 377 | val (term_names, _) = Variable.variant_fixes (map (prefix "t_") names) ctxt'' | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 378 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 46306 | 379 |     val genuine_only = Free (genuine_only_name, @{typ bool})    
 | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 380 |     val term_vars = map (fn n => Free (n, @{typ "unit => term"})) term_names
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 381 | fun lookup v = the (AList.lookup (op =) (names ~~ (frees ~~ term_vars)) v) | 
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 382 |     val return = mk_return (HOLogic.mk_list @{typ term}
 | 
| 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 383 |           (map (fn v => v $ @{term "()"}) term_vars @ map mk_safe_term eval_terms))
 | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 384 | fun mk_exhaustive_closure (free as Free (_, T), term_var) t = | 
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 385 |       if Sign.of_sort thy (T, @{sort check_all}) then
 | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 386 |         Const (@{const_name "Quickcheck_Exhaustive.check_all_class.check_all"}, check_allT T)
 | 
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 387 |           $ (HOLogic.split_const (T, @{typ "unit => term"}, resultT)
 | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 388 | $ lambda free (lambda term_var t)) | 
| 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 389 | else | 
| 42314 
8dfb7878a351
correcting constant name in exhaustive_generators
 bulwahn parents: 
42313diff
changeset | 390 |         Const (@{const_name "Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive"}, full_exhaustiveT T)
 | 
| 45722 | 391 |           $ (HOLogic.split_const (T, @{typ "unit => term"}, resultT)
 | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 392 | $ lambda free (lambda term_var t)) $ depth | 
| 45722 | 393 |     val none_t = Const (@{const_name "None"}, resultT)
 | 
| 45763 
3bb2bdf654f7
random reporting compilation returns if counterexample is genuine or potentially spurious, and takes genuine_only option as argument
 bulwahn parents: 
45761diff
changeset | 394 | val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t | 
| 48414 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 395 | fun mk_let safe _ (SOME (v, term_var)) t e = | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 396 | mk_safe_let_expr genuine_only none_t safe (v, t, | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 397 |           e #> subst_free [(term_var, absdummy @{typ unit} (mk_safe_term t))])
 | 
| 
43875bab3a4c
handling partiality in the case where the equality optimisation is applied
 bulwahn parents: 
48273diff
changeset | 398 | | mk_let safe v NONE t e = mk_safe_let_expr genuine_only none_t safe (v, t, e) | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 399 | val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_if mk_let none_t return ctxt | 
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 400 | in lambda genuine_only (lambda depth (mk_test_term t)) end | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 401 | |
| 42304 | 402 | fun mk_parametric_generator_expr mk_generator_expr = | 
| 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: 
42028diff
changeset | 403 | Quickcheck_Common.gen_mk_parametric_generator_expr | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 404 | ((mk_generator_expr, | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 405 |       absdummy @{typ bool} (absdummy @{typ "code_numeral"} (Const (@{const_name "None"}, resultT)))),
 | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 406 |       @{typ bool} --> @{typ "code_numeral"} --> resultT)
 | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 407 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 408 | fun mk_validator_expr ctxt t = | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 409 | let | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 410 |     fun bounded_forallT T = (T --> @{typ bool}) --> @{typ code_numeral} --> @{typ bool}
 | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 411 | val ctxt' = Variable.auto_fixes t ctxt | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 412 | val names = Term.add_free_names t [] | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 413 | val frees = map Free (Term.add_frees t []) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 414 | fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) | 
| 46306 | 415 | val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt' | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 416 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 417 | fun mk_bounded_forall (Free (s, T)) t = | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 418 |       Const (@{const_name "Quickcheck_Exhaustive.bounded_forall_class.bounded_forall"}, bounded_forallT T)
 | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 419 | $ lambda (Free (s, T)) t $ depth | 
| 45761 
90028fd2f1fa
exhaustive returns if a counterexample is genuine or potentially spurious in the presence of assumptions more correctly
 bulwahn parents: 
45754diff
changeset | 420 | fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine) | 
| 50046 | 421 | fun mk_let _ def v_opt t e = mk_let_expr (the_default def v_opt, t, e) | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 422 | val mk_test_term = | 
| 48013 
44de84112a67
added optimisation for equational premises in Quickcheck; added some Quickcheck examples; NEWS
 bulwahn parents: 
46565diff
changeset | 423 |       mk_test_term lookup mk_bounded_forall mk_safe_if mk_let @{term True} (K @{term False}) ctxt
 | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 424 | in lambda depth (mk_test_term t) end | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 425 | |
| 42391 | 426 | |
| 427 | fun mk_bounded_forall_generator_expr ctxt (t, eval_terms) = | |
| 428 | let | |
| 429 | val frees = Term.add_free_names t [] | |
| 430 |     val dummy_term = @{term "Code_Evaluation.Const (STR ''dummy_pattern'')
 | |
| 431 | (Typerep.Typerep (STR ''dummy'') [])"} | |
| 432 |     val return = @{term "Some :: term list => term list option"} $
 | |
| 433 |       (HOLogic.mk_list @{typ "term"}
 | |
| 434 | (replicate (length frees + length eval_terms) dummy_term)) | |
| 44241 | 435 |     val wrap = absdummy @{typ bool}
 | 
| 436 |       (@{term "If :: bool => term list option => term list option => term list option"} $
 | |
| 437 |         Bound 0 $ @{term "None :: term list option"} $ return)
 | |
| 42391 | 438 | in HOLogic.mk_comp (wrap, mk_validator_expr ctxt t) end | 
| 439 | ||
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 440 | (** generator compiliation **) | 
| 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 441 | |
| 46042 | 442 | (* FIXME just one data slot (record) per program unit *) | 
| 443 | ||
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41178diff
changeset | 444 | structure Counterexample = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41178diff
changeset | 445 | ( | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 446 | type T = unit -> int -> bool -> int -> (bool * term list) option | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41178diff
changeset | 447 | (* FIXME avoid user error with non-user text *) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 448 | fun init _ () = error "Counterexample" | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 449 | ); | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 450 | val put_counterexample = Counterexample.put; | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 451 | |
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 452 | structure Counterexample_Batch = Proof_Data | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 453 | ( | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 454 | type T = unit -> (int -> term list option) list | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 455 | (* FIXME avoid user error with non-user text *) | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 456 | fun init _ () = error "Counterexample" | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 457 | ); | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 458 | val put_counterexample_batch = Counterexample_Batch.put; | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 459 | |
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 460 | structure Validator_Batch = Proof_Data | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 461 | ( | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 462 | type T = unit -> (int -> bool) list | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 463 | (* FIXME avoid user error with non-user text *) | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 464 | fun init _ () = error "Counterexample" | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 465 | ); | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 466 | val put_validator_batch = Validator_Batch.put; | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 467 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 468 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 469 | val target = "Quickcheck"; | 
| 42391 | 470 | |
| 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: 
42028diff
changeset | 471 | fun compile_generator_expr 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: 
42028diff
changeset | 472 | let | 
| 42361 | 473 | val thy = Proof_Context.theory_of ctxt | 
| 42306 | 474 | val mk_generator_expr = | 
| 475 | if Config.get ctxt fast then mk_fast_generator_expr | |
| 42391 | 476 | else if Config.get ctxt bounded_forall then mk_bounded_forall_generator_expr | 
| 42306 | 477 | else if Config.get ctxt full_support then mk_full_generator_expr else mk_generator_expr | 
| 42304 | 478 | val t' = mk_parametric_generator_expr mk_generator_expr ctxt ts; | 
| 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: 
42028diff
changeset | 479 | val compile = Code_Runtime.dynamic_value_strict | 
| 
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: 
42028diff
changeset | 480 | (Counterexample.get, put_counterexample, "Exhaustive_Generators.put_counterexample") | 
| 
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: 
42028diff
changeset | 481 | thy (SOME target) (fn proc => fn g => | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 482 | fn card => fn genuine_only => fn size => g card genuine_only size | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 483 | |> (Option.map o apsnd o map) proc) t' [] | 
| 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: 
42028diff
changeset | 484 | in | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 485 | fn genuine_only => fn [card, size] => rpair NONE (compile card genuine_only size |> | 
| 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: 
42028diff
changeset | 486 | (if Config.get ctxt quickcheck_pretty then | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 487 | Option.map (apsnd (map Quickcheck_Common.post_process_term)) else I)) | 
| 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: 
42028diff
changeset | 488 | end; | 
| 42391 | 489 | |
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 490 | fun compile_generator_exprs ctxt ts = | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 491 | let | 
| 42361 | 492 | val thy = Proof_Context.theory_of ctxt | 
| 42028 
bd6515e113b2
passing a term with free variables to the quickcheck tester functions instead of an lambda expression because this is more natural with passing further evaluation terms; added output of evaluation terms; added evaluation of terms in the exhaustive testing
 bulwahn parents: 
42027diff
changeset | 493 | val ts' = map (fn t => mk_generator_expr ctxt (t, [])) ts; | 
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 494 | val compiles = Code_Runtime.dynamic_value_strict | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 495 | (Counterexample_Batch.get, put_counterexample_batch, | 
| 41918 | 496 | "Exhaustive_Generators.put_counterexample_batch") | 
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 497 | thy (SOME target) (fn proc => map (fn g => g #> (Option.map o map) proc)) | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 498 |       (HOLogic.mk_list @{typ "code_numeral => term list option"} ts') []
 | 
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 499 | in | 
| 41935 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41928diff
changeset | 500 | map (fn compile => fn size => compile size | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41928diff
changeset | 501 | |> Option.map (map Quickcheck_Common.post_process_term)) compiles | 
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 502 | end; | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 503 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 504 | fun compile_validator_exprs ctxt ts = | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 505 | let | 
| 42361 | 506 | val thy = Proof_Context.theory_of ctxt | 
| 42273 | 507 | val ts' = map (mk_validator_expr ctxt) ts | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 508 | in | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 509 | Code_Runtime.dynamic_value_strict | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 510 | (Validator_Batch.get, put_validator_batch, "Exhaustive_Generators.put_validator_batch") | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 511 |       thy (SOME target) (K I) (HOLogic.mk_list @{typ "code_numeral => bool"} ts') []
 | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 512 | end; | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 513 | |
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 514 | fun size_matters_for thy Ts = not (forall (fn T => Sign.of_sort thy (T,  @{sort check_all})) Ts)
 | 
| 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 515 | |
| 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 516 | val test_goals = | 
| 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 517 |   Quickcheck_Common.generator_test_goal_terms ("exhaustive", (size_matters_for, compile_generator_expr));
 | 
| 43875 
485d2ad43528
adding random, exhaustive and SML quickcheck as testers
 bulwahn parents: 
42616diff
changeset | 518 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 519 | (* setup *) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 520 | |
| 45484 
23871e17dddb
setting up exhaustive generators which are used for the smart generators
 bulwahn parents: 
45420diff
changeset | 521 | val setup_exhaustive_datatype_interpretation = | 
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45763diff
changeset | 522 |   Quickcheck_Common.datatype_interpretation (@{sort exhaustive}, instantiate_exhaustive_datatype)
 | 
| 45484 
23871e17dddb
setting up exhaustive generators which are used for the smart generators
 bulwahn parents: 
45420diff
changeset | 523 | |
| 48178 
0192811f0a96
exporting important function for the "many conjecture refutation" compilation of quickcheck
 bulwahn parents: 
48013diff
changeset | 524 | val setup_bounded_forall_datatype_interpretation = | 
| 
0192811f0a96
exporting important function for the "many conjecture refutation" compilation of quickcheck
 bulwahn parents: 
48013diff
changeset | 525 | Datatype.interpretation (Quickcheck_Common.ensure_sort | 
| 
0192811f0a96
exporting important function for the "many conjecture refutation" compilation of quickcheck
 bulwahn parents: 
48013diff
changeset | 526 |     (((@{sort type}, @{sort type}), @{sort bounded_forall}),
 | 
| 
0192811f0a96
exporting important function for the "many conjecture refutation" compilation of quickcheck
 bulwahn parents: 
48013diff
changeset | 527 | (Datatype.the_descr, instantiate_bounded_forall_datatype))) | 
| 
0192811f0a96
exporting important function for the "many conjecture refutation" compilation of quickcheck
 bulwahn parents: 
48013diff
changeset | 528 | |
| 43878 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43877diff
changeset | 529 | val active = Attrib.setup_config_bool @{binding quickcheck_exhaustive_active} (K true);
 | 
| 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43877diff
changeset | 530 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 531 | val setup = | 
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45763diff
changeset | 532 |   Quickcheck_Common.datatype_interpretation (@{sort full_exhaustive}, instantiate_full_exhaustive_datatype)
 | 
| 42316 
12635bb655fd
deactivating other compilations in quickcheck_exhaustive momentarily that only interesting for my benchmarks and experiments
 bulwahn parents: 
42315diff
changeset | 533 | (* #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype | 
| 42229 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 534 |       (((@{sort typerep}, @{sort term_of}), @{sort exhaustive}), instantiate_exhaustive_datatype))
 | 
| 42230 
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
 bulwahn parents: 
42229diff
changeset | 535 | #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype | 
| 42306 | 536 |       (((@{sort typerep}, @{sort term_of}), @{sort fast_exhaustive}), instantiate_fast_exhaustive_datatype))
 | 
| 42310 
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
 bulwahn parents: 
42309diff
changeset | 537 | #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype | 
| 42316 
12635bb655fd
deactivating other compilations in quickcheck_exhaustive momentarily that only interesting for my benchmarks and experiments
 bulwahn parents: 
42315diff
changeset | 538 |       (((@{sort type}, @{sort type}), @{sort bounded_forall}), instantiate_bounded_forall_datatype))*)
 | 
| 43878 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43877diff
changeset | 539 |   #> Context.theory_map (Quickcheck.add_tester ("exhaustive", (active, test_goals)))
 | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 540 |   #> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs))
 | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 541 |   #> Context.theory_map (Quickcheck.add_batch_validator ("exhaustive", compile_validator_exprs));
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 542 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 543 | end; |