| author | wenzelm |
| Thu, 17 Apr 2008 22:22:30 +0200 | |
| changeset 26717 | 2e1c3a0e7308 |
| parent 26711 | 3a478bfa1650 |
| child 28290 | 4cc2b6046258 |
| 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 |
| 11892 | 10 |
val make_thm: theory -> term -> thm |
11 |
val cheat_tac: theory -> tactic |
|
| 20289 | 12 |
val prove: Proof.context -> string list -> term list -> term -> |
13 |
({prems: thm list, context: Proof.context} -> tactic) -> thm
|
|
| 26530 | 14 |
val prove_global: theory -> string list -> term list -> term -> |
| 26711 | 15 |
({prems: thm list, context: Proof.context} -> tactic) -> thm
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
16 |
end; |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
17 |
|
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
18 |
structure SkipProof: SKIP_PROOF = |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
19 |
struct |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
20 |
|
| 11892 | 21 |
(* oracle setup *) |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
22 |
|
| 11892 | 23 |
exception SkipProof of term; |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
24 |
|
| 17362 | 25 |
fun skip_proof (_, SkipProof prop) = |
26 |
if ! quick_and_dirty then prop |
|
|
12318
72348ff7d4a0
skip_proof: do not require quick_and_dirty in interactive mode;
wenzelm
parents:
12244
diff
changeset
|
27 |
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
|
28 |
|
| 26463 | 29 |
val _ = Context.>> (Context.map_theory |
30 |
(Theory.add_oracle ("skip_proof", skip_proof)));
|
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
31 |
|
| 11892 | 32 |
|
| 11972 | 33 |
(* basic cheating *) |
| 11892 | 34 |
|
| 17362 | 35 |
fun make_thm thy prop = |
36 |
Thm.invoke_oracle_i thy "Pure.skip_proof" (thy, SkipProof prop); |
|
| 11892 | 37 |
|
| 8539 | 38 |
fun cheat_tac thy st = |
| 11892 | 39 |
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
|
40 |
|
| 20049 | 41 |
fun prove ctxt xs asms prop tac = |
|
24502
8d5326f0098b
prove: setmp quick_and_dirty (avoids race condition);
wenzelm
parents:
20289
diff
changeset
|
42 |
Goal.prove ctxt xs asms prop (fn args => fn st => |
|
8d5326f0098b
prove: setmp quick_and_dirty (avoids race condition);
wenzelm
parents:
20289
diff
changeset
|
43 |
if ! quick_and_dirty |
|
8d5326f0098b
prove: setmp quick_and_dirty (avoids race condition);
wenzelm
parents:
20289
diff
changeset
|
44 |
then setmp quick_and_dirty true (cheat_tac (ProofContext.theory_of ctxt)) st |
|
8d5326f0098b
prove: setmp quick_and_dirty (avoids race condition);
wenzelm
parents:
20289
diff
changeset
|
45 |
else tac args st); |
| 11972 | 46 |
|
| 26530 | 47 |
fun prove_global thy xs asms prop tac = |
| 26711 | 48 |
Drule.standard (prove (ProofContext.init thy) xs asms prop tac); |
| 26530 | 49 |
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
50 |
end; |