| author | eberlm |
| Fri, 26 Feb 2016 14:58:07 +0100 | |
| changeset 62425 | d0936b500bf5 |
| parent 61986 | 2461779da2b8 |
| child 62922 | 96691631c1eb |
| permissions | -rw-r--r-- |
| 52424 | 1 |
(* Title: HOL/Proofs/ex/XML_Data.thy |
2 |
Author: Makarius |
|
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
3 |
Author: Stefan Berghofer |
| 52424 | 4 |
|
5 |
XML data representation of proof terms. |
|
6 |
*) |
|
7 |
||
8 |
theory XML_Data |
|
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
9 |
imports "~~/src/HOL/Isar_Examples/Drinker" |
| 52424 | 10 |
begin |
11 |
||
| 61986 | 12 |
subsection \<open>Export and re-import of global proof terms\<close> |
| 52424 | 13 |
|
| 61986 | 14 |
ML \<open> |
| 60360 | 15 |
fun export_proof thy thm = |
| 52424 | 16 |
let |
| 61039 | 17 |
val (_, prop) = |
18 |
Logic.unconstrainT (Thm.shyps_of thm) |
|
19 |
(Logic.list_implies (Thm.hyps_of thm, Thm.prop_of thm)); |
|
| 52424 | 20 |
val prf = |
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
21 |
Proofterm.proof_of (Proofterm.strip_thm (Thm.proof_body_of thm)) |> |
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
22 |
Reconstruct.reconstruct_proof thy prop |> |
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
23 |
Reconstruct.expand_proof thy [("", NONE)] |>
|
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
24 |
Proofterm.rew_proof thy |> |
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
25 |
Proofterm.no_thm_proofs; |
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
26 |
in Proofterm.encode prf end; |
| 52424 | 27 |
|
28 |
fun import_proof thy xml = |
|
29 |
let |
|
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
30 |
val prf = Proofterm.decode xml; |
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
31 |
val (prf', _) = Proofterm.freeze_thaw_prf prf; |
|
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
32 |
in Drule.export_without_context (Proof_Checker.thm_of_proof thy prf') end; |
| 61986 | 33 |
\<close> |
| 52424 | 34 |
|
35 |
||
| 61986 | 36 |
subsection \<open>Examples\<close> |
| 52424 | 37 |
|
| 61986 | 38 |
ML \<open>val thy1 = @{theory}\<close>
|
| 52424 | 39 |
|
40 |
lemma ex: "A \<longrightarrow> A" .. |
|
41 |
||
| 61986 | 42 |
ML_val \<open> |
| 60360 | 43 |
val xml = export_proof @{theory} @{thm ex};
|
| 52424 | 44 |
val thm = import_proof thy1 xml; |
| 61986 | 45 |
\<close> |
| 52424 | 46 |
|
| 61986 | 47 |
ML_val \<open> |
| 60360 | 48 |
val xml = export_proof @{theory} @{thm de_Morgan};
|
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
49 |
val thm = import_proof thy1 xml; |
| 61986 | 50 |
\<close> |
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
51 |
|
| 61986 | 52 |
ML_val \<open> |
| 60360 | 53 |
val xml = export_proof @{theory} @{thm Drinker's_Principle};
|
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
54 |
val thm = import_proof thy1 xml; |
| 61986 | 55 |
\<close> |
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
56 |
|
| 61986 | 57 |
text \<open>Some fairly large proof:\<close> |
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
58 |
|
| 61986 | 59 |
ML_val \<open> |
| 60360 | 60 |
val xml = export_proof @{theory} @{thm abs_less_iff};
|
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
61 |
val thm = import_proof thy1 xml; |
|
60157
ccf9241af217
use smaller example that fits into 64MB string limit of Poly/ML x86 platform;
wenzelm
parents:
58372
diff
changeset
|
62 |
@{assert} (size (YXML.string_of_body xml) > 1000000);
|
| 61986 | 63 |
\<close> |
|
52630
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
wenzelm
parents:
52424
diff
changeset
|
64 |
|
| 52424 | 65 |
end |