author | wenzelm |
Fri, 31 Oct 2014 17:08:54 +0100 | |
changeset 58848 | fd0c85d7da38 |
parent 58837 | e84d900cd287 |
child 59498 | 50b60f501b05 |
permissions | -rw-r--r-- |
51551 | 1 |
(* Title: Pure/skip_proof.ML |
2 |
Author: Makarius |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
3 |
|
51551 | 4 |
Skip proof via oracle invocation. |
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 |
|
56438 | 7 |
val quick_and_dirty_raw = Config.declare_option ("quick_and_dirty", @{here}); |
8 |
val quick_and_dirty = Config.bool quick_and_dirty_raw; |
|
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 |
51551 | 12 |
val report: Proof.context -> unit |
42409 | 13 |
val make_thm_cterm: cterm -> thm |
11892 | 14 |
val make_thm: theory -> term -> thm |
51552
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
51551
diff
changeset
|
15 |
val cheat_tac: int -> tactic |
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 |
|
51551 | 21 |
(* report *) |
22 |
||
23 |
fun report ctxt = |
|
56294
85911b8a6868
prefer Context_Position where a context is available;
wenzelm
parents:
51552
diff
changeset
|
24 |
if Context_Position.is_visible ctxt then |
56333
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents:
56294
diff
changeset
|
25 |
Output.report [Markup.markup Markup.bad "Skipped proof"] |
56294
85911b8a6868
prefer Context_Position where a context is available;
wenzelm
parents:
51552
diff
changeset
|
26 |
else (); |
51551 | 27 |
|
28 |
||
11892 | 29 |
(* oracle setup *) |
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
30 |
|
42409 | 31 |
val (_, make_thm_cterm) = |
56436 | 32 |
Context.>>> |
33 |
(Context.map_theory_result (Thm.add_oracle (Binding.make ("skip_proof", @{here}), I))); |
|
42409 | 34 |
|
35 |
fun make_thm thy prop = make_thm_cterm (Thm.cterm_of thy prop); |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
36 |
|
11892 | 37 |
|
51551 | 38 |
(* cheat_tac *) |
11892 | 39 |
|
51552
c713c9505f68
clarified Skip_Proof.cheat_tac: more standard tactic;
wenzelm
parents:
51551
diff
changeset
|
40 |
fun cheat_tac i st = |
58837 | 41 |
resolve_tac [make_thm (Thm.theory_of_thm st) (Var (("A", 0), propT))] i st; |
26530 | 42 |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
diff
changeset
|
43 |
end; |