| author | nipkow | 
| Fri, 12 Aug 2011 17:01:30 +0200 | |
| changeset 44164 | 238c5ea1e2ce | 
| parent 42409 | 3e1e80df6a42 | 
| child 46045 | 332cb37cfcee | 
| 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  | 
| 42409 | 10  | 
val make_thm_cterm: cterm -> thm  | 
| 11892 | 11  | 
val make_thm: theory -> term -> thm  | 
12  | 
val cheat_tac: theory -> tactic  | 
|
| 20289 | 13  | 
val prove: Proof.context -> string list -> term list -> term ->  | 
14  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
|
| 26530 | 15  | 
val prove_global: theory -> string list -> term list -> term ->  | 
| 26711 | 16  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
end;  | 
| 
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
|
| 
32970
 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 
wenzelm 
parents: 
32966 
diff
changeset
 | 
19  | 
structure Skip_Proof: SKIP_PROOF =  | 
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
struct  | 
| 
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 11892 | 22  | 
(* oracle setup *)  | 
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
|
| 42409 | 24  | 
val (_, make_thm_cterm) =  | 
25  | 
Context.>>> (Context.map_theory_result (Thm.add_oracle (Binding.name "skip_proof", I)));  | 
|
26  | 
||
27  | 
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
 | 
28  | 
|
| 11892 | 29  | 
|
| 11972 | 30  | 
(* basic cheating *)  | 
| 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  | 
|
| 42360 | 39  | 
then cheat_tac (Proof_Context.theory_of ctxt) st  | 
| 28365 | 40  | 
else tac args st);  | 
| 11972 | 41  | 
|
| 26530 | 42  | 
fun prove_global thy xs asms prop tac =  | 
| 42360 | 43  | 
Drule.export_without_context (prove (Proof_Context.init_global 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;  |