| author | wenzelm | 
| Sat, 15 Oct 2011 17:00:17 +0200 | |
| changeset 45149 | 22ff7e226946 | 
| parent 43025 | 5a0dec7bc099 | 
| child 45666 | d83797ef0d2d | 
| 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 | 
| 40130 | 17 | val auto: bool Unsynchronized.ref | 
| 18 | val max_solutions: int Unsynchronized.ref | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 19 | val solve_direct: Proof.state -> bool * (string * Proof.state) | 
| 40130 | 20 | val setup: theory -> theory | 
| 30980 | 21 | end; | 
| 22 | ||
| 40130 | 23 | structure Solve_Direct: SOLVE_DIRECT = | 
| 30980 | 24 | struct | 
| 25 | ||
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 26 | datatype mode = Auto_Try | Try | Normal | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 27 | |
| 40129 | 28 | val solve_directN = "solve_direct"; | 
| 29 | ||
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 30 | val someN = "some"; | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 31 | val noneN = "none"; | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 32 | val unknownN = "none"; | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 33 | |
| 30980 | 34 | (* preferences *) | 
| 35 | ||
| 32740 | 36 | val auto = Unsynchronized.ref false; | 
| 39329 | 37 | val max_solutions = Unsynchronized.ref 5; | 
| 33889 
4328de748fb2
some rearangement of load order to keep preferences adjacent -- slightly fragile;
 wenzelm parents: 
33301diff
changeset | 38 | |
| 
4328de748fb2
some rearangement of load order to keep preferences adjacent -- slightly fragile;
 wenzelm parents: 
33301diff
changeset | 39 | val _ = | 
| 
4328de748fb2
some rearangement of load order to keep preferences adjacent -- slightly fragile;
 wenzelm parents: 
33301diff
changeset | 40 | ProofGeneralPgip.add_preference Preferences.category_tracing | 
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
39329diff
changeset | 41 | (Unsynchronized.setmp auto true (fn () => | 
| 30980 | 42 | Preferences.bool_pref auto | 
| 40130 | 43 | "auto-solve-direct" | 
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 44 |       ("Run " ^ quote solve_directN ^ " automatically.")) ());
 | 
| 30980 | 45 | |
| 40129 | 46 | |
| 47 | (* solve_direct command *) | |
| 48 | ||
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 49 | fun do_solve_direct mode state = | 
| 30980 | 50 | let | 
| 51 | val ctxt = Proof.context_of state; | |
| 52 | ||
| 33301 | 53 | val crits = [(true, Find_Theorems.Solves)]; | 
| 40129 | 54 | fun find g = | 
| 55 | snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! max_solutions)) false crits); | |
| 30980 | 56 | |
| 57 | fun prt_result (goal, results) = | |
| 58 | let | |
| 59 | val msg = | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 60 | (if mode = Auto_Try then "Auto " else "") ^ solve_directN ^ ": " ^ | 
| 30980 | 61 | (case goal of | 
| 62 | NONE => "The current goal" | |
| 63 | | SOME sg => Syntax.string_of_term ctxt (Thm.term_of sg)); | |
| 64 | in | |
| 65 | Pretty.big_list | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 66 | (msg ^ " can be solved directly with") | 
| 33301 | 67 | (map (Find_Theorems.pretty_thm ctxt) results) | 
| 30980 | 68 | end; | 
| 69 | ||
| 70 | fun seek_against_goal () = | |
| 33290 | 71 | (case try Proof.simple_goal state of | 
| 30980 | 72 | NONE => NONE | 
| 33290 | 73 |       | SOME {goal = st, ...} =>
 | 
| 30980 | 74 | 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 | 75 | val subgoals = Drule.strip_imp_prems (Thm.cprop_of st); | 
| 30980 | 76 | val rs = | 
| 77 | 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 | 78 | then [(NONE, find st)] | 
| 30980 | 79 | else map (fn sg => (SOME sg, find (Goal.init sg))) subgoals; | 
| 80 | val results = filter_out (null o snd) rs; | |
| 81 | in if null results then NONE else SOME results end); | |
| 40129 | 82 | 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 | 83 | in | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 84 | (case try seek_against_goal () of | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 85 | SOME (SOME results) => | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 86 | (someN, | 
| 40129 | 87 | state |> | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 88 | (if mode = Auto_Try then | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 89 | Proof.goal_message | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 90 | (fn () => | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 91 | Pretty.chunks | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 92 | [Pretty.str "", Pretty.markup Markup.hilite (message results)]) | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 93 | else | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 94 | tap (fn _ => | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 95 | Output.urgent_message (Pretty.string_of (Pretty.chunks (message results)))))) | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 96 | | SOME NONE => | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 97 | (if mode = Normal then Output.urgent_message "No proof found." | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 98 | else (); | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 99 | (noneN, state)) | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 100 | | NONE => | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 101 | (if mode = Normal then Output.urgent_message "An error occurred." | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 102 | else (); | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 103 | (unknownN, state))) | 
| 43020 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 104 | end | 
| 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 blanchet parents: 
43018diff
changeset | 105 | |> `(fn (outcome_code, _) => outcome_code = someN); | 
| 30980 | 106 | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 107 | val solve_direct = do_solve_direct Normal | 
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 108 | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 109 | val _ = | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 110 | Outer_Syntax.improper_command solve_directN | 
| 40129 | 111 | "try to solve conjectures directly with existing theorems" Keyword.diag | 
| 40130 | 112 | (Scan.succeed | 
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 113 | (Toplevel.keep (ignore o solve_direct o Toplevel.proof_of))); | 
| 40129 | 114 | |
| 40116 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 115 | |
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 116 | (* hook *) | 
| 
9ed3711366c8
introduced manual version of "Auto Solve" as "solve_direct"
 blanchet parents: 
39616diff
changeset | 117 | |
| 43025 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 118 | 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 | 119 | |
| 
5a0dec7bc099
handle non-auto try case gracefully in Solve Direct
 blanchet parents: 
43024diff
changeset | 120 | val setup = Try.register_tool (solve_directN, (10, auto, try_solve_direct)); | 
| 30980 | 121 | |
| 122 | end; |