wenzelm@6888: (* Title: Pure/Isar/skip_proof.ML wenzelm@6888: ID: $Id$ wenzelm@6888: Author: Markus Wenzel, TU Muenchen wenzelm@8807: License: GPL (GNU GENERAL PUBLIC LICENSE) wenzelm@6888: wenzelm@8539: Skipping proofs -- quick_and_dirty mode. wenzelm@6888: *) wenzelm@6888: wenzelm@6888: signature SKIP_PROOF = wenzelm@6888: sig wenzelm@8539: val quick_and_dirty: bool ref wenzelm@11892: val make_thm: theory -> term -> thm wenzelm@11892: val cheat_tac: theory -> tactic wenzelm@6984: val local_skip_proof: ({kind: string, name: string, thm: thm} -> unit) * (thm -> unit) wenzelm@6888: -> Proof.state -> Proof.state Seq.seq wenzelm@6888: val global_skip_proof: Proof.state -> theory * {kind: string, name: string, thm: thm} wenzelm@6888: val setup: (theory -> theory) list wenzelm@6888: end; wenzelm@6888: wenzelm@6888: structure SkipProof: SKIP_PROOF = wenzelm@6888: struct wenzelm@6888: wenzelm@6888: wenzelm@8539: (* quick_and_dirty *) wenzelm@8539: wenzelm@8539: (*if true then some packages will OMIT SOME PROOFS*) wenzelm@8539: val quick_and_dirty = ref false; wenzelm@8539: wenzelm@8539: wenzelm@11892: (* oracle setup *) wenzelm@6888: wenzelm@6888: val skip_proofN = "skip_proof"; wenzelm@6888: wenzelm@11892: exception SkipProof of term; wenzelm@6888: wenzelm@11892: fun skip_proof (_, SkipProof t) = wenzelm@11892: if ! quick_and_dirty then t wenzelm@6888: else error "Proofs may be skipped in quick_and_dirty mode only!"; wenzelm@6888: wenzelm@11892: val setup = [Theory.add_oracle (skip_proofN, skip_proof)]; wenzelm@6888: wenzelm@11892: wenzelm@11892: (* make_thm and cheat_tac *) wenzelm@11892: wenzelm@11892: fun make_thm thy t = wenzelm@11892: (*dynamic scoping of the oracle, cannot even qualify the name due to Pure/CPure!*) wenzelm@11892: Thm.invoke_oracle thy skip_proofN (Theory.sign_of thy, SkipProof t); wenzelm@11892: wenzelm@8539: fun cheat_tac thy st = wenzelm@11892: ALLGOALS (Tactic.rtac (make_thm thy (Var (("A", 0), propT)))) st; wenzelm@6888: wenzelm@6888: wenzelm@8539: (* "sorry" proof command *) wenzelm@8539: wenzelm@8539: fun cheating ctxt = Method.METHOD (K (cheat_tac (ProofContext.theory_of ctxt))); wenzelm@8539: wenzelm@8539: val local_skip_proof = Method.local_terminal_proof (Method.Basic cheating, None); wenzelm@8539: val global_skip_proof = Method.global_terminal_proof (Method.Basic cheating, None); wenzelm@8539: wenzelm@6888: end; wenzelm@8539: wenzelm@8539: wenzelm@8539: val quick_and_dirty = SkipProof.quick_and_dirty;