src/HOL/Tools/Predicate_Compile/code_prolog.ML
author wenzelm
Thu, 15 Mar 2012 20:07:00 +0100
changeset 47823 94aa7b81bcf6
parent 47485 165886a4fe64
child 47836 5c6955f487e5
permissions -rw-r--r--
prefer formally checked @{keyword} parser;
     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
     5 (a.k.a. Prolog).
     6 *)
     7 
     8 signature CODE_PROLOG =
     9 sig
    10   datatype prolog_system = SWI_PROLOG | YAP
    11   type code_options =
    12     {ensure_groundness : bool,
    13      limit_globally : int option,
    14      limited_types : (typ * int) list,
    15      limited_predicates : (string list * int) list,
    16      replacing : ((string * string) * string) list,
    17      manual_reorder : ((string * int) * int list) list}
    18   val set_ensure_groundness : code_options -> code_options
    19   val map_limit_predicates : ((string list * int) list -> (string list * int) list)
    20     -> code_options -> code_options
    21   val code_options_of : theory -> code_options 
    22   val map_code_options : (code_options -> code_options) -> theory -> theory
    23   
    24   datatype arith_op = Plus | Minus
    25   datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
    26     | Number of int | ArithOp of arith_op * prol_term list;
    27   datatype prem = Conj of prem list
    28     | Rel of string * prol_term list | NotRel of string * prol_term list
    29     | Eq of prol_term * prol_term | NotEq of prol_term * prol_term
    30     | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
    31     | Ground of string * typ;
    32 
    33   type clause = ((string * prol_term list) * prem);
    34   type logic_program = clause list;
    35   type constant_table = (string * string) list
    36   
    37   val generate : Predicate_Compile_Aux.mode option * bool ->
    38     Proof.context -> string -> (logic_program * constant_table)
    39   val write_program : logic_program -> string
    40   val run : (Time.time * prolog_system) -> logic_program -> (string * prol_term list) ->
    41     string list -> int option -> prol_term list list
    42   
    43   val active : bool Config.T
    44   val test_goals :
    45     Proof.context -> bool -> (string * typ) list -> (term * term list) list ->
    46       Quickcheck.result list
    47 
    48   val trace : bool Unsynchronized.ref
    49   
    50   val replace : ((string * string) * string) -> logic_program -> logic_program
    51 end;
    52 
    53 structure Code_Prolog : CODE_PROLOG =
    54 struct
    55 
    56 (* diagnostic tracing *)
    57 
    58 val trace = Unsynchronized.ref false
    59 
    60 fun tracing s = if !trace then Output.tracing s else () 
    61 
    62 (* code generation options *)
    63 
    64 
    65 type code_options =
    66   {ensure_groundness : bool,
    67    limit_globally : int option,
    68    limited_types : (typ * int) list,
    69    limited_predicates : (string list * int) list,
    70    replacing : ((string * string) * string) list,
    71    manual_reorder : ((string * int) * int list) list}
    72 
    73 
    74 fun set_ensure_groundness {ensure_groundness, limit_globally, limited_types, limited_predicates,
    75   replacing, manual_reorder} =
    76   {ensure_groundness = true, limit_globally = limit_globally, limited_types = limited_types,
    77    limited_predicates = limited_predicates, replacing = replacing,
    78    manual_reorder = manual_reorder}
    79 
    80 fun map_limit_predicates f {ensure_groundness, limit_globally, limited_types, limited_predicates,
    81   replacing, manual_reorder} =
    82   {ensure_groundness = ensure_groundness, limit_globally = limit_globally, limited_types = limited_types,
    83    limited_predicates = f limited_predicates, replacing = replacing,
    84    manual_reorder = manual_reorder}
    85 
    86 fun merge_global_limit (NONE, NONE) = NONE
    87   | merge_global_limit (NONE, SOME n) = SOME n
    88   | merge_global_limit (SOME n, NONE) = SOME n
    89   | merge_global_limit (SOME n, SOME m) = SOME (Int.max (n, m))  (* FIXME odd merge *)
    90    
    91 structure Options = Theory_Data
    92 (
    93   type T = code_options
    94   val empty = {ensure_groundness = false, limit_globally = NONE,
    95     limited_types = [], limited_predicates = [], replacing = [], manual_reorder = []}
    96   val extend = I;
    97   fun merge
    98     ({ensure_groundness = ensure_groundness1, limit_globally = limit_globally1,
    99       limited_types = limited_types1, limited_predicates = limited_predicates1,
   100       replacing = replacing1, manual_reorder = manual_reorder1},
   101      {ensure_groundness = ensure_groundness2, limit_globally = limit_globally2,
   102       limited_types = limited_types2, limited_predicates = limited_predicates2,
   103       replacing = replacing2, manual_reorder = manual_reorder2}) =
   104     {ensure_groundness = ensure_groundness1 orelse ensure_groundness2 (* FIXME odd merge *),
   105      limit_globally = merge_global_limit (limit_globally1, limit_globally2),
   106      limited_types = AList.merge (op =) (K true) (limited_types1, limited_types2),
   107      limited_predicates = AList.merge (op =) (K true) (limited_predicates1, limited_predicates2),
   108      manual_reorder = AList.merge (op =) (K true) (manual_reorder1, manual_reorder2),
   109      replacing = Library.merge (op =) (replacing1, replacing2)};
   110 );
   111 
   112 val code_options_of = Options.get
   113 
   114 val map_code_options = Options.map
   115 
   116 (* system configuration *)
   117 
   118 datatype prolog_system = SWI_PROLOG | YAP
   119 
   120 fun string_of_system SWI_PROLOG = "swiprolog"
   121   | string_of_system YAP = "yap"
   122 
   123 type system_configuration = {timeout : Time.time, prolog_system : prolog_system}
   124                                                 
   125 structure System_Config = Generic_Data
   126 (
   127   type T = system_configuration
   128   val empty = {timeout = seconds 10.0, prolog_system = SWI_PROLOG}
   129   val extend = I;
   130   fun merge (a, _) = a
   131 )
   132 
   133 (* general string functions *)
   134 
   135 val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o raw_explode;
   136 val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o raw_explode;
   137 
   138 (* internal program representation *)
   139 
   140 datatype arith_op = Plus | Minus
   141 
   142 datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
   143   | Number of int | ArithOp of arith_op * prol_term list;
   144 
   145 fun dest_Var (Var v) = v
   146 
   147 fun add_vars (Var v) = insert (op =) v
   148   | add_vars (ArithOp (_, ts)) = fold add_vars ts
   149   | add_vars (AppF (_, ts)) = fold add_vars ts
   150   | add_vars _ = I
   151 
   152 fun map_vars f (Var v) = Var (f v)
   153   | map_vars f (ArithOp (opr, ts)) = ArithOp (opr, map (map_vars f) ts)
   154   | map_vars f (AppF (fs, ts)) = AppF (fs, map (map_vars f) ts)
   155   | map_vars f t = t
   156   
   157 fun maybe_AppF (c, []) = Cons c
   158   | maybe_AppF (c, xs) = AppF (c, xs)
   159 
   160 fun is_Var (Var _) = true
   161   | is_Var _ = false
   162 
   163 fun is_arith_term (Var _) = true
   164   | is_arith_term (Number _) = true
   165   | is_arith_term (ArithOp (_, operands)) = forall is_arith_term operands
   166   | is_arith_term _ = false
   167 
   168 fun string_of_prol_term (Var s) = "Var " ^ s
   169   | string_of_prol_term (Cons s) = "Cons " ^ s
   170   | string_of_prol_term (AppF (f, args)) = f ^ "(" ^ commas (map string_of_prol_term args) ^ ")" 
   171   | string_of_prol_term (Number n) = "Number " ^ string_of_int n
   172 
   173 datatype prem = Conj of prem list
   174   | Rel of string * prol_term list | NotRel of string * prol_term list
   175   | Eq of prol_term * prol_term | NotEq of prol_term * prol_term
   176   | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
   177   | Ground of string * typ;
   178 
   179 fun dest_Rel (Rel (c, ts)) = (c, ts)
   180 
   181 fun map_term_prem f (Conj prems) = Conj (map (map_term_prem f) prems)
   182   | map_term_prem f (Rel (r, ts)) = Rel (r, map f ts)
   183   | map_term_prem f (NotRel (r, ts)) = NotRel (r, map f ts)
   184   | map_term_prem f (Eq (l, r)) = Eq (f l, f r)
   185   | map_term_prem f (NotEq (l, r)) = NotEq (f l, f r)
   186   | map_term_prem f (ArithEq (l, r)) = ArithEq (f l, f r)
   187   | map_term_prem f (NotArithEq (l, r)) = NotArithEq (f l, f r)
   188   | map_term_prem f (Ground (v, T)) = Ground (dest_Var (f (Var v)), T)
   189 
   190 fun fold_prem_terms f (Conj prems) = fold (fold_prem_terms f) prems
   191   | fold_prem_terms f (Rel (_, ts)) = fold f ts
   192   | fold_prem_terms f (NotRel (_, ts)) = fold f ts
   193   | fold_prem_terms f (Eq (l, r)) = f l #> f r
   194   | fold_prem_terms f (NotEq (l, r)) = f l #> f r
   195   | fold_prem_terms f (ArithEq (l, r)) = f l #> f r
   196   | fold_prem_terms f (NotArithEq (l, r)) = f l #> f r
   197   | fold_prem_terms f (Ground (v, T)) = f (Var v)
   198   
   199 type clause = ((string * prol_term list) * prem);
   200 
   201 type logic_program = clause list;
   202  
   203 (* translation from introduction rules to internal representation *)
   204 
   205 fun mk_conform f empty avoid name =
   206   let
   207     fun dest_Char (Symbol.Char c) = c
   208     val name' = space_implode "" (map (dest_Char o Symbol.decode)
   209       (filter (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s)
   210         (Symbol.explode name)))
   211     val name'' = f (if name' = "" then empty else name')
   212   in if member (op =) avoid name'' then singleton (Name.variant_list avoid) name'' else name'' end
   213 
   214 (** constant table **)
   215 
   216 type constant_table = (string * string) list
   217 
   218 fun declare_consts consts constant_table =
   219   let
   220     fun update' c table =
   221       if AList.defined (op =) table c then table else
   222         let
   223           val c' = mk_conform first_lower "pred" (map snd table) (Long_Name.base_name c)
   224         in
   225           AList.update (op =) (c, c') table
   226         end
   227   in
   228     fold update' consts constant_table
   229   end
   230   
   231 fun translate_const constant_table c =
   232   case AList.lookup (op =) constant_table c of
   233     SOME c' => c'
   234   | NONE => error ("No such constant: " ^ c)
   235 
   236 fun inv_lookup _ [] _ = NONE
   237   | inv_lookup eq ((key, value)::xs) value' =
   238       if eq (value', value) then SOME key
   239       else inv_lookup eq xs value';
   240 
   241 fun restore_const constant_table c =
   242   case inv_lookup (op =) constant_table c of
   243     SOME c' => c'
   244   | NONE => error ("No constant corresponding to "  ^ c)
   245 
   246 (** translation of terms, literals, premises, and clauses **)
   247 
   248 fun translate_arith_const @{const_name "Groups.plus_class.plus"} = SOME Plus
   249   | translate_arith_const @{const_name "Groups.minus_class.minus"} = SOME Minus
   250   | translate_arith_const _ = NONE
   251 
   252 fun mk_nat_term constant_table n =
   253   let
   254     val zero = translate_const constant_table @{const_name "Groups.zero_class.zero"}
   255     val Suc = translate_const constant_table @{const_name "Suc"}
   256   in funpow n (fn t => AppF (Suc, [t])) (Cons zero) end
   257 
   258 fun translate_term ctxt constant_table t =
   259   case try HOLogic.dest_number t of
   260     SOME (@{typ "int"}, n) => Number n
   261   | SOME (@{typ "nat"}, n) => mk_nat_term constant_table n
   262   | NONE =>
   263       (case strip_comb t of
   264         (Free (v, T), []) => Var v 
   265       | (Const (c, _), []) => Cons (translate_const constant_table c)
   266       | (Const (c, _), args) =>
   267         (case translate_arith_const c of
   268           SOME aop => ArithOp (aop, map (translate_term ctxt constant_table) args)
   269         | NONE =>                                                             
   270             AppF (translate_const constant_table c, map (translate_term ctxt constant_table) args))
   271       | _ => error ("illegal term for translation: " ^ Syntax.string_of_term ctxt t))
   272 
   273 fun translate_literal ctxt constant_table t =
   274   case strip_comb t of
   275     (Const (@{const_name HOL.eq}, _), [l, r]) =>
   276       let
   277         val l' = translate_term ctxt constant_table l
   278         val r' = translate_term ctxt constant_table r
   279       in
   280         (if is_Var l' andalso is_arith_term r' andalso not (is_Var r') then ArithEq else Eq) (l', r')
   281       end
   282   | (Const (c, _), args) =>
   283       Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args)
   284   | _ => error ("illegal literal for translation: " ^ Syntax.string_of_term ctxt t)
   285 
   286 fun NegRel_of (Rel lit) = NotRel lit
   287   | NegRel_of (Eq eq) = NotEq eq
   288   | NegRel_of (ArithEq eq) = NotArithEq eq
   289 
   290 fun mk_groundness_prems t = map Ground (Term.add_frees t [])
   291   
   292 fun translate_prem ensure_groundness ctxt constant_table t =  
   293     case try HOLogic.dest_not t of
   294       SOME t =>
   295         if ensure_groundness then
   296           Conj (mk_groundness_prems t @ [NegRel_of (translate_literal ctxt constant_table t)])
   297         else
   298           NegRel_of (translate_literal ctxt constant_table t)
   299     | NONE => translate_literal ctxt constant_table t
   300     
   301 fun imp_prems_conv cv ct =
   302   case Thm.term_of ct of
   303     Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
   304   | _ => Conv.all_conv ct
   305 
   306 fun Trueprop_conv cv ct =
   307   case Thm.term_of ct of
   308     Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
   309   | _ => raise Fail "Trueprop_conv"
   310 
   311 fun preprocess_intro thy rule =
   312   Conv.fconv_rule
   313     (imp_prems_conv
   314       (Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq}))))
   315     (Thm.transfer thy rule)
   316 
   317 fun translate_intros ensure_groundness ctxt gr const constant_table =
   318   let
   319     val intros = map (preprocess_intro (Proof_Context.theory_of ctxt)) (Graph.get_node gr const)
   320     val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
   321     val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
   322     fun translate_intro intro =
   323       let
   324         val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
   325         val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro)
   326         val prems' = Conj (map (translate_prem ensure_groundness ctxt' constant_table') prems)
   327         val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems')
   328       in clause end
   329   in
   330     (map translate_intro intros', constant_table')
   331   end
   332 
   333 fun depending_preds_of (key, intros) =
   334   fold Term.add_const_names (map Thm.prop_of intros) []
   335 
   336 fun add_edges edges_of key G =
   337   let
   338     fun extend' key (G, visited) = 
   339       case try (Graph.get_node G) key of
   340           SOME v =>
   341             let
   342               val new_edges = filter (fn k => is_some (try (Graph.get_node G) k)) (edges_of (key, v))
   343               val (G', visited') = fold extend'
   344                 (subtract (op =) (key :: visited) new_edges) (G, key :: visited)
   345             in
   346               (fold (Graph.add_edge o (pair key)) new_edges G', visited')
   347             end
   348         | NONE => (G, visited)
   349   in
   350     fst (extend' key (G, []))
   351   end
   352 
   353 fun print_intros ctxt gr consts =
   354   tracing (cat_lines (map (fn const =>
   355     "Constant " ^ const ^ "has intros:\n" ^
   356     cat_lines (map (Display.string_of_thm ctxt) (Graph.get_node gr const))) consts))
   357 
   358 (* translation of moded predicates *)
   359 
   360 (** generating graph of moded predicates **)
   361 
   362 (* could be moved to Predicate_Compile_Core *)
   363 fun requires_modes polarity cls =
   364   let
   365     fun req_mode_of pol (t, derivation) =
   366       (case fst (strip_comb t) of
   367         Const (c, _) => SOME (c, (pol, Predicate_Compile_Core.head_mode_of derivation))
   368       | _ => NONE)
   369     fun req (Predicate_Compile_Aux.Prem t, derivation) = req_mode_of polarity (t, derivation)
   370       | req (Predicate_Compile_Aux.Negprem t, derivation) = req_mode_of (not polarity) (t, derivation)
   371       | req _ = NONE
   372   in      
   373     maps (fn (_, prems) => map_filter req prems) cls
   374   end
   375  
   376 structure Mode_Graph = Graph(type key = string * (bool * Predicate_Compile_Aux.mode)
   377   val ord = prod_ord fast_string_ord (prod_ord bool_ord Predicate_Compile_Aux.mode_ord));
   378 
   379 fun mk_moded_clauses_graph ctxt scc gr =
   380   let
   381     val options = Predicate_Compile_Aux.default_options
   382     val mode_analysis_options =
   383       {use_generators = true, reorder_premises = true, infer_pos_and_neg_modes = true}
   384     fun infer prednames (gr, (pos_modes, neg_modes, random)) =
   385       let
   386         val (lookup_modes, lookup_neg_modes, needs_random) =
   387           ((fn s => the (AList.lookup (op =) pos_modes s)),
   388            (fn s => the (AList.lookup (op =) neg_modes s)),
   389            (fn s => member (op =) (the (AList.lookup (op =) random s))))
   390         val (preds, all_vs, param_vs, all_modes, clauses) =
   391           Predicate_Compile_Core.prepare_intrs options ctxt prednames
   392             (maps (Core_Data.intros_of ctxt) prednames)
   393         val ((moded_clauses, random'), _) =
   394           Mode_Inference.infer_modes mode_analysis_options options 
   395             (lookup_modes, lookup_neg_modes, needs_random) ctxt preds all_modes param_vs clauses
   396         val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
   397         val pos_modes' = map (apsnd (map_filter (fn (true, m) => SOME m | _ => NONE))) modes
   398         val neg_modes' = map (apsnd (map_filter (fn (false, m) => SOME m | _ => NONE))) modes
   399         val _ = tracing ("Inferred modes:\n" ^
   400           cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
   401             (fn (p, m) => Predicate_Compile_Aux.string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
   402         val gr' = gr
   403           |> fold (fn (p, mps) => fold (fn (mode, cls) =>
   404                 Mode_Graph.new_node ((p, mode), cls)) mps)
   405             moded_clauses
   406           |> fold (fn (p, mps) => fold (fn (mode, cls) => fold (fn req =>
   407               Mode_Graph.add_edge ((p, mode), req)) (requires_modes (fst mode) cls)) mps)
   408             moded_clauses
   409       in
   410         (gr', (AList.merge (op =) (op =) (pos_modes, pos_modes'),
   411           AList.merge (op =) (op =) (neg_modes, neg_modes'),
   412           AList.merge (op =) (op =) (random, random')))
   413       end
   414   in  
   415     fst (fold infer (rev scc) (Mode_Graph.empty, ([], [], []))) 
   416   end
   417 
   418 fun declare_moded_predicate moded_preds table =
   419   let
   420     fun update' (p as (pred, (pol, mode))) table =
   421       if AList.defined (op =) table p then table else
   422         let
   423           val name = Long_Name.base_name pred ^ (if pol then "p" else "n")
   424             ^ Predicate_Compile_Aux.ascii_string_of_mode mode
   425           val p' = mk_conform first_lower "pred" (map snd table) name
   426         in
   427           AList.update (op =) (p, p') table
   428         end
   429   in
   430     fold update' moded_preds table
   431   end
   432 
   433 fun mk_program ctxt moded_gr moded_preds (prog, (moded_pred_table, constant_table)) =
   434   let
   435     val moded_pred_table' = declare_moded_predicate moded_preds moded_pred_table
   436     fun mk_literal pol derivation constant_table' t =
   437       let
   438         val (p, args) = strip_comb t
   439         val mode = Predicate_Compile_Core.head_mode_of derivation 
   440         val name = fst (dest_Const p)
   441         
   442         val p' = the (AList.lookup (op =) moded_pred_table' (name, (pol, mode)))
   443         val args' = map (translate_term ctxt constant_table') args
   444       in
   445         Rel (p', args')
   446       end
   447     fun mk_prem pol (indprem, derivation) constant_table =
   448       case indprem of
   449         Predicate_Compile_Aux.Generator (s, T) => (Ground (s, T), constant_table)
   450       | _ =>
   451         declare_consts (Term.add_const_names (Predicate_Compile_Aux.dest_indprem indprem) []) constant_table
   452         |> (fn constant_table' =>
   453           (case indprem of Predicate_Compile_Aux.Negprem t =>
   454             NegRel_of (mk_literal (not pol) derivation constant_table' t)
   455           | _ =>
   456             mk_literal pol derivation constant_table' (Predicate_Compile_Aux.dest_indprem indprem), constant_table'))
   457     fun mk_clause pred_name pol (ts, prems) (prog, constant_table) =
   458     let
   459       val constant_table' = declare_consts (fold Term.add_const_names ts []) constant_table
   460       val args = map (translate_term ctxt constant_table') ts
   461       val (prems', constant_table'') = fold_map (mk_prem pol) prems constant_table'
   462     in
   463       (((pred_name, args), Conj prems') :: prog, constant_table'')
   464     end
   465     fun mk_clauses (pred, mode as (pol, _)) =
   466       let
   467         val clauses = Mode_Graph.get_node moded_gr (pred, mode)
   468         val pred_name = the (AList.lookup (op =) moded_pred_table' (pred, mode))
   469       in
   470         fold (mk_clause pred_name pol) clauses
   471       end
   472   in
   473     apsnd (pair moded_pred_table') (fold mk_clauses moded_preds (prog, constant_table))
   474   end
   475 
   476 fun generate (use_modes, ensure_groundness) ctxt const =
   477   let 
   478     fun strong_conn_of gr keys =
   479       Graph.strong_conn (Graph.restrict (member (op =) (Graph.all_succs gr keys)) gr)
   480     val gr = Core_Data.intros_graph_of ctxt
   481     val gr' = add_edges depending_preds_of const gr
   482     val scc = strong_conn_of gr' [const]
   483     val initial_constant_table = 
   484       declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}] []
   485   in
   486     case use_modes of
   487       SOME mode =>
   488         let
   489           val moded_gr = mk_moded_clauses_graph ctxt scc gr
   490           val moded_gr' = Mode_Graph.restrict
   491             (member (op =) (Mode_Graph.all_succs moded_gr [(const, (true, mode))])) moded_gr
   492           val scc = Mode_Graph.strong_conn moded_gr' 
   493         in
   494           apfst rev (apsnd snd
   495             (fold (mk_program ctxt moded_gr') (rev scc) ([], ([], initial_constant_table))))
   496         end
   497       | NONE =>
   498         let 
   499           val _ = print_intros ctxt gr (flat scc)
   500           val constant_table = declare_consts (flat scc) initial_constant_table
   501         in
   502           apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table)
   503         end
   504   end
   505   
   506 (* implementation for fully enumerating predicates and
   507   for size-limited predicates for enumerating the values of a datatype upto a specific size *)
   508 
   509 fun add_ground_typ (Conj prems) = fold add_ground_typ prems
   510   | add_ground_typ (Ground (_, T)) = insert (op =) T
   511   | add_ground_typ _ = I
   512 
   513 fun mk_relname (Type (Tcon, Targs)) =
   514   first_lower (Long_Name.base_name Tcon) ^ space_implode "_" (map mk_relname Targs)
   515   | mk_relname _ = raise Fail "unexpected type"
   516 
   517 fun mk_lim_relname T = "lim_" ^  mk_relname T
   518 
   519 (* This is copied from "pat_completeness.ML" *)
   520 fun inst_constrs_of thy (T as Type (name, _)) =
   521   map (fn (Cn,CT) =>
   522     Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
   523     (the (Datatype.get_constrs thy name))
   524   | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
   525 
   526 fun is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T
   527   
   528 fun mk_ground_impl ctxt limited_types (T as Type (Tcon, Targs)) (seen, constant_table) =
   529   if member (op =) seen T then ([], (seen, constant_table))
   530   else
   531     let
   532       val (limited, size) = case AList.lookup (op =) limited_types T of
   533         SOME s => (true, s)
   534       | NONE => (false, 0)      
   535       val rel_name = (if limited then mk_lim_relname else mk_relname) T
   536       fun mk_impl (Const (constr_name, cT), recursive) (seen, constant_table) =
   537         let
   538           val constant_table' = declare_consts [constr_name] constant_table
   539           val Ts = binder_types cT
   540           val (rec_clauses, (seen', constant_table'')) =
   541             fold_map (mk_ground_impl ctxt limited_types) Ts (seen, constant_table')
   542           val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto (length Ts))
   543           val lim_var =
   544             if limited then
   545               if recursive then [AppF ("suc", [Var "Lim"])]              
   546               else [Var "Lim"]
   547             else [] 
   548           fun mk_prem v T' =
   549             if limited andalso T' = T then Rel (mk_lim_relname T', [Var "Lim", v])
   550             else Rel (mk_relname T', [v])
   551           val clause =
   552             ((rel_name, lim_var @ [maybe_AppF (translate_const constant_table'' constr_name, vars)]),
   553              Conj (map2 mk_prem vars Ts))
   554         in
   555           (clause :: flat rec_clauses, (seen', constant_table''))
   556         end
   557       val constrs = inst_constrs_of (Proof_Context.theory_of ctxt) T
   558       val constrs' = (constrs ~~ map (is_recursive_constr T) constrs)
   559         |> (fn cs => filter_out snd cs @ filter snd cs)
   560       val (clauses, constant_table') =
   561         apfst flat (fold_map mk_impl constrs' (T :: seen, constant_table))
   562       val size_term = funpow size (fn t => AppF ("suc", [t])) (Cons "zero")
   563     in
   564       ((if limited then
   565         cons ((mk_relname T, [Var "x"]), Rel (mk_lim_relname T, [size_term, Var "x"]))
   566       else I) clauses, constant_table')
   567     end
   568  | mk_ground_impl ctxt _ T (seen, constant_table) =
   569    raise Fail ("unexpected type :" ^ Syntax.string_of_typ ctxt T)
   570 
   571 fun replace_ground (Conj prems) = Conj (map replace_ground prems)
   572   | replace_ground (Ground (x, T)) =
   573     Rel (mk_relname T, [Var x])  
   574   | replace_ground p = p
   575   
   576 fun add_ground_predicates ctxt limited_types (p, constant_table) =
   577   let
   578     val ground_typs = fold (add_ground_typ o snd) p []
   579     val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt limited_types) ground_typs ([], constant_table)
   580     val p' = map (apsnd replace_ground) p
   581   in
   582     ((flat grs) @ p', constant_table')
   583   end
   584 
   585 (* make depth-limited version of predicate *)
   586 
   587 fun mk_lim_rel_name rel_name = "lim_" ^ rel_name
   588 
   589 fun mk_depth_limited rel_names ((rel_name, ts), prem) =
   590   let
   591     fun has_positive_recursive_prems (Conj prems) = exists has_positive_recursive_prems prems
   592       | has_positive_recursive_prems (Rel (rel, ts)) = member (op =) rel_names rel
   593       | has_positive_recursive_prems _ = false
   594     fun mk_lim_prem (Conj prems) = Conj (map mk_lim_prem prems)
   595       | mk_lim_prem (p as Rel (rel, ts)) =
   596         if member (op =) rel_names rel then Rel (mk_lim_rel_name rel, Var "Lim" :: ts) else p
   597       | mk_lim_prem p = p
   598   in
   599     if has_positive_recursive_prems prem then
   600       ((mk_lim_rel_name rel_name, (AppF ("suc", [Var "Lim"]))  :: ts), mk_lim_prem prem)
   601     else
   602       ((mk_lim_rel_name rel_name, (Var "Lim") :: ts), prem)
   603   end
   604 
   605 fun nat_term_of n = funpow n (fn t => AppF ("suc", [t])) (Cons "zero")
   606 
   607 fun add_limited_predicates limited_predicates (p, constant_table) =
   608   let                                     
   609     fun add (rel_names, limit) p = 
   610       let
   611         val clauses = filter (fn ((rel, _), _) => member (op =) rel_names rel) p
   612         val clauses' = map (mk_depth_limited rel_names) clauses
   613         fun mk_entry_clause rel_name =
   614           let
   615             val nargs = length (snd (fst
   616               (the (find_first (fn ((rel, _), _) => rel = rel_name) clauses))))
   617             val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto nargs)        
   618           in
   619             (("limited_" ^ rel_name, vars), Rel ("lim_" ^ rel_name, nat_term_of limit :: vars))
   620           end
   621       in (p @ (map mk_entry_clause rel_names) @ clauses') end
   622   in
   623     (fold add limited_predicates p, constant_table)
   624   end
   625 
   626 
   627 (* replace predicates in clauses *)
   628 
   629 (* replace (A, B, C) p = replace A by B in clauses of C *)
   630 fun replace ((from, to), location) p =
   631   let
   632     fun replace_prem (Conj prems) = Conj (map replace_prem prems)
   633       | replace_prem (r as Rel (rel, ts)) =
   634           if rel = from then Rel (to, ts) else r
   635       | replace_prem r = r
   636   in
   637     map (fn ((rel, args), prem) => ((rel, args), (if rel = location then replace_prem else I) prem)) p
   638   end
   639 
   640   
   641 (* reorder manually : reorder premises of ith clause of predicate p by a permutation perm *)
   642 
   643 fun reorder_manually reorder p =
   644   let
   645     fun reorder' (clause as ((rel, args), prem)) seen =
   646       let
   647         val seen' = AList.map_default (op =) (rel, 0) (fn x => x + 1) seen
   648         val i = the (AList.lookup (op =) seen' rel)
   649         val perm = AList.lookup (op =) reorder (rel, i)
   650         val prem' = (case perm of 
   651           SOME p => (case prem of Conj prems => Conj (map (nth prems) p) | _ => prem)
   652         | NONE => prem)
   653       in (((rel, args), prem'), seen') end
   654   in
   655     fst (fold_map reorder' p [])
   656   end
   657 
   658 (* rename variables to prolog-friendly names *)
   659 
   660 fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v))
   661 
   662 fun rename_vars_prem renaming = map_term_prem (rename_vars_term renaming)
   663 
   664 fun is_prolog_conform v =
   665   forall (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s) (Symbol.explode v)
   666   
   667 fun mk_renaming v renaming =
   668   (v, mk_conform first_upper "Var" (map snd renaming) v) :: renaming
   669 
   670 fun rename_vars_clause ((rel, args), prem) =
   671   let
   672     val vars = fold_prem_terms add_vars prem (fold add_vars args [])
   673     val renaming = fold mk_renaming vars []
   674   in ((rel, map (rename_vars_term renaming) args), rename_vars_prem renaming prem) end
   675   
   676 val rename_vars_program = map rename_vars_clause
   677 
   678 (* limit computation globally by some threshold *)
   679 
   680 fun limit_globally ctxt limit const_name (p, constant_table) =
   681   let
   682     val rel_names = fold (fn ((r, _), _) => insert (op =) r) p []
   683     val p' = map (mk_depth_limited rel_names) p
   684     val rel_name = translate_const constant_table const_name
   685     val nargs = length (snd (fst
   686       (the (find_first (fn ((rel, _), _) => rel = rel_name) p))))
   687     val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto nargs)
   688     val entry_clause = ((rel_name, vars), Rel ("lim_" ^ rel_name, nat_term_of limit :: vars))
   689     val p'' = filter_out (fn ((rel, _), _) => rel = rel_name) p
   690   in
   691     (entry_clause :: p' @ p'', constant_table)
   692   end
   693 
   694 (* post processing of generated prolog program *)
   695 
   696 fun post_process ctxt options const_name (p, constant_table) =
   697   (p, constant_table)
   698   |> (case #limit_globally options of
   699         SOME limit => limit_globally ctxt limit const_name
   700       | NONE => I)
   701   |> (if #ensure_groundness options then
   702         add_ground_predicates ctxt (#limited_types options)
   703       else I)
   704   |> tap (fn _ => tracing "Adding limited predicates...")
   705   |> add_limited_predicates (#limited_predicates options)
   706   |> tap (fn _ => tracing "Replacing predicates...")
   707   |> apfst (fold replace (#replacing options))
   708   |> apfst (reorder_manually (#manual_reorder options))
   709   |> apfst rename_vars_program
   710 
   711 (* code printer *)
   712 
   713 fun write_arith_op Plus = "+"
   714   | write_arith_op Minus = "-"
   715 
   716 fun write_term (Var v) = v
   717   | write_term (Cons c) = c
   718   | write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
   719   | write_term (ArithOp (oper, [a1, a2])) = write_term a1 ^ " " ^ write_arith_op oper ^ " " ^ write_term a2
   720   | write_term (Number n) = string_of_int n
   721 
   722 fun write_rel (pred, args) =
   723   pred ^ "(" ^ space_implode ", " (map write_term args) ^ ")" 
   724 
   725 fun write_prem (Conj prems) = space_implode ", " (map write_prem prems)
   726   | write_prem (Rel p) = write_rel p  
   727   | write_prem (NotRel p) = "not(" ^ write_rel p ^ ")"
   728   | write_prem (Eq (l, r)) = write_term l ^ " = " ^ write_term r
   729   | write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r
   730   | write_prem (ArithEq (l, r)) = write_term l ^ " is " ^ write_term r
   731   | write_prem (NotArithEq (l, r)) = write_term l ^ " =\\= " ^ write_term r
   732   | write_prem _ = raise Fail "Not a valid prolog premise"
   733 
   734 fun write_clause (head, prem) =
   735   write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".")
   736 
   737 fun write_program p =
   738   cat_lines (map write_clause p) 
   739 
   740 (* query templates *)
   741 
   742 (** query and prelude for swi-prolog **)
   743 
   744 fun swi_prolog_query_first (rel, args) vnames =
   745   "eval :- once("  ^ rel ^ "(" ^ space_implode ", " (map write_term args) ^ ")),\n" ^
   746   "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
   747   "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
   748   
   749 fun swi_prolog_query_firstn n (rel, args) vnames =
   750   "eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
   751     rel ^ "(" ^ space_implode ", " (map write_term args) ^ "), Sols), writelist(Sols).\n" ^
   752     "writelist([]).\n" ^
   753     "writelist([(" ^ space_implode ", " vnames ^ ")|SolutionTail]) :- " ^
   754     "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
   755     "\\n', [" ^ space_implode ", " vnames ^ "]), writelist(SolutionTail).\n"
   756   
   757 val swi_prolog_prelude =
   758   "#!/usr/bin/swipl -q -t main -f\n\n" ^  (* FIXME hardwired executable!? *)
   759   ":- use_module(library('dialect/ciao/aggregates')).\n" ^
   760   ":- style_check(-singleton).\n" ^
   761   ":- style_check(-discontiguous).\n" ^
   762   ":- style_check(-atom).\n\n" ^
   763   "main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^
   764   "main :- halt(1).\n"
   765 
   766 (** query and prelude for yap **)
   767 
   768 fun yap_query_first (rel, args) vnames =
   769   "eval :- once(" ^ rel ^ "(" ^ space_implode ", " (map write_term args) ^ ")),\n" ^
   770   "format('" ^ space_implode ";" (map (fn v => v ^ " = ~w") vnames) ^
   771   "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
   772 
   773 val yap_prelude =
   774   "#!/usr/bin/yap -L\n\n" ^
   775   ":- initialization(eval).\n"
   776 
   777 (* system-dependent query, prelude and invocation *)
   778 
   779 fun query system nsols = 
   780   case system of
   781     SWI_PROLOG =>
   782       (case nsols of NONE => swi_prolog_query_first | SOME n => swi_prolog_query_firstn n)
   783   | YAP =>
   784       case nsols of NONE => yap_query_first | SOME n =>
   785         error "No support for querying multiple solutions in the prolog system yap"
   786 
   787 fun prelude system =
   788   case system of SWI_PROLOG => swi_prolog_prelude | YAP => yap_prelude
   789 
   790 fun invoke system file =
   791   let
   792     val (env_var, cmd) =
   793       (case system of
   794         SWI_PROLOG => ("ISABELLE_SWIPL", "\"$ISABELLE_SWIPL\" -f ")
   795       | YAP => ("ISABELLE_YAP", "\"$ISABELLE_YAP\" -L "))
   796   in
   797     if getenv env_var = "" then
   798       (warning (env_var ^ " not set; could not execute code for " ^ string_of_system system); "")
   799     else fst (Isabelle_System.bash_output (cmd ^ File.shell_path file))
   800   end
   801 
   802 
   803 (* parsing prolog solution *)
   804 
   805 val scan_number =
   806   Scan.many1 Symbol.is_ascii_digit
   807 
   808 val scan_atom =
   809   Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
   810 
   811 val scan_var =
   812   Scan.many1
   813     (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
   814 
   815 val scan_ident =
   816   Scan.repeat (Scan.one
   817     (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s))
   818 
   819 fun dest_Char (Symbol.Char s) = s
   820 
   821 val string_of = implode o map (dest_Char o Symbol.decode)
   822 
   823 val is_atom_ident = forall Symbol.is_ascii_lower
   824 
   825 val is_var_ident =
   826   forall (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
   827 
   828 fun int_of_symbol_list xs = fold (fn x => fn s => s * 10 + (ord x - ord "0")) xs 0
   829 
   830 fun scan_terms xs = (((scan_term --| $$ ",") ::: scan_terms)
   831   || (scan_term >> single)) xs
   832 and scan_term xs =
   833   ((scan_number >> (Number o int_of_symbol_list))
   834   || (scan_var >> (Var o string_of))
   835   || ((scan_atom -- ($$ "(" |-- scan_terms --| $$ ")"))
   836     >> (fn (f, ts) => AppF (string_of f, ts)))
   837   || (scan_atom >> (Cons o string_of))) xs
   838 
   839 val parse_term = fst o Scan.finite Symbol.stopper
   840     (Scan.error (!! (fn _ => raise Fail "parsing prolog output failed")) scan_term)
   841   o raw_explode
   842   
   843 fun parse_solutions sol =
   844   let
   845     fun dest_eq s = case space_explode "=" s of
   846         (l :: r :: []) => parse_term (unprefix " " r)
   847       | _ => raise Fail "unexpected equation in prolog output"
   848     fun parse_solution s = map dest_eq (space_explode ";" s)
   849     val sols = case space_explode "\n" sol of [] => [] | s => fst (split_last s)  
   850   in
   851     map parse_solution sols
   852   end 
   853   
   854 (* calling external interpreter and getting results *)
   855 
   856 fun run (timeout, system) p (query_rel, args) vnames nsols =
   857   let
   858     val renaming = fold mk_renaming (fold add_vars args vnames) [] 
   859     val vnames' = map (fn v => the (AList.lookup (op =) renaming v)) vnames
   860     val args' = map (rename_vars_term renaming) args
   861     val prog = prelude system ^ query system nsols (query_rel, args') vnames' ^ write_program p
   862     val _ = tracing ("Generated prolog program:\n" ^ prog)
   863     val solution = TimeLimit.timeLimit timeout (fn prog =>
   864       Isabelle_System.with_tmp_file "prolog_file" "" (fn prolog_file =>
   865         (File.write prolog_file prog; invoke system prolog_file))) prog
   866     val _ = tracing ("Prolog returned solution(s):\n" ^ solution)
   867     val tss = parse_solutions solution
   868   in
   869     tss
   870   end
   871 
   872 (* restoring types in terms *)
   873 
   874 fun restore_term ctxt constant_table (Var s, T) = Free (s, T)
   875   | restore_term ctxt constant_table (Number n, @{typ "int"}) = HOLogic.mk_number @{typ "int"} n
   876   | restore_term ctxt constant_table (Number n, _) = raise (Fail "unexpected type for number") 
   877   | restore_term ctxt constant_table (Cons s, T) = Const (restore_const constant_table s, T)
   878   | restore_term ctxt constant_table (AppF (f, args), T) =
   879     let
   880       val thy = Proof_Context.theory_of ctxt
   881       val c = restore_const constant_table f
   882       val cT = Sign.the_const_type thy c
   883       val (argsT, resT) = strip_type cT
   884       val subst = Sign.typ_match thy (resT, T) Vartab.empty
   885       val argsT' = map (Envir.subst_type subst) argsT
   886     in
   887       list_comb (Const (c, Envir.subst_type subst cT),
   888         map (restore_term ctxt constant_table) (args ~~ argsT'))
   889     end
   890 
   891     
   892 (* restore numerals in natural numbers *)
   893 
   894 fun restore_nat_numerals t =
   895   if fastype_of t = @{typ nat} andalso is_some (try HOLogic.dest_nat t) then
   896     HOLogic.mk_number @{typ nat} (HOLogic.dest_nat t)
   897   else
   898     (case t of
   899         t1 $ t2 => restore_nat_numerals t1 $ restore_nat_numerals t2
   900       | t => t)
   901   
   902 (* values command *)
   903 
   904 val preprocess_options = Predicate_Compile_Aux.Options {
   905   expected_modes = NONE,
   906   proposed_modes = [],
   907   proposed_names = [],
   908   show_steps = false,
   909   show_intermediate_results = false,
   910   show_proof_trace = false,
   911   show_modes = false,
   912   show_mode_inference = false,
   913   show_compilation = false,
   914   show_caught_failures = false,
   915   show_invalid_clauses = false,
   916   skip_proof = true,
   917   no_topmost_reordering = false,
   918   function_flattening = true,
   919   specialise = false,
   920   fail_safe_function_flattening = false,
   921   no_higher_order_predicate = [],
   922   inductify = false,
   923   detect_switches = true,
   924   smart_depth_limiting = true,
   925   compilation = Predicate_Compile_Aux.Pred
   926 }
   927 
   928 fun values ctxt soln t_compr =
   929   let
   930     val options = code_options_of (Proof_Context.theory_of ctxt)
   931     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
   932       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
   933     val (body, Ts, fp) = HOLogic.strip_psplits split;
   934     val output_names = Name.variant_list (Term.add_free_names body [])
   935       (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
   936     val output_frees = rev (map2 (curry Free) output_names Ts)
   937     val body = subst_bounds (output_frees, body)
   938     val (pred as Const (name, T), all_args) =
   939       case strip_comb body of
   940         (Const (name, T), all_args) => (Const (name, T), all_args)
   941       | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
   942     val _ = tracing "Preprocessing specification..."
   943     val T = Sign.the_const_type (Proof_Context.theory_of ctxt) name
   944     val t = Const (name, T)
   945     val thy' =
   946       Theory.copy (Proof_Context.theory_of ctxt)
   947       |> Predicate_Compile.preprocess preprocess_options t
   948     val ctxt' = Proof_Context.init_global thy'
   949     val _ = tracing "Generating prolog program..."
   950     val (p, constant_table) = generate (NONE, #ensure_groundness options) ctxt' name (* FIXME *)
   951       |> post_process ctxt' options name
   952     val constant_table' = declare_consts (fold Term.add_const_names all_args []) constant_table
   953     val args' = map (translate_term ctxt constant_table') all_args
   954     val _ = tracing "Running prolog program..."
   955     val system_config = System_Config.get (Context.Proof ctxt)
   956     val tss = run (#timeout system_config, #prolog_system system_config)
   957       p (translate_const constant_table' name, args') output_names soln
   958     val _ = tracing "Restoring terms..."
   959     val empty = Const("Orderings.bot_class.bot", fastype_of t_compr)
   960     fun mk_insert x S =
   961       Const (@{const_name "Set.insert"}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S 
   962     fun mk_set_compr in_insert [] xs =
   963        rev ((Free ("dots", fastype_of t_compr)) ::  (* FIXME proper name!? *)
   964         (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs))
   965       | mk_set_compr in_insert (t :: ts) xs =
   966         let
   967           val frees = Term.add_frees t []
   968         in
   969           if null frees then
   970             mk_set_compr (t :: in_insert) ts xs
   971           else
   972             let
   973               val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt' [t]) ("uu", fastype_of t)
   974               val set_compr =
   975                 HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t))
   976                   frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"})))
   977             in
   978               mk_set_compr [] ts
   979                 (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs))  
   980             end
   981         end
   982   in
   983       foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr []
   984         (map (fn ts => HOLogic.mk_tuple 
   985           (map (restore_nat_numerals o restore_term ctxt' constant_table) (ts ~~ Ts))) tss) [])
   986   end
   987 
   988 fun values_cmd print_modes soln raw_t state =
   989   let
   990     val ctxt = Toplevel.context_of state
   991     val t = Syntax.read_term ctxt raw_t
   992     val t' = values ctxt soln t
   993     val ty' = Term.type_of t'
   994     val ctxt' = Variable.auto_fixes t' ctxt
   995     val _ = tracing "Printing terms..."
   996     val p = Print_Mode.with_modes print_modes (fn () =>
   997       Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
   998         Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
   999   in Pretty.writeln p end;
  1000 
  1001 
  1002 (* renewing the values command for Prolog queries *)
  1003 
  1004 val opt_print_modes =
  1005   Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) [];
  1006 
  1007 val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag
  1008   (opt_print_modes -- Scan.optional (Parse.nat >> SOME) NONE -- Parse.term
  1009    >> (fn ((print_modes, soln), t) => Toplevel.keep
  1010         (values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*)
  1011 
  1012 (* quickcheck generator *)
  1013 
  1014 (* FIXME: a small clone of Predicate_Compile_Quickcheck - maybe refactor out commons *)
  1015 
  1016 val active = Attrib.setup_config_bool @{binding quickcheck_prolog_active} (K true);
  1017 
  1018 fun test_term ctxt (t, eval_terms) =
  1019   let
  1020     val t' = fold_rev absfree (Term.add_frees t []) t
  1021     val options = code_options_of (Proof_Context.theory_of ctxt)
  1022     val thy = Theory.copy (Proof_Context.theory_of ctxt)
  1023     val ((((full_constname, constT), vs'), intro), thy1) =
  1024       Predicate_Compile_Aux.define_quickcheck_predicate t' thy
  1025     val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
  1026     val thy3 = Predicate_Compile.preprocess preprocess_options (Const (full_constname, constT)) thy2
  1027     val ctxt' = Proof_Context.init_global thy3
  1028     val _ = tracing "Generating prolog program..."
  1029     val (p, constant_table) = generate (NONE, true) ctxt' full_constname
  1030       |> post_process ctxt' (set_ensure_groundness options) full_constname
  1031     val _ = tracing "Running prolog program..."
  1032     val system_config = System_Config.get (Context.Proof ctxt)
  1033     val tss = run (#timeout system_config, #prolog_system system_config)
  1034       p (translate_const constant_table full_constname, map (Var o fst) vs') (map fst vs') (SOME 1)
  1035     val _ = tracing "Restoring terms..."
  1036     val counterexample =
  1037       case tss of
  1038         [ts] => SOME (map (restore_term ctxt' constant_table) (ts ~~ map snd vs'))
  1039       | _ => NONE
  1040   in
  1041     Quickcheck.Result
  1042       {counterexample = Option.map (pair true o curry (op ~~) (Term.add_free_names t [])) counterexample,
  1043        evaluation_terms = Option.map (K []) counterexample, timings = [], reports = []}
  1044   end;
  1045 
  1046 fun test_goals ctxt _ insts goals =
  1047   let
  1048     val correct_inst_goals = Quickcheck_Common.instantiate_goals ctxt insts goals
  1049   in
  1050     Quickcheck_Common.collect_results (test_term ctxt) (maps (map snd) correct_inst_goals) []
  1051   end
  1052   
  1053   
  1054 end;