4 Goals in tactical theorem proving, with support for forked proofs.
9 val skip_proofs: bool Unsynchronized.ref
10 val parallel_proofs: int Unsynchronized.ref
11 val parallel_subproofs_saturation: int Unsynchronized.ref
12 val parallel_subproofs_threshold: real Unsynchronized.ref
13 val SELECT_GOAL: tactic -> int -> tactic
14 val CONJUNCTS: tactic -> int -> tactic
15 val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic
16 val PARALLEL_CHOICE: tactic list -> tactic
17 val PARALLEL_GOALS: tactic -> tactic
23 val init: cterm -> thm
24 val protect: thm -> thm
25 val conclude: thm -> thm
26 val check_finished: Proof.context -> thm -> thm
27 val finish: Proof.context -> thm -> thm
28 val norm_result: thm -> thm
29 val fork_params: {name: string, pos: Position.T, pri: int} -> (unit -> 'a) -> 'a future
30 val fork: int -> (unit -> 'a) -> 'a future
31 val peek_futures: serial -> unit future list
32 val reset_futures: unit -> Future.group list
33 val shutdown_futures: unit -> unit
34 val future_enabled_level: int -> bool
35 val future_enabled: unit -> bool
36 val future_enabled_nested: int -> bool
37 val future_enabled_timing: Time.time -> bool
38 val future_result: Proof.context -> thm future -> term -> thm
39 val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
40 val is_schematic: term -> bool
41 val prove_multi: Proof.context -> string list -> term list -> term list ->
42 ({prems: thm list, context: Proof.context} -> tactic) -> thm list
43 val prove_future: Proof.context -> string list -> term list -> term ->
44 ({prems: thm list, context: Proof.context} -> tactic) -> thm
45 val prove: Proof.context -> string list -> term list -> term ->
46 ({prems: thm list, context: Proof.context} -> tactic) -> thm
47 val prove_global_future: theory -> string list -> term list -> term ->
48 ({prems: thm list, context: Proof.context} -> tactic) -> thm
49 val prove_global: theory -> string list -> term list -> term ->
50 ({prems: thm list, context: Proof.context} -> tactic) -> thm
51 val prove_sorry: Proof.context -> string list -> term list -> term ->
52 ({prems: thm list, context: Proof.context} -> tactic) -> thm
53 val prove_sorry_global: theory -> string list -> term list -> term ->
54 ({prems: thm list, context: Proof.context} -> tactic) -> thm
55 val extract: int -> int -> thm -> thm Seq.seq
56 val retrofit: int -> int -> thm -> thm -> thm Seq.seq
57 val conjunction_tac: int -> tactic
58 val precise_conjunction_tac: int -> int -> tactic
59 val recover_conjunction_tac: tactic
60 val norm_hhf_tac: int -> tactic
61 val compose_hhf_tac: thm -> int -> tactic
62 val assume_rule_tac: Proof.context -> int -> tactic
65 structure Goal: GOAL =
75 let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI))
76 in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;
83 fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI;
87 ---------------- (conclude)
90 fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD;
97 fun check_finished ctxt th =
98 if Thm.no_prems th then th
100 raise THM ("Proof failed.\n" ^
101 Pretty.string_of (Goal_Display.pretty_goal {main = true, limit = false} ctxt th), 0, [th]);
103 fun finish ctxt = check_finished ctxt #> conclude;
112 Drule.flexflex_unique
113 #> Raw_Simplifier.norm_hhf_protect
115 #> Drule.zero_var_indexes;
123 Synchronized.var "forked_proofs"
124 (0, []: Future.group list, Inttab.empty: unit future list Inttab.table);
127 Synchronized.change forked_proofs (fn (m, groups, tab) =>
130 val _ = Future.forked_proofs := n;
131 in (n, groups, tab) end);
133 fun register_forked id future =
134 Synchronized.change forked_proofs (fn (m, groups, tab) =>
136 val groups' = Task_Queue.group_of_task (Future.task_of future) :: groups;
137 val tab' = Inttab.cons_list (id, Future.map (K ()) future) tab;
138 in (m, groups', tab') end);
140 fun status task markups =
141 let val props = Markup.properties [(Markup.taskN, Task_Queue.str_of_task task)]
142 in Output.status (implode (map (Markup.markup_only o props) markups)) end;
146 fun fork_params {name, pos, pri} e =
147 uninterruptible (fn _ => Position.setmp_thread_data pos (fn () =>
149 val id = the_default 0 (Position.parse_id pos);
150 val _ = count_forked 1;
153 (singleton o Future.forks)
154 {name = name, group = NONE, deps = [], pri = pri, interrupts = false}
157 val task = the (Future.worker_task ());
158 val _ = status task [Markup.running];
160 Exn.capture (Future.interruptible_task e) ()
161 |> Future.identify_result pos;
162 val _ = status task [Markup.finished, Markup.joined];
167 if id = 0 orelse Exn.is_interrupt exn then ()
169 (status task [Markup.failed];
170 Output.report (Markup.markup_only Markup.bad);
171 List.app (Future.error_msg pos) (ML_Compiler.exn_messages_ids exn)));
172 val _ = count_forked ~1;
173 in Exn.release result end);
174 val _ = status (Future.task_of future) [Markup.forked];
175 val _ = register_forked id future;
179 fork_params {name = "Goal.fork", pos = Position.thread_data (), pri = pri} e;
181 fun forked_count () = #1 (Synchronized.value forked_proofs);
183 fun peek_futures id =
184 Inttab.lookup_list (#3 (Synchronized.value forked_proofs)) id;
186 fun reset_futures () =
187 Synchronized.change_result forked_proofs (fn (_, groups, _) =>
188 (Future.forked_proofs := 0; (groups, (0, [], Inttab.empty))));
190 fun shutdown_futures () =
192 (case map_filter Task_Queue.group_status (reset_futures ()) of
194 | exns => raise Par_Exn.make exns));
199 (* scheduling parameters *)
201 val skip_proofs = Unsynchronized.ref false;
203 val parallel_proofs = Unsynchronized.ref 1;
204 val parallel_subproofs_saturation = Unsynchronized.ref 100;
205 val parallel_subproofs_threshold = Unsynchronized.ref 0.01;
207 fun future_enabled_level n =
208 Multithreading.enabled () andalso ! parallel_proofs >= n andalso
209 is_some (Future.worker_task ());
211 fun future_enabled () = future_enabled_level 1;
213 fun future_enabled_nested n =
214 future_enabled_level n andalso
215 forked_count () < ! parallel_subproofs_saturation * Multithreading.max_threads_value ();
217 fun future_enabled_timing t =
218 future_enabled () andalso Time.toReal t >= ! parallel_subproofs_threshold;
223 fun future_result ctxt result prop =
225 val thy = Proof_Context.theory_of ctxt;
226 val _ = Context.reject_draft thy;
227 val cert = Thm.cterm_of thy;
228 val certT = Thm.ctyp_of thy;
230 val assms = Assumption.all_assms_of ctxt;
231 val As = map Thm.term_of assms;
233 val xs = map Free (fold Term.add_frees (prop :: As) []);
234 val fixes = map cert xs;
236 val tfrees = fold Term.add_tfrees (prop :: As) [];
237 val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees;
240 cert (Logic.varify_types_global (fold_rev Logic.all xs (Logic.list_implies (As, prop))))
241 |> Thm.weaken_sorts (Variable.sorts_of ctxt);
242 val global_result = result |> Future.map
243 (Drule.flexflex_unique #>
244 Thm.adjust_maxidx_thm ~1 #>
245 Drule.implies_intr_list assms #>
246 Drule.forall_intr_list fixes #>
247 Thm.generalize (map #1 tfrees, []) 0 #>
250 Thm.future global_result global_prop
251 |> Thm.close_derivation
252 |> Thm.instantiate (instT, [])
253 |> Drule.forall_elim_list fixes
254 |> fold (Thm.elim_implies o Thm.assume) assms;
259 (** tactical theorem proving **)
261 (* prove_internal -- minimal checks, no normalization of result! *)
263 fun prove_internal casms cprop tac =
264 (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
265 SOME th => Drule.implies_intr_list casms
266 (finish (Syntax.init_pretty_global (Thm.theory_of_thm th)) th)
267 | NONE => error "Tactic failed");
270 (* prove variations *)
273 Term.exists_subterm Term.is_Var t orelse
274 Term.exists_type (Term.exists_subtype Term.is_TVar) t;
276 fun prove_common immediate ctxt xs asms props tac =
278 val thy = Proof_Context.theory_of ctxt;
279 val string_of_term = Syntax.string_of_term ctxt;
281 val schematic = exists is_schematic props;
282 val future = future_enabled ();
283 val skip = not immediate andalso not schematic andalso future andalso ! skip_proofs;
285 val pos = Position.thread_data ();
286 fun err msg = cat_error msg
287 ("The error(s) above occurred for the goal statement:\n" ^
288 string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)) ^
289 (case Position.here pos of "" => "" | s => "\n" ^ s));
291 fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))
292 handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;
293 val casms = map cert_safe asms;
294 val cprops = map cert_safe props;
296 val (prems, ctxt') = ctxt
297 |> Variable.add_fixes_direct xs
298 |> fold Variable.declare_term (asms @ props)
299 |> Assumption.add_assumes casms
300 ||> Variable.set_body true;
301 val sorts = Variable.sorts_of ctxt';
303 val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops);
306 if skip then ALLGOALS Skip_Proof.cheat_tac st before Skip_Proof.report ctxt
309 (case SINGLE (tac' {prems = prems, context = ctxt'}) (init stmt) of
310 NONE => err "Tactic failed"
312 let val res = finish ctxt' st handle THM (msg, _, _) => err msg in
313 if Unify.matches_list thy [Thm.term_of stmt] [Thm.prop_of res]
314 then Thm.check_shyps sorts res
315 else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
318 if immediate orelse schematic orelse not future orelse skip
320 else future_result ctxt' (fork ~1 result) (Thm.term_of stmt);
322 Conjunction.elim_balanced (length props) res
323 |> map (Assumption.export false ctxt' ctxt)
324 |> Variable.export ctxt' ctxt
325 |> map Drule.zero_var_indexes
328 val prove_multi = prove_common true;
330 fun prove_future ctxt xs asms prop tac = hd (prove_common false ctxt xs asms [prop] tac);
332 fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac);
334 fun prove_global_future thy xs asms prop tac =
335 Drule.export_without_context (prove_future (Proof_Context.init_global thy) xs asms prop tac);
337 fun prove_global thy xs asms prop tac =
338 Drule.export_without_context (prove (Proof_Context.init_global thy) xs asms prop tac);
340 fun prove_sorry ctxt xs asms prop tac =
341 if ! quick_and_dirty then
342 prove ctxt xs asms prop (fn _ => ALLGOALS Skip_Proof.cheat_tac)
343 else (if future_enabled () then prove_future else prove) ctxt xs asms prop tac;
345 fun prove_sorry_global thy xs asms prop tac =
346 Drule.export_without_context
347 (prove_sorry (Proof_Context.init_global thy) xs asms prop tac);
351 (** goal structure **)
356 (if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty
357 else if n = 1 then Seq.single (Thm.cprem_of st i)
359 Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1))))
360 |> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);
362 fun retrofit i n st' st =
364 else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))
365 |> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;
367 fun SELECT_GOAL tac i st =
368 if Thm.nprems_of st = 1 andalso i = 1 then tac st
369 else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;
374 fun precise_conjunction_tac 0 i = eq_assume_tac i
375 | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i
376 | precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n));
378 val adhoc_conjunction_tac = REPEAT_ALL_NEW
379 (SUBGOAL (fn (goal, i) =>
380 if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i
383 val conjunction_tac = SUBGOAL (fn (goal, i) =>
384 precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE
385 TRY (adhoc_conjunction_tac i));
387 val recover_conjunction_tac = PRIMITIVE (fn th =>
388 Conjunction.uncurry_balanced (Thm.nprems_of th) th);
390 fun PRECISE_CONJUNCTS n tac =
391 SELECT_GOAL (precise_conjunction_tac n 1
393 THEN recover_conjunction_tac);
396 SELECT_GOAL (conjunction_tac 1
398 THEN recover_conjunction_tac);
401 (* hhf normal form *)
404 rtac Drule.asm_rl (*cheap approximation -- thanks to builtin Logic.flatten_params*)
405 THEN' SUBGOAL (fn (t, i) =>
406 if Drule.is_norm_hhf t then all_tac
407 else rewrite_goal_tac Drule.norm_hhf_eqs i);
409 fun compose_hhf_tac th i st =
410 PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st;
413 (* non-atomic goal assumptions *)
415 fun non_atomic (Const ("==>", _) $ _ $ _) = true
416 | non_atomic (Const ("all", _) $ _) = true
417 | non_atomic _ = false;
419 fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>
421 val ((_, goal'), ctxt') = Variable.focus_cterm goal ctxt;
422 val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
423 val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
424 val tacs = Rs |> map (fn R =>
425 Tactic.etac (Raw_Simplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);
426 in fold_rev (curry op APPEND') tacs (K no_tac) i end);
429 (* parallel tacticals *)
431 (*parallel choice of single results*)
432 fun PARALLEL_CHOICE tacs st =
433 (case Par_List.get_some (fn tac => SINGLE tac st) tacs of
435 | SOME st' => Seq.single st');
437 (*parallel refinement of non-schematic goal by single results*)
438 exception FAILED of unit;
439 fun PARALLEL_GOALS tac =
440 Thm.adjust_maxidx_thm ~1 #>
442 if not (Multithreading.enabled ()) orelse Thm.maxidx_of st >= 0 orelse Thm.nprems_of st <= 1
447 (case SINGLE tac g of
448 NONE => raise FAILED ()
451 val goals = Drule.strip_imp_prems (Thm.cprop_of st);
452 val results = Par_List.map (try_tac o init) goals;
453 in ALLGOALS (fn i => retrofit i 1 (nth results (i - 1))) st end
454 handle FAILED () => Seq.empty);
458 structure Basic_Goal: BASIC_GOAL = Goal;