author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Wed, 04 Nov 2009 15:27:32 +0100 | |
changeset 283 | 5470176d6730 |
parent 183 | 6acf9e001038 |
child 319 | 0ae9d9e66cb7 |
permissions | -rw-r--r-- |
6 | 1 |
theory Prove |
83 | 2 |
imports Main |
6 | 3 |
begin |
4 |
||
5 |
ML {* |
|
183
6acf9e001038
proved the two lemmas in QuotScript (reformulated them without leading forall)
Christian Urban <urbanc@in.tum.de>
parents:
83
diff
changeset
|
6 |
val r = Unsynchronized.ref (NONE:(unit -> term) option) |
6 | 7 |
*} |
46 | 8 |
|
6 | 9 |
ML {* |
10 |
let |
|
48 | 11 |
fun after_qed thm_name thms lthy = |
12 |
LocalTheory.note Thm.theoremK (thm_name, (flat thms)) lthy |> snd |
|
13 |
||
14 |
fun setup_proof (name_spec, (txt, pos)) lthy = |
|
15 |
let |
|
16 |
val trm = ML_Context.evaluate lthy true ("r", r) txt |
|
17 |
in |
|
18 |
Proof.theorem_i NONE (after_qed name_spec) [[(trm,[])]] lthy |
|
57 | 19 |
end |
6 | 20 |
|
48 | 21 |
val parser = SpecParse.opt_thm_name ":" -- OuterParse.ML_source |
6 | 22 |
in |
46 | 23 |
OuterSyntax.local_theory_to_proof "prove" "proving a proposition" |
48 | 24 |
OuterKeyword.thy_goal (parser >> setup_proof) |
6 | 25 |
end |
26 |
*} |
|
27 |
||
28 |
end |