| author | isatest |
| Fri, 22 Apr 2005 00:34:15 +0200 | |
| changeset 15807 | 98bff95f88ba |
| parent 15801 | d2f5ca3c048d |
| child 15831 | aa58e4ec3a1f |
| 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:
12244
diff
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:
12244
diff
changeset
|
16 |
val global_skip_proof: |
|
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
wenzelm
parents:
12244
diff
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 |
|
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
32 |
val skip_proofN = "skip_proof"; |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
33 |
|
| 11892 | 34 |
exception SkipProof of term; |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
35 |
|
| 11892 | 36 |
fun skip_proof (_, SkipProof t) = |
37 |
if ! quick_and_dirty then t |
|
|
12318
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
wenzelm
parents:
12244
diff
changeset
|
38 |
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
|
39 |
|
| 15801 | 40 |
val _ = Context.add_setup |
41 |
[PureThy.global_path, Theory.add_oracle (skip_proofN, skip_proof), PureThy.local_path]; |
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
42 |
|
| 11892 | 43 |
|
| 11972 | 44 |
(* basic cheating *) |
| 11892 | 45 |
|
46 |
fun make_thm thy t = |
|
| 15667 | 47 |
Thm.invoke_oracle_i thy skip_proofN (Theory.sign_of thy, SkipProof t); |
| 11892 | 48 |
|
| 8539 | 49 |
fun cheat_tac thy st = |
| 11892 | 50 |
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
|
51 |
|
| 11972 | 52 |
fun prove thy xs asms prop tac = |
53 |
Tactic.prove (Theory.sign_of thy) xs asms prop |
|
54 |
(if ! quick_and_dirty then (K (cheat_tac thy)) else tac); |
|
55 |
||
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
56 |
|
| 8539 | 57 |
(* "sorry" proof command *) |
58 |
||
|
12318
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
wenzelm
parents:
12244
diff
changeset
|
59 |
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:
12244
diff
changeset
|
60 |
(cheat_tac (ProofContext.theory_of ctxt)))); |
| 8539 | 61 |
|
| 15531 | 62 |
fun local_skip_proof x int = Method.local_terminal_proof (Method.Basic (cheating int), NONE) x; |
63 |
fun global_skip_proof int = Method.global_terminal_proof (Method.Basic (cheating int), NONE); |
|
| 8539 | 64 |
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
65 |
end; |
| 8539 | 66 |
|
67 |
||
68 |
val quick_and_dirty = SkipProof.quick_and_dirty; |