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