| author | wenzelm | 
| Wed, 29 Jun 2005 15:13:36 +0200 | |
| changeset 16606 | e45c9a95a554 | 
| parent 15831 | aa58e4ec3a1f | 
| child 16842 | 5979c46853d1 | 
| permissions | -rw-r--r-- | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/Isar/skip_proof.ML | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 4 | |
| 8539 | 5 | Skipping proofs -- quick_and_dirty mode. | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 6 | *) | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 7 | |
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 8 | signature SKIP_PROOF = | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 9 | sig | 
| 8539 | 10 | val quick_and_dirty: bool ref | 
| 11892 | 11 | val make_thm: theory -> term -> thm | 
| 12 | val cheat_tac: theory -> tactic | |
| 11972 | 13 | val prove: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm | 
| 12148 | 14 | val local_skip_proof: (Proof.context -> string * (string * thm list) list -> unit) * | 
| 12318 
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
 wenzelm parents: 
12244diff
changeset | 15 | (Proof.context -> thm -> unit) -> bool -> Proof.state -> Proof.state Seq.seq | 
| 
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
 wenzelm parents: 
12244diff
changeset | 16 | val global_skip_proof: | 
| 
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
 wenzelm parents: 
12244diff
changeset | 17 | bool -> Proof.state -> theory * ((string * string) * (string * thm list) list) | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 18 | end; | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 19 | |
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 20 | structure SkipProof: SKIP_PROOF = | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 21 | struct | 
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 22 | |
| 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 23 | |
| 8539 | 24 | (* quick_and_dirty *) | 
| 25 | ||
| 26 | (*if true then some packages will OMIT SOME PROOFS*) | |
| 27 | val quick_and_dirty = ref false; | |
| 28 | ||
| 29 | ||
| 11892 | 30 | (* oracle setup *) | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 31 | |
| 11892 | 32 | exception SkipProof of term; | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 33 | |
| 11892 | 34 | fun skip_proof (_, SkipProof t) = | 
| 35 | if ! quick_and_dirty then t | |
| 12318 
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
 wenzelm parents: 
12244diff
changeset | 36 | else error "Proof may be skipped in quick_and_dirty mode only!"; | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 37 | |
| 15831 | 38 | val _ = Context.add_setup [Theory.add_oracle ("skip_proof", skip_proof)];
 | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 39 | |
| 11892 | 40 | |
| 11972 | 41 | (* basic cheating *) | 
| 11892 | 42 | |
| 43 | fun make_thm thy t = | |
| 15831 | 44 | Thm.invoke_oracle_i thy "Pure.skip_proof" (Theory.sign_of thy, SkipProof t); | 
| 11892 | 45 | |
| 8539 | 46 | fun cheat_tac thy st = | 
| 11892 | 47 |   ALLGOALS (Tactic.rtac (make_thm thy (Var (("A", 0), propT)))) st;
 | 
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 48 | |
| 11972 | 49 | fun prove thy xs asms prop tac = | 
| 50 | Tactic.prove (Theory.sign_of thy) xs asms prop | |
| 51 | (if ! quick_and_dirty then (K (cheat_tac thy)) else tac); | |
| 52 | ||
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 53 | |
| 8539 | 54 | (* "sorry" proof command *) | 
| 55 | ||
| 12318 
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
 wenzelm parents: 
12244diff
changeset | 56 | fun cheating int ctxt = Method.METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty) | 
| 
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
 wenzelm parents: 
12244diff
changeset | 57 | (cheat_tac (ProofContext.theory_of ctxt)))); | 
| 8539 | 58 | |
| 15531 | 59 | fun local_skip_proof x int = Method.local_terminal_proof (Method.Basic (cheating int), NONE) x; | 
| 60 | fun global_skip_proof int = Method.global_terminal_proof (Method.Basic (cheating int), NONE); | |
| 8539 | 61 | |
| 6888 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 wenzelm parents: diff
changeset | 62 | end; | 
| 8539 | 63 | |
| 64 | ||
| 65 | val quick_and_dirty = SkipProof.quick_and_dirty; |