src/Pure/Tools/legacy_xml_syntax.ML
author wenzelm
Mon, 26 Nov 2012 16:28:22 +0100
changeset 50217 ce1f0602f48e
parent 38228 src/Pure/Tools/xml_syntax.ML@ada3ab6b9085
permissions -rw-r--r--
clarified status of Legacy_XML_Syntax, despite lack of Proofterm_XML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50217
ce1f0602f48e clarified status of Legacy_XML_Syntax, despite lack of Proofterm_XML;
wenzelm
parents: 38228
diff changeset
     1
(*  Title:      Pure/Tools/legacy_xml_syntax.ML
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
     3
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
     4
Input and output of types, terms, and proofs in XML format.
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
     5
See isabelle.xsd for a description of the syntax.
50217
ce1f0602f48e clarified status of Legacy_XML_Syntax, despite lack of Proofterm_XML;
wenzelm
parents: 38228
diff changeset
     6
ce1f0602f48e clarified status of Legacy_XML_Syntax, despite lack of Proofterm_XML;
wenzelm
parents: 38228
diff changeset
     7
Legacy module, see Pure/term_xml.ML etc.
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
     8
*)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
     9
50217
ce1f0602f48e clarified status of Legacy_XML_Syntax, despite lack of Proofterm_XML;
wenzelm
parents: 38228
diff changeset
    10
signature LEGACY_XML_SYNTAX =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    11
sig
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    12
  val xml_of_type: typ -> XML.tree
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    13
  val xml_of_term: term -> XML.tree
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    14
  val xml_of_proof: Proofterm.proof -> XML.tree
21645
4af699cdfe47 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20658
diff changeset
    15
  val write_to_file: Path.T -> string -> XML.tree -> unit
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    16
  exception XML of string * XML.tree
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    17
  val type_of_xml: XML.tree -> typ
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    18
  val term_of_xml: XML.tree -> term
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    19
  val proof_of_xml: XML.tree -> Proofterm.proof
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    20
end;
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    21
50217
ce1f0602f48e clarified status of Legacy_XML_Syntax, despite lack of Proofterm_XML;
wenzelm
parents: 38228
diff changeset
    22
