src/HOL/Proofs/ex/XML_Data.thy
author wenzelm
Sun, 23 Jun 2013 16:47:45 +0200
changeset 52424 77075c576d4c
child 52630 fe411c1dc180
permissions -rw-r--r--
support for XML data representation of proof terms;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52424
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Proofs/ex/XML_Data.thy
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     3
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     4
XML data representation of proof terms.
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     5
*)
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     6
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     7
theory XML_Data
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     8
imports Main
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
     9
begin
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    10
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    11
subsection {* Export and re-import of global proof terms *}
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    12
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    13
ML {*
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    14
  fun export_proof thm0 =
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    15
    let
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    16
      val thy = Thm.theory_of_thm thm0;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    17
      val ctxt0 = Proof_Context.init_global thy;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    18
      val thy = Proof_Context.theory_of ctxt0;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    19
      val ((_, [thm]), ctxt) = Variable.import true [Thm.transfer thy thm0] ctxt0;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    20
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    21
      val prop = Thm.prop_of thm;  (* FIXME proper prop (wrt. import / strip) (!??) *)
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    22
      val prf =
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    23
        Proofterm.proof_of (Proofterm.strip_thm (Thm.proof_body_of thm))
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    24
        |> Proofterm.no_thm_proofs;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    25
    in XML.Encode.pair Term_XML.Encode.term Proofterm.encode (prop, prf) end;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    26
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    27
  fun import_proof thy xml =
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    28
    let
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    29
      val (prop, prf) = XML.Decode.pair Term_XML.Decode.term Proofterm.decode xml;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    30
      val full_prf = Reconstruct.reconstruct_proof thy prop prf;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    31
    in Drule.export_without_context (Proof_Checker.thm_of_proof thy full_prf) end;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    32
*}
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    33
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    34
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    35
subsection {* Example *}
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    36
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    37
ML {* val thy1 = @{theory} *}
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    38
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    39
lemma ex: "A \<longrightarrow> A" ..
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    40
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    41
ML {*
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    42
  val xml = export_proof @{thm ex};
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    43
  val thm = import_proof thy1 xml;
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    44
*}
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    45
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    46
end
77075c576d4c support for XML data representation of proof terms;
wenzelm
parents:
diff changeset
    47