| author | wenzelm | 
| Sun, 27 Sep 2009 21:06:43 +0200 | |
| changeset 32710 | fa46afc8c05f | 
| parent 32571 | d4bb776874b8 | 
| child 32740 | 9dd0a2f83429 | 
| permissions | -rw-r--r-- | 
| 32327 
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
 wenzelm parents: 
32091diff
changeset | 1 | (* Title: HOL/Tools/ATP_Manager/atp_minimal.ML | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 2 | Author: Philipp Meyer, TU Muenchen | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 3 | |
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 4 | Minimalization of theorem list for metis by using an external automated theorem prover | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 5 | *) | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 6 | |
| 32525 | 7 | signature ATP_MINIMAL = | 
| 8 | sig | |
| 9 | val minimalize: AtpManager.prover -> string -> int -> Proof.state -> | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 10 | (string * thm list) list -> ((string * thm list) list * int) option * string | 
| 32525 | 11 | end | 
| 12 | ||
| 13 | structure AtpMinimal: ATP_MINIMAL = | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 14 | struct | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 15 | |
| 31236 | 16 | (* output control *) | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 17 | |
| 31236 | 18 | fun debug str = Output.debug (fn () => str) | 
| 19 | fun debug_fn f = if ! Output.debugging then f () else () | |
| 20 | fun answer str = Output.writeln str | |
| 21 | fun println str = Output.priority str | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 22 | |
| 31236 | 23 | fun order_unique name_list = OrdList.make (String.collate Char.compare) name_list | 
| 24 | fun length_string namelist = Int.toString (length namelist) | |
| 25 | ||
| 26 | fun print_names name_thms_pairs = | |
| 27 | let | |
| 28 | val names = map fst name_thms_pairs | |
| 29 | val ordered = order_unique names | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 30 | in | 
| 31236 | 31 |     app (fn name => (debug ("  " ^ name))) ordered
 | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 32 | end | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 33 | |
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 34 | |
| 31236 | 35 | (* minimalization algorithm *) | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 36 | |
| 31236 | 37 | local | 
| 38 | fun isplit (l,r) [] = (l,r) | |
| 39 | | isplit (l,r) (h::[]) = (h::l, r) | |
| 40 | | isplit (l,r) (h1::h2::t) = isplit (h1::l, h2::r) t | |
| 41 | in | |
| 42 | fun split lst = isplit ([],[]) lst | |
| 43 | end | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 44 | |
| 31236 | 45 | local | 
| 46 | fun min p sup [] = raise Empty | |
| 47 | | min p sup [s0] = [s0] | |
| 48 | | min p sup s = | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 49 | let | 
| 31236 | 50 | val (l0, r0) = split s | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 51 | in | 
| 31236 | 52 | if p (sup @ l0) | 
| 53 | then min p sup l0 | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 54 | else | 
| 31236 | 55 | if p (sup @ r0) | 
| 56 | then min p sup r0 | |
| 57 | else | |
| 58 | let | |
| 59 | val l = min p (sup @ r0) l0 | |
| 60 | val r = min p (sup @ l) r0 | |
| 61 | in | |
| 62 | l @ r | |
| 63 | end | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 64 | end | 
| 31236 | 65 | in | 
| 66 | (* return a minimal subset v of s that satisfies p | |
| 67 | @pre p(s) & ~p([]) & monotone(p) | |
| 68 | @post v subset s & p(v) & | |
| 69 | forall e in v. ~p(v \ e) | |
| 70 | *) | |
| 32538 
86035c5f61b5
Fixed "minimal" to cover the case that "p []" holds (excluded in the article by Bradley & Manna)
 nipkow parents: 
32525diff
changeset | 71 | fun minimal p s = | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 72 | let val c = ref 0 | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 73 | fun pc xs = (c := !c + 1; p xs) | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 74 | in | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 75 | (case min pc [] s of | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 76 | [x] => if pc [] then [] else [x] | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 77 | | m => m, | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 78 | !c) | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 79 | end | 
| 31236 | 80 | end | 
| 81 | ||
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 82 | |
| 31236 | 83 | (* failure check and producing answer *) | 
| 84 | ||
| 85 | datatype 'a prove_result = Success of 'a | Failure | Timeout | Error | |
| 86 | ||
| 87 | val string_of_result = | |
| 88 | fn Success _ => "Success" | |
| 89 | | Failure => "Failure" | |
| 90 | | Timeout => "Timeout" | |
| 91 | | Error => "Error" | |
| 92 | ||
| 93 | val failure_strings = | |
| 94 |   [("SPASS beiseite: Ran out of time.", Timeout),
 | |
| 95 |    ("Timeout", Timeout),
 | |
| 96 |    ("time limit exceeded", Timeout),
 | |
| 97 |    ("# Cannot determine problem status within resource limit", Timeout),
 | |
| 98 |    ("Error", Error)]
 | |
