| author | wenzelm | 
| Mon, 29 Jul 2013 15:01:44 +0200 | |
| changeset 52769 | 0827b6f5de44 | 
| parent 52630 | fe411c1dc180 | 
| child 58372 | bfd497f2f4c2 | 
| 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: 
52424diff
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: 
52424diff
changeset | 9 | imports "~~/src/HOL/Isar_Examples/Drinker" | 
| 52424 | 10 | begin | 
| 11 | ||
| 12 | subsection {* Export and re-import of global proof terms *}
 | |
| 13 | ||
| 14 | ML {*
 | |
| 52630 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 15 | fun export_proof thm = | 
| 52424 | 16 | let | 
| 52630 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 17 | val thy = Thm.theory_of_thm thm; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 18 |       val {prop, hyps, shyps, ...} = Thm.rep_thm thm;
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 19 | val (_, prop) = Logic.unconstrainT shyps (Logic.list_implies (hyps, prop)); | 
| 52424 | 20 | val prf = | 
| 52630 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
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: 
52424diff
changeset | 22 | Reconstruct.reconstruct_proof thy prop |> | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 23 |         Reconstruct.expand_proof thy [("", NONE)] |>
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 24 | Proofterm.rew_proof thy |> | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 25 | Proofterm.no_thm_proofs; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
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: 
52424diff
changeset | 30 | val prf = Proofterm.decode xml; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 31 | val (prf', _) = Proofterm.freeze_thaw_prf prf; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 32 | in Drule.export_without_context (Proof_Checker.thm_of_proof thy prf') end; | 
| 52424 | 33 | *} | 
| 34 | ||
| 35 | ||
| 52630 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 36 | subsection {* Examples *}
 | 
| 52424 | 37 | |
| 38 | ML {* val thy1 = @{theory} *}
 | |
| 39 | ||
| 40 | lemma ex: "A \<longrightarrow> A" .. | |
| 41 | ||
| 52630 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 42 | ML_val {*
 | 
| 52424 | 43 |   val xml = export_proof @{thm ex};
 | 
| 44 | val thm = import_proof thy1 xml; | |
| 45 | *} | |
| 46 | ||
| 52630 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 47 | ML_val {*
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 48 |   val xml = export_proof @{thm de_Morgan};
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 49 | val thm = import_proof thy1 xml; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 50 | *} | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 51 | |
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 52 | ML_val {*
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 53 |   val xml = export_proof @{thm Drinker's_Principle};
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 54 | val thm = import_proof thy1 xml; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 55 | *} | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 56 | |
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 57 | text {* Some fairly large proof: *}
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 58 | |
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 59 | ML_val {*
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 60 |   val xml = export_proof @{thm Int.times_int.abs_eq};
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 61 | val thm = import_proof thy1 xml; | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 62 |   @{assert} (size (YXML.string_of_body xml) > 50000000);
 | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 63 | *} | 
| 
fe411c1dc180
more robust proof export / import due to Stefan Berghofer;
 wenzelm parents: 
52424diff
changeset | 64 | |
| 52424 | 65 | end | 
| 66 |