| author | wenzelm | 
| Mon, 09 Mar 2009 22:43:25 +0100 | |
| changeset 30399 | a4772a650b4e | 
| parent 30242 | aea5d7fa7ef5 | 
| child 30785 | 15f64e05e703 | 
| permissions | -rw-r--r-- | 
| 30147 | 1 | (* Title: Pure/Tools/auto_solve.ML | 
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 2 | Author: Timothy Bourke and Gerwin Klein, NICTA | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 3 | |
| 30147 | 4 | Check whether a newly stated theorem can be solved directly by an | 
| 5 | existing theorem. Duplicate lemmas can be detected in this way. | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 6 | |
| 30147 | 7 | The implemenation is based in part on Berghofer and Haftmann's | 
| 8 | Pure/codegen.ML. It relies critically on the FindTheorems solves | |
| 9 | feature. | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 10 | *) | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 11 | |
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 12 | signature AUTO_SOLVE = | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 13 | sig | 
| 30147 | 14 | val auto : bool ref | 
| 15 | val auto_time_limit : int ref | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 16 | |
| 30147 | 17 | val seek_solution : bool -> Proof.state -> Proof.state | 
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 18 | end; | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 19 | |
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 20 | structure AutoSolve : AUTO_SOLVE = | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 21 | struct | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 22 | |
| 30147 | 23 | val auto = ref false; | 
| 24 | val auto_time_limit = ref 2500; | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 25 | |
| 30147 | 26 | fun seek_solution int state = | 
| 27 | let | |
| 28 | val ctxt = Proof.context_of state; | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 29 | |
| 30147 | 30 | fun conj_to_list [] = [] | 
| 31 | | conj_to_list (t::ts) = | |
| 32 | (Conjunction.dest_conjunction t | |
| 33 | |> (fn (t1, t2) => conj_to_list (t1::t2::ts))) | |
| 34 | handle TERM _ => t::conj_to_list ts; | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 35 | |
| 30147 | 36 | val crits = [(true, FindTheorems.Solves)]; | 
| 37 | fun find g = (NONE, FindTheorems.find_theorems ctxt g true crits); | |
| 38 | fun find_cterm g = (SOME g, FindTheorems.find_theorems ctxt | |
| 39 | (SOME (Goal.init g)) true crits); | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 40 | |
| 30147 | 41 | fun prt_result (goal, results) = | 
| 42 | let | |
| 43 | val msg = case goal of | |
| 44 | NONE => "The current goal" | |
| 45 | | SOME g => Syntax.string_of_term ctxt (term_of g); | |
| 46 | in | |
| 47 | Pretty.big_list (msg ^ " could be solved directly with:") | |
| 30186 
1f836e949ac2
replaced archaic Display.pretty_fact by FindTheorems.pretty_thm, which observes the context properly (as did the former prt_fact already);
 wenzelm parents: 
30147diff
changeset | 48 | (map (FindTheorems.pretty_thm ctxt) results) | 
| 30147 | 49 | end; | 
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 50 | |
| 30147 | 51 | fun seek_against_goal () = | 
| 52 | let | |
| 53 | val goal = try Proof.get_goal state | |
| 54 | |> Option.map (#2 o #2); | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 55 | |
| 30147 | 56 | val goals = goal | 
| 57 | |> Option.map (fn g => cprem_of g 1) | |
| 58 | |> the_list | |
| 59 | |> conj_to_list; | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 60 | |
| 30147 | 61 | val rs = if length goals = 1 | 
| 62 | then [find goal] | |
| 63 | else map find_cterm goals; | |
| 64 | val frs = filter_out (null o snd) rs; | |
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 65 | |
| 30147 | 66 | in if null frs then NONE else SOME frs end; | 
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 67 | |
| 30147 | 68 | fun go () = | 
| 69 | let | |
| 70 | val res = TimeLimit.timeLimit | |
| 71 | (Time.fromMilliseconds (! auto_time_limit)) | |
| 72 | (try seek_against_goal) (); | |
| 73 | in | |
| 74 | case Option.join res of | |
| 75 | NONE => state | |
| 76 | | SOME results => (Proof.goal_message | |
| 77 | (fn () => Pretty.chunks [Pretty.str "", | |
| 78 | Pretty.markup Markup.hilite | |
| 79 | (Library.separate (Pretty.brk 0) | |
| 80 | (map prt_result results))]) | |
| 81 | state) | |
| 82 | end handle TimeLimit.TimeOut => (warning "AutoSolve: timeout."; state); | |
| 83 | in | |
| 84 | if int andalso ! auto andalso not (! Toplevel.quiet) | |
| 85 | then go () | |
| 86 | else state | |
| 87 | end; | |
| 88 | ||
| 29858 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 89 | end; | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 90 | |
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 91 | val _ = Context.>> (Specification.add_theorem_hook AutoSolve.seek_solution); | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 92 | |
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 93 | val auto_solve = AutoSolve.auto; | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 94 | val auto_solve_time_limit = AutoSolve.auto_time_limit; | 
| 
c8cee17d7e50
Autosolve feature for detecting duplicate theorems; patch by Timothy Bourke
 kleing parents: diff
changeset | 95 |