author | bulwahn |
Wed, 24 Mar 2010 17:40:44 +0100 | |
changeset 35956 | f53dc0b539fa |
parent 35021 | c839a4c670c6 |
child 36610 | bafd82950e24 |
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 |
|
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32966
diff
changeset
|
4 |
Skipping proofs -- via oracle (in quick and dirty mode) or by forking |
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32966
diff
changeset
|
5 |
(parallel 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 |
|
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32966
diff
changeset
|
18 |
structure Skip_Proof: SKIP_PROOF = |
6888
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 |
|
28290 | 23 |
val (_, skip_proof) = Context.>>> (Context.map_theory_result |
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32966
diff
changeset
|
24 |
(Thm.add_oracle (Binding.name "skip_proof", fn (thy, prop) => Thm.cterm_of thy prop))); |
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
25 |
|
11892 | 26 |
|
11972 | 27 |
(* basic cheating *) |
11892 | 28 |
|
28290 | 29 |
fun make_thm thy prop = skip_proof (thy, prop); |
11892 | 30 |
|
8539 | 31 |
fun cheat_tac thy st = |
11892 | 32 |
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
|
33 |
|
20049 | 34 |
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
|
35 |
(if Goal.future_enabled () then Goal.prove_future else Goal.prove) ctxt xs asms prop |
28365 | 36 |
(fn args => fn st => |
37 |
if ! quick_and_dirty |
|
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32966
diff
changeset
|
38 |
then cheat_tac (ProofContext.theory_of ctxt) st |
28365 | 39 |
else tac args st); |
11972 | 40 |
|
26530 | 41 |
fun prove_global thy xs asms prop tac = |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
32970
diff
changeset
|
42 |
Drule.export_without_context (prove (ProofContext.init thy) xs asms prop tac); |
26530 | 43 |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
44 |
end; |