| 99 | ||
| 32510 | 100 | fun produce_answer (success, message, _, result_string, thm_name_vec, filtered) = | 
| 31236 | 101 | if success then | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 102 | (Success (Vector.foldr op:: [] thm_name_vec, filtered), result_string) | 
| 31236 | 103 | else | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 104 | let | 
| 31236 | 105 | val failure = failure_strings |> get_first (fn (s, t) => | 
| 106 | if String.isSubstring s result_string then SOME (t, result_string) else NONE) | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 107 | in | 
| 31236 | 108 | if is_some failure then | 
| 109 | the failure | |
| 110 | else | |
| 111 | (Failure, result_string) | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 112 | end | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 113 | |
| 31236 | 114 | |
| 115 | (* wrapper for calling external prover *) | |
| 116 | ||
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 117 | fun sh_test_thms prover prover_name time_limit subgoalno state filtered name_thms_pairs = | 
| 31236 | 118 | let | 
| 119 |     val _ = println ("Testing " ^ (length_string name_thms_pairs) ^ " theorems... ")
 | |
| 32525 | 120 | val name_thm_pairs = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs | 
| 31236 | 121 | val _ = debug_fn (fn () => print_names name_thm_pairs) | 
| 122 | val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs | |
| 123 | val (result, proof) = | |
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
31236diff
changeset | 124 | produce_answer | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 125 | (prover time_limit (SOME axclauses) filtered prover_name subgoalno (Proof.get_goal state)) | 
| 31236 | 126 | val _ = println (string_of_result result) | 
| 127 | val _ = debug proof | |
| 128 | in | |
| 129 | (result, proof) | |
| 130 | end | |
| 131 | ||
| 132 | ||
| 133 | (* minimalization of thms *) | |
| 134 | ||
| 135 | fun minimalize prover prover_name time_limit state name_thms_pairs = | |
| 136 | let | |
| 137 | val _ = | |
| 138 |       println ("Minimize called with " ^ length_string name_thms_pairs ^ " theorems, prover: "
 | |
| 139 | ^ prover_name ^ ", time limit: " ^ Int.toString time_limit ^ " seconds") | |
| 140 | val _ = debug_fn (fn () => app (fn (n, tl) => | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31752diff
changeset | 141 | (debug n; app (fn t => | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31752diff
changeset | 142 |           debug ("  " ^ Display.string_of_thm (Proof.context_of state) t)) tl)) name_thms_pairs)
 | 
| 31236 | 143 | val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 144 | fun test_thms filtered thms = | 
| 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 145 | case test_thms_fun filtered thms of (Success _, _) => true | _ => false | 
| 32525 | 146 | val answer' = pair and answer'' = pair NONE | 
| 31236 | 147 | in | 
| 148 | (* try prove first to check result and get used theorems *) | |
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
31236diff
changeset | 149 | (case test_thms_fun NONE name_thms_pairs of | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 150 | (Success (used, filtered), _) => | 
| 31236 | 151 | let | 
| 152 | val ordered_used = order_unique used | |
| 153 | val to_use = | |
| 154 | if length ordered_used < length name_thms_pairs then | |
| 155 | filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs | |
| 156 | else | |
| 157 | name_thms_pairs | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 158 | val (min_thms, n) = if null to_use then ([], 0) | 
| 32525 | 159 | else minimal (test_thms (SOME filtered)) to_use | 
| 31236 | 160 | val min_names = order_unique (map fst min_thms) | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 161 |           val _ = println ("Interations: " ^ string_of_int n)
 | 
| 31236 | 162 |           val _ = println ("Minimal " ^ (length_string min_thms) ^ " theorems")
 | 
| 163 | val _ = debug_fn (fn () => print_names min_thms) | |
| 164 | in | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 165 |           answer' (SOME(min_thms,n)) ("Try this command: " ^
 | 
| 31236 | 166 |             Markup.markup Markup.sendback ("apply (metis " ^ space_implode " " min_names ^ ")"))
 | 
