--- a/src/Pure/Isar/skip_proof.ML Mon Mar 20 18:47:47 2000 +0100
+++ b/src/Pure/Isar/skip_proof.ML Mon Mar 20 18:48:12 2000 +0100
@@ -2,14 +2,16 @@
ID: $Id$
Author: Markus Wenzel, TU Muenchen
-Skip subproofs (for quick_and_dirty mode only).
+Skipping proofs -- quick_and_dirty mode.
*)
signature SKIP_PROOF =
sig
+ val quick_and_dirty: bool ref
val local_skip_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit)
-> Proof.state -> Proof.state Seq.seq
val global_skip_proof: Proof.state -> theory * {kind: string, name: string, thm: thm}
+ val prove_goalw_cterm: theory -> thm list -> cterm -> (thm list -> tactic list) -> thm
val setup: (theory -> theory) list
end;
@@ -17,7 +19,13 @@
struct
-(* oracle *)
+(* quick_and_dirty *)
+
+(*if true then some packages will OMIT SOME PROOFS*)
+val quick_and_dirty = ref false;
+
+
+(* cheat_tac *)
val skip_proofN = "skip_proof";
@@ -29,20 +37,24 @@
else error "Proofs may be skipped in quick_and_dirty mode only!";
-(* proof command *)
-
-fun cheat_meth ctxt =
- let
- val thy = ProofContext.theory_of ctxt;
- (*too bad -- dynamic scoping of the oracle, cannot even qualify the name*)
- val thm = Thm.invoke_oracle thy skip_proofN (Theory.sign_of thy, SkipProof);
- in Method.METHOD (fn _ => ALLGOALS (Tactic.rtac thm)) end;
-
-val local_skip_proof = Method.local_terminal_proof (Method.Basic cheat_meth, None);
-val global_skip_proof = Method.global_terminal_proof (Method.Basic cheat_meth, None);
+(*too bad -- dynamic scoping of the oracle, cannot even qualify the name*)
+fun cheat_tac thy st =
+ ALLGOALS (Tactic.rtac (Thm.invoke_oracle thy skip_proofN (Theory.sign_of thy, SkipProof))) st;
-(* proof command *)
+(* "sorry" proof command *)
+
+fun cheating ctxt = Method.METHOD (K (cheat_tac (ProofContext.theory_of ctxt)));
+
+val local_skip_proof = Method.local_terminal_proof (Method.Basic cheating, None);
+val global_skip_proof = Method.global_terminal_proof (Method.Basic cheating, None);
+
+
+(* conditional prove_goalw_cterm *)
+
+fun prove_goalw_cterm thy rews ct tacs =
+ Goals.prove_goalw_cterm rews ct
+ (if ! quick_and_dirty then (K [cheat_tac thy]) else tacs);
(* theory setup *)
@@ -51,3 +63,6 @@
end;
+
+
+val quick_and_dirty = SkipProof.quick_and_dirty;