38073
|
1 |
(* Title: HOL/Tools/Predicate_Compile/code_prolog.ML
|
|
2 |
Author: Lukas Bulwahn, TU Muenchen
|
|
3 |
|
|
4 |
Prototype of an code generator for logic programming languages (a.k.a. Prolog)
|
|
5 |
*)
|
|
6 |
|
|
7 |
signature CODE_PROLOG =
|
|
8 |
sig
|
|
9 |
datatype term = Var of string * typ | Cons of string | AppF of string * term list;
|
|
10 |
datatype prem = Conj of prem list | NotRel of string * term list | Rel of string * term list | Eq of term * term | NotEq of term * term;
|
|
11 |
type clause = ((string * term list) * prem);
|
|
12 |
type logic_program = clause list;
|
|
13 |
|
|
14 |
val generate : Proof.context -> string list -> logic_program
|
|
15 |
val write_program : logic_program -> string
|
38075
|
16 |
val run : logic_program -> string -> string list -> term list
|
38073
|
17 |
|
|
18 |
end;
|
|
19 |
|
|
20 |
structure Code_Prolog : CODE_PROLOG =
|
|
21 |
struct
|
|
22 |
|
|
23 |
(* general string functions *)
|
|
24 |
|
|
25 |
val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
|
|
26 |
val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
|
|
27 |
|
|
28 |
(* internal program representation *)
|
|
29 |
|
|
30 |
datatype term = Var of string * typ | Cons of string | AppF of string * term list;
|
|
31 |
|
38075
|
32 |
fun string_of_prol_term (Var (s, T)) = "Var " ^ s
|
|
33 |
| string_of_prol_term (Cons s) = "Cons " ^ s
|
|
34 |
| string_of_prol_term (AppF (f, args)) = f ^ "(" ^ commas (map string_of_prol_term args) ^ ")"
|
|
35 |
|
38073
|
36 |
datatype prem = Conj of prem list | NotRel of string * term list | Rel of string * term list | Eq of term * term | NotEq of term * term;
|
|
37 |
|
|
38 |
fun dest_Rel (Rel (c, ts)) = (c, ts)
|
|
39 |
|
|
40 |
type clause = ((string * term list) * prem);
|
|
41 |
|
|
42 |
type logic_program = clause list;
|
|
43 |
|
|
44 |
(* translation from introduction rules to internal representation *)
|
|
45 |
|
|
46 |
(* assuming no clashing *)
|
|
47 |
fun translate_const c = Long_Name.base_name c
|
|
48 |
|
|
49 |
fun translate_term ctxt t =
|
|
50 |
case strip_comb t of
|
|
51 |
(Free (v, T), []) => Var (v, T)
|
|
52 |
| (Const (c, _), []) => Cons (translate_const c)
|
|
53 |
| (Const (c, _), args) => AppF (translate_const c, map (translate_term ctxt) args)
|
|
54 |
| _ => error ("illegal term for translation: " ^ Syntax.string_of_term ctxt t)
|
|
55 |
|
|
56 |
|
|
57 |
fun translate_literal ctxt t =
|
|
58 |
case strip_comb t of
|
|
59 |
(Const (@{const_name "op ="}, _), [l, r]) => Eq (pairself (translate_term ctxt) (l, r))
|
|
60 |
| (Const (c, _), args) => Rel (translate_const c, map (translate_term ctxt) args)
|
|
61 |
| _ => error ("illegal literal for translation: " ^ Syntax.string_of_term ctxt t)
|
|
62 |
|
|
63 |
fun NegRel_of (Rel lit) = NotRel lit
|
|
64 |
| NegRel_of (Eq eq) = NotEq eq
|
|
65 |
|
|
66 |
fun translate_prem ctxt t =
|
|
67 |
case try HOLogic.dest_not t of
|
|
68 |
SOME t => NegRel_of (translate_literal ctxt t)
|
|
69 |
| NONE => translate_literal ctxt t
|
|
70 |
|
|
71 |
fun translate_intros ctxt gr const =
|
|
72 |
let
|
|
73 |
val intros = Graph.get_node gr const
|
|
74 |
val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
|
|
75 |
fun translate_intro intro =
|
|
76 |
let
|
|
77 |
val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
|
|
78 |
val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro)
|
|
79 |
val prems' = Conj (map (translate_prem ctxt') prems)
|
|
80 |
val clause = (dest_Rel (translate_literal ctxt' head), prems')
|
|
81 |
in clause end
|
|
82 |
in map translate_intro intros' end
|
|
83 |
|
|
84 |
fun generate ctxt const =
|
|
85 |
let
|
|
86 |
fun strong_conn_of gr keys =
|
|
87 |
Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
|
|
88 |
val gr = Predicate_Compile_Core.intros_graph_of ctxt
|
|
89 |
val scc = strong_conn_of gr const
|
|
90 |
val _ = tracing (commas (flat scc))
|
|
91 |
in
|
|
92 |
maps (translate_intros ctxt gr) (flat scc)
|
|
93 |
end
|
|
94 |
|
|
95 |
(* transform logic program *)
|
|
96 |
|
|
97 |
(** ensure groundness of terms before negation **)
|
|
98 |
|
|
99 |
fun add_vars (Var x) vs = insert (op =) x vs
|
|
100 |
| add_vars (Cons c) vs = vs
|
|
101 |
| add_vars (AppF (f, args)) vs = fold add_vars args vs
|
|
102 |
|
|
103 |
fun string_of_typ (Type (s, Ts)) = Long_Name.base_name s
|
|
104 |
|
|
105 |
fun mk_groundness_prems ts =
|
|
106 |
let
|
|
107 |
val vars = fold add_vars ts []
|
|
108 |
fun mk_ground (v, T) =
|
|
109 |
Rel ("ground_" ^ string_of_typ T, [Var (v, T)])
|
|
110 |
in
|
|
111 |
map mk_ground vars
|
|
112 |
end
|
|
113 |
|
|
114 |
fun ensure_groundness_prem (NotRel (c, ts)) = Conj (mk_groundness_prems ts @ [NotRel (c, ts)])
|
|
115 |
| ensure_groundness_prem (NotEq (l, r)) = Conj (mk_groundness_prems [l, r] @ [NotEq (l, r)])
|
|
116 |
| ensure_groundness_prem (Conj ps) = Conj (map ensure_groundness_prem ps)
|
|
117 |
| ensure_groundness_prem p = p
|
|
118 |
|
|
119 |
fun ensure_groundness_before_negation p =
|
|
120 |
map (apsnd ensure_groundness_prem) p
|
|
121 |
|
|
122 |
(* code printer *)
|
|
123 |
|
|
124 |
fun write_term (Var (v, _)) = first_upper v
|
|
125 |
| write_term (Cons c) = first_lower c
|
|
126 |
| write_term (AppF (f, args)) = first_lower f ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
|
|
127 |
|
|
128 |
fun write_rel (pred, args) =
|
|
129 |
pred ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
|
|
130 |
|
|
131 |
fun write_prem (Conj prems) = space_implode ", " (map write_prem prems)
|
|
132 |
| write_prem (Rel p) = write_rel p
|
|
133 |
| write_prem (NotRel p) = "not(" ^ write_rel p ^ ")"
|
|
134 |
| write_prem (Eq (l, r)) = write_term l ^ " = " ^ write_term r
|
|
135 |
| write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r
|
|
136 |
|
|
137 |
fun write_clause (head, prem) =
|
|
138 |
write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".")
|
|
139 |
|
|
140 |
fun write_program p =
|
|
141 |
cat_lines (map write_clause p)
|
|
142 |
|
|
143 |
fun query_first rel vnames =
|
|
144 |
"eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
|
|
145 |
"writef('" ^ implode (map (fn v => v ^ " = %w\\n") vnames) ^"', [" ^ space_implode ", " vnames ^ "]).\n"
|
|
146 |
|
|
147 |
val prelude =
|
|
148 |
"#!/usr/bin/swipl -q -t main -f\n\n" ^
|
|
149 |
":- style_check(-singleton).\n\n" ^
|
|
150 |
"main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^
|
|
151 |
"main :- halt(1).\n"
|
38075
|
152 |
|
|
153 |
(* parsing prolog solution *)
|
|
154 |
|
|
155 |
val scan_atom =
|
|
156 |
Scan.repeat (Scan.one (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_quasi s))
|
|
157 |
|
|
158 |
val scan_var =
|
|
159 |
Scan.repeat (Scan.one
|
|
160 |
(fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s))
|
|
161 |
|
|
162 |
fun dest_Char (Symbol.Char s) = s
|
|
163 |
|
|
164 |
val string_of = concat o map (dest_Char o Symbol.decode)
|
|
165 |
|
|
166 |
val scan_term =
|
|
167 |
scan_atom >> (Cons o string_of) || scan_var >> (Var o rpair dummyT o string_of)
|
|
168 |
|
|
169 |
val parse_term = fst o Scan.finite Symbol.stopper
|
|
170 |
(Scan.error (!! (fn _ => raise Fail "parsing prolog output failed"))
|
|
171 |
scan_term)
|
|
172 |
o explode
|
|
173 |
|
|
174 |
fun parse_solution sol =
|
|
175 |
let
|
|
176 |
fun dest_eq s = (tracing s; tracing(commas (space_explode "=" s)); case space_explode "=" s of
|
|
177 |
(l :: r :: []) => parse_term (unprefix " " r)
|
|
178 |
| _ => raise Fail "unexpected equation in prolog output")
|
|
179 |
in
|
|
180 |
map dest_eq (fst (split_last (space_explode "\n" sol)))
|
|
181 |
end
|
38073
|
182 |
|
|
183 |
(* calling external interpreter and getting results *)
|
|
184 |
|
|
185 |
fun run p query_rel vnames =
|
|
186 |
let
|
|
187 |
val cmd = Path.named_root
|
|
188 |
val prog = prelude ^ query_first query_rel vnames ^ write_program p
|
|
189 |
val prolog_file = File.tmp_path (Path.basic "prolog_file")
|
|
190 |
val _ = File.write prolog_file prog
|
|
191 |
val (solution, _) = bash_output ("/usr/bin/swipl -f " ^ File.shell_path prolog_file)
|
38075
|
192 |
val ts = parse_solution solution
|
|
193 |
val _ = tracing (commas (map string_of_prol_term ts))
|
38073
|
194 |
in
|
38075
|
195 |
ts
|
38073
|
196 |
end
|
|
197 |
|
38075
|
198 |
(* values command *)
|
|
199 |
|
|
200 |
fun mk_term (Var (s, T)) = Free (s, T)
|
|
201 |
| mk_term (Cons s) = Const (s, dummyT)
|
|
202 |
| mk_term (AppF (f, args)) = list_comb (Const (f, dummyT), map mk_term args)
|
|
203 |
|
|
204 |
fun values ctxt soln t_compr =
|
|
205 |
let
|
|
206 |
val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
|
|
207 |
| _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
|
|
208 |
val (body, Ts, fp) = HOLogic.strip_psplits split;
|
|
209 |
val output_names = Name.variant_list (Term.add_free_names body [])
|
|
210 |
(map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
|
|
211 |
val output_frees = map2 (curry Free) output_names (rev Ts)
|
|
212 |
val body = subst_bounds (output_frees, body)
|
|
213 |
val (pred as Const (name, T), all_args) =
|
|
214 |
case strip_comb body of
|
|
215 |
(Const (name, T), all_args) => (Const (name, T), all_args)
|
|
216 |
| (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
|
|
217 |
val vnames =
|
|
218 |
case try (map (fst o dest_Free)) all_args of
|
|
219 |
SOME vs => vs
|
|
220 |
| NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args))
|
|
221 |
val ts = run (generate ctxt [name]) (translate_const name) (map first_upper vnames)
|
|
222 |
in
|
|
223 |
HOLogic.mk_tuple (map mk_term ts)
|
|
224 |
end
|
|
225 |
|
|
226 |
fun values_cmd print_modes soln raw_t state =
|
|
227 |
let
|
|
228 |
val ctxt = Toplevel.context_of state
|
|
229 |
val t = Syntax.read_term ctxt raw_t
|
|
230 |
val t' = values ctxt soln t
|
|
231 |
val ty' = Term.type_of t'
|
|
232 |
val ctxt' = Variable.auto_fixes t' ctxt
|
|
233 |
val p = Print_Mode.with_modes print_modes (fn () =>
|
|
234 |
Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
|
|
235 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
|
|
236 |
in Pretty.writeln p end;
|
|
237 |
|
|
238 |
|
|
239 |
(* renewing the values command for Prolog queries *)
|
|
240 |
|
|
241 |
val opt_print_modes =
|
|
242 |
Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
|
|
243 |
|
|
244 |
val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag
|
|
245 |
(opt_print_modes -- Scan.optional Parse.nat ~1 -- Parse.term
|
|
246 |
>> (fn ((print_modes, soln), t) => Toplevel.keep
|
|
247 |
(values_cmd print_modes soln t)));
|
|
248 |
|
38073
|
249 |
end;
|