author | bulwahn |
Fri, 08 Apr 2011 16:31:14 +0200 | |
changeset 42308 | e2abd1ca8d01 |
parent 42307 | 72e2fabb4bc2 |
child 42309 | 74ea14d13247 |
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: |
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
|
10 |
Proof.context -> (term * term list) list -> int list -> term list option * Quickcheck.report option |
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 |
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
|
13 |
val put_counterexample: (unit -> int -> int -> 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 |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
20 |
val quickcheck_pretty : bool Config.T |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
21 |
val setup: theory -> theory |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
22 |
end; |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
23 |
|
41918 | 24 |
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
|
25 |
struct |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
26 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
27 |
(* basics *) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
28 |
|
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
29 |
(** dynamic options **) |
40907 | 30 |
|
31 |
val (smart_quantifier, setup_smart_quantifier) = |
|
32 |
Attrib.config_bool "quickcheck_smart_quantifier" (K true) |
|
33 |
||
42306 | 34 |
val (fast, setup_fast) = |
35 |
Attrib.config_bool "quickcheck_fast" (K true) |
|
36 |
||
42304 | 37 |
val (full_support, setup_full_support) = |
38 |
Attrib.config_bool "quickcheck_full_support" (K true) |
|
39 |
||
41903
39fd77f0ae59
fixing postprocessing; adding a configuration to enable and disable pretty presentation of quickcheck's counterexample
bulwahn
parents:
41902
diff
changeset
|
40 |
val (quickcheck_pretty, setup_quickcheck_pretty) = |
39fd77f0ae59
fixing postprocessing; adding a configuration to enable and disable pretty presentation of quickcheck's counterexample
bulwahn
parents:
41902
diff
changeset
|
41 |
Attrib.config_bool "quickcheck_pretty" (K true) |
39fd77f0ae59
fixing postprocessing; adding a configuration to enable and disable pretty presentation of quickcheck's counterexample
bulwahn
parents:
41902
diff
changeset
|
42 |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
43 |
(** general term functions **) |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
44 |
|
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
45 |
fun mk_measure f = |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
46 |
let |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
47 |
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
|
48 |
in |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
49 |
Const (@{const_name Wellfounded.measure}, |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
50 |
(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
|
51 |
$ f |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
52 |
end |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
53 |
|
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
54 |
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
|
55 |
let |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
56 |
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
|
57 |
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
|
58 |
in |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
59 |
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
|
60 |
end |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
61 |
| 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
|
62 |
|
40901
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
63 |
fun mk_undefined T = Const(@{const_name undefined}, T) |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
64 |
|
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
65 |
(** abstract syntax **) |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
66 |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
67 |
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
|
68 |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
69 |
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
|
70 |
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
|
71 |
val size_ge_zero = @{term "(i :: code_numeral) > 0"} |
42304 | 72 |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
73 |
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
|
74 |
let |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
75 |
val (T as Type(@{type_name "option"}, [T'])) = fastype_of x |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
76 |
in |
42214
9ca13615c619
refactoring generator definition in quickcheck and removing clone
bulwahn
parents:
42195
diff
changeset
|
77 |
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
|
78 |
end |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
79 |
|
42306 | 80 |
fun mk_unit_let (x, y) = |
81 |
Const (@{const_name "Let"}, @{typ "unit => (unit => unit) => unit"}) $ x $ (absdummy (@{typ unit}, y)) |
|
82 |
||
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
83 |
(* handling inductive datatypes *) |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
84 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
85 |
(** constructing generator instances **) |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
86 |
|
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
87 |
exception FUNCTION_TYPE; |
42306 | 88 |
|
89 |
exception Counterexample of term list |
|
90 |
||
41918 | 91 |
val exhaustiveN = "exhaustive"; |
42304 | 92 |
val full_exhaustiveN = "full_exhaustive"; |
42306 | 93 |
val fast_exhaustiveN = "fast_exhaustive"; |
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
94 |
val bounded_forallN = "bounded_forall"; |
42306 | 95 |
|
96 |
fun fast_exhaustiveT T = (T --> @{typ unit}) |
|
97 |
--> @{typ code_numeral} --> @{typ unit} |
|
98 |
||
42304 | 99 |
fun exhaustiveT T = (T --> @{typ "Code_Evaluation.term list option"}) |
100 |
--> @{typ code_numeral} --> @{typ "Code_Evaluation.term list option"} |
|
101 |
||
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
102 |
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
|
103 |
|
42304 | 104 |
fun full_exhaustiveT T = (termifyT T --> @{typ "Code_Evaluation.term list option"}) |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
105 |
--> @{typ code_numeral} --> @{typ "Code_Evaluation.term list option"} |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40913
diff
changeset
|
106 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40913
diff
changeset
|
107 |
fun check_allT T = (termifyT T --> @{typ "Code_Evaluation.term list option"}) |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40913
diff
changeset
|
108 |
--> @{typ "Code_Evaluation.term list option"} |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40913
diff
changeset
|
109 |
|
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
110 |
fun mk_equation_terms generics (descr, vs, Ts) = |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
111 |
let |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
112 |
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
|
113 |
val rhss = |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
114 |
Datatype_Aux.interpret_construction descr vs |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
115 |
{ atyp = mk_call, dtyp = mk_aux_call } |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
116 |
|> (map o apfst) Type |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
117 |
|> 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
|
118 |
|> map mk_rhs |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
119 |
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
|
120 |
in |
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
121 |
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
|
122 |
end |
42306 | 123 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
124 |
fun gen_mk_call c T = (T, fn t => c T $ absdummy (T, t) $ size_pred) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
125 |
|
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
126 |
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
|
127 |
let |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
128 |
val T = Type (tyco, Ts) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
129 |
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
|
130 |
in |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
131 |
(T, fn t => nth functerms k $ absdummy (T, t) $ size_pred) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
132 |
end |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
133 |
|
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
134 |
fun gen_mk_consexpr test_function functerms simpleT (c, xs) = |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
135 |
let |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
136 |
val (Ts, fns) = split_list xs |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
137 |
val constr = Const (c, Ts ---> simpleT) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
138 |
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
|
139 |
val term_bounds = map (fn x => Bound (2 * x)) (((length xs) - 1) downto 0) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
140 |
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
|
141 |
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
|
142 |
|
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
143 |
fun mk_fast_equations functerms = |
42306 | 144 |
let |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
145 |
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
|
146 |
val mk_call = gen_mk_call (fn T => |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
147 |
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
|
148 |
fast_exhaustiveT T)) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
149 |
val mk_aux_call = gen_mk_aux_call functerms |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
150 |
val mk_consexpr = gen_mk_consexpr test_function functerms |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
151 |
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
|
152 |
$ size_ge_zero $ (foldr1 mk_unit_let exprs) $ @{term "()"} |
42306 | 153 |
in |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
154 |
mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) |
42306 | 155 |
end |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
156 |
|
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
157 |
fun mk_equations functerms = |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
158 |
let |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
159 |
fun test_function T = Free ("f", T --> @{typ "term list option"}) |
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
160 |
val mk_call = gen_mk_call (fn T => |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
161 |
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
|
162 |
val mk_aux_call = gen_mk_aux_call functerms |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
163 |
val mk_consexpr = gen_mk_consexpr test_function functerms |
42304 | 164 |
fun mk_rhs exprs = |
165 |
@{term "If :: bool => term list option => term list option => term list option"} |
|
166 |
$ size_ge_zero $ (foldr1 mk_none_continuation exprs) $ @{term "None :: term list option"} |
|
167 |
in |
|
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
168 |
mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) |
42304 | 169 |
end |
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
170 |
|
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
171 |
fun mk_bounded_forall_equations functerms = |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
172 |
let |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
173 |
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
|
174 |
val mk_call = gen_mk_call (fn T => |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
175 |
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
|
176 |
bounded_forallT T)) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
177 |
val mk_aux_call = gen_mk_aux_call functerms |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
178 |
val mk_consexpr = gen_mk_consexpr test_function functerms |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
179 |
fun mk_rhs exprs = |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
180 |
@{term "If :: bool => bool => bool => bool"} $ size_ge_zero $ |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
181 |
(foldr1 HOLogic.mk_conj exprs) $ @{term "True"} |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
182 |
in |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
183 |
mk_equation_terms (mk_call, mk_aux_call, mk_consexpr, mk_rhs, test_function, functerms) |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
184 |
end |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
185 |
|
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
186 |
fun mk_full_equations functerms = |
42304 | 187 |
let |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
188 |
fun test_function T = Free ("f", termifyT T --> @{typ "term list option"}) |
42304 | 189 |
fun mk_call T = |
190 |
let |
|
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
191 |
val full_exhaustive = |
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
192 |
Const (@{const_name "Quickcheck_Exhaustive.exhaustive_class.full_exhaustive"}, |
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
193 |
full_exhaustiveT T) |
42304 | 194 |
in |
195 |
(T, (fn t => full_exhaustive $ |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
196 |
(HOLogic.split_const (T, @{typ "unit => Code_Evaluation.term"}, @{typ "Code_Evaluation.term list option"}) |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
197 |
$ absdummy (T, absdummy (@{typ "unit => Code_Evaluation.term"}, t))) $ size_pred)) |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
198 |
end |
41918 | 199 |
fun mk_aux_call fTs (k, _) (tyco, Ts) = |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
200 |
let |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
201 |
val T = Type (tyco, Ts) |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
202 |
val _ = if not (null fTs) then raise FUNCTION_TYPE else () |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
203 |
in |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
204 |
(T, (fn t => nth functerms k $ |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
205 |
(HOLogic.split_const (T, @{typ "unit => Code_Evaluation.term"}, @{typ "Code_Evaluation.term list option"}) |
41918 | 206 |
$ 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
|
207 |
end |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
208 |
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
|
209 |
let |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
210 |
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
|
211 |
val constr = Const (c, Ts ---> simpleT) |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
212 |
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
|
213 |
val term_bounds = map (fn x => Bound (2 * x)) (((length xs) - 1) downto 0) |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
214 |
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
|
215 |
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
|
216 |
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
|
217 |
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
|
218 |
val start_term = test_function simpleT $ |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
219 |
(HOLogic.pair_const simpleT @{typ "unit => Code_Evaluation.term"} |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
220 |
$ (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
|
221 |
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
|
222 |
fun mk_rhs exprs = |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
223 |
@{term "If :: bool => term list option => term list option => term list option"} |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
224 |
$ size_ge_zero $ (foldr1 mk_none_continuation exprs) $ @{term "None :: term list option"} |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
225 |
in |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
226 |
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
|
227 |
end |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
228 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
229 |
(** 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
|
230 |
|
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
231 |
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
|
232 |
|
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
233 |
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
|
234 |
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
|
235 |
|
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
236 |
fun mk_termination_measure T = |
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
237 |
let |
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
238 |
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
|
239 |
in |
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
240 |
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
|
241 |
end |
40901
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
242 |
|
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
243 |
fun termination_tac ctxt = |
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
244 |
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
|
245 |
THEN rtac @{thm wf_measure} 1 |
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
246 |
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
|
247 |
(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
|
248 |
@{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
|
249 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
250 |
(** instantiating generator classes **) |
40901
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
251 |
|
41918 | 252 |
fun instantiate_exhaustive_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
253 |
let |
42230
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
254 |
val _ = Datatype_Aux.message config "Creating exhaustive generators..."; |
42304 | 255 |
val exhaustivesN = map (prefix (exhaustiveN ^ "_")) (names @ auxnames) |
256 |
val full_exhaustivesN = map (prefix (full_exhaustiveN ^ "_")) (names @ auxnames) |
|
40901
8fdfa9c4e7ed
smallvalue_generator are defined quick via oracle or sound via function package
bulwahn
parents:
40899
diff
changeset
|
257 |
in |
40641
5615cc557120
moving the error handling to the right scope in smallvalue_generators
bulwahn
parents:
40640
diff
changeset
|
258 |
thy |
41918 | 259 |
|> Class.instantiation (tycos, vs, @{sort exhaustive}) |
42214
9ca13615c619
refactoring generator definition in quickcheck and removing clone
bulwahn
parents:
42195
diff
changeset
|
260 |
|> Quickcheck_Common.define_functions |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
261 |
(fn functerms => mk_equations functerms (descr, vs, Ts @ Us), SOME termination_tac) |
42214
9ca13615c619
refactoring generator definition in quickcheck and removing clone
bulwahn
parents:
42195
diff
changeset
|
262 |
prfx ["f", "i"] exhaustivesN (map exhaustiveT (Ts @ Us)) |
42304 | 263 |
|> Quickcheck_Common.define_functions |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
264 |
(fn functerms => mk_full_equations functerms (descr, vs, Ts @ Us), SOME termination_tac) |
42304 | 265 |
prfx ["f", "i"] full_exhaustivesN (map full_exhaustiveT (Ts @ Us)) |
40641
5615cc557120
moving the error handling to the right scope in smallvalue_generators
bulwahn
parents:
40640
diff
changeset
|
266 |
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) |
5615cc557120
moving the error handling to the right scope in smallvalue_generators
bulwahn
parents:
40640
diff
changeset
|
267 |
end handle FUNCTION_TYPE => |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
268 |
(Datatype_Aux.message config |
41963
d8c3b26b3da4
tuned exhaustive generator compilation; added narrowing generator compilation; removed exec as does not work properly here (reverting changeset 994d088fbfbc)
bulwahn
parents:
41935
diff
changeset
|
269 |
"Creation of exhaustive generators failed because the datatype contains a function type"; |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40420
diff
changeset
|
270 |
thy) |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
271 |
|
42306 | 272 |
fun instantiate_fast_exhaustive_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = |
273 |
let |
|
274 |
val _ = Datatype_Aux.message config "Creating fast exhaustive generators..."; |
|
275 |
val fast_exhaustivesN = map (prefix (fast_exhaustiveN ^ "_")) (names @ auxnames) |
|
276 |
in |
|
277 |
thy |
|
278 |
|> Class.instantiation (tycos, vs, @{sort fast_exhaustive}) |
|
279 |
|> Quickcheck_Common.define_functions |
|
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
280 |
(fn functerms => mk_fast_equations functerms (descr, vs, Ts @ Us), SOME termination_tac) |
42306 | 281 |
prfx ["f", "i"] fast_exhaustivesN (map fast_exhaustiveT (Ts @ Us)) |
282 |
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) |
|
283 |
end handle FUNCTION_TYPE => |
|
284 |
(Datatype_Aux.message config |
|
285 |
"Creation of exhaustive generators failed because the datatype contains a function type"; |
|
286 |
thy) |
|
42230
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
287 |
|
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
288 |
fun instantiate_bounded_forall_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
289 |
let |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
290 |
val _ = Datatype_Aux.message config "Creating bounded universal quantifiers..."; |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
291 |
val bounded_forallsN = map (prefix (bounded_forallN ^ "_")) (names @ auxnames); |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
292 |
in |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
293 |
thy |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
294 |
|> Class.instantiation (tycos, vs, @{sort bounded_forall}) |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
295 |
|> Quickcheck_Common.define_functions |
42307
72e2fabb4bc2
creating a general mk_equation_terms for the different compilations
bulwahn
parents:
42306
diff
changeset
|
296 |
(fn functerms => mk_bounded_forall_equations functerms (descr, vs, Ts @ Us), NONE) |
42230
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
297 |
prfx ["P", "i"] bounded_forallsN (map bounded_forallT (Ts @ Us)) |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
298 |
|> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
299 |
end handle FUNCTION_TYPE => |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
300 |
(Datatype_Aux.message config |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
301 |
"Creation of bounded universal quantifiers failed because the datatype contains a function type"; |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
302 |
thy) |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
303 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
304 |
(* 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
|
305 |
|
42306 | 306 |
fun mk_fast_generator_expr ctxt (t, eval_terms) = |
307 |
let |
|
308 |
val thy = ProofContext.theory_of ctxt |
|
309 |
val ctxt' = Variable.auto_fixes t ctxt |
|
310 |
val names = Term.add_free_names t [] |
|
311 |
val frees = map Free (Term.add_frees t []) |
|
312 |
val ([depth_name], ctxt'') = Variable.variant_fixes ["depth"] ctxt' |
|
313 |
val depth = Free (depth_name, @{typ code_numeral}) |
|
314 |
val return = @{term "throw_Counterexample :: term list => unit"} $ |
|
315 |
(HOLogic.mk_list @{typ "term"} |
|
316 |
(map (fn t => HOLogic.mk_term_of (fastype_of t) t) (frees @ eval_terms))) |
|
317 |
fun mk_exhaustive_closure (free as Free (_, T)) t = |
|
318 |
Const (@{const_name "Quickcheck_Exhaustive.fast_exhaustive_class.fast_exhaustive"}, fast_exhaustiveT T) |
|
319 |
$ lambda free t $ depth |
|
320 |
val none_t = @{term "()"} |
|
321 |
fun mk_safe_if (cond, then_t, else_t) = |
|
322 |
@{term "If :: bool => unit => unit => unit"} $ cond $ then_t $ else_t |
|
323 |
fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) |
|
324 |
fun mk_naive_test_term t = |
|
325 |
fold_rev mk_exhaustive_closure frees (mk_safe_if (t, none_t, return)) |
|
326 |
fun mk_smart_test_term' concl bound_vars assms = |
|
327 |
let |
|
328 |
fun vars_of t = subtract (op =) bound_vars (Term.add_free_names t []) |
|
329 |
val (vars, check) = |
|
330 |
case assms of [] => (vars_of concl, (concl, none_t, return)) |
|
331 |
| assm :: assms => (vars_of assm, (assm, |
|
332 |
mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms, none_t)) |
|
333 |
in |
|
334 |
fold_rev mk_exhaustive_closure (map lookup vars) (mk_safe_if check) |
|
335 |
end |
|
336 |
fun mk_smart_test_term t = |
|
337 |
let |
|
338 |
val (assms, concl) = Quickcheck_Common.strip_imp t |
|
339 |
in |
|
340 |
mk_smart_test_term' concl [] assms |
|
341 |
end |
|
342 |
val mk_test_term = |
|
343 |
if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term |
|
344 |
in lambda depth (@{term "catch_Counterexample :: unit => term list option"} $ mk_test_term t) end |
|
345 |
||
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
|
346 |
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
|
347 |
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
|
348 |
val thy = ProofContext.theory_of 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
|
349 |
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
|
350 |
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
|
351 |
val frees = map Free (Term.add_frees 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
|
352 |
val ([depth_name], ctxt'') = Variable.variant_fixes ["depth"] ctxt' |
42304 | 353 |
val depth = Free (depth_name, @{typ code_numeral}) |
354 |
val return = @{term "Some :: term list => term list option"} $ |
|
355 |
(HOLogic.mk_list @{typ "term"} |
|
356 |
(map (fn t => HOLogic.mk_term_of (fastype_of t) t) (frees @ eval_terms))) |
|
357 |
fun mk_exhaustive_closure (free as Free (_, T)) t = |
|
358 |
Const (@{const_name "Quickcheck_Exhaustive.exhaustive_class.exhaustive"}, exhaustiveT T) |
|
359 |
$ lambda free t $ depth |
|
360 |
val none_t = @{term "None :: term list option"} |
|
361 |
fun mk_safe_if (cond, then_t, else_t) = |
|
362 |
@{term "Quickcheck_Exhaustive.catch_match :: term list option => term list option => term list option"} $ |
|
363 |
(@{term "If :: bool => term list option => term list option => term list option"} |
|
364 |
$ cond $ then_t $ else_t) $ none_t; |
|
365 |
fun lookup v = the (AList.lookup (op =) (names ~~ frees) v) |
|
366 |
fun mk_naive_test_term t = |
|
367 |
fold_rev mk_exhaustive_closure frees (mk_safe_if (t, none_t, return)) |
|
368 |
fun mk_smart_test_term' concl bound_vars assms = |
|
369 |
let |
|
370 |
fun vars_of t = subtract (op =) bound_vars (Term.add_free_names t []) |
|
371 |
val (vars, check) = |
|
372 |
case assms of [] => (vars_of concl, (concl, none_t, return)) |
|
373 |
| assm :: assms => (vars_of assm, (assm, |
|
374 |
mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms, none_t)) |
|
375 |
in |
|
376 |
fold_rev mk_exhaustive_closure (map lookup vars) (mk_safe_if check) |
|
377 |
end |
|
378 |
fun mk_smart_test_term t = |
|
379 |
let |
|
380 |
val (assms, concl) = Quickcheck_Common.strip_imp t |
|
381 |
in |
|
382 |
mk_smart_test_term' concl [] assms |
|
383 |
end |
|
384 |
val mk_test_term = |
|
385 |
if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term |
|
386 |
in lambda depth (mk_test_term t) end |
|
387 |
||
388 |
fun mk_full_generator_expr ctxt (t, eval_terms) = |
|
389 |
let |
|
390 |
val thy = ProofContext.theory_of ctxt |
|
391 |
val ctxt' = Variable.auto_fixes t ctxt |
|
392 |
val names = Term.add_free_names t [] |
|
393 |
val frees = map Free (Term.add_frees t []) |
|
394 |
val ([depth_name], ctxt'') = Variable.variant_fixes ["depth"] 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
|
395 |
val (term_names, ctxt''') = Variable.variant_fixes (map (prefix "t_") names) 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
|
396 |
val depth = Free (depth_name, @{typ code_numeral}) |
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
|
397 |
val term_vars = map (fn n => Free (n, @{typ "unit => term"})) term_names |
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
|
398 |
val terms = HOLogic.mk_list @{typ term} (map (fn v => v $ @{term "()"}) term_vars) |
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
|
399 |
val appendC = @{term "List.append :: term list => term list => term list"} |
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
|
400 |
val return = @{term "Some :: term list => term list option"} $ (appendC $ 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
|
401 |
(HOLogic.mk_list @{typ "term"} (map (fn t => HOLogic.mk_term_of (fastype_of t) 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
|
402 |
fun mk_exhaustive_closure (free as Free (_, T), 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
|
403 |
if Sign.of_sort thy (T, @{sort enum}) then |
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
|
404 |
Const (@{const_name "Quickcheck_Exhaustive.check_all_class.check_all"}, check_allT 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
|
405 |
$ (HOLogic.split_const (T, @{typ "unit => term"}, @{typ "term list option"}) |
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
|
406 |
$ 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
|
407 |
else |
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
|
408 |
Const (@{const_name "Quickcheck_Exhaustive.exhaustive_class.exhaustive"}, exhaustiveT 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
|
409 |
$ (HOLogic.split_const (T, @{typ "unit => term"}, @{typ "term list option"}) |
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
|
410 |
$ lambda free (lambda term_var t)) $ depth |
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
|
411 |
val none_t = @{term "None :: term list option"} |
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
|
412 |
fun mk_safe_if (cond, then_t, else_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
|
413 |
@{term "Quickcheck_Exhaustive.catch_match :: term list option => term list option => term list option"} $ |
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
|
414 |
(@{term "If :: bool => term list option => term list option => term list option"} |
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
|
415 |
$ cond $ then_t $ else_t) $ none_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
|
416 |
fun lookup v = the (AList.lookup (op =) (names ~~ (frees ~~ term_vars)) v) |
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
|
417 |
fun mk_naive_test_term 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
|
418 |
fold_rev mk_exhaustive_closure (frees ~~ term_vars) (mk_safe_if (t, none_t, return)) |
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
|
419 |
fun mk_smart_test_term' concl bound_vars assms = |
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
|
420 |
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
|
421 |
fun vars_of t = subtract (op =) bound_vars (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
|
422 |
val (vars, check) = |
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 |
case assms of [] => (vars_of concl, (concl, none_t, return)) |
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
|
424 |
| assm :: assms => (vars_of assm, (assm, |
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
|
425 |
mk_smart_test_term' concl (union (op =) (vars_of assm) bound_vars) assms, none_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
|
426 |
in |
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 |
fold_rev mk_exhaustive_closure (map lookup vars) (mk_safe_if check) |
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
|
428 |
end |
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 |
fun mk_smart_test_term 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
|
430 |
let |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
431 |
val (assms, concl) = Quickcheck_Common.strip_imp t |
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
|
432 |
in |
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
|
433 |
mk_smart_test_term' concl [] assms |
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
|
434 |
end |
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 |
val mk_test_term = |
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
|
436 |
if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term |
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
|
437 |
in lambda depth (mk_test_term t) end |
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
|
438 |
|
42304 | 439 |
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
|
440 |
Quickcheck_Common.gen_mk_parametric_generator_expr |
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
|
441 |
((mk_generator_expr, absdummy (@{typ "code_numeral"}, @{term "None :: term list option"})), |
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
|
442 |
@{typ "code_numeral => term list option"}) |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
443 |
|
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
444 |
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
|
445 |
let |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
446 |
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
|
447 |
val thy = ProofContext.theory_of ctxt |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
448 |
val ctxt' = Variable.auto_fixes t ctxt |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
449 |
val ([depth_name], ctxt'') = Variable.variant_fixes ["depth"] ctxt' |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
450 |
val depth = Free (depth_name, @{typ code_numeral}) |
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_bounded_forall (s, T) 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 |
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
|
453 |
$ lambda (Free (s, T)) t $ depth |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
454 |
fun mk_implies (cond, then_t) = |
42273 | 455 |
@{term "If :: bool => bool => bool => bool"} $ cond $ then_t $ @{term True} |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
456 |
fun mk_naive_test_term t = fold_rev mk_bounded_forall (Term.add_frees t []) t |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
457 |
fun mk_smart_test_term' concl bound_frees assms = |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
458 |
let |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
459 |
fun vars_of t = subtract (op =) bound_frees (Term.add_frees t []) |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
460 |
val (vars, check) = |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
461 |
case assms of [] => (vars_of concl, concl) |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
462 |
| assm :: assms => (vars_of assm, mk_implies (assm, |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
463 |
mk_smart_test_term' concl (union (op =) (vars_of assm) bound_frees) assms)) |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
464 |
in |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
465 |
fold_rev mk_bounded_forall vars check |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
466 |
end |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
467 |
fun mk_smart_test_term t = |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
468 |
let |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
469 |
val (assms, concl) = Quickcheck_Common.strip_imp t |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
470 |
in |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
471 |
mk_smart_test_term' concl [] assms |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
472 |
end |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
473 |
val mk_test_term = |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
474 |
if Config.get ctxt smart_quantifier then mk_smart_test_term else mk_naive_test_term |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
475 |
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
|
476 |
|
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
477 |
|
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
|
478 |
(** 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
|
479 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41178
diff
changeset
|
480 |
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
|
481 |
( |
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
|
482 |
type T = unit -> int -> int -> 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
|
483 |
(* 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
|
484 |
fun init _ () = error "Counterexample" |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
485 |
); |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
486 |
val put_counterexample = Counterexample.put; |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
487 |
|
41861
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
bulwahn
parents:
41472
diff
changeset
|
488 |
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
|
489 |
( |
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
bulwahn
parents:
41472
diff
changeset
|
490 |
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
|
491 |
(* 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
|
492 |
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
|
493 |
); |
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
bulwahn
parents:
41472
diff
changeset
|
494 |
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
|
495 |
|
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
496 |
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
|
497 |
( |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
498 |
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
|
499 |
(* 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
|
500 |
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
|
501 |
); |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
502 |
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
|
503 |
|
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
504 |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
505 |
val target = "Quickcheck"; |
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
|
506 |
|
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
|
507 |
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
|
508 |
let |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
509 |
val thy = ProofContext.theory_of ctxt |
42306 | 510 |
val mk_generator_expr = |
511 |
if Config.get ctxt fast then mk_fast_generator_expr |
|
512 |
else if Config.get ctxt full_support then mk_full_generator_expr else mk_generator_expr |
|
42304 | 513 |
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
|
514 |
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
|
515 |
(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
|
516 |
thy (SOME target) (fn proc => fn g => |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
517 |
fn card => fn size => g card size |> (Option.map 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
|
518 |
in |
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
|
519 |
fn [card, size] => rpair NONE (compile card size |> |
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
|
520 |
(if Config.get ctxt quickcheck_pretty then |
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
|
521 |
Option.map (map Quickcheck_Common.post_process_term) else I)) |
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 |
end; |
234ec7011e5d
generalizing compilation scheme of quickcheck generators to multiple arguments; changing random and exhaustive tester to use one code invocation for polymorphic instances with multiple cardinalities
bulwahn
parents:
42028
diff
changeset
|
523 |
|
41861
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
bulwahn
parents:
41472
diff
changeset
|
524 |
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
|
525 |
let |
77d87dc50e5a
adding a function to compile a batch of terms for quickcheck with one code generation invocation
bulwahn
parents:
41472
diff
changeset
|
526 |
val thy = ProofContext.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
|
527 |
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
|
528 |
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
|
529 |
(Counterexample_Batch.get, put_counterexample_batch, |
41918 | 530 |
"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
|
531 |
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
|
532 |
(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
|
533 |
in |
41935
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents:
41928
diff
changeset
|
534 |
map (fn compile => fn size => compile size |
d786a8a3dc47
minor corrections for renaming; moved postprocessing of terms to Quickcheck_Common
bulwahn
parents:
41928
diff
changeset
|
535 |
|> 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
|
536 |
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
|
537 |
|
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
538 |
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
|
539 |
let |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
540 |
val thy = ProofContext.theory_of ctxt |
42273 | 541 |
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
|
542 |
in |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
543 |
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
|
544 |
(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
|
545 |
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
|
546 |
end; |
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
547 |
|
42308
e2abd1ca8d01
revisiting mk_equation functions and refactoring them in exhaustive quickcheck
bulwahn
parents:
42307
diff
changeset
|
548 |
(* setup *) |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
549 |
|
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
550 |
val setup = |
42229
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents:
42214
diff
changeset
|
551 |
Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype |
1491b7209e76
generalizing ensure_sort_datatype for bounded_forall instances
bulwahn
parents:
42214
diff
changeset
|
552 |
(((@{sort typerep}, @{sort term_of}), @{sort exhaustive}), instantiate_exhaustive_datatype)) |
42230
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
553 |
#> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype |
594480d25aaa
deriving bounded_forall instances in quickcheck_exhaustive
bulwahn
parents:
42229
diff
changeset
|
554 |
(((@{sort type}, @{sort type}), @{sort bounded_forall}), instantiate_bounded_forall_datatype)) |
42306 | 555 |
#> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype |
556 |
(((@{sort typerep}, @{sort term_of}), @{sort fast_exhaustive}), instantiate_fast_exhaustive_datatype)) |
|
40907 | 557 |
#> setup_smart_quantifier |
42304 | 558 |
#> setup_full_support |
42306 | 559 |
#> setup_fast |
41903
39fd77f0ae59
fixing postprocessing; adding a configuration to enable and disable pretty presentation of quickcheck's counterexample
bulwahn
parents:
41902
diff
changeset
|
560 |
#> setup_quickcheck_pretty |
41900 | 561 |
#> Context.theory_map (Quickcheck.add_generator ("exhaustive", compile_generator_expr)) |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42176
diff
changeset
|
562 |
#> 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
|
563 |
#> 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
|
564 |
|
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
565 |
end; |