| author | wenzelm | 
| Tue, 03 Apr 2012 16:49:05 +0200 | |
| changeset 47293 | 052cd5f1a591 | 
| parent 46565 | ad21900e0ee9 | 
| child 48013 | 44de84112a67 | 
| 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 | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 20 | val quickcheck_pretty : bool Config.T | 
| 45484 
23871e17dddb
setting up exhaustive generators which are used for the smart generators
 bulwahn parents: 
45420diff
changeset | 21 | val setup_exhaustive_datatype_interpretation : theory -> theory | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 22 | val setup: theory -> theory | 
| 45924 
f03dd48829d8
exporting instantiation functions in quickcheck for their usage in abstract generators
 bulwahn parents: 
45923diff
changeset | 23 | |
| 
f03dd48829d8
exporting instantiation functions in quickcheck for their usage in abstract generators
 bulwahn parents: 
45923diff
changeset | 24 | 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 | 25 | (string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory | 
| 46565 | 26 | val instantiate_exhaustive_datatype : Datatype_Aux.config -> Datatype_Aux.descr -> | 
| 27 | (string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory | |
| 28 | ||
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 29 | end; | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 30 | |
| 41918 | 31 | 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 | 32 | struct | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 33 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 34 | (* basics *) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 35 | |
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 36 | (** dynamic options **) | 
| 40907 | 37 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 38 | val smart_quantifier = Attrib.setup_config_bool @{binding quickcheck_smart_quantifier} (K true)
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42391diff
changeset | 39 | 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 | 40 | 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 | 41 | 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 | 42 | 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 | 43 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 44 | (** general term functions **) | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 45 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 46 | fun mk_measure f = | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 47 | let | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 48 |     val Type ("fun", [T, @{typ nat}]) = fastype_of f 
 | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 49 | in | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 50 |     Const (@{const_name Wellfounded.measure},
 | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 51 |       (T --> @{typ nat}) --> HOLogic.mk_prodT (T, T) --> @{typ bool})
 | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 52 | $ f | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 53 | end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 54 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 55 | fun mk_sumcases rT f (Type (@{type_name Sum_Type.sum}, [TL, TR])) =
 | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 56 | let | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 57 | val lt = mk_sumcases rT f TL | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 58 | val rt = mk_sumcases rT f TR | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 59 | in | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 60 | SumTree.mk_sumcase TL TR rT lt rt | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 61 | end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 62 | | mk_sumcases _ f T = f T | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 63 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 64 | (** abstract syntax **) | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 65 | |
| 40639 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 66 | 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 | 67 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 68 | 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 | 69 | 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 | 70 | val size_ge_zero = @{term "(i :: code_numeral) > 0"}
 | 
| 42304 | 71 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 72 | 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 | 73 | let | 
| 46306 | 74 |     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 | 75 | in | 
| 42214 
9ca13615c619
refactoring generator definition in quickcheck and removing clone
 bulwahn parents: 
42195diff
changeset | 76 |     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 | 77 | end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 78 | |
| 42306 | 79 | fun mk_unit_let (x, y) = | 
| 44241 | 80 |   Const (@{const_name "Let"}, @{typ "unit => (unit => unit) => unit"}) $ x $ absdummy @{typ unit} y
 | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 81 | |
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 82 | 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 | 83 | let | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 84 | 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 | 85 |   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 | 86 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 87 | (* handling inductive datatypes *) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 88 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 89 | (** constructing generator instances **) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 90 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 91 | exception FUNCTION_TYPE; | 
| 42306 | 92 | |
| 93 | exception Counterexample of term list | |
| 94 | ||
| 45722 | 95 | val resultT =  @{typ "(bool * term list) option"};
 | 
| 96 | ||
| 41918 | 97 | val exhaustiveN = "exhaustive"; | 
| 42304 | 98 | val full_exhaustiveN = "full_exhaustive"; | 
| 42306 | 99 | val fast_exhaustiveN = "fast_exhaustive"; | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 100 | val bounded_forallN = "bounded_forall"; | 
| 42306 | 101 | |
| 102 | fun fast_exhaustiveT T = (T --> @{typ unit})
 | |
| 103 |   --> @{typ code_numeral} --> @{typ unit}
 | |
| 104 | ||
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 105 | fun exhaustiveT T = (T --> resultT) --> @{typ code_numeral} --> resultT
 | 
| 42304 | 106 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 107 | 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 | 108 | |
| 45722 | 109 | 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 | 110 | |
| 45722 | 111 | fun check_allT T = (termifyT T --> resultT) --> resultT | 
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
40913diff
changeset | 112 | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 113 | fun mk_equation_terms generics (descr, vs, Ts) = | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 114 | let | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 115 | 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 | 116 | val rhss = | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 117 | Datatype_Aux.interpret_construction descr vs | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 118 |         { atyp = mk_call, dtyp = mk_aux_call }
 | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 119 | |> (map o apfst) Type | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 120 | |> map (fn (T, cs) => map (mk_consexpr T) cs) | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 121 | |> map mk_rhs | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 122 | 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 | 123 | in | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 124 | 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 | 125 | end | 
| 42306 | 126 | |
| 44241 | 127 | 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 | 128 | |
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 129 | 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 | 130 | let | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 131 | val T = Type (tyco, Ts) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 132 | 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 | 133 | in | 
| 44241 | 134 | (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 | 135 | end | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 136 | |
| 46306 | 137 | 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 | 138 | let | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 139 | val (Ts, fns) = split_list xs | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 140 | val constr = Const (c, Ts ---> simpleT) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 141 | val bounds = map Bound (((length xs) - 1) downto 0) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 142 | 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 | 143 | in fold_rev (fn f => fn t => f t) fns start_term end | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 144 | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 145 | fun mk_fast_equations functerms = | 
| 42306 | 146 | let | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 147 |     fun test_function T = Free ("f", T --> @{typ "unit"})
 | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 148 | val mk_call = gen_mk_call (fn T => | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 149 |       Const (@{const_name "Quickcheck_Exhaustive.fast_exhaustive_class.fast_exhaustive"},
 | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 150 | fast_exhaustiveT T)) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 151 | val mk_aux_call = gen_mk_aux_call functerms | 
| 46306 | 152 | val mk_consexpr = gen_mk_consexpr test_function | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 153 |     fun mk_rhs exprs = @{term "If :: bool => unit => unit => unit"}
 | 
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 154 |         $ size_ge_zero $ (foldr1 mk_unit_let exprs) $ @{term "()"}
 | 
| 42306 | 155 | in | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 156 | mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) | 
| 42306 | 157 | end | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 158 | |
| 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 159 | fun mk_equations functerms = | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 160 | let | 
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 161 |     fun test_function T = Free ("f", T --> resultT)
 | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 162 | val mk_call = gen_mk_call (fn T => | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 163 |       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 | 164 | val mk_aux_call = gen_mk_aux_call functerms | 
| 46306 | 165 | val mk_consexpr = gen_mk_consexpr test_function | 
| 42304 | 166 | fun mk_rhs exprs = | 
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 167 |       mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, Const (@{const_name "None"}, resultT))
 | 
| 42304 | 168 | in | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 169 | mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) | 
| 42304 | 170 | end | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 171 | |
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 172 | fun mk_bounded_forall_equations functerms = | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 173 | let | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 174 |     fun test_function T = Free ("P", T --> @{typ bool})
 | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 175 | val mk_call = gen_mk_call (fn T => | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 176 |       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 | 177 | bounded_forallT T)) | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 178 | val mk_aux_call = gen_mk_aux_call functerms | 
