| author | wenzelm | 
| Wed, 31 Dec 2008 00:01:07 +0100 | |
| changeset 29262 | 3ee4656b9e0c | 
| parent 27869 | af1f79cbc198 | 
| child 29606 | fedb8be05f24 | 
| permissions | -rw-r--r-- | 
| 27389 | 1  | 
(* Title: Pure/ML/ml_thms.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Makarius  | 
|
4  | 
||
5  | 
Isar theorem values within ML.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature ML_THMS =  | 
|
9  | 
sig  | 
|
10  | 
val the_thms: Proof.context -> int -> thm list  | 
|
11  | 
val the_thm: Proof.context -> int -> thm  | 
|
12  | 
end;  | 
|
13  | 
||
14  | 
structure ML_Thms: ML_THMS =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
(* auxiliary facts table *)  | 
|
18  | 
||
19  | 
structure AuxFacts = ProofDataFun  | 
|
20  | 
(  | 
|
21  | 
type T = thm list Inttab.table;  | 
|
22  | 
fun init _ = Inttab.empty;  | 
|
23  | 
);  | 
|
24  | 
||
25  | 
val put_thms = AuxFacts.map o Inttab.update;  | 
|
26  | 
||
27  | 
fun the_thms ctxt name = the (Inttab.lookup (AuxFacts.get ctxt) name);  | 
|
28  | 
fun the_thm ctxt name = the_single (the_thms ctxt name);  | 
|
29  | 
||
30  | 
fun thm_bind kind a i =  | 
|
31  | 
"val " ^ a ^ " = ML_Thms.the_" ^ kind ^  | 
|
32  | 
" (ML_Context.the_local_context ()) " ^ string_of_int i ^ ";\n";  | 
|
33  | 
||
34  | 
||
35  | 
(* fact references *)  | 
|
36  | 
||
37  | 
fun thm_antiq kind scan = ML_Context.add_antiq kind  | 
|
| 27869 | 38  | 
  (fn _ => scan >> (fn ths => fn {struct_name, background} =>
 | 
| 27389 | 39  | 
let  | 
40  | 
val i = serial ();  | 
|
41  | 
val (a, background') = background  | 
|
42  | 
|> ML_Antiquote.variant kind ||> put_thms (i, ths);  | 
|
43  | 
val ml = (thm_bind kind a i, struct_name ^ "." ^ a);  | 
|
44  | 
in (K ml, background') end));  | 
|
45  | 
||
46  | 
val _ = thm_antiq "thm" (Attrib.thm >> single);  | 
|
47  | 
val _ = thm_antiq "thms" Attrib.thms;  | 
|
48  | 
||
49  | 
||
50  | 
(* ad-hoc goals *)  | 
|
51  | 
||
| 
27809
 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27521 
diff
changeset
 | 
52  | 
val by = Args.$$$ "by";  | 
| 
 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27521 
diff
changeset
 | 
53  | 
val goal = Scan.unless (Scan.lift by) Args.prop;  | 
| 27389 | 54  | 
|
55  | 
val _ = ML_Context.add_antiq "lemma"  | 
|
| 27869 | 56  | 
(fn pos => Args.context -- Args.mode "open" -- Scan.repeat1 goal --  | 
| 
27809
 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
27521 
diff
changeset
 | 
57  | 
Scan.lift (by |-- Method.parse -- Scan.option Method.parse) >>  | 
| 
27521
 
44081396d735
@{lemma}: allow terminal method, close derivation unless (open) mode is given;
 
wenzelm 
parents: 
27389 
diff
changeset
 | 
58  | 
    (fn (((ctxt, is_open), props), methods) => fn {struct_name, background} =>
 | 
| 27389 | 59  | 
let  | 
60  | 
val i = serial ();  | 
|
| 27869 | 61  | 
val prep_result =  | 
62  | 
Goal.norm_result #> Thm.default_position pos #> not is_open ? Thm.close_derivation;  | 
|
| 27389 | 63  | 
fun after_qed [res] goal_ctxt =  | 
| 
27521
 
44081396d735
@{lemma}: allow terminal method, close derivation unless (open) mode is given;
 
wenzelm 
parents: 
27389 
diff
changeset
 | 
64  | 
put_thms (i, map prep_result (ProofContext.export goal_ctxt ctxt res)) goal_ctxt;  | 
| 27389 | 65  | 
val ctxt' = ctxt  | 
66  | 
|> Proof.theorem_i NONE after_qed [map (rpair []) props]  | 
|
| 
27521
 
44081396d735
@{lemma}: allow terminal method, close derivation unless (open) mode is given;
 
wenzelm 
parents: 
27389 
diff
changeset
 | 
67  | 
|> Proof.global_terminal_proof methods;  | 
| 27389 | 68  | 
val (a, background') = background  | 
69  | 
|> ML_Antiquote.variant "lemma" ||> put_thms (i, the_thms ctxt' i);  | 
|
70  | 
val ml = (thm_bind (if length props = 1 then "thm" else "thms") a i, struct_name ^ "." ^ a);  | 
|
71  | 
in (K ml, background') end));  | 
|
72  | 
||
73  | 
end;  | 
|
74  |