src/Pure/skip_proof.ML
author wenzelm
Wed, 19 Aug 2015 16:21:10 +0200
changeset 60975 5f3d6e16ea78
parent 60820 d0a88a2182a8
child 61527 d05f3d86a758
permissions -rw-r--r--
avoid ambiguities on native Windows, such as / vs. /cygdrive/c/cygwin;

(*  Title:      Pure/skip_proof.ML
    Author:     Makarius

Skip proof via oracle invocation.
*)

val quick_and_dirty_raw = Config.declare_option ("quick_and_dirty", @{here});
val quick_and_dirty = Config.bool quick_and_dirty_raw;

signature SKIP_PROOF =
sig
  val report: Proof.context -> unit
  val make_thm_cterm: cterm -> thm
  val make_thm: theory -> term -> thm
  val cheat_tac: Proof.context -> int -> tactic
end;

structure Skip_Proof: SKIP_PROOF =
struct

(* report *)

fun report ctxt =
  if Context_Position.is_visible ctxt then
    Output.report [Markup.markup Markup.bad "Skipped proof"]
  else ();


(* oracle setup *)

val (_, make_thm_cterm) =
  Context.>>>
    (Context.map_theory_result (Thm.add_oracle (Binding.make ("skip_proof", @{here}), I)));

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


(* cheat_tac *)

fun cheat_tac ctxt i st =
  resolve_tac ctxt [make_thm (Proof_Context.theory_of ctxt) (Var (("A", 0), propT))] i st;

end;