| author | wenzelm |
| Wed, 02 Sep 2009 14:11:45 +0200 | |
| changeset 32490 | 6a48db3e627c |
| parent 30288 | a32700e45ab3 |
| child 32966 | 5b21661fe618 |
| 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 |
Author: Markus Wenzel, TU Muenchen |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
3 |
|
| 8539 | 4 |
Skipping proofs -- quick_and_dirty mode. |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
5 |
*) |
|
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 |
signature SKIP_PROOF = |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
8 |
sig |
| 11892 | 9 |
val make_thm: theory -> term -> thm |
10 |
val cheat_tac: theory -> tactic |
|
| 20289 | 11 |
val prove: Proof.context -> string list -> term list -> term -> |
12 |
({prems: thm list, context: Proof.context} -> tactic) -> thm
|
|
| 26530 | 13 |
val prove_global: theory -> string list -> term list -> term -> |
| 26711 | 14 |
({prems: thm list, context: Proof.context} -> tactic) -> thm
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
15 |
end; |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
16 |
|
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
17 |
structure SkipProof: SKIP_PROOF = |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
18 |
struct |
|
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
19 |
|
| 11892 | 20 |
(* oracle setup *) |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
21 |
|
| 28290 | 22 |
val (_, skip_proof) = Context.>>> (Context.map_theory_result |
|
30288
a32700e45ab3
Thm.add_oracle interface: replaced old bstring by binding;
wenzelm
parents:
29606
diff
changeset
|
23 |
(Thm.add_oracle (Binding.name "skip_proof", fn (thy, prop) => |
| 28290 | 24 |
if ! quick_and_dirty then Thm.cterm_of thy prop |
25 |
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
|
26 |
|
| 11892 | 27 |
|
| 11972 | 28 |
(* basic cheating *) |
| 11892 | 29 |
|
| 28290 | 30 |
fun make_thm thy prop = skip_proof (thy, prop); |
| 11892 | 31 |
|
| 8539 | 32 |
fun cheat_tac thy st = |
| 11892 | 33 |
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
|
34 |
|
| 20049 | 35 |
fun prove ctxt xs asms prop tac = |
|
29448
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
36 |
(if Goal.future_enabled () then Goal.prove_future else Goal.prove) ctxt xs asms prop |
| 28365 | 37 |
(fn args => fn st => |
38 |
if ! quick_and_dirty |
|
39 |
then setmp quick_and_dirty true (cheat_tac (ProofContext.theory_of ctxt)) st |
|
40 |
else tac args st); |
|
| 11972 | 41 |
|
| 26530 | 42 |
fun prove_global thy xs asms prop tac = |
| 26711 | 43 |
Drule.standard (prove (ProofContext.init thy) xs asms prop tac); |
| 26530 | 44 |
|
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
45 |
end; |