| author | wenzelm | 
| Wed, 19 Apr 2017 16:25:26 +0200 | |
| changeset 65512 | 9fd620f2fa7d | 
| parent 63167 | 0909deb8059b | 
| child 66453 | cc19f7ca2ed6 | 
| 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  | 
|
| 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,  | 
| 
39184
 
71f3f194b962
adding a List example (challenge from Tobias) for counterexample search
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
   limited_types = [(@{typ nat}, 2), (@{typ "nat list"}, 4)],
 | 
| 
 
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: 
43937 
diff
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  |