bulwahn@38073: (* Title: HOL/Tools/Predicate_Compile/code_prolog.ML bulwahn@38073: Author: Lukas Bulwahn, TU Muenchen bulwahn@38073: bulwahn@38073: Prototype of an code generator for logic programming languages (a.k.a. Prolog) bulwahn@38073: *) bulwahn@38073: bulwahn@38073: signature CODE_PROLOG = bulwahn@38073: sig bulwahn@38727: type code_options = {ensure_groundness : bool} bulwahn@38727: val options : code_options ref bulwahn@38727: bulwahn@38113: datatype arith_op = Plus | Minus bulwahn@38112: datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list bulwahn@38113: | Number of int | ArithOp of arith_op * prol_term list; bulwahn@38113: datatype prem = Conj of prem list bulwahn@38113: | Rel of string * prol_term list | NotRel of string * prol_term list bulwahn@38113: | Eq of prol_term * prol_term | NotEq of prol_term * prol_term bulwahn@38727: | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term bulwahn@38727: | Ground of string * typ; bulwahn@38727: bulwahn@38079: type clause = ((string * prol_term list) * prem); bulwahn@38073: type logic_program = clause list; bulwahn@38079: type constant_table = (string * string) list bulwahn@38727: bulwahn@38727: val generate : code_options -> Proof.context -> string list -> (logic_program * constant_table) bulwahn@38079: val write_program : logic_program -> string bulwahn@38079: val run : logic_program -> string -> string list -> int option -> prol_term list list bulwahn@38073: bulwahn@38079: val trace : bool Unsynchronized.ref bulwahn@38073: end; bulwahn@38073: bulwahn@38073: structure Code_Prolog : CODE_PROLOG = bulwahn@38073: struct bulwahn@38073: bulwahn@38079: (* diagnostic tracing *) bulwahn@38079: bulwahn@38079: val trace = Unsynchronized.ref false bulwahn@38079: bulwahn@38079: fun tracing s = if !trace then Output.tracing s else () bulwahn@38727: bulwahn@38727: (* code generation options *) bulwahn@38727: bulwahn@38727: type code_options = {ensure_groundness : bool} bulwahn@38727: bulwahn@38727: val options = Unsynchronized.ref {ensure_groundness = false}; bulwahn@38727: bulwahn@38073: (* general string functions *) bulwahn@38073: bulwahn@38073: val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode; bulwahn@38073: val first_lower = implode o nth_map 0 Symbol.to_ascii_lower o explode; bulwahn@38073: bulwahn@38073: (* internal program representation *) bulwahn@38073: bulwahn@38113: datatype arith_op = Plus | Minus bulwahn@38113: bulwahn@38112: datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list bulwahn@38113: | Number of int | ArithOp of arith_op * prol_term list; bulwahn@38113: bulwahn@38113: fun is_Var (Var _) = true bulwahn@38113: | is_Var _ = false bulwahn@38113: bulwahn@38113: fun is_arith_term (Var _) = true bulwahn@38113: | is_arith_term (Number _) = true bulwahn@38113: | is_arith_term (ArithOp (_, operands)) = forall is_arith_term operands bulwahn@38113: | is_arith_term _ = false bulwahn@38073: bulwahn@38081: fun string_of_prol_term (Var s) = "Var " ^ s bulwahn@38075: | string_of_prol_term (Cons s) = "Cons " ^ s bulwahn@38075: | string_of_prol_term (AppF (f, args)) = f ^ "(" ^ commas (map string_of_prol_term args) ^ ")" bulwahn@38112: | string_of_prol_term (Number n) = "Number " ^ string_of_int n bulwahn@38075: bulwahn@38113: datatype prem = Conj of prem list bulwahn@38113: | Rel of string * prol_term list | NotRel of string * prol_term list bulwahn@38113: | Eq of prol_term * prol_term | NotEq of prol_term * prol_term bulwahn@38727: | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term bulwahn@38727: | Ground of string * typ; bulwahn@38727: bulwahn@38073: fun dest_Rel (Rel (c, ts)) = (c, ts) bulwahn@38073: bulwahn@38079: type clause = ((string * prol_term list) * prem); bulwahn@38073: bulwahn@38073: type logic_program = clause list; bulwahn@38073: bulwahn@38073: (* translation from introduction rules to internal representation *) bulwahn@38073: bulwahn@38079: (** constant table **) bulwahn@38079: bulwahn@38079: type constant_table = (string * string) list bulwahn@38079: bulwahn@38073: (* assuming no clashing *) bulwahn@38079: fun mk_constant_table consts = bulwahn@38079: AList.make (first_lower o Long_Name.base_name) consts bulwahn@38079: bulwahn@38079: fun declare_consts consts constant_table = bulwahn@38079: fold (fn c => AList.update (op =) (c, first_lower (Long_Name.base_name c))) consts constant_table bulwahn@38079: bulwahn@38079: fun translate_const constant_table c = bulwahn@38079: case AList.lookup (op =) constant_table c of bulwahn@38079: SOME c' => c' bulwahn@38079: | NONE => error ("No such constant: " ^ c) bulwahn@38073: bulwahn@38079: fun inv_lookup _ [] _ = NONE bulwahn@38079: | inv_lookup eq ((key, value)::xs) value' = bulwahn@38079: if eq (value', value) then SOME key bulwahn@38079: else inv_lookup eq xs value'; bulwahn@38079: bulwahn@38079: fun restore_const constant_table c = bulwahn@38079: case inv_lookup (op =) constant_table c of bulwahn@38079: SOME c' => c' bulwahn@38079: | NONE => error ("No constant corresponding to " ^ c) bulwahn@38727: bulwahn@38079: (** translation of terms, literals, premises, and clauses **) bulwahn@38079: bulwahn@38113: fun translate_arith_const @{const_name "Groups.plus_class.plus"} = SOME Plus bulwahn@38113: | translate_arith_const @{const_name "Groups.minus_class.minus"} = SOME Minus bulwahn@38113: | translate_arith_const _ = NONE bulwahn@38113: bulwahn@38079: fun translate_term ctxt constant_table t = bulwahn@38112: case try HOLogic.dest_number t of bulwahn@38112: SOME (@{typ "int"}, n) => Number n bulwahn@38112: | NONE => bulwahn@38112: (case strip_comb t of bulwahn@38112: (Free (v, T), []) => Var v bulwahn@38112: | (Const (c, _), []) => Cons (translate_const constant_table c) bulwahn@38112: | (Const (c, _), args) => bulwahn@38113: (case translate_arith_const c of bulwahn@38113: SOME aop => ArithOp (aop, map (translate_term ctxt constant_table) args) bulwahn@38113: | NONE => bulwahn@38113: AppF (translate_const constant_table c, map (translate_term ctxt constant_table) args)) bulwahn@38112: | _ => error ("illegal term for translation: " ^ Syntax.string_of_term ctxt t)) bulwahn@38073: bulwahn@38079: fun translate_literal ctxt constant_table t = bulwahn@38073: case strip_comb t of bulwahn@38079: (Const (@{const_name "op ="}, _), [l, r]) => bulwahn@38113: let bulwahn@38113: val l' = translate_term ctxt constant_table l bulwahn@38113: val r' = translate_term ctxt constant_table r bulwahn@38113: in bulwahn@38113: (if is_Var l' andalso is_arith_term r' then ArithEq else Eq) (l', r') bulwahn@38113: end bulwahn@38079: | (Const (c, _), args) => bulwahn@38079: Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args) bulwahn@38073: | _ => error ("illegal literal for translation: " ^ Syntax.string_of_term ctxt t) bulwahn@38073: bulwahn@38073: fun NegRel_of (Rel lit) = NotRel lit bulwahn@38073: | NegRel_of (Eq eq) = NotEq eq bulwahn@38113: | NegRel_of (ArithEq eq) = NotArithEq eq bulwahn@38113: bulwahn@38727: fun mk_groundness_prems t = map Ground (Term.add_frees t []) bulwahn@38727: bulwahn@38727: fun translate_prem options ctxt constant_table t = bulwahn@38073: case try HOLogic.dest_not t of bulwahn@38727: SOME t => bulwahn@38727: if #ensure_groundness options then bulwahn@38727: Conj (mk_groundness_prems t @ [NegRel_of (translate_literal ctxt constant_table t)]) bulwahn@38727: else bulwahn@38727: NegRel_of (translate_literal ctxt constant_table t) bulwahn@38079: | NONE => translate_literal ctxt constant_table t bulwahn@38114: bulwahn@38114: fun imp_prems_conv cv ct = bulwahn@38114: case Thm.term_of ct of bulwahn@38114: Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct bulwahn@38114: | _ => Conv.all_conv ct bulwahn@38114: bulwahn@38114: fun Trueprop_conv cv ct = bulwahn@38114: case Thm.term_of ct of haftmann@38558: Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct bulwahn@38114: | _ => raise Fail "Trueprop_conv" bulwahn@38114: bulwahn@38114: fun preprocess_intro thy rule = bulwahn@38114: Conv.fconv_rule bulwahn@38114: (imp_prems_conv bulwahn@38114: (Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq})))) bulwahn@38114: (Thm.transfer thy rule) bulwahn@38114: bulwahn@38727: fun translate_intros options ctxt gr const constant_table = bulwahn@38073: let bulwahn@38114: val intros = map (preprocess_intro (ProofContext.theory_of ctxt)) (Graph.get_node gr const) bulwahn@38073: val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt bulwahn@38079: val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table bulwahn@38073: fun translate_intro intro = bulwahn@38073: let bulwahn@38073: val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro) bulwahn@38727: val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro) bulwahn@38727: val prems' = Conj (map (translate_prem options ctxt' constant_table') prems) bulwahn@38079: val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems') bulwahn@38073: in clause end bulwahn@38079: in (map translate_intro intros', constant_table') end bulwahn@38073: bulwahn@38727: fun generate options ctxt const = bulwahn@38073: let bulwahn@38073: fun strong_conn_of gr keys = bulwahn@38073: Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr) bulwahn@38073: val gr = Predicate_Compile_Core.intros_graph_of ctxt bulwahn@38073: val scc = strong_conn_of gr const bulwahn@38079: val constant_table = mk_constant_table (flat scc) bulwahn@38073: in bulwahn@38727: apfst flat (fold_map (translate_intros options ctxt gr) (flat scc) constant_table) bulwahn@38073: end bulwahn@38727: bulwahn@38727: (* add implementation for ground predicates *) bulwahn@38073: bulwahn@38727: fun add_ground_typ (Conj prems) = fold add_ground_typ prems bulwahn@38727: | add_ground_typ (Ground (_, T)) = insert (op =) T bulwahn@38727: | add_ground_typ _ = I bulwahn@38073: bulwahn@38727: fun mk_ground_impl ctxt (Type (Tcon, [])) constant_table = bulwahn@38073: let bulwahn@38727: fun mk_impl (constr_name, T) constant_table = bulwahn@38727: if binder_types T = [] then bulwahn@38727: let bulwahn@38727: val constant_table' = declare_consts [constr_name] constant_table bulwahn@38727: val clause = (("is_" ^ first_lower (Long_Name.base_name Tcon), bulwahn@38727: [Cons (translate_const constant_table' constr_name)]), Conj []) bulwahn@38727: in bulwahn@38727: (clause, constant_table') bulwahn@38727: end bulwahn@38727: else raise Fail "constructor with arguments" bulwahn@38727: val constrs = the (Datatype.get_constrs (ProofContext.theory_of ctxt) Tcon) bulwahn@38727: in fold_map mk_impl constrs constant_table end bulwahn@38727: | mk_ground_impl ctxt (Type (Tcon, _)) constant_table = bulwahn@38727: raise Fail "type constructor with type arguments" bulwahn@38727: bulwahn@38727: fun replace_ground (Conj prems) = Conj (map replace_ground prems) bulwahn@38727: | replace_ground (Ground (x, Type (Tcon, []))) = bulwahn@38727: Rel ("is_" ^ first_lower (Long_Name.base_name Tcon), [Var x]) bulwahn@38727: | replace_ground p = p bulwahn@38727: bulwahn@38727: fun add_ground_predicates ctxt (p, constant_table) = bulwahn@38727: let bulwahn@38727: val ground_typs = fold (add_ground_typ o snd) p [] bulwahn@38727: val (grs, constant_table') = fold_map (mk_ground_impl ctxt) ground_typs constant_table bulwahn@38727: val p' = map (apsnd replace_ground) p bulwahn@38073: in bulwahn@38727: ((flat grs) @ p', constant_table') bulwahn@38073: end bulwahn@38727: bulwahn@38073: (* code printer *) bulwahn@38073: bulwahn@38113: fun write_arith_op Plus = "+" bulwahn@38113: | write_arith_op Minus = "-" bulwahn@38113: bulwahn@38081: fun write_term (Var v) = first_upper v bulwahn@38079: | write_term (Cons c) = c bulwahn@38113: | write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")" bulwahn@38113: | write_term (ArithOp (oper, [a1, a2])) = write_term a1 ^ " " ^ write_arith_op oper ^ " " ^ write_term a2 bulwahn@38112: | write_term (Number n) = string_of_int n bulwahn@38073: bulwahn@38073: fun write_rel (pred, args) = bulwahn@38073: pred ^ "(" ^ space_implode ", " (map write_term args) ^ ")" bulwahn@38073: bulwahn@38073: fun write_prem (Conj prems) = space_implode ", " (map write_prem prems) bulwahn@38073: | write_prem (Rel p) = write_rel p bulwahn@38073: | write_prem (NotRel p) = "not(" ^ write_rel p ^ ")" bulwahn@38073: | write_prem (Eq (l, r)) = write_term l ^ " = " ^ write_term r bulwahn@38073: | write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r bulwahn@38113: | write_prem (ArithEq (l, r)) = write_term l ^ " is " ^ write_term r bulwahn@38113: | write_prem (NotArithEq (l, r)) = write_term l ^ " =\\= " ^ write_term r bulwahn@38073: bulwahn@38073: fun write_clause (head, prem) = bulwahn@38073: write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".") bulwahn@38073: bulwahn@38073: fun write_program p = bulwahn@38073: cat_lines (map write_clause p) bulwahn@38073: bulwahn@38078: (** query templates **) bulwahn@38078: bulwahn@38073: fun query_first rel vnames = bulwahn@38073: "eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^ bulwahn@38082: "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^ bulwahn@38082: "\\n', [" ^ space_implode ", " vnames ^ "]).\n" bulwahn@38077: bulwahn@38077: fun query_firstn n rel vnames = bulwahn@38077: "eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^ bulwahn@38077: rel ^ "(" ^ space_implode ", " vnames ^ "), Sols), writelist(Sols).\n" ^ bulwahn@38077: "writelist([]).\n" ^ bulwahn@38077: "writelist([(" ^ space_implode ", " vnames ^ ")|T]) :- " ^ bulwahn@38079: "writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^ bulwahn@38079: "\\n', [" ^ space_implode ", " vnames ^ "]), writelist(T).\n" bulwahn@38077: bulwahn@38073: val prelude = bulwahn@38073: "#!/usr/bin/swipl -q -t main -f\n\n" ^ bulwahn@38077: ":- use_module(library('dialect/ciao/aggregates')).\n" ^ bulwahn@38073: ":- style_check(-singleton).\n\n" ^ bulwahn@38073: "main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^ bulwahn@38073: "main :- halt(1).\n" bulwahn@38075: bulwahn@38075: (* parsing prolog solution *) bulwahn@38115: val scan_number = bulwahn@38115: Scan.many1 Symbol.is_ascii_digit bulwahn@38075: bulwahn@38075: val scan_atom = bulwahn@38078: Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_quasi s) bulwahn@38075: bulwahn@38075: val scan_var = bulwahn@38078: Scan.many1 bulwahn@38078: (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s) bulwahn@38075: bulwahn@38076: val scan_ident = bulwahn@38076: Scan.repeat (Scan.one bulwahn@38076: (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)) bulwahn@38076: bulwahn@38075: fun dest_Char (Symbol.Char s) = s bulwahn@38075: bulwahn@38075: val string_of = concat o map (dest_Char o Symbol.decode) bulwahn@38075: bulwahn@38076: val is_atom_ident = forall Symbol.is_ascii_lower bulwahn@38076: bulwahn@38076: val is_var_ident = bulwahn@38076: forall (fn s => Symbol.is_ascii_upper s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s) bulwahn@38078: bulwahn@38115: fun int_of_symbol_list xs = fold (fn x => fn s => s * 10 + (ord x - ord "0")) xs 0 bulwahn@38115: bulwahn@38078: fun scan_terms xs = (((scan_term --| $$ ",") ::: scan_terms) bulwahn@38078: || (scan_term >> single)) xs bulwahn@38078: and scan_term xs = bulwahn@38115: ((scan_number >> (Number o int_of_symbol_list)) bulwahn@38115: || (scan_var >> (Var o string_of)) bulwahn@38078: || ((scan_atom -- ($$ "(" |-- scan_terms --| $$ ")")) bulwahn@38079: >> (fn (f, ts) => AppF (string_of f, ts))) bulwahn@38078: || (scan_atom >> (Cons o string_of))) xs bulwahn@38079: bulwahn@38075: val parse_term = fst o Scan.finite Symbol.stopper bulwahn@38077: (Scan.error (!! (fn _ => raise Fail "parsing prolog output failed")) scan_term) bulwahn@38075: o explode bulwahn@38075: bulwahn@38079: fun parse_solutions sol = bulwahn@38075: let bulwahn@38077: fun dest_eq s = case space_explode "=" s of bulwahn@38075: (l :: r :: []) => parse_term (unprefix " " r) bulwahn@38078: | _ => raise Fail "unexpected equation in prolog output" bulwahn@38079: fun parse_solution s = map dest_eq (space_explode ";" s) bulwahn@38075: in bulwahn@38079: map parse_solution (fst (split_last (space_explode "\n" sol))) bulwahn@38075: end bulwahn@38073: bulwahn@38073: (* calling external interpreter and getting results *) bulwahn@38073: bulwahn@38077: fun run p query_rel vnames nsols = bulwahn@38073: let bulwahn@38073: val cmd = Path.named_root bulwahn@38077: val query = case nsols of NONE => query_first | SOME n => query_firstn n bulwahn@38077: val prog = prelude ^ query query_rel vnames ^ write_program p bulwahn@38079: val _ = tracing ("Generated prolog program:\n" ^ prog) bulwahn@38073: val prolog_file = File.tmp_path (Path.basic "prolog_file") bulwahn@38073: val _ = File.write prolog_file prog bulwahn@38077: val (solution, _) = bash_output ("/usr/local/bin/swipl -f " ^ File.shell_path prolog_file) bulwahn@38079: val _ = tracing ("Prolog returned solution(s):\n" ^ solution) bulwahn@38079: val tss = parse_solutions solution bulwahn@38073: in bulwahn@38079: tss bulwahn@38073: end bulwahn@38073: bulwahn@38075: (* values command *) bulwahn@38075: bulwahn@38081: fun restore_term ctxt constant_table (Var s, T) = Free (s, T) bulwahn@38115: | restore_term ctxt constant_table (Number n, @{typ "int"}) = HOLogic.mk_number @{typ "int"} n bulwahn@38115: | restore_term ctxt constant_table (Number n, _) = raise (Fail "unexpected type for number") bulwahn@38079: | restore_term ctxt constant_table (Cons s, T) = Const (restore_const constant_table s, T) bulwahn@38079: | restore_term ctxt constant_table (AppF (f, args), T) = bulwahn@38079: let bulwahn@38079: val thy = ProofContext.theory_of ctxt bulwahn@38079: val c = restore_const constant_table f bulwahn@38079: val cT = Sign.the_const_type thy c bulwahn@38079: val (argsT, resT) = strip_type cT bulwahn@38079: val subst = Sign.typ_match thy (resT, T) Vartab.empty bulwahn@38079: val argsT' = map (Envir.subst_type subst) argsT bulwahn@38079: in bulwahn@38079: list_comb (Const (c, Envir.subst_type subst cT), bulwahn@38079: map (restore_term ctxt constant_table) (args ~~ argsT')) bulwahn@38079: end bulwahn@38079: bulwahn@38075: fun values ctxt soln t_compr = bulwahn@38075: let bulwahn@38727: val options = !options bulwahn@38075: val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t bulwahn@38075: | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr); bulwahn@38075: val (body, Ts, fp) = HOLogic.strip_psplits split; bulwahn@38075: val output_names = Name.variant_list (Term.add_free_names body []) bulwahn@38075: (map (fn i => "x" ^ string_of_int i) (1 upto length Ts)) bulwahn@38080: val output_frees = rev (map2 (curry Free) output_names Ts) bulwahn@38075: val body = subst_bounds (output_frees, body) bulwahn@38075: val (pred as Const (name, T), all_args) = bulwahn@38075: case strip_comb body of bulwahn@38075: (Const (name, T), all_args) => (Const (name, T), all_args) bulwahn@38075: | (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head) bulwahn@38075: val vnames = bulwahn@38075: case try (map (fst o dest_Free)) all_args of bulwahn@38075: SOME vs => vs bulwahn@38075: | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args)) bulwahn@38079: val _ = tracing "Generating prolog program..." bulwahn@38727: val (p, constant_table) = generate options ctxt [name] bulwahn@38727: |> (if #ensure_groundness options then add_ground_predicates ctxt else I) bulwahn@38079: val _ = tracing "Running prolog program..." bulwahn@38079: val tss = run p (translate_const constant_table name) (map first_upper vnames) soln bulwahn@38079: val _ = tracing "Restoring terms..." bulwahn@38115: val empty = Const("Orderings.bot_class.bot", fastype_of t_compr) bulwahn@38115: fun mk_insert x S = bulwahn@38115: Const (@{const_name "Set.insert"}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S bulwahn@38115: fun mk_set_compr in_insert [] xs = bulwahn@38115: rev ((Free ("...", fastype_of t_compr)) :: bulwahn@38115: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)) bulwahn@38115: | mk_set_compr in_insert (t :: ts) xs = bulwahn@38115: let bulwahn@38115: val frees = Term.add_frees t [] bulwahn@38115: in bulwahn@38115: if null frees then bulwahn@38115: mk_set_compr (t :: in_insert) ts xs bulwahn@38115: else bulwahn@38115: let bulwahn@38115: val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt [t]) ("uu", fastype_of t) bulwahn@38115: val set_compr = bulwahn@38115: HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t)) bulwahn@38115: frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"}))) bulwahn@38115: in bulwahn@38115: set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs) bulwahn@38115: end bulwahn@38115: end bulwahn@38075: in bulwahn@38115: foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr [] bulwahn@38115: (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt constant_table) (ts ~~ Ts))) tss) []) bulwahn@38075: end bulwahn@38075: bulwahn@38075: fun values_cmd print_modes soln raw_t state = bulwahn@38075: let bulwahn@38075: val ctxt = Toplevel.context_of state bulwahn@38075: val t = Syntax.read_term ctxt raw_t bulwahn@38075: val t' = values ctxt soln t bulwahn@38075: val ty' = Term.type_of t' bulwahn@38075: val ctxt' = Variable.auto_fixes t' ctxt bulwahn@38115: val _ = tracing "Printing terms..." bulwahn@38075: val p = Print_Mode.with_modes print_modes (fn () => bulwahn@38075: Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, bulwahn@38075: Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); bulwahn@38075: in Pretty.writeln p end; bulwahn@38075: bulwahn@38075: bulwahn@38075: (* renewing the values command for Prolog queries *) bulwahn@38075: bulwahn@38075: val opt_print_modes = bulwahn@38075: Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) []; bulwahn@38075: bulwahn@38075: val _ = Outer_Syntax.improper_command "values" "enumerate and print comprehensions" Keyword.diag bulwahn@38077: (opt_print_modes -- Scan.optional (Parse.nat >> SOME) NONE -- Parse.term bulwahn@38075: >> (fn ((print_modes, soln), t) => Toplevel.keep haftmann@38504: (values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*) bulwahn@38075: bulwahn@38073: end;