author | wenzelm |
Sat, 23 Apr 2011 13:00:19 +0200 | |
changeset 42463 | f270e3e18be5 |
parent 41956 | c15ef1b85035 |
child 43937 | 768c70befd59 |
permissions | -rw-r--r-- |
39184
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
1 |
theory List_Examples |
41956 | 2 |
imports |
3 |
Main |
|
4 |
"~~/src/HOL/Library/Predicate_Compile_Quickcheck" |
|
5 |
"~~/src/HOL/Library/Code_Prolog" |
|
39184
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
6 |
begin |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
7 |
|
39463 | 8 |
setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *} |
39184
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
9 |
|
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
10 |
setup {* Code_Prolog.map_code_options (K |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
11 |
{ensure_groundness = true, |
39800 | 12 |
limit_globally = NONE, |
39184
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
13 |
limited_types = [(@{typ nat}, 2), (@{typ "nat list"}, 4)], |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
14 |
limited_predicates = [(["appendP"], 4), (["revP"], 4)], |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
15 |
replacing = |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
16 |
[(("appendP", "limited_appendP"), "quickcheck"), |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
17 |
(("revP", "limited_revP"), "quickcheck"), |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
18 |
(("appendP", "limited_appendP"), "lim_revP")], |
39463 | 19 |
manual_reorder = []}) *} |
39184
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
20 |
|
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
21 |
lemma "(xs :: nat list) = ys @ ys --> rev xs = xs" |
40924 | 22 |
quickcheck[tester = random, iterations = 10000] |
23 |
quickcheck[tester = predicate_compile_wo_ff, iterations = 1, expect = counterexample] |
|
24 |
quickcheck[tester = prolog, expect = counterexample] |
|
39184
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
25 |
oops |
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
26 |
|
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
bulwahn
parents:
diff
changeset
|
27 |
end |