Theory Code_Prolog_Examples
theory Code_Prolog_Examples
imports "HOL-Library.Code_Prolog"
begin
section ‹Example append›
inductive append
where
"append [] ys ys"
| "append xs ys zs ==> append (x # xs) ys (x # zs)"
setup ‹Code_Prolog.map_code_options (K
{ensure_groundness = false,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
manual_reorder = []})›
values_prolog "{(x, y, z). append x y z}"
values_prolog 4 "{(z, x, y). append x y ((1::nat) # (2 # (3 # z)))}"
values_prolog 3 "{(x, y, z). append x y z}"
setup ‹Code_Prolog.map_code_options (K
{ensure_groundness = false,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
manual_reorder = []})›
values_prolog "{(x, y, z). append x y z}"
setup ‹Code_Prolog.map_code_options (K
{ensure_groundness = false,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
manual_reorder = []})›
section ‹Example queens›
inductive nodiag :: "int => int => int list => bool"
where
"nodiag B D []"
| "D ≠ N - B ==> D ≠ B - N ==> Da = D + 1 ==> nodiag B Da L ==> nodiag B D (N # L)"
text ‹
qdelete(A, [A|L], L).
qdelete(X, [A|Z], [A|R]) :-
qdelete(X, Z, R).
›
inductive qdelete :: "int => int list => int list => bool"
where
"qdelete A (A # L) L"
| "qdelete X Z R ==> qdelete X (A # Z) (A # R)"
text ‹
qperm([], []).
qperm([X|Y], K) :-
qdelete(U, [X|Y], Z),
K = [U|V],
qperm(Z, V).
›
inductive qperm :: "int list => int list => bool"
where
"qperm [] []"
| "qdelete U (X # Y) Z ==> qperm Z V ==> qperm (X # Y) (U # V)"
text ‹
safe([]).
safe([N|L]) :-
nodiag(N, 1, L),
safe(L).
›
inductive safe :: "int list => bool"
where
"safe []"
| "nodiag N 1 L ==> safe L ==> safe (N # L)"
text ‹
queen(Data, Out) :-
qperm(Data, Out),
safe(Out)
›
inductive queen :: "int list => int list => bool"
where
"qperm Data Out ==> safe Out ==> queen Data Out"
inductive queen_9
where
"queen [1,2,3,4,5,6,7,8,9] ys ==> queen_9 ys"
values_prolog 10 "{ys. queen_9 ys}"
section ‹Example symbolic derivation›
hide_const Pow