| author | wenzelm | 
| Tue, 26 Jan 2021 23:34:40 +0100 | |
| changeset 73194 | c0d6d57a9a31 | 
| parent 69597 | ff784d5a5bfb | 
| 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 | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63167diff
changeset | 4 | "HOL-Library.Predicate_Compile_Quickcheck" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63167diff
changeset | 5 | "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 | |
| 63167 | 8 | setup \<open> | 
| 52666 | 9 | Context.theory_map | 
| 10 |     (Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
 | |
| 63167 | 11 | \<close> | 
| 39184 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 12 | |
| 63167 | 13 | setup \<open>Code_Prolog.map_code_options (K | 
| 39184 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 14 |   {ensure_groundness = true,
 | 
| 39800 | 15 | limit_globally = NONE, | 
| 69597 | 16 | limited_types = [(\<^typ>\<open>nat\<close>, 2), (\<^typ>\<open>nat list\<close>, 4)], | 
| 39184 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 17 | limited_predicates = [(["appendP"], 4), (["revP"], 4)], | 
| 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 18 | replacing = | 
| 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 19 |      [(("appendP", "limited_appendP"), "quickcheck"),
 | 
| 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 20 |       (("revP", "limited_revP"), "quickcheck"),
 | 
| 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 21 |       (("appendP", "limited_appendP"), "lim_revP")],
 | 
| 63167 | 22 | manual_reorder = []})\<close> | 
| 39184 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 23 | |
| 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 24 | lemma "(xs :: nat list) = ys @ ys --> rev xs = xs" | 
| 40924 | 25 | quickcheck[tester = random, iterations = 10000] | 
| 45451 
74515e8e6046
renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
 bulwahn parents: 
43937diff
changeset | 26 | quickcheck[tester = smart_exhaustive, iterations = 1, expect = counterexample] | 
| 40924 | 27 | quickcheck[tester = prolog, expect = counterexample] | 
| 39184 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 28 | oops | 
| 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 bulwahn parents: diff
changeset | 29 | |
| 62390 | 30 | end |