20658
|
1 |
(* Title: Pure/Tools/xml_syntax.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Stefan Berghofer, TU Muenchen
|
|
4 |
|
|
5 |
Input and output of types, terms, and proofs in XML format.
|
|
6 |
See isabelle.xsd for a description of the syntax.
|
|
7 |
*)
|
|
8 |
|
|
9 |
signature XML_SYNTAX =
|
|
10 |
sig
|
|
11 |
val xml_of_type: typ -> XML.tree
|
|
12 |
val xml_of_term: term -> XML.tree
|
|
13 |
val xml_of_proof: Proofterm.proof -> XML.tree
|
|
14 |
val write_to_file: string -> string -> XML.tree -> unit
|
|
15 |
exception XML of string * XML.tree
|
|
16 |
val type_of_xml: XML.tree -> typ
|
|
17 |
val term_of_xml: XML.tree -> term
|
|
18 |
val proof_of_xml: XML.tree -> Proofterm.proof
|
|
19 |
end;
|
|
20 |
|
|
21 |
structure XMLSyntax : XML_SYNTAX =
|
|
22 |
struct
|
|
23 |
|
|
24 |
(**** XML output ****)
|
|
25 |
|
|
26 |
fun xml_of_class name = XML.Elem ("class", [("name", name)], []);
|
|
27 |
|
|
28 |
fun xml_of_type (TVar ((s, i), S)) = XML.Elem ("TVar",
|
|
29 |
("name", s) :: (if i=0 then [] else [("index", string_of_int i)]),
|
|
30 |
map xml_of_class S)
|
|
31 |
| xml_of_type (TFree (s, S)) =
|
|
32 |
XML.Elem ("TFree", [("name", s)], map xml_of_class S)
|
|
33 |
| xml_of_type (Type (s, Ts)) =
|
|
34 |
XML.Elem ("Type", [("name", s)], map xml_of_type Ts);
|
|
35 |
|
|
36 |
fun xml_of_term (Bound i) =
|
|
37 |
XML.Elem ("Bound", [("index", string_of_int i)], [])
|
|
38 |
| xml_of_term (Free (s, T)) =
|
|
39 |
XML.Elem ("Free", [("name", s)], [xml_of_type T])
|
|
40 |
| xml_of_term (Var ((s, i), T)) = XML.Elem ("Var",
|
|
41 |
("name", s) :: (if i=0 then [] else [("index", string_of_int i)]),
|
|
42 |
[xml_of_type T])
|
|
43 |
| xml_of_term (Const (s, T)) =
|
|
44 |
XML.Elem ("Const", [("name", s)], [xml_of_type T])
|
|
45 |
| xml_of_term (t $ u) =
|
|
46 |
XML.Elem ("App", [], [xml_of_term t, xml_of_term u])
|
|
47 |
| xml_of_term (Abs (s, T, t)) =
|
|
48 |
XML.Elem ("Abs", [("vname", s)], [xml_of_type T, xml_of_term t]);
|
|
49 |
|
|
50 |
fun xml_of_opttypes NONE = []
|
|
51 |
| xml_of_opttypes (SOME Ts) = [XML.Elem ("types", [], map xml_of_type Ts)];
|
|
52 |
|
|
53 |
(* FIXME: the t argument of PThm and PAxm is actually redundant, since *)
|
|
54 |
(* it can be looked up in the theorem database. Thus, it could be *)
|
|
55 |
(* omitted from the xml representation. *)
|
|
56 |
|
|
57 |
fun xml_of_proof (PBound i) =
|
|
58 |
XML.Elem ("PBound", [("index", string_of_int i)], [])
|
|
59 |
| xml_of_proof (Abst (s, optT, prf)) =
|
|
60 |
XML.Elem ("Abst", [("vname", s)],
|
|
61 |
(case optT of NONE => [] | SOME T => [xml_of_type T]) @
|
|
62 |
[xml_of_proof prf])
|
|
63 |
| xml_of_proof (AbsP (s, optt, prf)) =
|
|
64 |
XML.Elem ("AbsP", [("vname", s)],
|
|
65 |
(case optt of NONE => [] | SOME t => [xml_of_term t]) @
|
|
66 |
[xml_of_proof prf])
|
|
67 |
| xml_of_proof (prf % optt) =
|
|
68 |
XML.Elem ("Appt", [],
|
|
69 |
xml_of_proof prf ::
|
|
70 |
(case optt of NONE => [] | SOME t => [xml_of_term t]))
|
|
71 |
| xml_of_proof (prf %% prf') =
|
|
72 |
XML.Elem ("AppP", [], [xml_of_proof prf, xml_of_proof prf'])
|
|
73 |
| xml_of_proof (Hyp t) = XML.Elem ("Hyp", [], [xml_of_term t])
|
|
74 |
| xml_of_proof (PThm ((s, _), _, t, optTs)) =
|
|
75 |
XML.Elem ("PThm", [("name", s)], xml_of_term t :: xml_of_opttypes optTs)
|
|
76 |
| xml_of_proof (PAxm (s, t, optTs)) =
|
|
77 |
XML.Elem ("PAxm", [("name", s)], xml_of_term t :: xml_of_opttypes optTs)
|
|
78 |
| xml_of_proof (Oracle (s, t, optTs)) =
|
|
79 |
XML.Elem ("Oracle", [("name", s)], xml_of_term t :: xml_of_opttypes optTs)
|
|
80 |
| xml_of_proof (MinProof prfs) =
|
|
81 |
XML.Elem ("MinProof", [], []);
|
|
82 |
|
|
83 |
(* useful for checking the output against a schema file *)
|
|
84 |
|
|
85 |
fun write_to_file fname elname x =
|
|
86 |
File.write (Path.unpack fname)
|
|
87 |
("<?xml version=\"1.0\" encoding=\"UTF-8\"?>\n" ^
|
|
88 |
XML.string_of_tree (XML.Elem (elname,
|
|
89 |
[("xmlns:xsi", "http://www.w3.org/2001/XMLSchema-instance"),
|
|
90 |
("xsi:noNamespaceSchemaLocation", "isabelle.xsd")],
|
|
91 |
[x])));
|
|
92 |
|
|
93 |
|
|
94 |
(**** XML input ****)
|
|
95 |
|
|
96 |
exception XML of string * XML.tree;
|
|
97 |
|
|
98 |
fun class_of_xml (XML.Elem ("class", [("name", name)], [])) = name
|
|
99 |
| class_of_xml tree = raise XML ("class_of_xml: bad tree", tree);
|
|
100 |
|
|
101 |
fun index_of_string s tree idx = (case Int.fromString idx of
|
|
102 |
NONE => raise XML (s ^ ": bad index", tree) | SOME i => i);
|
|
103 |
|
|
104 |
fun type_of_xml (tree as XML.Elem ("TVar", atts, classes)) = TVar
|
|
105 |
((case AList.lookup op = atts "name" of
|
|
106 |
NONE => raise XML ("type_of_xml: name of TVar missing", tree)
|
|
107 |
| SOME name => name,
|
|
108 |
the_default 0 (Option.map (index_of_string "type_of_xml" tree)
|
|
109 |
(AList.lookup op = atts "index"))),
|
|
110 |
map class_of_xml classes)
|
|
111 |
| type_of_xml (XML.Elem ("TFree", [("name", s)], classes)) =
|
|
112 |
TFree (s, map class_of_xml classes)
|
|
113 |
| type_of_xml (XML.Elem ("Type", [("name", s)], types)) =
|
|
114 |
Type (s, map type_of_xml types)
|
|
115 |
| type_of_xml tree = raise XML ("type_of_xml: bad tree", tree);
|
|
116 |
|
|
117 |
fun term_of_xml (tree as XML.Elem ("Bound", [("index", idx)], [])) =
|
|
118 |
Bound (index_of_string "bad variable index" tree idx)
|
|
119 |
| term_of_xml (XML.Elem ("Free", [("name", s)], [typ])) =
|
|
120 |
Free (s, type_of_xml typ)
|
|
121 |
| term_of_xml (tree as XML.Elem ("Var", atts, [typ])) = Var
|
|
122 |
((case AList.lookup op = atts "name" of
|
|
123 |
NONE => raise XML ("type_of_xml: name of Var missing", tree)
|
|
124 |
| SOME name => name,
|
|
125 |
the_default 0 (Option.map (index_of_string "term_of_xml" tree)
|
|
126 |
(AList.lookup op = atts "index"))),
|
|
127 |
type_of_xml typ)
|
|
128 |
| term_of_xml (XML.Elem ("Const", [("name", s)], [typ])) =
|
|
129 |
Const (s, type_of_xml typ)
|
|
130 |
| term_of_xml (XML.Elem ("App", [], [term, term'])) =
|
|
131 |
term_of_xml term $ term_of_xml term'
|
|
132 |
| term_of_xml (XML.Elem ("Abs", [("vname", s)], [typ, term])) =
|
|
133 |
Abs (s, type_of_xml typ, term_of_xml term)
|
|
134 |
| term_of_xml tree = raise XML ("term_of_xml: bad tree", tree);
|
|
135 |
|
|
136 |
fun opttypes_of_xml [] = NONE
|
|
137 |
| opttypes_of_xml [XML.Elem ("types", [], types)] =
|
|
138 |
SOME (map type_of_xml types)
|
|
139 |
| opttypes_of_xml (tree :: _) = raise XML ("opttypes_of_xml: bad tree", tree);
|
|
140 |
|
|
141 |
fun proof_of_xml (tree as XML.Elem ("PBound", [("index", idx)], [])) =
|
|
142 |
PBound (index_of_string "proof_of_xml" tree idx)
|
|
143 |
| proof_of_xml (XML.Elem ("Abst", [("vname", s)], [proof])) =
|
|
144 |
Abst (s, NONE, proof_of_xml proof)
|
|
145 |
| proof_of_xml (XML.Elem ("Abst", [("vname", s)], [typ, proof])) =
|
|
146 |
Abst (s, SOME (type_of_xml typ), proof_of_xml proof)
|
|
147 |
| proof_of_xml (XML.Elem ("AbsP", [("vname", s)], [proof])) =
|
|
148 |
AbsP (s, NONE, proof_of_xml proof)
|
|
149 |
| proof_of_xml (XML.Elem ("AbsP", [("vname", s)], [term, proof])) =
|
|
150 |
AbsP (s, SOME (term_of_xml term), proof_of_xml proof)
|
|
151 |
| proof_of_xml (XML.Elem ("Appt", [], [proof])) =
|
|
152 |
proof_of_xml proof % NONE
|
|
153 |
| proof_of_xml (XML.Elem ("Appt", [], [proof, term])) =
|
|
154 |
proof_of_xml proof % SOME (term_of_xml term)
|
|
155 |
| proof_of_xml (XML.Elem ("AppP", [], [proof, proof'])) =
|
|
156 |
proof_of_xml proof %% proof_of_xml proof'
|
|
157 |
| proof_of_xml (XML.Elem ("PThm", [("name", s)], term :: opttypes)) =
|
|
158 |
PThm ((s, []), MinProof ([], [], []), (* FIXME? *)
|
|
159 |
term_of_xml term, opttypes_of_xml opttypes)
|
|
160 |
| proof_of_xml (XML.Elem ("PAxm", [("name", s)], term :: opttypes)) =
|
|
161 |
PAxm (s, term_of_xml term, opttypes_of_xml opttypes)
|
|
162 |
| proof_of_xml (XML.Elem ("Oracle", [("name", s)], term :: opttypes)) =
|
|
163 |
Oracle (s, term_of_xml term, opttypes_of_xml opttypes)
|
|
164 |
| proof_of_xml (XML.Elem ("MinProof", _, _)) = MinProof ([], [], [])
|
|
165 |
| proof_of_xml tree = raise XML ("proof_of_xml: bad tree", tree);
|
|
166 |
|
|
167 |
end;
|