wenzelm@6888: (* Title: Pure/Isar/skip_proof.ML wenzelm@6888: Author: Markus Wenzel, TU Muenchen wenzelm@6888: wenzelm@8539: Skipping proofs -- quick_and_dirty mode. wenzelm@6888: *) wenzelm@6888: wenzelm@6888: signature SKIP_PROOF = wenzelm@6888: sig wenzelm@11892: val make_thm: theory -> term -> thm wenzelm@11892: val cheat_tac: theory -> tactic wenzelm@20289: val prove: Proof.context -> string list -> term list -> term -> wenzelm@20289: ({prems: thm list, context: Proof.context} -> tactic) -> thm berghofe@26530: val prove_global: theory -> string list -> term list -> term -> wenzelm@26711: ({prems: thm list, context: Proof.context} -> tactic) -> thm wenzelm@6888: end; wenzelm@6888: wenzelm@6888: structure SkipProof: SKIP_PROOF = wenzelm@6888: struct wenzelm@6888: wenzelm@11892: (* oracle setup *) wenzelm@6888: wenzelm@28290: val (_, skip_proof) = Context.>>> (Context.map_theory_result wenzelm@28290: (Thm.add_oracle ("skip_proof", fn (thy, prop) => wenzelm@28290: if ! quick_and_dirty then Thm.cterm_of thy prop wenzelm@28290: else error "Proof may be skipped in quick_and_dirty mode only!"))); wenzelm@6888: wenzelm@11892: wenzelm@11972: (* basic cheating *) wenzelm@11892: wenzelm@28290: fun make_thm thy prop = skip_proof (thy, prop); 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@20049: fun prove ctxt xs asms prop tac = wenzelm@29448: (if Goal.future_enabled () then Goal.prove_future else Goal.prove) ctxt xs asms prop wenzelm@28365: (fn args => fn st => wenzelm@28365: if ! quick_and_dirty wenzelm@28365: then setmp quick_and_dirty true (cheat_tac (ProofContext.theory_of ctxt)) st wenzelm@28365: else tac args st); wenzelm@11972: berghofe@26530: fun prove_global thy xs asms prop tac = wenzelm@26711: Drule.standard (prove (ProofContext.init thy) xs asms prop tac); berghofe@26530: wenzelm@6888: end;