blanchet@35866: (* Title: HOL/Sledgehammer/sledgehammer_isar.ML blanchet@35866: Author: Jasmin Blanchette, TU Muenchen blanchet@35866: blanchet@35866: Adds "sledgehammer" and related commands to Isabelle/Isar's outer syntax. blanchet@35866: *) blanchet@35866: blanchet@35961: signature SLEDGEHAMMER_ISAR = blanchet@35961: sig blanchet@35961: type params = ATP_Manager.params blanchet@35961: blanchet@35961: val default_params : theory -> (string * string) list -> params blanchet@35961: end; blanchet@35961: blanchet@35866: structure Sledgehammer_Isar : sig end = blanchet@35866: struct blanchet@35866: blanchet@35866: open ATP_Manager blanchet@35866: open ATP_Minimal blanchet@35963: open ATP_Wrapper blanchet@35961: open Sledgehammer_Util blanchet@35866: blanchet@35866: structure K = OuterKeyword and P = OuterParse blanchet@35866: blanchet@35963: datatype facta = Facta of {add: Facts.ref list, del: Facts.ref list, only: bool} blanchet@35963: blanchet@35963: fun add_to_facta ns = Facta {add = ns, del = [], only = false}; blanchet@35963: fun del_from_facta ns = Facta {add = [], del = ns, only = false}; blanchet@35963: fun only_facta ns = Facta {add = ns, del = [], only = true}; blanchet@35963: val default_facta = add_to_facta [] blanchet@35963: fun merge_facta_pairwise (Facta f1) (Facta f2) = blanchet@35963: Facta {add = #add f1 @ #add f2, del = #del f1 @ #del f2, blanchet@35963: only = #only f1 orelse #only f2} blanchet@35963: fun merge_facta fs = fold merge_facta_pairwise fs default_facta blanchet@35963: blanchet@35961: type raw_param = string * string list blanchet@35866: blanchet@35961: val default_default_params = blanchet@35961: [("debug", "false"), blanchet@35961: ("verbose", "false"), blanchet@35961: ("relevance_threshold", "0.5"), blanchet@35961: ("higher_order", "smart"), blanchet@35961: ("respect_no_atp", "true"), blanchet@35961: ("follow_defs", "false"), blanchet@35961: ("isar_proof", "false"), blanchet@35961: ("minimize_timeout", "5 s")] blanchet@35866: blanchet@35961: val negated_params = blanchet@35961: [("no_debug", "debug"), blanchet@35961: ("quiet", "verbose"), blanchet@35961: ("partial_types", "full_types"), blanchet@35961: ("first_order", "higher_order"), blanchet@35961: ("ignore_no_atp", "respect_no_atp"), blanchet@35961: ("dont_follow_defs", "follow_defs"), blanchet@35961: ("metis_proof", "isar_proof")] blanchet@35866: blanchet@35961: val property_affected_params = ["atps", "full_types", "timeout"] blanchet@35866: blanchet@35961: fun is_known_raw_param s = blanchet@35961: AList.defined (op =) default_default_params s orelse blanchet@35961: AList.defined (op =) negated_params s orelse blanchet@35961: member (op =) property_affected_params s blanchet@35866: blanchet@35961: fun check_raw_param (s, _) = blanchet@35961: if is_known_raw_param s then () blanchet@35961: else error ("Unknown parameter: " ^ quote s ^ ".") blanchet@35961: blanchet@35961: fun unnegate_raw_param (name, value) = blanchet@35961: case AList.lookup (op =) negated_params name of blanchet@35961: SOME name' => (name', case value of blanchet@35961: ["false"] => ["true"] blanchet@35961: | ["true"] => ["false"] blanchet@35961: | [] => ["false"] blanchet@35961: | _ => value) blanchet@35961: | NONE => (name, value) blanchet@35961: blanchet@35961: structure Data = Theory_Data( blanchet@35961: type T = raw_param list blanchet@35961: val empty = default_default_params |> map (apsnd single) blanchet@35961: val extend = I blanchet@35961: fun merge p : T = AList.merge (op =) (K true) p) blanchet@35961: blanchet@35961: val set_default_raw_param = Data.map o AList.update (op =) o unnegate_raw_param blanchet@35961: fun default_raw_params thy = blanchet@35961: [("atps", [!atps]), blanchet@35961: ("full_types", [if !full_types then "true" else "false"]), blanchet@35961: ("timeout", [string_of_int (!timeout) ^ " s"])] @ blanchet@35961: Data.get thy blanchet@35961: blanchet@35961: val infinity_time_in_secs = 60 * 60 * 24 * 365 blanchet@35961: val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs) blanchet@35961: blanchet@35961: fun extract_params thy default_params override_params = blanchet@35961: let blanchet@35961: val override_params = map unnegate_raw_param override_params blanchet@35961: val raw_params = rev override_params @ rev default_params blanchet@35961: val lookup = Option.map (space_implode " ") o AList.lookup (op =) raw_params blanchet@35961: val lookup_string = the_default "" o lookup blanchet@35961: fun general_lookup_bool option default_value name = blanchet@35961: case lookup name of blanchet@35961: SOME s => s |> parse_bool_option option name blanchet@35961: | NONE => default_value blanchet@35961: val lookup_bool = the o general_lookup_bool false (SOME false) blanchet@35961: val lookup_bool_option = general_lookup_bool true NONE blanchet@35961: fun lookup_time name = blanchet@35961: the_timeout (case lookup name of blanchet@35961: NONE => NONE blanchet@35961: | SOME s => parse_time_option name s) blanchet@35961: fun lookup_real name = blanchet@35961: case lookup name of blanchet@35961: NONE => 0.0 blanchet@35961: | SOME s => 0.0 (* FIXME ### *) blanchet@35961: val debug = lookup_bool "debug" blanchet@35961: val verbose = debug orelse lookup_bool "verbose" blanchet@35961: val atps = lookup_string "atps" |> space_explode " " blanchet@35961: val full_types = lookup_bool "full_types" blanchet@35961: val relevance_threshold = lookup_real "relevance_threshold" blanchet@35961: val higher_order = lookup_bool_option "higher_order" blanchet@35961: val respect_no_atp = lookup_bool "respect_no_atp" blanchet@35961: val follow_defs = lookup_bool "follow_defs" blanchet@35961: val isar_proof = lookup_bool "isar_proof" blanchet@35961: val timeout = lookup_time "timeout" blanchet@35961: val minimize_timeout = lookup_time "minimize_timeout" blanchet@35961: in blanchet@35961: {debug = debug, verbose = verbose, atps = atps, full_types = full_types, blanchet@35961: relevance_threshold = relevance_threshold, higher_order = higher_order, blanchet@35961: respect_no_atp = respect_no_atp, follow_defs = follow_defs, blanchet@35961: isar_proof = isar_proof, timeout = timeout, blanchet@35961: minimize_timeout = minimize_timeout} blanchet@35961: end blanchet@35961: blanchet@35963: fun get_params thy = extract_params thy (default_raw_params thy) blanchet@35963: fun default_params thy = get_params thy o map (apsnd single) blanchet@35866: blanchet@35963: fun atp_minimize_command override_params old_style_args fact_refs state = blanchet@35866: let blanchet@35961: val thy = Proof.theory_of state blanchet@35961: val ctxt = Proof.context_of state blanchet@35963: fun theorems_from_refs ctxt = blanchet@35963: map (fn fact_ref => blanchet@35866: let blanchet@35963: val ths = ProofContext.get_fact ctxt fact_ref blanchet@35963: val name' = Facts.string_of_ref fact_ref blanchet@35866: in (name', ths) end) blanchet@35961: fun get_time_limit_arg s = blanchet@35961: (case Int.fromString s of blanchet@35961: SOME t => Time.fromSeconds t blanchet@35961: | NONE => error ("Invalid time limit: " ^ quote s)) blanchet@35866: fun get_opt (name, a) (p, t) = blanchet@35866: (case name of blanchet@35866: "time" => (p, get_time_limit_arg a) blanchet@35866: | "atp" => (a, t) blanchet@35866: | n => error ("Invalid argument: " ^ n)) blanchet@35963: val {atps, minimize_timeout, ...} = get_params thy override_params blanchet@35963: val (atp, timeout) = fold get_opt old_style_args (hd atps, minimize_timeout) blanchet@35961: val params = blanchet@35963: get_params thy blanchet@35963: [("atps", [atp]), blanchet@35963: ("minimize_timeout", blanchet@35963: [string_of_int (Time.toSeconds timeout) ^ " s"])] blanchet@35866: val prover = blanchet@35961: (case get_prover thy atp of blanchet@35866: SOME prover => prover blanchet@35961: | NONE => error ("Unknown ATP: " ^ quote atp)) blanchet@35963: val name_thms_pairs = theorems_from_refs ctxt fact_refs blanchet@35866: in blanchet@35961: writeln (#2 (minimize_theorems params linear_minimize prover atp state blanchet@35961: name_thms_pairs)) blanchet@35866: end blanchet@35866: blanchet@35963: val runN = "run" blanchet@35963: val minimizeN = "minimize" blanchet@35963: val messagesN = "messages" blanchet@35963: val available_atpsN = "available_atps" blanchet@35963: val running_atpsN = "running_atps" blanchet@35963: val kill_atpsN = "kill_atps" blanchet@35963: val refresh_tptpN = "refresh_tptp" blanchet@35963: val helpN = "help" blanchet@35866: blanchet@35963: val addN = "add" blanchet@35963: val delN = "del" blanchet@35963: val onlyN = "only" blanchet@35963: blanchet@35963: fun hammer_away override_params subcommand opt_i (Facta f) state = blanchet@35961: let blanchet@35961: val thy = Proof.theory_of state blanchet@35961: val _ = List.app check_raw_param override_params blanchet@35963: in blanchet@35963: if subcommand = runN then blanchet@35963: sledgehammer (get_params thy override_params) (the_default 1 opt_i) state blanchet@35963: else if subcommand = minimizeN then blanchet@35963: atp_minimize_command override_params [] (#add f) state blanchet@35963: else if subcommand = messagesN then blanchet@35963: messages opt_i blanchet@35963: else if subcommand = available_atpsN then blanchet@35963: available_atps thy blanchet@35963: else if subcommand = running_atpsN then blanchet@35963: running_atps () blanchet@35963: else if subcommand = kill_atpsN then blanchet@35963: kill_atps () blanchet@35963: else if subcommand = refresh_tptpN then blanchet@35963: refresh_systems_on_tptp () blanchet@35963: else blanchet@35963: error ("Unknown subcommand: " ^ quote subcommand ^ ".") blanchet@35963: end blanchet@35961: blanchet@35963: fun sledgehammer_trans (((params, subcommand), opt_i), facta) = blanchet@35963: Toplevel.keep (hammer_away params subcommand opt_i facta o Toplevel.proof_of) blanchet@35961: blanchet@35961: fun string_for_raw_param (name, value) = name ^ " = " ^ space_implode " " value blanchet@35961: blanchet@35963: fun sledgehammer_params_trans params = blanchet@35961: Toplevel.theory blanchet@35963: (fold set_default_raw_param params blanchet@35961: #> tap (fn thy => blanchet@35961: writeln ("Default parameters for Sledgehammer:\n" ^ blanchet@35961: (case rev (default_raw_params thy) of blanchet@35961: [] => "none" blanchet@35961: | params => blanchet@35961: (map check_raw_param params; blanchet@35961: params |> map string_for_raw_param blanchet@35961: |> sort_strings |> cat_lines))))) blanchet@35961: blanchet@35961: val parse_key = Scan.repeat1 P.typ_group >> space_implode " " blanchet@35961: val parse_value = blanchet@35961: Scan.repeat1 (P.minus >> single blanchet@35961: || Scan.repeat1 (Scan.unless P.minus P.name)) >> flat blanchet@35963: val parse_param = parse_key -- Scan.optional (P.$$$ "=" |-- parse_value) [] blanchet@35963: val parse_params = Scan.optional (Args.bracks (P.list parse_param)) [] blanchet@35963: val parse_fact_refs = blanchet@35963: Scan.repeat (P.xname -- Scan.option Attrib.thm_sel blanchet@35963: >> (fn (name, interval) => blanchet@35963: Facts.Named ((name, Position.none), interval))) blanchet@35963: val parse_slew = blanchet@35963: (Args.add |-- Args.colon |-- parse_fact_refs >> add_to_facta) blanchet@35963: || (Args.del |-- Args.colon |-- parse_fact_refs >> del_from_facta) blanchet@35963: || (P.$$$ onlyN |-- Args.colon |-- parse_fact_refs >> only_facta) blanchet@35963: val parse_facta = blanchet@35963: Args.parens (Scan.optional (parse_fact_refs >> only_facta) default_facta blanchet@35963: -- Scan.repeat parse_slew) >> op :: >> merge_facta blanchet@35961: val parse_sledgehammer_command = blanchet@35963: (parse_params -- Scan.optional P.name runN -- Scan.option P.nat blanchet@35963: -- parse_facta) blanchet@35963: #>> sledgehammer_trans blanchet@35961: val parse_sledgehammer_params_command = blanchet@35961: parse_params #>> sledgehammer_params_trans blanchet@35961: blanchet@35963: val parse_minimize_args = blanchet@35963: Scan.optional (Args.bracks (P.list (P.xname --| P.$$$ "=" -- P.xname))) [] blanchet@35963: blanchet@35961: val _ = blanchet@35961: OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag blanchet@35963: (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill_atps)) blanchet@35961: val _ = blanchet@35961: OuterSyntax.improper_command "atp_info" blanchet@35961: "print information about managed provers" K.diag blanchet@35963: (Scan.succeed (Toplevel.no_timing o Toplevel.imperative running_atps)) blanchet@35961: val _ = blanchet@35961: OuterSyntax.improper_command "atp_messages" blanchet@35961: "print recent messages issued by managed provers" K.diag blanchet@35961: (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >> blanchet@35961: (fn limit => Toplevel.no_timing blanchet@35961: o Toplevel.imperative (fn () => messages limit))) blanchet@35961: val _ = blanchet@35961: OuterSyntax.improper_command "print_atps" "print external provers" K.diag blanchet@35961: (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o blanchet@35963: Toplevel.keep (available_atps o Toplevel.theory_of))) blanchet@35963: val _ = blanchet@35963: OuterSyntax.improper_command "atp_minimize" blanchet@35963: "minimize theorem list with external prover" K.diag blanchet@35963: (parse_minimize_args -- parse_fact_refs >> (fn (args, fact_refs) => blanchet@35963: Toplevel.no_timing o Toplevel.unknown_proof o blanchet@35963: Toplevel.keep (atp_minimize_command [] args fact_refs blanchet@35963: o Toplevel.proof_of))) blanchet@35963: blanchet@35961: val _ = blanchet@35961: OuterSyntax.improper_command "sledgehammer" blanchet@35961: "search for first-order proof using automatic theorem provers" K.diag blanchet@35961: parse_sledgehammer_command blanchet@35961: val _ = blanchet@35961: OuterSyntax.command "sledgehammer_params" blanchet@35961: "set and display the default parameters for Sledgehammer" K.thy_decl blanchet@35961: parse_sledgehammer_params_command blanchet@35866: blanchet@35866: end;