68171
|
1 |
/* Title: Pure/Thy/export_theory.scala
|
|
2 |
Author: Makarius
|
|
3 |
|
|
4 |
Export foundational theory content.
|
|
5 |
*/
|
|
6 |
|
|
7 |
package isabelle
|
|
8 |
|
|
9 |
|
|
10 |
object Export_Theory
|
|
11 |
{
|
68206
|
12 |
/** session content **/
|
|
13 |
|
|
14 |
sealed case class Session(name: String, theory_graph: Graph[String, Theory])
|
|
15 |
{
|
|
16 |
override def toString: String = name
|
|
17 |
|
|
18 |
def theory(theory_name: String): Theory =
|
|
19 |
if (theory_graph.defined(theory_name)) theory_graph.get_node(theory_name)
|
|
20 |
else error("Bad theory " + quote(theory_name))
|
|
21 |
|
|
22 |
def theories: List[Theory] =
|
|
23 |
theory_graph.topological_order.map(theory_graph.get_node(_))
|
|
24 |
}
|
|
25 |
|
68209
|
26 |
def read_session(store: Sessions.Store,
|
|
27 |
session_name: String,
|
68206
|
28 |
types: Boolean = true,
|
68264
|
29 |
consts: Boolean = true,
|
|
30 |
axioms: Boolean = true,
|
|
31 |
facts: Boolean = true,
|
|
32 |
classes: Boolean = true,
|
68267
|
33 |
typedefs: Boolean = true,
|
68295
|
34 |
classrel: Boolean = true,
|
|
35 |
arities: Boolean = true,
|
68267
|
36 |
cache: Term.Cache = Term.make_cache()): Session =
|
68206
|
37 |
{
|
|
38 |
val thys =
|
68210
|
39 |
using(store.open_database(session_name))(db =>
|
68206
|
40 |
{
|
|
41 |
db.transaction {
|
68222
|
42 |
Export.read_theory_names(db, session_name).map((theory_name: String) =>
|
68418
|
43 |
read_theory(Export.Provider.database(db, session_name, theory_name),
|
|
44 |
session_name, theory_name, types = types, consts = consts,
|
68267
|
45 |
axioms = axioms, facts = facts, classes = classes, typedefs = typedefs,
|
|
46 |
cache = Some(cache)))
|
68206
|
47 |
}
|
|
48 |
})
|
|
49 |
|
|
50 |
val graph0 =
|
|
51 |
(Graph.string[Theory] /: thys) { case (g, thy) => g.new_node(thy.name, thy) }
|
|
52 |
val graph1 =
|
|
53 |
(graph0 /: thys) { case (g0, thy) =>
|
|
54 |
(g0 /: thy.parents) { case (g1, parent) =>
|
|
55 |
g1.default_node(parent, empty_theory(parent)).add_edge_acyclic(parent, thy.name) } }
|
|
56 |
|
|
57 |
Session(session_name, graph1)
|
|
58 |
}
|
|
59 |
|
|
60 |
|
|
61 |
|
68203
|
62 |
/** theory content **/
|
|
63 |
|
68346
|
64 |
val export_prefix: String = "theory/"
|
|
65 |
|
68208
|
66 |
sealed case class Theory(name: String, parents: List[String],
|
|
67 |
types: List[Type],
|
|
68 |
consts: List[Const],
|
68726
|
69 |
axioms: List[Fact_Single],
|
|
70 |
facts: List[Fact_Multi],
|
68264
|
71 |
classes: List[Class],
|
68295
|
72 |
typedefs: List[Typedef],
|
|
73 |
classrel: List[Classrel],
|
|
74 |
arities: List[Arity])
|
68206
|
75 |
{
|
|
76 |
override def toString: String = name
|
68267
|
77 |
|
68711
|
78 |
lazy val entities: Set[Long] =
|
|
79 |
Set.empty[Long] ++
|
|
80 |
types.iterator.map(_.entity.serial) ++
|
|
81 |
consts.iterator.map(_.entity.serial) ++
|
|
82 |
axioms.iterator.map(_.entity.serial) ++
|
|
83 |
facts.iterator.map(_.entity.serial) ++
|
|
84 |
classes.iterator.map(_.entity.serial)
|
|
85 |
|
68267
|
86 |
def cache(cache: Term.Cache): Theory =
|
|
87 |
Theory(cache.string(name),
|
|
88 |
parents.map(cache.string(_)),
|
|
89 |
types.map(_.cache(cache)),
|
|
90 |
consts.map(_.cache(cache)),
|
|
91 |
axioms.map(_.cache(cache)),
|
|
92 |
facts.map(_.cache(cache)),
|
|
93 |
classes.map(_.cache(cache)),
|
68295
|
94 |
typedefs.map(_.cache(cache)),
|
|
95 |
classrel.map(_.cache(cache)),
|
|
96 |
arities.map(_.cache(cache)))
|
68206
|
97 |
}
|
|
98 |
|
68295
|
99 |
def empty_theory(name: String): Theory =
|
|
100 |
Theory(name, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil, Nil)
|
68203
|
101 |
|
68418
|
102 |
def read_theory(provider: Export.Provider, session_name: String, theory_name: String,
|
68203
|
103 |
types: Boolean = true,
|
68208
|
104 |
consts: Boolean = true,
|
68232
|
105 |
axioms: Boolean = true,
|
68264
|
106 |
facts: Boolean = true,
|
|
107 |
classes: Boolean = true,
|
68267
|
108 |
typedefs: Boolean = true,
|
68295
|
109 |
classrel: Boolean = true,
|
|
110 |
arities: Boolean = true,
|
68267
|
111 |
cache: Option[Term.Cache] = None): Theory =
|
68203
|
112 |
{
|
68206
|
113 |
val parents =
|
68418
|
114 |
provider(export_prefix + "parents") match {
|
68231
|
115 |
case Some(entry) => split_lines(entry.uncompressed().text)
|
68206
|
116 |
case None =>
|
|
117 |
error("Missing theory export in session " + quote(session_name) + ": " +
|
|
118 |
quote(theory_name))
|
|
119 |
}
|
68267
|
120 |
val theory =
|
|
121 |
Theory(theory_name, parents,
|
68418
|
122 |
if (types) read_types(provider) else Nil,
|
|
123 |
if (consts) read_consts(provider) else Nil,
|
|
124 |
if (axioms) read_axioms(provider) else Nil,
|
|
125 |
if (facts) read_facts(provider) else Nil,
|
|
126 |
if (classes) read_classes(provider) else Nil,
|
|
127 |
if (typedefs) read_typedefs(provider) else Nil,
|
|
128 |
if (classrel) read_classrel(provider) else Nil,
|
|
129 |
if (arities) read_arities(provider) else Nil)
|
68267
|
130 |
if (cache.isDefined) theory.cache(cache.get) else theory
|
68203
|
131 |
}
|
|
132 |
|
68711
|
133 |
def read_pure_theory(store: Sessions.Store, cache: Option[Term.Cache] = None): Theory =
|
68710
|
134 |
{
|
|
135 |
val session_name = Thy_Header.PURE
|
|
136 |
val theory_name = Thy_Header.PURE
|
|
137 |
|
|
138 |
using(store.open_database(session_name))(db =>
|
|
139 |
{
|
|
140 |
db.transaction {
|
|
141 |
read_theory(Export.Provider.database(db, session_name, theory_name),
|
68711
|
142 |
session_name, theory_name, cache = cache)
|
68710
|
143 |
}
|
|
144 |
})
|
|
145 |
}
|
|
146 |
|
68203
|
147 |
|
68171
|
148 |
/* entities */
|
|
149 |
|
68714
|
150 |
object Kind extends Enumeration
|
68171
|
151 |
{
|
68714
|
152 |
val TYPE = Value("type")
|
|
153 |
val CONST = Value("const")
|
|
154 |
val AXIOM = Value("axiom")
|
|
155 |
val FACT = Value("fact")
|
|
156 |
val CLASS = Value("class")
|
|
157 |
}
|
|
158 |
|
|
159 |
sealed case class Entity(kind: Kind.Value, name: String, serial: Long, pos: Position.T)
|
|
160 |
{
|
68718
|
161 |
override def toString: String = kind.toString + " " + quote(name)
|
68267
|
162 |
|
|
163 |
def cache(cache: Term.Cache): Entity =
|
68714
|
164 |
Entity(kind, cache.string(name), serial, cache.position(pos))
|
68171
|
165 |
}
|
|
166 |
|
68714
|
167 |
def decode_entity(kind: Kind.Value, tree: XML.Tree): (Entity, XML.Body) =
|
68171
|
168 |
{
|
|
169 |
def err(): Nothing = throw new XML.XML_Body(List(tree))
|
|
170 |
|
|
171 |
tree match {
|
|
172 |
case XML.Elem(Markup(Markup.ENTITY, props), body) =>
|
|
173 |
val name = Markup.Name.unapply(props) getOrElse err()
|
|
174 |
val serial = Markup.Serial.unapply(props) getOrElse err()
|
|
175 |
val pos = props.filter({ case (a, _) => Markup.POSITION_PROPERTIES(a) })
|
68714
|
176 |
(Entity(kind, name, serial, pos), body)
|
68171
|
177 |
case _ => err()
|
|
178 |
}
|
|
179 |
}
|
|
180 |
|
|
181 |
|
|
182 |
/* types */
|
|
183 |
|
|
184 |
sealed case class Type(entity: Entity, args: List[String], abbrev: Option[Term.Typ])
|
68267
|
185 |
{
|
|
186 |
def cache(cache: Term.Cache): Type =
|
|
187 |
Type(entity.cache(cache),
|
|
188 |
args.map(cache.string(_)),
|
|
189 |
abbrev.map(cache.typ(_)))
|
|
190 |
}
|
68171
|
191 |
|
68418
|
192 |
def read_types(provider: Export.Provider): List[Type] =
|
|
193 |
provider.uncompressed_yxml(export_prefix + "types").map((tree: XML.Tree) =>
|
|
194 |
{
|
68714
|
195 |
val (entity, body) = decode_entity(Kind.TYPE, tree)
|
68418
|
196 |
val (args, abbrev) =
|
68203
|
197 |
{
|
68418
|
198 |
import XML.Decode._
|
|
199 |
pair(list(string), option(Term_XML.Decode.typ))(body)
|
|
200 |
}
|
|
201 |
Type(entity, args, abbrev)
|
|
202 |
})
|
68171
|
203 |
|
|
204 |
|
|
205 |
/* consts */
|
|
206 |
|
68173
|
207 |
sealed case class Const(
|
|
208 |
entity: Entity, typargs: List[String], typ: Term.Typ, abbrev: Option[Term.Term])
|
68267
|
209 |
{
|
|
210 |
def cache(cache: Term.Cache): Const =
|
|
211 |
Const(entity.cache(cache),
|
|
212 |
typargs.map(cache.string(_)),
|
|
213 |
cache.typ(typ),
|
|
214 |
abbrev.map(cache.term(_)))
|
|
215 |
}
|
68171
|
216 |
|
68418
|
217 |
def read_consts(provider: Export.Provider): List[Const] =
|
|
218 |
provider.uncompressed_yxml(export_prefix + "consts").map((tree: XML.Tree) =>
|
|
219 |
{
|
68714
|
220 |
val (entity, body) = decode_entity(Kind.CONST, tree)
|
68418
|
221 |
val (args, typ, abbrev) =
|
68203
|
222 |
{
|
68418
|
223 |
import XML.Decode._
|
|
224 |
triple(list(string), Term_XML.Decode.typ, option(Term_XML.Decode.term))(body)
|
|
225 |
}
|
|
226 |
Const(entity, args, typ, abbrev)
|
|
227 |
})
|
68208
|
228 |
|
|
229 |
|
68726
|
230 |
/* facts */
|
68232
|
231 |
|
68726
|
232 |
sealed case class Prop(
|
|
233 |
typargs: List[(String, Term.Sort)],
|
|
234 |
args: List[(String, Term.Typ)],
|
|
235 |
term: Term.Term)
|
68232
|
236 |
{
|
68726
|
237 |
def cache(cache: Term.Cache): Prop =
|
|
238 |
Prop(
|
|
239 |
typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
|
|
240 |
args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }),
|
|
241 |
cache.term(term))
|
68232
|
242 |
}
|
68208
|
243 |
|
68726
|
244 |
def decode_prop(body: XML.Body): Prop =
|
68267
|
245 |
{
|
68726
|
246 |
val (typargs, args, t) =
|
|
247 |
{
|
|
248 |
import XML.Decode._
|
|
249 |
import Term_XML.Decode._
|
|
250 |
triple(list(pair(string, sort)), list(pair(string, typ)), term)(body)
|
|
251 |
}
|
|
252 |
Prop(typargs, args, t)
|
68267
|
253 |
}
|
68208
|
254 |
|
68726
|
255 |
sealed case class Fact_Single(entity: Entity, prop: Prop)
|
|
256 |
{
|
|
257 |
def cache(cache: Term.Cache): Fact_Single =
|
|
258 |
Fact_Single(entity.cache(cache), prop.cache(cache))
|
|
259 |
}
|
|
260 |
|
|
261 |
sealed case class Fact_Multi(entity: Entity, props: List[Prop])
|
|
262 |
{
|
|
263 |
def cache(cache: Term.Cache): Fact_Multi =
|
|
264 |
Fact_Multi(entity.cache(cache), props.map(_.cache(cache)))
|
|
265 |
|
|
266 |
def split: List[Fact_Single] =
|
|
267 |
props match {
|
|
268 |
case List(prop) => List(Fact_Single(entity, prop))
|
|
269 |
case _ =>
|
|
270 |
for ((prop, i) <- props.zipWithIndex)
|
|
271 |
yield Fact_Single(entity.copy(name = entity.name + "(" + (i + 1) + ")"), prop)
|
|
272 |
}
|
|
273 |
}
|
|
274 |
|
|
275 |
def read_axioms(provider: Export.Provider): List[Fact_Single] =
|
68418
|
276 |
provider.uncompressed_yxml(export_prefix + "axioms").map((tree: XML.Tree) =>
|
|
277 |
{
|
68714
|
278 |
val (entity, body) = decode_entity(Kind.AXIOM, tree)
|
68726
|
279 |
val prop = decode_prop(body)
|
|
280 |
Fact_Single(entity, prop)
|
68418
|
281 |
})
|
68232
|
282 |
|
68726
|
283 |
def read_facts(provider: Export.Provider): List[Fact_Multi] =
|
68418
|
284 |
provider.uncompressed_yxml(export_prefix + "facts").map((tree: XML.Tree) =>
|
|
285 |
{
|
68714
|
286 |
val (entity, body) = decode_entity(Kind.FACT, tree)
|
68726
|
287 |
val props = XML.Decode.list(decode_prop)(body)
|
|
288 |
Fact_Multi(entity, props)
|
68418
|
289 |
})
|
68264
|
290 |
|
|
291 |
|
|
292 |
/* type classes */
|
|
293 |
|
|
294 |
sealed case class Class(
|
|
295 |
entity: Entity, params: List[(String, Term.Typ)], axioms: List[Term.Term])
|
68267
|
296 |
{
|
|
297 |
def cache(cache: Term.Cache): Class =
|
|
298 |
Class(entity.cache(cache),
|
|
299 |
params.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }),
|
|
300 |
axioms.map(cache.term(_)))
|
|
301 |
}
|
68264
|
302 |
|
68418
|
303 |
def read_classes(provider: Export.Provider): List[Class] =
|
|
304 |
provider.uncompressed_yxml(export_prefix + "classes").map((tree: XML.Tree) =>
|
|
305 |
{
|
68714
|
306 |
val (entity, body) = decode_entity(Kind.CLASS, tree)
|
68418
|
307 |
val (params, axioms) =
|
68264
|
308 |
{
|
68418
|
309 |
import XML.Decode._
|
|
310 |
import Term_XML.Decode._
|
|
311 |
pair(list(pair(string, typ)), list(term))(body)
|
|
312 |
}
|
|
313 |
Class(entity, params, axioms)
|
|
314 |
})
|
68264
|
315 |
|
|
316 |
|
68295
|
317 |
/* sort algebra */
|
|
318 |
|
|
319 |
sealed case class Classrel(class_name: String, super_names: List[String])
|
|
320 |
{
|
|
321 |
def cache(cache: Term.Cache): Classrel =
|
|
322 |
Classrel(cache.string(class_name), super_names.map(cache.string(_)))
|
|
323 |
}
|
|
324 |
|
68418
|
325 |
def read_classrel(provider: Export.Provider): List[Classrel] =
|
|
326 |
{
|
|
327 |
val body = provider.uncompressed_yxml(export_prefix + "classrel")
|
|
328 |
val classrel =
|
|
329 |
{
|
|
330 |
import XML.Decode._
|
|
331 |
list(pair(string, list(string)))(body)
|
|
332 |
}
|
|
333 |
for ((c, cs) <- classrel) yield Classrel(c, cs)
|
|
334 |
}
|
68295
|
335 |
|
|
336 |
sealed case class Arity(type_name: String, domain: List[Term.Sort], codomain: String)
|
|
337 |
{
|
|
338 |
def cache(cache: Term.Cache): Arity =
|
|
339 |
Arity(cache.string(type_name), domain.map(cache.sort(_)), cache.string(codomain))
|
|
340 |
}
|
|
341 |
|
68418
|
342 |
def read_arities(provider: Export.Provider): List[Arity] =
|
|
343 |
{
|
|
344 |
val body = provider.uncompressed_yxml(export_prefix + "arities")
|
|
345 |
val arities =
|
|
346 |
{
|
|
347 |
import XML.Decode._
|
|
348 |
import Term_XML.Decode._
|
|
349 |
list(triple(string, list(sort), string))(body)
|
|
350 |
}
|
|
351 |
for ((a, b, c) <- arities) yield Arity(a, b, c)
|
|
352 |
}
|
68295
|
353 |
|
|
354 |
|
68264
|
355 |
/* HOL typedefs */
|
|
356 |
|
|
357 |
sealed case class Typedef(name: String,
|
|
358 |
rep_type: Term.Typ, abs_type: Term.Typ, rep_name: String, abs_name: String, axiom_name: String)
|
68267
|
359 |
{
|
|
360 |
def cache(cache: Term.Cache): Typedef =
|
|
361 |
Typedef(cache.string(name),
|
|
362 |
cache.typ(rep_type),
|
|
363 |
cache.typ(abs_type),
|
|
364 |
cache.string(rep_name),
|
|
365 |
cache.string(abs_name),
|
|
366 |
cache.string(axiom_name))
|
|
367 |
}
|
68264
|
368 |
|
68418
|
369 |
def read_typedefs(provider: Export.Provider): List[Typedef] =
|
|
370 |
{
|
|
371 |
val body = provider.uncompressed_yxml(export_prefix + "typedefs")
|
|
372 |
val typedefs =
|
|
373 |
{
|
|
374 |
import XML.Decode._
|
|
375 |
import Term_XML.Decode._
|
|
376 |
list(pair(string, pair(typ, pair(typ, pair(string, pair(string, string))))))(body)
|
|
377 |
}
|
|
378 |
for { (name, (rep_type, (abs_type, (rep_name, (abs_name, axiom_name))))) <- typedefs }
|
|
379 |
yield Typedef(name, rep_type, abs_type, rep_name, abs_name, axiom_name)
|
|
380 |
}
|
68171
|
381 |
}
|