| 46306 | 179 | val mk_consexpr = gen_mk_consexpr test_function | 
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 180 | 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 | 181 |       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 | 182 | in | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 183 | 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 | 184 | end | 
| 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 185 | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 186 | fun mk_full_equations functerms = | 
| 42304 | 187 | let | 
| 45722 | 188 |     fun test_function T = Free ("f", termifyT T --> resultT)
 | 
| 189 |     fun split_const T = HOLogic.split_const (T, @{typ "unit => Code_Evaluation.term"}, resultT)
 | |
| 42304 | 190 | fun mk_call T = | 
| 191 | let | |
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 192 | val full_exhaustive = | 
| 42310 
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
 bulwahn parents: 
42309diff
changeset | 193 |           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 | 194 | 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 | 195 | in | 
| 44241 | 196 | (T, fn t => full_exhaustive $ | 
| 45722 | 197 |           (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 | 198 | end | 
| 41918 | 199 | 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 | 200 | let | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 201 | val T = Type (tyco, Ts) | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 202 | 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 | 203 | in | 
| 44241 | 204 | (T, fn t => nth functerms k $ | 
| 45722 | 205 |           (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 | 206 | end | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 207 | 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 | 208 | let | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 209 | 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 | 210 | val constr = Const (c, Ts ---> simpleT) | 
| 40639 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 211 | 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 | 212 |         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 | 213 |         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 | 214 |         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 | 215 | 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 | 216 | val start_term = test_function simpleT $ | 
| 40639 
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
 bulwahn parents: 
40420diff
changeset | 217 |         (HOLogic.pair_const simpleT @{typ "unit => Code_Evaluation.term"}
 | 
| 44241 | 218 |           $ (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 | 219 | 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 | 220 | 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 | 221 | mk_if (size_ge_zero, foldr1 mk_none_continuation exprs, | 
| 45722 | 222 |         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 | 223 | in | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 224 | 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 | 225 | end | 
| 42307 
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
 bulwahn parents: 
42306diff
changeset | 226 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 227 | (** foundational definition with the function package **) | 
| 40901 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 228 | |
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 229 | val less_int_pred = @{lemma "i > 0 ==> Code_Numeral.nat_of ((i :: code_numeral) - 1) < Code_Numeral.nat_of i" by auto}
 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 230 | |
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 231 | fun mk_single_measure T = HOLogic.mk_comp (@{term "Code_Numeral.nat_of"},
 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 232 |     Const (@{const_name "Product_Type.snd"}, T --> @{typ "code_numeral"}))
 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 233 | |
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 234 | fun mk_termination_measure T = | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 235 | let | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 236 | val T' = fst (HOLogic.dest_prodT (HOLogic.dest_setT T)) | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 237 | in | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 238 |     mk_measure (mk_sumcases @{typ nat} mk_single_measure T')
 | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 239 | end | 
| 40901 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 240 | |
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 241 | fun termination_tac ctxt = | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 242 | Function_Relation.relation_tac ctxt mk_termination_measure 1 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 243 |   THEN rtac @{thm wf_measure} 1
 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 244 | THEN (REPEAT_DETERM (Simplifier.asm_full_simp_tac | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 245 |     (HOL_basic_ss addsimps [@{thm in_measure}, @{thm o_def}, @{thm snd_conv},
 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 246 |      @{thm nat_mono_iff}, less_int_pred] @ @{thms sum.cases}) 1))
 | 
| 
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
 bulwahn parents: 
40899diff
changeset | 247 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 248 | (** instantiating generator classes **) | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 249 | |
| 42325 | 250 | fun contains_recursive_type_under_function_types xs = | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 251 | exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT => | 
| 42325 | 252 | (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 | 253 | |
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 254 | 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 | 255 | 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 | 256 | if not (contains_recursive_type_under_function_types descr) then | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 257 | let | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 258 |       val _ = Datatype_Aux.message config ("Creating " ^ name ^ "...")
 | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 259 | val fullnames = map (prefix (constprfx ^ "_")) (names @ auxnames) | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 260 | in | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 261 | thy | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 262 | |> Class.instantiation (tycos, vs, sort) | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 263 | |> Quickcheck_Common.define_functions | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 264 | (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 | 265 | prfx argnames fullnames (map mk_T (Ts @ Us)) | 
| 42312 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 266 | |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 267 | end | 
| 
5bf3b9612e43
ensuring datatype limitations before the instantiation in quickcheck_exhaustive
 bulwahn parents: 
42310diff
changeset | 268 | else | 
| 42230 
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
 bulwahn parents: 
42229diff
changeset | 269 | (Datatype_Aux.message config | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 270 |       ("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 | 271 | thy) | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 272 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 273 | val instantiate_bounded_forall_datatype = instantiate_datatype | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 274 |  ("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 | 275 | 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 | 276 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 277 | val instantiate_fast_exhaustive_datatype = instantiate_datatype | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 278 |  ("fast exhaustive generators", fast_exhaustiveN, @{sort fast_exhaustive},
 | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 279 | mk_fast_equations, fast_exhaustiveT, ["f", "i"]) | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 280 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 281 | val instantiate_exhaustive_datatype = instantiate_datatype | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 282 |   ("exhaustive generators", exhaustiveN, @{sort exhaustive},
 | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 283 | mk_equations, exhaustiveT, ["f", "i"]) | 
| 42315 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 284 | |
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 285 | val instantiate_full_exhaustive_datatype = instantiate_datatype | 
| 
95dfa082065a
generalizing instantiate_datatype in quickcheck_exhaustive to remove clones for different compilations
 bulwahn parents: 
42314diff
changeset | 286 |   ("full exhaustive generators", full_exhaustiveN, @{sort full_exhaustive},
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 287 | mk_full_equations, full_exhaustiveT, ["f", "i"]) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 288 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 289 | (* 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 | 290 | |
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 291 | |
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 292 | fun mk_test_term lookup mk_closure mk_if none_t return ctxt = | 
| 42306 | 293 | let | 
| 294 | 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 | 295 | 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 | 296 | (mk_if (t, none_t, return) true) | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 297 | fun mk_smart_test_term' concl bound_vars assms genuine = | 
| 42306 | 298 | let | 
| 299 | fun vars_of t = subtract (op =) bound_vars (Term.add_free_names t []) | |
| 300 | val (vars, check) = | |
| 301 | case assms of [] => (vars_of concl, (concl, none_t, return)) | |
| 45683 
805a2acf47fd
quickcheck returns counterexamples that are potentially spurious due to underspecified code equations and match exceptions
 bulwahn parents: 
45484diff
changeset | 302 | | assm :: assms => (vars_of assm, (HOLogic.mk_not assm, none_t, | 
| 
805a2acf47fd
quickcheck returns counterexamples that are potentially spurious due to underspecified code equations and match exceptions
 bulwahn parents: 
45484diff
changeset | 303 | mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms)) | 
| 42306 | 304 | in | 
| 45761 
90028fd2f1fa
exhaustive returns if a counterexample is genuine or potentially spurious in the presence of assumptions more correctly
 bulwahn parents: 
45754diff
changeset | 305 | fold_rev mk_closure (map lookup vars) (mk_if check genuine) | 
| 42306 | 306 | end | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 307 | 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 | 308 | 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 | 309 | in | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 310 | if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 311 | end | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 312 | |
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 313 | fun mk_fast_generator_expr ctxt (t, eval_terms) = | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 314 | let | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 315 | val ctxt' = Variable.auto_fixes t ctxt | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 316 | val names = Term.add_free_names t [] | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 317 | val frees = map Free (Term.add_frees t []) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 318 | fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) | 
| 46306 | 319 | val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt' | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 320 |     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 | 321 |     fun return _ = @{term "throw_Counterexample :: term list => unit"} $
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 322 |       (HOLogic.mk_list @{typ "term"}
 | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 323 | (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 | 324 | fun mk_exhaustive_closure (free as Free (_, T)) t = | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 325 |       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 | 326 | fast_exhaustiveT T) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 327 | $ lambda free t $ depth | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 328 |     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 | 329 | fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine) | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 330 | val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_safe_if none_t return ctxt | 
| 42306 | 331 |   in lambda depth (@{term "catch_Counterexample :: unit => term list option"} $ mk_test_term t) end
 | 
| 332 | ||
| 45688 
d4ac5e090cd9
also potential counterexamples in the simple exhaustive testing in quickcheck
 bulwahn parents: 
45684diff
changeset | 333 | 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 | 334 | |
| 45732 
ac5775bbc748
removing catch_match' now that catch_match is polymorphic
 bulwahn parents: 
45728diff
changeset | 335 | 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 | 336 | $ (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 | 337 | |
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 338 | 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 | 339 |   (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 | 340 | |
| 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 | 341 | 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 | 342 | 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 | 343 | 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 | 344 | 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 | 345 | val frees = map Free (Term.add_frees t []) | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 346 | fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) | 
| 46306 | 347 | val ([depth_name, genuine_only_name], _) = | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 348 | Variable.variant_fixes ["depth", "genuine_only"] ctxt' | 
| 42304 | 349 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 350 |     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 | 351 |     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 | 352 | (map (fn t => HOLogic.mk_term_of (fastype_of t) t) frees @ map mk_safe_term eval_terms)) | 
| 42304 | 353 | fun mk_exhaustive_closure (free as Free (_, T)) t = | 
| 354 |       Const (@{const_name "Quickcheck_Exhaustive.exhaustive_class.exhaustive"}, exhaustiveT T)
 | |
| 355 | $ lambda free t $ depth | |
| 45724 
1f5fc44254d7
the simple exhaustive compilation also indicates if counterexample is potentially spurious;
 bulwahn parents: 
45722diff
changeset | 356 |     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 | 357 | val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t | 
| 45718 
8979b2463fc8
quickcheck random can also find potential counterexamples;
 bulwahn parents: 
45688diff
changeset | 358 | val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_if none_t return ctxt | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 359 | in lambda genuine_only (lambda depth (mk_test_term t)) end | 
| 42304 | 360 | |
| 361 | fun mk_full_generator_expr ctxt (t, eval_terms) = | |
| 362 | let | |
| 42361 | 363 | val thy = Proof_Context.theory_of ctxt | 
| 42304 | 364 | val ctxt' = Variable.auto_fixes t ctxt | 
| 365 | val names = Term.add_free_names t [] | |
| 366 | val frees = map Free (Term.add_frees t []) | |
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 367 | val ([depth_name, genuine_only_name], ctxt'') = | 
| 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 368 | Variable.variant_fixes ["depth", "genuine_only"] ctxt' | 
| 46306 | 369 | 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 | 370 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 46306 | 371 |     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 | 372 |     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 | 373 | 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 | 374 |     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 | 375 |           (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 | 376 | 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 | 377 |       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 | 378 |         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 | 379 |           $ (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 | 380 | $ 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 | 381 | else | 
| 42314 
8dfb7878a351
correcting constant name in exhaustive_generators
 bulwahn parents: 
42313diff
changeset | 382 |         Const (@{const_name "Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive"}, full_exhaustiveT T)
 | 
| 45722 | 383 |           $ (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 | 384 | $ lambda free (lambda term_var t)) $ depth | 
| 45722 | 385 |     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 | 386 | val mk_if = Quickcheck_Common.mk_safe_if genuine_only none_t | 
| 45718 
8979b2463fc8
quickcheck random can also find potential counterexamples;
 bulwahn parents: 
45688diff
changeset | 387 | val mk_test_term = mk_test_term lookup mk_exhaustive_closure mk_if none_t return ctxt | 
| 45753 
196697f71488
indicating where the restart should occur; making safe_if dynamic
 bulwahn parents: 
45732diff
changeset | 388 | 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 | 389 | |
| 42304 | 390 | 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 | 391 | Quickcheck_Common.gen_mk_parametric_generator_expr | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 392 | ((mk_generator_expr, | 
| 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 393 |       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 | 394 |       @{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 | 395 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 396 | 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 | 397 | let | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 398 |     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 | 399 | val ctxt' = Variable.auto_fixes t ctxt | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 400 | val names = Term.add_free_names t [] | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 401 | val frees = map Free (Term.add_frees t []) | 
| 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 402 | fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) | 
| 46306 | 403 | 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 | 404 |     val depth = Free (depth_name, @{typ code_numeral})
 | 
| 42390 
be7af468f7b3
creating generic test_term function; corrected instantiate_exhaustive_datatype; tuned
 bulwahn parents: 
42361diff
changeset | 405 | 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 | 406 |       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 | 407 | $ 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 | 408 | fun mk_safe_if (cond, then_t, else_t) genuine = mk_if (cond, then_t, else_t genuine) | 
| 45721 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 409 | val mk_test_term = | 
| 
d1fb55c2ed65
quickcheck's compilation returns if it is genuine counterexample or a counterexample due to a match exception
 bulwahn parents: 
45718diff
changeset | 410 |       mk_test_term lookup mk_bounded_forall mk_safe_if @{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 | 411 | 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 | 412 | |
| 42391 | 413 | |
| 414 | fun mk_bounded_forall_generator_expr ctxt (t, eval_terms) = | |
| 415 | let | |
| 416 | val frees = Term.add_free_names t [] | |
| 417 |     val dummy_term = @{term "Code_Evaluation.Const (STR ''dummy_pattern'')
 | |
| 418 | (Typerep.Typerep (STR ''dummy'') [])"} | |
| 419 |     val return = @{term "Some :: term list => term list option"} $
 | |
| 420 |       (HOLogic.mk_list @{typ "term"}
 | |
| 421 | (replicate (length frees + length eval_terms) dummy_term)) | |
| 44241 | 422 |     val wrap = absdummy @{typ bool}
 | 
| 423 |       (@{term "If :: bool => term list option => term list option => term list option"} $
 | |
| 424 |         Bound 0 $ @{term "None :: term list option"} $ return)
 | |
| 42391 | 425 | in HOLogic.mk_comp (wrap, mk_validator_expr ctxt t) end | 
| 426 | ||
| 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 | 427 | (** 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 | 428 | |
| 46042 | 429 | (* FIXME just one data slot (record) per program unit *) | 
| 430 | ||
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41178diff
changeset | 431 | structure Counterexample = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41178diff
changeset | 432 | ( | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 433 | 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 | 434 | (* 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 | 435 | fun init _ () = error "Counterexample" | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 436 | ); | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 437 | val put_counterexample = Counterexample.put; | 
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 438 | |
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 439 | 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 | 440 | ( | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 441 | 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 | 442 | (* 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 | 443 | 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 | 444 | ); | 
| 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 445 | 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 | 446 | |
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 447 | 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 | 448 | ( | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 449 | 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 | 450 | (* 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 | 451 | 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 | 452 | ); | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 453 | 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 | 454 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 455 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 456 | val target = "Quickcheck"; | 
| 42391 | 457 | |
| 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 | 458 | 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 | 459 | let | 
| 42361 | 460 | val thy = Proof_Context.theory_of ctxt | 
| 42306 | 461 | val mk_generator_expr = | 
| 462 | if Config.get ctxt fast then mk_fast_generator_expr | |
| 42391 | 463 | else if Config.get ctxt bounded_forall then mk_bounded_forall_generator_expr | 
| 42306 | 464 | else if Config.get ctxt full_support then mk_full_generator_expr else mk_generator_expr | 
| 42304 | 465 | 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 | 466 | 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 | 467 | (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 | 468 | thy (SOME target) (fn proc => fn g => | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 469 | 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 | 470 | |> (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 | 471 | in | 
| 45754 
394ecd91434a
dynamic genuine_flag in compilation of random and exhaustive
 bulwahn parents: 
45753diff
changeset | 472 | 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 | 473 | (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 | 474 | 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 | 475 | end; | 
| 42391 | 476 | |
| 41861 
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
 bulwahn parents: 
41472diff
changeset | 477 | 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 | 478 | let | 
| 42361 | 479 | 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 | 480 | 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 | 481 | 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 | 482 | (Counterexample_Batch.get, put_counterexample_batch, | 
| 41918 | 483 | "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 | 484 | 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 | 485 |       (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 | 486 | in | 
| 41935 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41928diff
changeset | 487 | map (fn compile => fn size => compile size | 
| 
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
 bulwahn parents: 
41928diff
changeset | 488 | |> 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 | 489 | end; | 
| 42195 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 490 | |
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 491 | 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 | 492 | let | 
| 42361 | 493 | val thy = Proof_Context.theory_of ctxt | 
| 42273 | 494 | 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 | 495 | in | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 496 | 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 | 497 | (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 | 498 |       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 | 499 | end; | 
| 
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
 bulwahn parents: 
42176diff
changeset | 500 | |
| 46331 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 501 | 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 | 502 | |
| 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 503 | val test_goals = | 
| 
f5598b604a54
generalizing check if size matters because it is different for random and exhaustive testing
 bulwahn parents: 
46306diff
changeset | 504 |   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 | 505 | |
| 42308 
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
 bulwahn parents: 
42307diff
changeset | 506 | (* setup *) | 
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 507 | |
| 45484 
23871e17dddb
setting up exhaustive generators which are used for the smart generators
 bulwahn parents: 
45420diff
changeset | 508 | 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 | 509 |   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 | 510 | |
| 43878 
eeb10fdd9535
changed every tester to have a configuration in quickcheck; enabling parallel testing of different testers in quickcheck
 bulwahn parents: 
43877diff
changeset | 511 | 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 | 512 | |
| 40420 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 513 | val setup = | 
| 45923 
473b744c23f2
generalize ensure_sort_datatype to ensure_sort  in quickcheck_common to allow generators for abstract types;
 bulwahn parents: 
45763diff
changeset | 514 |   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 | 515 | (* #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype | 
| 42229 
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
 bulwahn parents: 
42214diff
changeset | 516 |       (((@{sort typerep}, @{sort term_of}), @{sort exhaustive}), instantiate_exhaustive_datatype))
 | 
| 42230 
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
 bulwahn parents: 
42229diff
changeset | 517 | #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype | 
| 42306 | 518 |       (((@{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 | 519 | #> 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 | 520 |       (((@{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 | 521 |   #> 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 | 522 |   #> 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 | 523 |   #> 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 | 524 | |
| 
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
 bulwahn parents: diff
changeset | 525 | end; |