author | wenzelm |
Thu, 25 Sep 2008 20:34:20 +0200 | |
changeset 28365 | 6249297461cb |
parent 28290 | 4cc2b6046258 |
child 28446 | a01de3b3fa2e |
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 |
|
28365 | 8 |
val future_scheduler = ref false; |
9 |
||
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
10 |
signature SKIP_PROOF = |
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
11 |
sig |
11892 | 12 |
val make_thm: theory -> term -> thm |
13 |
val cheat_tac: theory -> tactic |
|
20289 | 14 |
val prove: Proof.context -> string list -> term list -> term -> |
15 |
({prems: thm list, context: Proof.context} -> tactic) -> thm |
|
26530 | 16 |
val prove_global: theory -> string list -> term list -> term -> |
26711 | 17 |
({prems: thm list, context: Proof.context} -> tactic) -> thm |
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 |
|
11892 | 23 |
(* oracle setup *) |
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
24 |
|
28290 | 25 |
val (_, skip_proof) = Context.>>> (Context.map_theory_result |
26 |
(Thm.add_oracle ("skip_proof", fn (thy, prop) => |
|
27 |
if ! quick_and_dirty then Thm.cterm_of thy prop |
|
28 |
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
|
29 |
|
11892 | 30 |
|
11972 | 31 |
(* basic cheating *) |
11892 | 32 |
|
28290 | 33 |
fun make_thm thy prop = skip_proof (thy, prop); |
11892 | 34 |
|
8539 | 35 |
fun cheat_tac thy st = |
11892 | 36 |
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
|
37 |
|
20049 | 38 |
fun prove ctxt xs asms prop tac = |
28365 | 39 |
(if ! future_scheduler then Goal.prove_promise else Goal.prove) ctxt xs asms prop |
40 |
(fn args => fn st => |
|
41 |
if ! quick_and_dirty |
|
42 |
then setmp quick_and_dirty true (cheat_tac (ProofContext.theory_of ctxt)) st |
|
43 |
else tac args st); |
|
11972 | 44 |
|
26530 | 45 |
fun prove_global thy xs asms prop tac = |
26711 | 46 |
Drule.standard (prove (ProofContext.init thy) xs asms prop tac); |
26530 | 47 |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
48 |
end; |