added lemmas sinl_defined_iff sinr_defined_iff, sinl_eq_sinr, sinr_eq_sinl; added more simp rules; cleaned up
(*  Title:      Pure/Isar/skip_proof.ML
    ID:         $Id$
    Author:     Markus Wenzel, TU Muenchen
Skipping proofs -- quick_and_dirty mode.
*)
signature SKIP_PROOF =
sig
  val quick_and_dirty: bool ref
  val make_thm: theory -> term -> thm
  val cheat_tac: theory -> tactic
  val prove: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm
  val local_skip_proof: (Proof.context -> string * (string * thm list) list -> unit) *
    (Proof.context -> thm -> unit) -> bool -> Proof.state -> Proof.state Seq.seq
  val global_skip_proof:
    bool -> Proof.state -> theory * ((string * string) * (string * thm list) list)
end;
structure SkipProof: SKIP_PROOF =
struct
(* quick_and_dirty *)
(*if true then some packages will OMIT SOME PROOFS*)
val quick_and_dirty = ref false;
(* oracle setup *)
exception SkipProof of term;
fun skip_proof (_, SkipProof t) =
  if ! quick_and_dirty then t
  else error "Proof may be skipped in quick_and_dirty mode only!";
val _ = Context.add_setup [Theory.add_oracle ("skip_proof", skip_proof)];
(* basic cheating *)
fun make_thm thy t =
  Thm.invoke_oracle_i thy "Pure.skip_proof" (Theory.sign_of thy, SkipProof t);
fun cheat_tac thy st =
  ALLGOALS (Tactic.rtac (make_thm thy (Var (("A", 0), propT)))) st;
fun prove thy xs asms prop tac =
  Tactic.prove (Theory.sign_of thy) xs asms prop
    (if ! quick_and_dirty then (K (cheat_tac thy)) else tac);
(* "sorry" proof command *)
fun cheating int ctxt = Method.METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty)
    (cheat_tac (ProofContext.theory_of ctxt))));
fun local_skip_proof x int = Method.local_terminal_proof (Method.Basic (cheating int), NONE) x;
fun global_skip_proof int = Method.global_terminal_proof (Method.Basic (cheating int), NONE);
end;
val quick_and_dirty = SkipProof.quick_and_dirty;