| author | wenzelm | 
| Mon, 18 Nov 2024 11:12:51 +0100 | |
| changeset 81479 | d9e8f594487e | 
| parent 74508 | 3315c551fe6e | 
| permissions | -rw-r--r-- | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 1 | (* Title: Tools/solve_direct.ML | 
| 30980 | 2 | Author: Timothy Bourke and Gerwin Klein, NICTA | 
| 3 | ||
| 4 | Check whether a newly stated theorem can be solved directly by an | |
| 5 | existing theorem. Duplicate lemmas can be detected in this way. | |
| 6 | ||
| 39329 | 7 | The implementation relies critically on the Find_Theorems solves | 
| 30980 | 8 | feature. | 
| 9 | *) | |
| 10 | ||
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 11 | signature SOLVE_DIRECT = | 
| 30980 | 12 | sig | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 13 | val solve_directN: string | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 14 | val someN: string | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 15 | val noneN: string | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 16 | val unknownN: string | 
| 64012 | 17 | val max_solutions: int Config.T | 
| 64013 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 18 | val strict_warnings: bool Config.T | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58843diff
changeset | 19 | val solve_direct: Proof.state -> bool * (string * string list) | 
| 30980 | 20 | end; | 
| 21 | ||
| 40130 | 22 | structure Solve_Direct: SOLVE_DIRECT = | 
| 30980 | 23 | struct | 
| 24 | ||
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 25 | datatype mode = Auto_Try | Try | Normal | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 26 | |
| 40129 | 27 | val solve_directN = "solve_direct"; | 
| 28 | ||
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 29 | val someN = "some"; | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 30 | val noneN = "none"; | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 31 | val unknownN = "none"; | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 32 | |
| 52701 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 wenzelm parents: 
52645diff
changeset | 33 | |
| 30980 | 34 | (* preferences *) | 
| 35 | ||
| 67149 | 36 | val max_solutions = Attrib.setup_config_int \<^binding>\<open>solve_direct_max_solutions\<close> (K 5); | 
| 37 | val strict_warnings = Attrib.setup_config_bool \<^binding>\<open>solve_direct_strict_warnings\<close> (K false); | |
| 33889 
4328de748fb2
some rearangement of load order to keep preferences adjacent -- slightly fragile;
 wenzelm parents: 
33301diff
changeset | 38 | |
| 40129 | 39 | |
| 40 | (* solve_direct command *) | |
| 41 | ||
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 42 | fun do_solve_direct mode state = | 
| 30980 | 43 | let | 
| 44 | val ctxt = Proof.context_of state; | |
| 45 | ||
| 64982 | 46 | fun find goal = | 
| 64983 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 wenzelm parents: 
64982diff
changeset | 47 | #2 (Find_Theorems.find_theorems ctxt (SOME goal) | 
| 
481b2855ee9a
suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
 wenzelm parents: 
64982diff
changeset | 48 | (SOME (Config.get ctxt max_solutions)) false [(true, Find_Theorems.Solves)]); | 
| 30980 | 49 | |
| 64985 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 50 | fun prt_result (subgoal, results) = | 
| 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 51 | Pretty.big_list | 
| 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 52 | ((if mode = Auto_Try then "Auto " else "") ^ solve_directN ^ ": " ^ | 
| 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 53 | (case subgoal of NONE => "the current goal" | SOME i => "subgoal #" ^ string_of_int i) ^ | 
| 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 54 | " can be solved directly with") | 
| 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 55 | (map (Find_Theorems.pretty_thm ctxt) results); | 
| 30980 | 56 | |
| 57 | fun seek_against_goal () = | |
| 64984 | 58 | (case try (#goal o Proof.simple_goal) state of | 
| 30980 | 59 | NONE => NONE | 
| 64984 | 60 | | SOME goal => | 
| 30980 | 61 | let | 
| 64982 | 62 | val subgoals = Drule.strip_imp_prems (Thm.cprop_of goal); | 
| 30980 | 63 | val rs = | 
| 64 | if length subgoals = 1 | |
| 64982 | 65 | then [(NONE, find goal)] | 
| 64985 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 66 | else map_index (fn (i, sg) => (SOME (i + 1), find (Goal.init sg))) subgoals; | 
| 64982 | 67 | val results = filter_out (null o #2) rs; | 
| 30980 | 68 | in if null results then NONE else SOME results end); | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 69 | in | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 70 | (case try seek_against_goal () of | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 71 | SOME (SOME results) => | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 72 | (someN, | 
| 64013 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 73 | let | 
| 64985 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 74 | val chunks = separate (Pretty.str "") (map prt_result results); | 
| 
69592ac04836
clarified message: print subgoal number instead of potentially large term;
 wenzelm parents: 
64984diff
changeset | 75 | val msg = Pretty.string_of (Pretty.chunks chunks); | 
| 64013 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 76 | in | 
| 64982 | 77 | if Config.get ctxt strict_warnings then (warning msg; []) | 
| 78 | else if mode = Auto_Try then [msg] | |
| 79 | else (writeln msg; []) | |
| 64013 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 80 | end) | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 81 | | SOME NONE => | 
| 64982 | 82 | (if mode = Normal then writeln "No proof found" else (); | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58843diff
changeset | 83 | (noneN, [])) | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 84 | | NONE => | 
| 64982 | 85 | (if mode = Normal then writeln "An error occurred" else (); | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58843diff
changeset | 86 | (unknownN, []))) | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 87 | end | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 88 | |> `(fn (outcome_code, _) => outcome_code = someN); | 
| 30980 | 89 | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 90 | val solve_direct = do_solve_direct Normal | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 91 | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 92 | val _ = | 
| 67149 | 93 | Outer_Syntax.command \<^command_keyword>\<open>solve_direct\<close> | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45666diff
changeset | 94 | "try to solve conjectures directly with existing theorems" | 
| 60190 
906de96ba68a
allow diagnostic proof commands with skip_proofs;
 wenzelm parents: 
60094diff
changeset | 95 | (Scan.succeed (Toplevel.keep_proof (ignore o solve_direct o Toplevel.proof_of))); | 
| 40129 | 96 | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 97 | |
| 74508 | 98 | (* 'try' setup *) | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 99 | |
| 64982 | 100 | val _ = | 
| 74508 | 101 | Try.tool_setup | 
| 102 |    {name = solve_directN, weight = 10, auto_option = \<^system_option>\<open>auto_solve_direct\<close>,
 | |
| 103 | body = fn auto => do_solve_direct (if auto then Auto_Try else Try)}; | |
| 30980 | 104 | |
| 105 | end; |