structure Legacy_XML_Syntax : LEGACY_XML_SYNTAX =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    23
struct
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    24
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    25
(**** XML output ****)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    26
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    27
fun xml_of_class name = XML.Elem (("class", [("name", name)]), []);
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    28
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    29
fun xml_of_type (TVar ((s, i), S)) =
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    30
      XML.Elem (("TVar", ("name", s) :: (if i=0 then [] else [("index", string_of_int i)])),
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    31
        map xml_of_class S)
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    32
  | xml_of_type (TFree (s, S)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    33
      XML.Elem (("TFree", [("name", s)]), map xml_of_class S)
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    34
  | xml_of_type (Type (s, Ts)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    35
      XML.Elem (("Type", [("name", s)]), map xml_of_type Ts);
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    36
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    37
fun xml_of_term (Bound i) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    38
      XML.Elem (("Bound", [("index", string_of_int i)]), [])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    39
  | xml_of_term (Free (s, T)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    40
      XML.Elem (("Free", [("name", s)]), [xml_of_type T])
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    41
  | xml_of_term (Var ((s, i), T)) =
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    42
      XML.Elem (("Var", ("name", s) :: (if i=0 then [] else [("index", string_of_int i)])),
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    43
        [xml_of_type T])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    44
  | xml_of_term (Const (s, T)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    45
      XML.Elem (("Const", [("name", s)]), [xml_of_type T])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    46
  | xml_of_term (t $ u) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    47
      XML.Elem (("App", []), [xml_of_term t, xml_of_term u])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    48
  | xml_of_term (Abs (s, T, t)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    49
      XML.Elem (("Abs", [("vname", s)]), [xml_of_type T, xml_of_term t]);
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    50
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    51
fun xml_of_opttypes NONE = []
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    52
  | xml_of_opttypes (SOME Ts) = [XML.Elem (("types", []), map xml_of_type Ts)];
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    53
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    54
(* FIXME: the t argument of PThm and PAxm is actually redundant, since *)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    55
(* it can be looked up in the theorem database. Thus, it could be      *)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    56
(* omitted from the xml representation.                                *)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    57
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    58
(* FIXME not exhaustive *)
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    59
fun xml_of_proof (PBound i) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    60
      XML.Elem (("PBound", [("index", string_of_int i)]), [])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    61
  | xml_of_proof (Abst (s, optT, prf)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    62
      XML.Elem (("Abst", [("vname", s)]),
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    63
        (case optT of NONE => [] | SOME T => [xml_of_type T]) @ [xml_of_proof prf])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    64
  | xml_of_proof (AbsP (s, optt, prf)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    65
      XML.Elem (("AbsP", [("vname", s)]),
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    66
        (case optt of NONE => [] | SOME t => [xml_of_term t]) @ [xml_of_proof prf])
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    67
  | xml_of_proof (prf % optt) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    68
      XML.Elem (("Appt", []),
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    69
        xml_of_proof prf :: (case optt of NONE => [] | SOME t => [xml_of_term t]))
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    70
  | xml_of_proof (prf %% prf') =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    71
      XML.Elem (("AppP", []), [xml_of_proof prf, xml_of_proof prf'])
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    72
  | xml_of_proof (Hyp t) = XML.Elem (("Hyp", []), [xml_of_term t])
28811
aa36d05926ec adapted PThm and MinProof;
wenzelm
parents: 28017
diff changeset
    73
  | xml_of_proof (PThm (_, ((s, t, optTs), _))) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    74
      XML.Elem (("PThm", [("name", s)]), xml_of_term t :: xml_of_opttypes optTs)
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    75
  | xml_of_proof (PAxm (s, t, optTs)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    76
      XML.Elem (("PAxm", [("name", s)]), xml_of_term t :: xml_of_opttypes optTs)
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    77
  | xml_of_proof (Oracle (s, t, optTs)) =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    78
      XML.Elem (("Oracle", [("name", s)]), xml_of_term t :: xml_of_opttypes optTs)
28811
aa36d05926ec adapted PThm and MinProof;
wenzelm
parents: 28017
diff changeset
    79
  | xml_of_proof MinProof =
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    80
      XML.Elem (("MinProof", []), []);
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    81
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    82
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    83
(* useful for checking the output against a schema file *)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    84
21645
4af699cdfe47 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20658
diff changeset
    85
fun write_to_file path elname x =
4af699cdfe47 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 20658
diff changeset
    86
  File.write path
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    87
    ("<?xml version=\"1.0\" encoding=\"UTF-8\"?>\n" ^
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    88
     XML.string_of (XML.Elem ((elname,
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    89
         [("xmlns:xsi", "http://www.w3.org/2001/XMLSchema-instance"),
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    90
          ("xsi:noNamespaceSchemaLocation", "isabelle.xsd")]),
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    91
       [x])));
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    92
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    93
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    94
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    95
(**** XML input ****)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    96
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    97
exception XML of string * XML.tree;
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
    98
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
    99
fun class_of_xml (XML.Elem (("class", [("name", name)]), [])) = name
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   100
  | class_of_xml tree = raise XML ("class_of_xml: bad tree", tree);
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   101
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   102
fun index_of_string s tree idx =
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   103
  (case Int.fromString idx of
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   104
    NONE => raise XML (s ^ ": bad index", tree)
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   105
  | SOME i => i);
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   106
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   107
fun type_of_xml (tree as XML.Elem (("TVar", atts), classes)) = TVar
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 26544
diff changeset
   108
      ((case Properties.get atts "name" of
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   109
          NONE => raise XML ("type_of_xml: name of TVar missing", tree)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   110
        | SOME name => name,
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   111
        the_default 0 (Option.map (index_of_string "type_of_xml" tree)
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 26544
diff changeset
   112
          (Properties.get atts "index"))),
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   113
       map class_of_xml classes)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   114
  | type_of_xml (XML.Elem (("TFree", [("name", s)]), classes)) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   115
      TFree (s, map class_of_xml classes)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   116
  | type_of_xml (XML.Elem (("Type", [("name", s)]), types)) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   117
      Type (s, map type_of_xml types)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   118
  | type_of_xml tree = raise XML ("type_of_xml: bad tree", tree);
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   119
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   120
fun term_of_xml (tree as XML.Elem (("Bound", [("index", idx)]), [])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   121
      Bound (index_of_string "bad variable index" tree idx)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   122
  | term_of_xml (XML.Elem (("Free", [("name", s)]), [typ])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   123
      Free (s, type_of_xml typ)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   124
  | term_of_xml (tree as XML.Elem (("Var", atts), [typ])) = Var
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 26544
diff changeset
   125
      ((case Properties.get atts "name" of
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   126
          NONE => raise XML ("type_of_xml: name of Var missing", tree)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   127
        | SOME name => name,
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   128
        the_default 0 (Option.map (index_of_string "term_of_xml" tree)
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 26544
diff changeset
   129
          (Properties.get atts "index"))),
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   130
       type_of_xml typ)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   131
  | term_of_xml (XML.Elem (("Const", [("name", s)]), [typ])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   132
      Const (s, type_of_xml typ)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   133
  | term_of_xml (XML.Elem (("App", []), [term, term'])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   134
      term_of_xml term $ term_of_xml term'
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   135
  | term_of_xml (XML.Elem (("Abs", [("vname", s)]), [typ, term])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   136
      Abs (s, type_of_xml typ, term_of_xml term)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   137
  | term_of_xml tree = raise XML ("term_of_xml: bad tree", tree);
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   138
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   139
fun opttypes_of_xml [] = NONE
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   140
  | opttypes_of_xml [XML.Elem (("types", []), types)] =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   141
      SOME (map type_of_xml types)
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   142
  | opttypes_of_xml (tree :: _) = raise XML ("opttypes_of_xml: bad tree", tree);
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   143
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   144
fun proof_of_xml (tree as XML.Elem (("PBound", [("index", idx)]), [])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   145
      PBound (index_of_string "proof_of_xml" tree idx)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   146
  | proof_of_xml (XML.Elem (("Abst", [("vname", s)]), [proof])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   147
      Abst (s, NONE, proof_of_xml proof)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   148
  | proof_of_xml (XML.Elem (("Abst", [("vname", s)]), [typ, proof])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   149
      Abst (s, SOME (type_of_xml typ), proof_of_xml proof)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   150
  | proof_of_xml (XML.Elem (("AbsP", [("vname", s)]), [proof])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   151
      AbsP (s, NONE, proof_of_xml proof)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   152
  | proof_of_xml (XML.Elem (("AbsP", [("vname", s)]), [term, proof])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   153
      AbsP (s, SOME (term_of_xml term), proof_of_xml proof)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   154
  | proof_of_xml (XML.Elem (("Appt", []), [proof])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   155
      proof_of_xml proof % NONE
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   156
  | proof_of_xml (XML.Elem (("Appt", []), [proof, term])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   157
      proof_of_xml proof % SOME (term_of_xml term)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   158
  | proof_of_xml (XML.Elem (("AppP", []), [proof, proof'])) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   159
      proof_of_xml proof %% proof_of_xml proof'
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   160
  | proof_of_xml (XML.Elem (("Hyp", []), [term])) =
23831
64e6e5c738a1 Added clause for hypotheses to proof_of_xml function.
berghofe
parents: 21645
diff changeset
   161
      Hyp (term_of_xml term)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   162
  | proof_of_xml (XML.Elem (("PThm", [("name", s)]), term :: opttypes)) =
28811
aa36d05926ec adapted PThm and MinProof;
wenzelm
parents: 28017
diff changeset
   163
      (* FIXME? *)
aa36d05926ec adapted PThm and MinProof;
wenzelm
parents: 28017
diff changeset
   164
      PThm (serial (), ((s, term_of_xml term, opttypes_of_xml opttypes),
30718
15041c7e51e4 Proofterm.approximate_proof_body;
wenzelm
parents: 29635
diff changeset
   165
        Future.value (Proofterm.approximate_proof_body MinProof)))
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   166
  | proof_of_xml (XML.Elem (("PAxm", [("name", s)]), term :: opttypes)) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   167
      PAxm (s, term_of_xml term, opttypes_of_xml opttypes)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   168
  | proof_of_xml (XML.Elem (("Oracle", [("name", s)]), term :: opttypes)) =
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   169
      Oracle (s, term_of_xml term, opttypes_of_xml opttypes)
38228
ada3ab6b9085 simplified type XML.tree: embed Markup.T directly, avoid slightly odd triple;
wenzelm
parents: 33392
diff changeset
   170
  | proof_of_xml (XML.Elem (("MinProof", _), _)) = MinProof
20658
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   171
  | proof_of_xml tree = raise XML ("proof_of_xml: bad tree", tree);
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   172
2586df9fb95a XML syntax for types, terms, and proofs.
berghofe
parents:
diff changeset
   173
end;