8 inductive append |
8 inductive append |
9 where |
9 where |
10 "append [] ys ys" |
10 "append [] ys ys" |
11 | "append xs ys zs ==> append (x # xs) ys (x # zs)" |
11 | "append xs ys zs ==> append (x # xs) ys (x # zs)" |
12 |
12 |
13 ML {* Code_Prolog.options := |
13 setup {* Code_Prolog.map_code_options (K {ensure_groundness = false, |
|
14 limited_types = [], limited_predicates = [], replacing = [], |
|
15 prolog_system = Code_Prolog.SWI_PROLOG}) *} |
|
16 |
|
17 values "{(x, y, z). append x y z}" |
|
18 |
|
19 values 3 "{(x, y, z). append x y z}" |
|
20 |
|
21 setup {* Code_Prolog.map_code_options (K |
14 {ensure_groundness = false, |
22 {ensure_groundness = false, |
15 limited_types = [], limited_predicates = [], replacing = [], |
23 limited_types = [], limited_predicates = [], replacing = [], |
16 prolog_system = Code_Prolog.SWI_PROLOG} *} |
24 prolog_system = Code_Prolog.YAP}) *} |
17 |
25 |
18 values "{(x, y, z). append x y z}" |
26 values "{(x, y, z). append x y z}" |
19 |
27 |
20 values 3 "{(x, y, z). append x y z}" |
28 setup {* Code_Prolog.map_code_options (K |
21 |
|
22 ML {* Code_Prolog.options := |
|
23 {ensure_groundness = false, |
29 {ensure_groundness = false, |
24 limited_types = [], limited_predicates = [], replacing = [], |
30 limited_types = [], limited_predicates = [], replacing = [], |
25 prolog_system = Code_Prolog.YAP} *} |
31 prolog_system = Code_Prolog.SWI_PROLOG}) *} |
26 |
|
27 values "{(x, y, z). append x y z}" |
|
28 |
|
29 ML {* Code_Prolog.options := |
|
30 {ensure_groundness = false, |
|
31 limited_types = [], limited_predicates = [], replacing = [], |
|
32 prolog_system = Code_Prolog.SWI_PROLOG} *} |
|
33 |
32 |
34 |
33 |
35 section {* Example queens *} |
34 section {* Example queens *} |
36 |
35 |
37 inductive nodiag :: "int => int => int list => bool" |
36 inductive nodiag :: "int => int => int list => bool" |
190 |
189 |
191 inductive notB :: "abc => bool" |
190 inductive notB :: "abc => bool" |
192 where |
191 where |
193 "y \<noteq> B \<Longrightarrow> notB y" |
192 "y \<noteq> B \<Longrightarrow> notB y" |
194 |
193 |
195 ML {* Code_Prolog.options := {ensure_groundness = true, |
194 setup {* Code_Prolog.map_code_options (K {ensure_groundness = true, |
196 limited_types = [], |
195 limited_types = [], |
197 limited_predicates = [], |
196 limited_predicates = [], |
198 replacing = [], |
197 replacing = [], |
199 prolog_system = Code_Prolog.SWI_PROLOG} *} |
198 prolog_system = Code_Prolog.SWI_PROLOG}) *} |
200 |
199 |
201 values 2 "{y. notB y}" |
200 values 2 "{y. notB y}" |
202 |
201 |
203 inductive notAB :: "abc * abc \<Rightarrow> bool" |
202 inductive notAB :: "abc * abc \<Rightarrow> bool" |
204 where |
203 where |