| author | wenzelm | 
| Wed, 14 Dec 2016 10:40:25 +0100 | |
| changeset 64563 | 20e361014f55 | 
| parent 64013 | 048b7dbfdfa3 | 
| child 64982 | c515464b4652 | 
| 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 | ||
| 64012 | 36 | val max_solutions = Attrib.setup_config_int @{binding solve_direct_max_solutions} (K 5);
 | 
| 64013 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 37 | val strict_warnings = Attrib.setup_config_bool @{binding solve_direct_strict_warnings} (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; | |
| 52701 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 wenzelm parents: 
52645diff
changeset | 45 | val find_ctxt = if mode = Auto_Try then Context_Position.set_visible false ctxt else ctxt; | 
| 30980 | 46 | |
| 33301 | 47 | val crits = [(true, Find_Theorems.Solves)]; | 
| 40129 | 48 | fun find g = | 
| 64012 | 49 | snd (Find_Theorems.find_theorems find_ctxt (SOME g) | 
| 50 | (SOME (Config.get find_ctxt max_solutions)) false crits); | |
| 30980 | 51 | |
| 52 | fun prt_result (goal, results) = | |
| 53 | let | |
| 54 | val msg = | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 55 | (if mode = Auto_Try then "Auto " else "") ^ solve_directN ^ ": " ^ | 
| 30980 | 56 | (case goal of | 
| 57 | NONE => "The current goal" | |
| 58 | | SOME sg => Syntax.string_of_term ctxt (Thm.term_of sg)); | |
| 59 | in | |
| 60 | Pretty.big_list | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 61 | (msg ^ " can be solved directly with") | 
| 33301 | 62 | (map (Find_Theorems.pretty_thm ctxt) results) | 
| 30980 | 63 | end; | 
| 64 | ||
| 65 | fun seek_against_goal () = | |
| 33290 | 66 | (case try Proof.simple_goal state of | 
| 30980 | 67 | NONE => NONE | 
| 33290 | 68 |       | SOME {goal = st, ...} =>
 | 
| 30980 | 69 | let | 
| 32860 
a4ab5d0cccd1
replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
 wenzelm parents: 
32740diff
changeset | 70 | val subgoals = Drule.strip_imp_prems (Thm.cprop_of st); | 
| 30980 | 71 | val rs = | 
| 72 | if length subgoals = 1 | |
| 32860 
a4ab5d0cccd1
replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
 wenzelm parents: 
32740diff
changeset | 73 | then [(NONE, find st)] | 
| 30980 | 74 | else map (fn sg => (SOME sg, find (Goal.init sg))) subgoals; | 
| 75 | val results = filter_out (null o snd) rs; | |
| 76 | in if null results then NONE else SOME results end); | |
| 40129 | 77 | fun message results = separate (Pretty.brk 0) (map prt_result results); | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 78 | in | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 79 | (case try seek_against_goal () of | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 80 | SOME (SOME results) => | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 81 | (someN, | 
| 64013 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 82 | let | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 83 | val msg = Pretty.string_of (Pretty.chunks (message results)) | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 84 | in | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 85 | if Config.get ctxt strict_warnings | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 86 | then (warning msg; []) | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 87 | else if mode = Auto_Try | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 88 | then [msg] | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 89 | else (writeln msg; []) | 
| 
048b7dbfdfa3
option to report results of solve_direct as explicit warnings
 haftmann parents: 
64012diff
changeset | 90 | end) | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 91 | | SOME NONE => | 
| 63728 | 92 | (if mode = Normal then writeln "No proof found" | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 93 | else (); | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58843diff
changeset | 94 | (noneN, [])) | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 95 | | NONE => | 
| 63728 | 96 | (if mode = Normal then writeln "An error occurred" | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 97 | else (); | 
| 58892 
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
 wenzelm parents: 
58843diff
changeset | 98 | (unknownN, []))) | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 99 | end | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 100 | |> `(fn (outcome_code, _) => outcome_code = someN); | 
| 30980 | 101 | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 102 | val solve_direct = do_solve_direct Normal | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 103 | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 104 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59184diff
changeset | 105 |   Outer_Syntax.command @{command_keyword solve_direct}
 | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
45666diff
changeset | 106 | "try to solve conjectures directly with existing theorems" | 
| 60190 
906de96ba68a
allow diagnostic proof commands with skip_proofs;
 wenzelm parents: 
60094diff
changeset | 107 | (Scan.succeed (Toplevel.keep_proof (ignore o solve_direct o Toplevel.proof_of))); | 
| 40129 | 108 | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 109 | |
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 110 | (* hook *) | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 111 | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 112 | fun try_solve_direct auto = do_solve_direct (if auto then Auto_Try else Try) | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 113 | |
| 56467 | 114 | val _ = Try.tool_setup (solve_directN, (10, @{system_option auto_solve_direct}, try_solve_direct));
 | 
| 30980 | 115 | |
| 116 | end; |