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