src/Pure/Isar/skip_proof.ML
author wenzelm
Wed, 19 Oct 2011 16:45:46 +0200
changeset 45197 b6c527c64789
parent 42409 3e1e80df6a42
child 46045 332cb37cfcee
permissions -rw-r--r--
more robust toplevel_error reporting (NB: Context.proof_of on a stale theory crashes ungracefully);

(*  Title:      Pure/Isar/skip_proof.ML
    Author:     Markus Wenzel, TU Muenchen

Skipping proofs -- via oracle (in quick and dirty mode) or by forking
(parallel mode).
*)

signature SKIP_PROOF =
sig
  val make_thm_cterm: cterm -> thm
  val make_thm: theory -> term -> thm
  val cheat_tac: theory -> tactic
  val prove: Proof.context -> string list -> term list -> term ->
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
  val prove_global: theory -> string list -> term list -> term ->
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
end;

structure Skip_Proof: SKIP_PROOF =
struct

(* oracle setup *)

val (_, make_thm_cterm) =
  Context.>>> (Context.map_theory_result (Thm.add_oracle (Binding.name "skip_proof", I)));

fun make_thm thy prop = make_thm_cterm (Thm.cterm_of thy prop);


(* basic cheating *)

fun cheat_tac thy st =
  ALLGOALS (Tactic.rtac (make_thm thy (Var (("A", 0), propT)))) st;

fun prove ctxt xs asms prop tac =
  (if Goal.future_enabled () then Goal.prove_future else Goal.prove) ctxt xs asms prop
    (fn args => fn st =>
      if ! quick_and_dirty
      then cheat_tac (Proof_Context.theory_of ctxt) st
      else tac args st);

fun prove_global thy xs asms prop tac =
  Drule.export_without_context (prove (Proof_Context.init_global thy) xs asms prop tac);

end;