src/HOL/Tools/Predicate_Compile/code_prolog.ML
author bulwahn
Thu, 29 Jul 2010 18:16:35 +0200
changeset 38328 61280b97b761
parent 38327 8b02ce312893
child 38358 cf08f4780938
permissions -rw-r--r--
adapting output for first solution
     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 prol_term = Var of string | Cons of string | AppF of string * prol_term list;
    10   datatype prem = Conj of prem list | NotRel of string * prol_term list
    11     | Rel of string * prol_term list | Eq of prol_term * prol_term | NotEq of prol_term * prol_term;
    12   type clause = ((string * prol_term list) * prem);
    13   type logic_program = clause list;
    14   type constant_table = (string * string) list
    15   
    16   val generate : Proof.context -> string list -> (logic_program * constant_table)
    17   val write_program : logic_program -> string
    18   val run : logic_program -> string -> string list -> int option -> prol_term list list
    19 
    20   val trace : bool Unsynchronized.ref
    21 end;
    22 
    23 structure Code_Prolog : CODE_PROLOG =
    24 struct
    25 
    26 (* diagnostic tracing *)
    27 
    28 val trace = Unsynchronized.ref false
    29 
    30 fun tracing s = if !trace then Output.tracing s else () 
    31 (* general string functions *)
    32 
    33 val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
    34 val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode;
    35 
    36 (* internal program representation *)
    37 
    38 datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list;
    39 
    40 fun string_of_prol_term (Var s) = "Var " ^ s
    41   | string_of_prol_term (Cons s) = "Cons " ^ s
    42   | string_of_prol_term (AppF (f, args)) = f ^ "(" ^ commas (map string_of_prol_term args) ^ ")" 
    43 
    44 datatype prem = Conj of prem list | NotRel of string * prol_term list
    45     | Rel of string * prol_term list | Eq of prol_term * prol_term | NotEq of prol_term * prol_term;
    46 
    47 fun dest_Rel (Rel (c, ts)) = (c, ts)
    48  
    49 type clause = ((string * prol_term list) * prem);
    50 
    51 type logic_program = clause list;
    52 
    53 (* translation from introduction rules to internal representation *)
    54 
    55 (** constant table **)
    56 
    57 type constant_table = (string * string) list
    58 
    59 (* assuming no clashing *)
    60 fun mk_constant_table consts =
    61   AList.make (first_lower o Long_Name.base_name) consts
    62 
    63 fun declare_consts consts constant_table =
    64   fold (fn c => AList.update (op =) (c, first_lower (Long_Name.base_name c))) consts constant_table
    65   
    66 fun translate_const constant_table c =
    67   case AList.lookup (op =) constant_table c of
    68     SOME c' => c'
    69   | NONE => error ("No such constant: " ^ c)
    70 
    71 fun inv_lookup _ [] _ = NONE
    72   | inv_lookup eq ((key, value)::xs) value' =
    73       if eq (value', value) then SOME key
    74       else inv_lookup eq xs value';
    75 
    76 fun restore_const constant_table c =
    77   case inv_lookup (op =) constant_table c of
    78     SOME c' => c'
    79   | NONE => error ("No constant corresponding to "  ^ c)
    80   
    81 (** translation of terms, literals, premises, and clauses **)
    82 
    83 fun translate_term ctxt constant_table t =
    84   case strip_comb t of
    85     (Free (v, T), []) => Var v 
    86   | (Const (c, _), []) => Cons (translate_const constant_table c)
    87   | (Const (c, _), args) =>
    88     AppF (translate_const constant_table c, map (translate_term ctxt constant_table) args)
    89   | _ => error ("illegal term for translation: " ^ Syntax.string_of_term ctxt t)
    90 
    91 fun translate_literal ctxt constant_table t =
    92   case strip_comb t of
    93     (Const (@{const_name "op ="}, _), [l, r]) =>
    94       Eq (pairself (translate_term ctxt constant_table) (l, r))
    95   | (Const (c, _), args) =>
    96       Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args)
    97   | _ => error ("illegal literal for translation: " ^ Syntax.string_of_term ctxt t)
    98 
    99 fun NegRel_of (Rel lit) = NotRel lit
   100   | NegRel_of (Eq eq) = NotEq eq
   101   
   102 fun translate_prem ctxt constant_table t =  
   103     case try HOLogic.dest_not t of
   104       SOME t => NegRel_of (translate_literal ctxt constant_table t)
   105     | NONE => translate_literal ctxt constant_table t
   106 
   107 fun translate_intros ctxt gr const constant_table =
   108   let
   109     val intros = Graph.get_node gr const
   110     val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
   111     val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
   112     fun translate_intro intro =
   113       let
   114         val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
   115         val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro) 
   116         val prems' = Conj (map (translate_prem ctxt' constant_table') prems)
   117         val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems')
   118       in clause end
   119   in (map translate_intro intros', constant_table') end
   120 
   121 fun generate ctxt const =
   122   let 
   123      fun strong_conn_of gr keys =
   124       Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
   125     val gr = Predicate_Compile_Core.intros_graph_of ctxt
   126     val scc = strong_conn_of gr const
   127     val constant_table = mk_constant_table (flat scc)
   128   in
   129     apfst flat (fold_map (translate_intros ctxt gr) (flat scc) constant_table)
   130   end
   131 
   132 (* transform logic program *)
   133 
   134 (** ensure groundness of terms before negation **)
   135 
   136 fun add_vars (Var x) vs = insert (op =) x vs
   137   | add_vars (Cons c) vs = vs
   138   | add_vars (AppF (f, args)) vs = fold add_vars args vs
   139 
   140 fun string_of_typ (Type (s, Ts)) = Long_Name.base_name s
   141 
   142 fun mk_groundness_prems ts =
   143   let
   144     val vars = fold add_vars ts []
   145     fun mk_ground v =
   146       Rel ("ground", [Var v])
   147   in
   148     map mk_ground vars
   149   end
   150 
   151 fun ensure_groundness_prem (NotRel (c, ts)) = Conj (mk_groundness_prems ts @ [NotRel (c, ts)]) 
   152   | ensure_groundness_prem (NotEq (l, r)) = Conj (mk_groundness_prems [l, r] @ [NotEq (l, r)])
   153   | ensure_groundness_prem (Conj ps) = Conj (map ensure_groundness_prem ps)
   154   | ensure_groundness_prem p = p
   155 
   156 fun ensure_groundness_before_negation p =
   157   map (apsnd ensure_groundness_prem) p
   158 
   159 (* code printer *)
   160 
   161 fun write_term (Var v) = first_upper v
   162   | write_term (Cons c) = c
   163   | write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")" 
   164 
   165 fun write_rel (pred, args) =
   166   pred ^ "(" ^ space_implode ", " (map write_term args) ^ ")" 
   167 
   168 fun write_prem (Conj prems) = space_implode ", " (map write_prem prems)
   169   | write_prem (Rel p) = write_rel p  
   170   | write_prem (NotRel p) = "not(" ^ write_rel p ^ ")"
   171   | write_prem (Eq (l, r)) = write_term l ^ " = " ^ write_term r
   172   | write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r
   173 
   174 fun write_clause (head, prem) =
   175   write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".")
   176 
   177 fun write_program p =
   178   cat_lines (map write_clause p) 
   179 
   180 (** query templates **)
   181 
   182 fun query_first rel vnames =
   183   "eval :- once("  ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
   184   "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
   185   "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
   186   
   187 fun query_firstn n rel vnames =
   188   "eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
   189     rel ^ "(" ^ space_implode ", " vnames ^ "), Sols), writelist(Sols).\n" ^
   190     "writelist([]).\n" ^
   191     "writelist([(" ^ space_implode ", " vnames ^ ")|T]) :- " ^
   192     "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
   193     "\\n', [" ^ space_implode ", " vnames ^ "]), writelist(T).\n"
   194   
   195 val prelude =
   196   "#!/usr/bin/swipl -q -t main -f\n\n" ^
   197   ":- use_module(library('dialect/ciao/aggregates')).\n" ^
   198   ":- style_check(-singleton).\n\n" ^
   199   "main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^
   200   "main :- halt(1).\n"
   201 
   202 (* parsing prolog solution *)
   203 
   204 val scan_atom =
   205   Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_quasi s)
   206 
   207 val scan_var =
   208   Scan.many1
   209     (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
   210 
   211 val scan_ident =
   212   Scan.repeat (Scan.one
   213     (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s))
   214 
   215 fun dest_Char (Symbol.Char s) = s
   216 
   217 val string_of = concat o map (dest_Char o Symbol.decode)
   218 
   219 val is_atom_ident = forall Symbol.is_ascii_lower
   220 
   221 val is_var_ident =
   222   forall (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
   223 
   224 fun scan_terms xs = (((scan_term --| $$ ",") ::: scan_terms)
   225   || (scan_term >> single)) xs
   226 and scan_term xs =
   227   ((scan_var >> (Var o string_of))
   228   || ((scan_atom -- ($$ "(" |-- scan_terms --| $$ ")"))
   229     >> (fn (f, ts) => AppF (string_of f, ts)))
   230   || (scan_atom >> (Cons o string_of))) xs
   231 
   232 val parse_term = fst o Scan.finite Symbol.stopper
   233     (Scan.error (!! (fn _ => raise Fail "parsing prolog output failed")) scan_term)
   234   o explode
   235   
   236 fun parse_solutions sol =
   237   let
   238     fun dest_eq s = case space_explode "=" s of
   239         (l :: r :: []) => parse_term (unprefix " " r)
   240       | _ => raise Fail "unexpected equation in prolog output"
   241     fun parse_solution s = map dest_eq (space_explode ";" s)
   242   in
   243     map parse_solution (fst (split_last (space_explode "\n" sol)))
   244   end 
   245   
   246 (* calling external interpreter and getting results *)
   247 
   248 fun run p query_rel vnames nsols =
   249   let
   250     val cmd = Path.named_root
   251     val query = case nsols of NONE => query_first | SOME n => query_firstn n 
   252     val prog = prelude ^ query query_rel vnames ^ write_program p
   253     val _ = tracing ("Generated prolog program:\n" ^ prog)
   254     val prolog_file = File.tmp_path (Path.basic "prolog_file")
   255     val _ = File.write prolog_file prog
   256     val (solution, _) = bash_output ("/usr/local/bin/swipl -f " ^ File.shell_path prolog_file)
   257     val _ = tracing ("Prolog returned solution(s):\n" ^ solution)
   258     val tss = parse_solutions solution
   259   in
   260     tss
   261   end
   262 
   263 (* values command *)
   264 
   265 fun restore_term ctxt constant_table (Var s, T) = Free (s, T)
   266   | restore_term ctxt constant_table (Cons s, T) = Const (restore_const constant_table s, T)
   267   | restore_term ctxt constant_table (AppF (f, args), T) =
   268     let
   269       val thy = ProofContext.theory_of ctxt
   270       val c = restore_const constant_table f
   271       val cT = Sign.the_const_type thy c
   272       val (argsT, resT) = strip_type cT
   273       val subst = Sign.typ_match thy (resT, T) Vartab.empty
   274       val argsT' = map (Envir.subst_type subst) argsT
   275     in
   276       list_comb (Const (c, Envir.subst_type subst cT),
   277         map (restore_term ctxt constant_table) (args ~~ argsT'))
   278     end
   279 
   280 fun values ctxt soln t_compr =
   281   let
   282     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
   283       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
   284     val (body, Ts, fp) = HOLogic.strip_psplits split;
   285     val output_names = Name.variant_list (Term.add_free_names body [])
   286       (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
   287     val output_frees = rev (map2 (curry Free) output_names Ts)
   288     val body = subst_bounds (output_frees, body)
   289     val (pred as Const (name, T), all_args) =
   290       case strip_comb body of
   291         (Const (name, T), all_args) => (Const (name, T), all_args)
   292       | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
   293     val vnames =
   294       case try (map (fst o dest_Free)) all_args of
   295         SOME vs => vs
   296       | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args))
   297     val _ = tracing "Generating prolog program..."
   298     val (p, constant_table) = generate ctxt [name]
   299     val _ = tracing "Running prolog program..."
   300     val tss = run p (translate_const constant_table name) (map first_upper vnames) soln
   301     val _ = tracing "Restoring terms..."
   302     fun mk_set_comprehension t =
   303       let
   304         val frees = Term.add_frees t []
   305         val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt [t]) ("uu", fastype_of t)
   306       in HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t))
   307         frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"}))) end
   308     val set_comprs = map (fn ts =>
   309       mk_set_comprehension (HOLogic.mk_tuple (map (restore_term ctxt constant_table) (ts ~~ Ts)))) tss
   310   in
   311     foldl1 (HOLogic.mk_binop @{const_name sup}) (set_comprs @ [Free ("...", fastype_of t_compr)])
   312   end
   313 
   314 fun values_cmd print_modes soln raw_t state =
   315   let
   316     val ctxt = Toplevel.context_of state
   317     val t = Syntax.read_term ctxt raw_t
   318     val t' = values ctxt soln t
   319     val ty' = Term.type_of t'
   320     val ctxt' = Variable.auto_fixes t' ctxt
   321     val p = Print_Mode.with_modes print_modes (fn () =>
   322       Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
   323         Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
   324   in Pretty.writeln p end;
   325 
   326 
   327 (* renewing the values command for Prolog queries *)
   328 
   329 val opt_print_modes =
   330   Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
   331 
   332 val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag
   333   (opt_print_modes -- Scan.optional (Parse.nat >> SOME) NONE -- Parse.term
   334    >> (fn ((print_modes, soln), t) => Toplevel.keep
   335         (values_cmd print_modes soln t)));
   336 
   337 end;