| 167 | end | |
| 168 | | (Timeout, _) => | |
| 32525 | 169 |         answer'' ("Timeout: You may need to increase the time limit of " ^
 | 
| 31236 | 170 | Int.toString time_limit ^ " seconds. Call atp_minimize [time=...] ") | 
| 171 | | (Error, msg) => | |
| 32525 | 172 |         answer'' ("Error in prover: " ^ msg)
 | 
| 31236 | 173 | | (Failure, _) => | 
| 32525 | 174 | answer'' "Failure: No proof with the theorems supplied") | 
| 31236 | 175 | handle ResHolClause.TOO_TRIVIAL => | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32538diff
changeset | 176 |         answer' (SOME ([],0)) ("Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
 | 
| 31236 | 177 | | ERROR msg => | 
| 32525 | 178 |         answer'' ("Error: " ^ msg)
 | 
| 31236 | 179 | end | 
| 180 | ||
| 181 | ||
| 182 | (* Isar command and parsing input *) | |
| 183 | ||
| 184 | local structure K = OuterKeyword and P = OuterParse and T = OuterLex in | |
| 185 | ||
| 186 | fun get_thms context = | |
| 187 | map (fn (name, interval) => | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 188 | let | 
| 31236 | 189 | val thmref = Facts.Named ((name, Position.none), interval) | 
| 190 | val ths = ProofContext.get_fact context thmref | |
| 191 | val name' = Facts.string_of_ref thmref | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 192 | in | 
| 31236 | 193 | (name', ths) | 
| 194 | end) | |
| 195 | ||
| 196 | val default_prover = "remote_vampire" | |
| 197 | val default_time_limit = 5 | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 198 | |
| 31236 | 199 | fun get_time_limit_arg time_string = | 
| 200 | (case Int.fromString time_string of | |
| 201 | SOME t => t | |
| 202 |   | NONE => error ("Invalid time limit: " ^ quote time_string))
 | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 203 | |
| 31236 | 204 | val get_options = | 
| 205 | let | |
| 206 | val def = (default_prover, default_time_limit) | |
| 207 | in | |
| 208 | foldl (fn ((name, a), (p, t)) => | |
| 209 | (case name of | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 210 | "time" => (p, (get_time_limit_arg a)) | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 211 | | "atp" => (a, t) | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 212 |       | n => error ("Invalid argument: " ^ n))) def
 | 
| 31236 | 213 | end | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 214 | |
| 31236 | 215 | fun sh_min_command args thm_names state = | 
| 216 | let | |
| 217 | val (prover_name, time_limit) = get_options args | |
| 218 | val prover = | |
| 219 | case AtpManager.get_prover prover_name (Proof.theory_of state) of | |
| 220 | SOME prover => prover | |
| 221 |       | NONE => error ("Unknown prover: " ^ quote prover_name)
 | |
| 222 | val name_thms_pairs = get_thms (Proof.context_of state) thm_names | |
| 32525 | 223 | fun print_answer (_, msg) = answer msg | 
| 31236 | 224 | in | 
| 32525 | 225 | print_answer (minimalize prover prover_name time_limit state name_thms_pairs) | 
| 31236 | 226 | end | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 227 | |
| 31236 | 228 | val parse_args = Scan.optional (Args.bracks (P.list (P.xname --| P.$$$ "=" -- P.xname))) [] | 
| 229 | val parse_thm_names = Scan.repeat (P.xname -- Scan.option Attrib.thm_sel) | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 230 | |
| 31236 | 231 | val _ = | 
| 232 | OuterSyntax.command "atp_minimize" "minimize theorem list with external prover" K.diag | |
| 233 | (parse_args -- parse_thm_names >> (fn (args, thm_names) => | |
| 234 | Toplevel.no_timing o Toplevel.unknown_proof o | |
| 235 | Toplevel.keep (sh_min_command args thm_names o Toplevel.proof_of))) | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 236 | |
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 237 | end | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 238 | |
| 31236 | 239 | end | 
| 240 |