author | wenzelm |
Thu, 17 Oct 2019 21:03:59 +0200 | |
changeset 70896 | 8017d382a0d7 |
parent 70891 | f21a76a4d01f |
child 70899 | 5f6dea6a7a4c |
permissions | -rw-r--r-- |
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 |
||
70789 | 14 |
sealed case class Session(name: String, theory_graph: Graph[String, Option[Theory]]) |
68206 | 15 |
{ |
16 |
override def toString: String = name |
|
17 |
||
70789 | 18 |
def theory(theory_name: String): Option[Theory] = |
68206 | 19 |
if (theory_graph.defined(theory_name)) theory_graph.get_node(theory_name) |
70789 | 20 |
else None |
68206 | 21 |
|
22 |
def theories: List[Theory] = |
|
70789 | 23 |
theory_graph.topological_order.flatMap(theory(_)) |
68206 | 24 |
} |
25 |
||
70790
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
26 |
def read_session( |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
27 |
store: Sessions.Store, |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
28 |
sessions_structure: Sessions.Structure, |
68209 | 29 |
session_name: String, |
70790
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
30 |
progress: Progress = No_Progress, |
68206 | 31 |
types: Boolean = true, |
68264 | 32 |
consts: Boolean = true, |
33 |
axioms: Boolean = true, |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
34 |
thms: Boolean = true, |
68264 | 35 |
classes: Boolean = true, |
68862 | 36 |
locales: Boolean = true, |
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
37 |
locale_dependencies: Boolean = true, |
68295 | 38 |
classrel: Boolean = true, |
39 |
arities: Boolean = true, |
|
68862 | 40 |
typedefs: Boolean = true, |
68267 | 41 |
cache: Term.Cache = Term.make_cache()): Session = |
68206 | 42 |
{ |
43 |
val thys = |
|
70790
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
44 |
sessions_structure.build_requirements(List(session_name)).flatMap(session => |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
45 |
using(store.open_database(session))(db => |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
46 |
{ |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
47 |
db.transaction { |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
48 |
for (theory <- Export.read_theory_names(db, session)) |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
49 |
yield { |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
50 |
progress.echo("Reading theory " + theory) |
70891 | 51 |
read_theory(Export.Provider.database(db, session, theory), |
52 |
session, theory, types = types, consts = consts, |
|
53 |
axioms = axioms, thms = thms, classes = classes, locales = locales, |
|
54 |
locale_dependencies = locale_dependencies, classrel = classrel, arities = arities, |
|
55 |
typedefs = typedefs, cache = Some(cache)) |
|
70790
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
56 |
} |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
57 |
} |
73514ccad7a6
clarified signature: read full session requirements;
wenzelm
parents:
70789
diff
changeset
|
58 |
})) |
68206 | 59 |
|
60 |
val graph0 = |
|
70890 | 61 |
(Graph.string[Option[Theory]] /: thys) { |
62 |
case (g, thy) => g.default_node(thy.name, Some(thy)) } |
|
68206 | 63 |
val graph1 = |
64 |
(graph0 /: thys) { case (g0, thy) => |
|
65 |
(g0 /: thy.parents) { case (g1, parent) => |
|
70789 | 66 |
g1.default_node(parent, None).add_edge_acyclic(parent, thy.name) } } |
68206 | 67 |
|
68 |
Session(session_name, graph1) |
|
69 |
} |
|
70 |
||
71 |
||
72 |
||
68203 | 73 |
/** theory content **/ |
74 |
||
68346 | 75 |
val export_prefix: String = "theory/" |
70539
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
wenzelm
parents:
70534
diff
changeset
|
76 |
val export_prefix_proofs: String = "proofs/" |
68346 | 77 |
|
68208 | 78 |
sealed case class Theory(name: String, parents: List[String], |
79 |
types: List[Type], |
|
80 |
consts: List[Const], |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
81 |
axioms: List[Axiom], |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
82 |
thms: List[Thm], |
68264 | 83 |
classes: List[Class], |
68862 | 84 |
locales: List[Locale], |
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
85 |
locale_dependencies: List[Locale_Dependency], |
68295 | 86 |
classrel: List[Classrel], |
68862 | 87 |
arities: List[Arity], |
88 |
typedefs: List[Typedef]) |
|
68206 | 89 |
{ |
90 |
override def toString: String = name |
|
68267 | 91 |
|
70789 | 92 |
def entity_iterator: Iterator[Entity] = |
93 |
types.iterator.map(_.entity) ++ |
|
94 |
consts.iterator.map(_.entity) ++ |
|
95 |
axioms.iterator.map(_.entity) ++ |
|
96 |
thms.iterator.map(_.entity) ++ |
|
97 |
classes.iterator.map(_.entity) ++ |
|
98 |
locales.iterator.map(_.entity) ++ |
|
99 |
locale_dependencies.iterator.map(_.entity) |
|
68711 | 100 |
|
68267 | 101 |
def cache(cache: Term.Cache): Theory = |
102 |
Theory(cache.string(name), |
|
103 |
parents.map(cache.string(_)), |
|
104 |
types.map(_.cache(cache)), |
|
105 |
consts.map(_.cache(cache)), |
|
106 |
axioms.map(_.cache(cache)), |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
107 |
thms.map(_.cache(cache)), |
68267 | 108 |
classes.map(_.cache(cache)), |
68862 | 109 |
locales.map(_.cache(cache)), |
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
110 |
locale_dependencies.map(_.cache(cache)), |
68295 | 111 |
classrel.map(_.cache(cache)), |
68862 | 112 |
arities.map(_.cache(cache)), |
113 |
typedefs.map(_.cache(cache))) |
|
68206 | 114 |
} |
115 |
||
68418 | 116 |
def read_theory(provider: Export.Provider, session_name: String, theory_name: String, |
68203 | 117 |
types: Boolean = true, |
68208 | 118 |
consts: Boolean = true, |
68232 | 119 |
axioms: Boolean = true, |
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
120 |
thms: Boolean = true, |
68264 | 121 |
classes: Boolean = true, |
68862 | 122 |
locales: Boolean = true, |
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
123 |
locale_dependencies: Boolean = true, |
68295 | 124 |
classrel: Boolean = true, |
125 |
arities: Boolean = true, |
|
68862 | 126 |
typedefs: Boolean = true, |
68267 | 127 |
cache: Option[Term.Cache] = None): Theory = |
68203 | 128 |
{ |
68206 | 129 |
val parents = |
70891 | 130 |
if (theory_name == Thy_Header.PURE) Nil |
131 |
else { |
|
132 |
provider(export_prefix + "parents") match { |
|
133 |
case Some(entry) => split_lines(entry.uncompressed().text) |
|
134 |
case None => |
|
135 |
error("Missing theory export in session " + quote(session_name) + ": " + |
|
136 |
quote(theory_name)) |
|
137 |
} |
|
68206 | 138 |
} |
68267 | 139 |
val theory = |
140 |
Theory(theory_name, parents, |
|
68418 | 141 |
if (types) read_types(provider) else Nil, |
142 |
if (consts) read_consts(provider) else Nil, |
|
143 |
if (axioms) read_axioms(provider) else Nil, |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
144 |
if (thms) read_thms(provider) else Nil, |
68418 | 145 |
if (classes) read_classes(provider) else Nil, |
68862 | 146 |
if (locales) read_locales(provider) else Nil, |
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
147 |
if (locale_dependencies) read_locale_dependencies(provider) else Nil, |
68418 | 148 |
if (classrel) read_classrel(provider) else Nil, |
68862 | 149 |
if (arities) read_arities(provider) else Nil, |
150 |
if (typedefs) read_typedefs(provider) else Nil) |
|
68267 | 151 |
if (cache.isDefined) theory.cache(cache.get) else theory |
68203 | 152 |
} |
153 |
||
70883 | 154 |
def read_pure[A](store: Sessions.Store, read: (Export.Provider, String, String) => A): A = |
68710 | 155 |
{ |
156 |
val session_name = Thy_Header.PURE |
|
157 |
val theory_name = Thy_Header.PURE |
|
158 |
||
159 |
using(store.open_database(session_name))(db => |
|
160 |
{ |
|
161 |
db.transaction { |
|
70883 | 162 |
read(Export.Provider.database(db, session_name, theory_name), session_name, theory_name) |
68710 | 163 |
} |
164 |
}) |
|
165 |
} |
|
166 |
||
70883 | 167 |
def read_pure_theory(store: Sessions.Store, cache: Option[Term.Cache] = None): Theory = |
168 |
read_pure(store, read_theory(_, _, _, cache = cache)) |
|
169 |
||
70884 | 170 |
def read_pure_proof(store: Sessions.Store, id: Thm_Id, cache: Option[Term.Cache] = None): Proof = |
171 |
read_pure(store, (provider, _, _) => read_proof(provider, id, cache = cache)) |
|
70883 | 172 |
|
68203 | 173 |
|
68171 | 174 |
/* entities */ |
175 |
||
68714 | 176 |
object Kind extends Enumeration |
68171 | 177 |
{ |
68714 | 178 |
val TYPE = Value("type") |
179 |
val CONST = Value("const") |
|
180 |
val AXIOM = Value("axiom") |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
181 |
val THM = Value("thm") |
70882 | 182 |
val PROOF = Value("proof") |
68714 | 183 |
val CLASS = Value("class") |
68862 | 184 |
val LOCALE = Value("locale") |
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
185 |
val LOCALE_DEPENDENCY = Value("locale_dependency") |
68714 | 186 |
} |
187 |
||
68835 | 188 |
sealed case class Entity( |
68997 | 189 |
kind: Kind.Value, name: String, xname: String, pos: Position.T, id: Option[Long], serial: Long) |
68714 | 190 |
{ |
68718 | 191 |
override def toString: String = kind.toString + " " + quote(name) |
68267 | 192 |
|
193 |
def cache(cache: Term.Cache): Entity = |
|
68997 | 194 |
Entity(kind, cache.string(name), cache.string(xname), cache.position(pos), id, serial) |
68171 | 195 |
} |
196 |
||
68714 | 197 |
def decode_entity(kind: Kind.Value, tree: XML.Tree): (Entity, XML.Body) = |
68171 | 198 |
{ |
199 |
def err(): Nothing = throw new XML.XML_Body(List(tree)) |
|
200 |
||
201 |
tree match { |
|
202 |
case XML.Elem(Markup(Markup.ENTITY, props), body) => |
|
203 |
val name = Markup.Name.unapply(props) getOrElse err() |
|
68997 | 204 |
val xname = Markup.XName.unapply(props) getOrElse err() |
68830
44ec6fdaacf8
retain original id, which is command_id/exec_id for PIDE;
wenzelm
parents:
68726
diff
changeset
|
205 |
val pos = props.filter({ case (a, _) => Markup.POSITION_PROPERTIES(a) && a != Markup.ID }) |
68835 | 206 |
val id = Position.Id.unapply(props) |
68171 | 207 |
val serial = Markup.Serial.unapply(props) getOrElse err() |
68997 | 208 |
(Entity(kind, name, xname, pos, id, serial), body) |
68171 | 209 |
case _ => err() |
210 |
} |
|
211 |
} |
|
212 |
||
213 |
||
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
214 |
/* approximative syntax */ |
69003 | 215 |
|
216 |
object Assoc extends Enumeration |
|
217 |
{ |
|
218 |
val NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC = Value |
|
219 |
} |
|
220 |
||
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
221 |
sealed abstract class Syntax |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
222 |
case object No_Syntax extends Syntax |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
223 |
case class Prefix(delim: String) extends Syntax |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
224 |
case class Infix(assoc: Assoc.Value, delim: String, pri: Int) extends Syntax |
69003 | 225 |
|
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
226 |
def decode_syntax: XML.Decode.T[Syntax] = |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
227 |
XML.Decode.variant(List( |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
228 |
{ case (Nil, Nil) => No_Syntax }, |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
229 |
{ case (List(delim), Nil) => Prefix(delim) }, |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
230 |
{ case (Nil, body) => |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
231 |
import XML.Decode._ |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
232 |
val (ass, delim, pri) = triple(int, string, int)(body) |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
233 |
Infix(Assoc(ass), delim, pri) })) |
69003 | 234 |
|
235 |
||
68171 | 236 |
/* types */ |
237 |
||
69003 | 238 |
sealed case class Type( |
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
239 |
entity: Entity, syntax: Syntax, args: List[String], abbrev: Option[Term.Typ]) |
68267 | 240 |
{ |
241 |
def cache(cache: Term.Cache): Type = |
|
242 |
Type(entity.cache(cache), |
|
69003 | 243 |
syntax, |
68267 | 244 |
args.map(cache.string(_)), |
245 |
abbrev.map(cache.typ(_))) |
|
246 |
} |
|
68171 | 247 |
|
68418 | 248 |
def read_types(provider: Export.Provider): List[Type] = |
249 |
provider.uncompressed_yxml(export_prefix + "types").map((tree: XML.Tree) => |
|
250 |
{ |
|
68714 | 251 |
val (entity, body) = decode_entity(Kind.TYPE, tree) |
69003 | 252 |
val (syntax, args, abbrev) = |
68203 | 253 |
{ |
68418 | 254 |
import XML.Decode._ |
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
255 |
triple(decode_syntax, list(string), option(Term_XML.Decode.typ))(body) |
68418 | 256 |
} |
69003 | 257 |
Type(entity, syntax, args, abbrev) |
68418 | 258 |
}) |
68171 | 259 |
|
260 |
||
261 |
/* consts */ |
|
262 |
||
68173 | 263 |
sealed case class Const( |
69003 | 264 |
entity: Entity, |
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
265 |
syntax: Syntax, |
69003 | 266 |
typargs: List[String], |
267 |
typ: Term.Typ, |
|
69988 | 268 |
abbrev: Option[Term.Term], |
69992
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
wenzelm
parents:
69988
diff
changeset
|
269 |
propositional: Boolean, |
69996
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
270 |
primrec_types: List[String], |
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
271 |
corecursive: Boolean) |
68267 | 272 |
{ |
273 |
def cache(cache: Term.Cache): Const = |
|
274 |
Const(entity.cache(cache), |
|
69003 | 275 |
syntax, |
68267 | 276 |
typargs.map(cache.string(_)), |
277 |
cache.typ(typ), |
|
69988 | 278 |
abbrev.map(cache.term(_)), |
69992
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
wenzelm
parents:
69988
diff
changeset
|
279 |
propositional, |
69996
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
280 |
primrec_types.map(cache.string(_)), |
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
281 |
corecursive) |
68267 | 282 |
} |
68171 | 283 |
|
68418 | 284 |
def read_consts(provider: Export.Provider): List[Const] = |
285 |
provider.uncompressed_yxml(export_prefix + "consts").map((tree: XML.Tree) => |
|
286 |
{ |
|
68714 | 287 |
val (entity, body) = decode_entity(Kind.CONST, tree) |
69996
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
288 |
val (syntax, (args, (typ, (abbrev, (propositional, (primrec_types, corecursive)))))) = |
68203 | 289 |
{ |
68418 | 290 |
import XML.Decode._ |
69992
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
wenzelm
parents:
69988
diff
changeset
|
291 |
pair(decode_syntax, |
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
wenzelm
parents:
69988
diff
changeset
|
292 |
pair(list(string), |
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
wenzelm
parents:
69988
diff
changeset
|
293 |
pair(Term_XML.Decode.typ, |
69996
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
294 |
pair(option(Term_XML.Decode.term), pair(bool, |
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
295 |
pair(list(string), bool))))))(body) |
68418 | 296 |
} |
69996
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
wenzelm
parents:
69992
diff
changeset
|
297 |
Const(entity, syntax, args, typ, abbrev, propositional, primrec_types, corecursive) |
68418 | 298 |
}) |
68208 | 299 |
|
300 |
||
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
301 |
/* axioms */ |
68232 | 302 |
|
68726 | 303 |
sealed case class Prop( |
304 |
typargs: List[(String, Term.Sort)], |
|
305 |
args: List[(String, Term.Typ)], |
|
306 |
term: Term.Term) |
|
68232 | 307 |
{ |
68726 | 308 |
def cache(cache: Term.Cache): Prop = |
309 |
Prop( |
|
310 |
typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), |
|
311 |
args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), |
|
312 |
cache.term(term)) |
|
68232 | 313 |
} |
68208 | 314 |
|
68726 | 315 |
def decode_prop(body: XML.Body): Prop = |
68267 | 316 |
{ |
68726 | 317 |
val (typargs, args, t) = |
318 |
{ |
|
319 |
import XML.Decode._ |
|
320 |
import Term_XML.Decode._ |
|
321 |
triple(list(pair(string, sort)), list(pair(string, typ)), term)(body) |
|
322 |
} |
|
323 |
Prop(typargs, args, t) |
|
68267 | 324 |
} |
68208 | 325 |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
326 |
sealed case class Axiom(entity: Entity, prop: Prop) |
70534
fb876ebbf5a7
export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents:
70384
diff
changeset
|
327 |
{ |
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
328 |
def cache(cache: Term.Cache): Axiom = |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
329 |
Axiom(entity.cache(cache), prop.cache(cache)) |
70534
fb876ebbf5a7
export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents:
70384
diff
changeset
|
330 |
} |
fb876ebbf5a7
export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
wenzelm
parents:
70384
diff
changeset
|
331 |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
332 |
def read_axioms(provider: Export.Provider): List[Axiom] = |
68418 | 333 |
provider.uncompressed_yxml(export_prefix + "axioms").map((tree: XML.Tree) => |
334 |
{ |
|
68714 | 335 |
val (entity, body) = decode_entity(Kind.AXIOM, tree) |
68726 | 336 |
val prop = decode_prop(body) |
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
337 |
Axiom(entity, prop) |
68418 | 338 |
}) |
68232 | 339 |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
340 |
|
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
341 |
/* theorems */ |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
342 |
|
70884 | 343 |
sealed case class Thm_Id(serial: Long, theory_name: String) |
344 |
{ |
|
345 |
def pure: Boolean = theory_name == Thy_Header.PURE |
|
346 |
||
347 |
def cache(cache: Term.Cache): Thm_Id = |
|
348 |
Thm_Id(serial, cache.string(theory_name)) |
|
349 |
} |
|
350 |
||
351 |
sealed case class Thm( |
|
352 |
entity: Entity, |
|
353 |
prop: Prop, |
|
354 |
deps: List[String], |
|
70896 | 355 |
proof_boxes: List[Thm_Id], |
356 |
proof: Term.Proof) |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
357 |
{ |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
358 |
def cache(cache: Term.Cache): Thm = |
70884 | 359 |
Thm( |
360 |
entity.cache(cache), |
|
361 |
prop.cache(cache), |
|
362 |
deps.map(cache.string _), |
|
70896 | 363 |
proof_boxes.map(_.cache(cache)), |
364 |
cache.proof(proof)) |
|
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
365 |
} |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
366 |
|
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
367 |
def read_thms(provider: Export.Provider): List[Thm] = |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
368 |
provider.uncompressed_yxml(export_prefix + "thms").map((tree: XML.Tree) => |
68418 | 369 |
{ |
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
370 |
val (entity, body) = decode_entity(Kind.THM, tree) |
70896 | 371 |
val (prop, (deps, (prf_boxes, prf))) = |
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
372 |
{ |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
373 |
import XML.Decode._ |
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
374 |
import Term_XML.Decode._ |
70884 | 375 |
def thm_id(body: XML.Body): Thm_Id = |
376 |
{ |
|
377 |
val (serial, theory_name) = pair(long, string)(body) |
|
378 |
Thm_Id(serial, theory_name) |
|
379 |
} |
|
70896 | 380 |
pair(decode_prop _, pair(list(string), pair(list(thm_id), proof)))(body) |
70579
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
wenzelm
parents:
70539
diff
changeset
|
381 |
} |
70896 | 382 |
Thm(entity, prop, deps, prf_boxes, prf) |
68418 | 383 |
}) |
68264 | 384 |
|
70881 | 385 |
sealed case class Proof( |
386 |
typargs: List[(String, Term.Sort)], |
|
387 |
args: List[(String, Term.Typ)], |
|
388 |
term: Term.Term, |
|
389 |
proof: Term.Proof) |
|
390 |
{ |
|
70883 | 391 |
def prop: Prop = Prop(typargs, args, term) |
392 |
||
70881 | 393 |
def cache(cache: Term.Cache): Proof = |
394 |
Proof( |
|
395 |
typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), |
|
396 |
args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), |
|
397 |
cache.term(term), |
|
398 |
cache.proof(proof)) |
|
399 |
} |
|
400 |
||
70884 | 401 |
def read_proof(provider: Export.Provider, id: Thm_Id, cache: Option[Term.Cache] = None): Proof = |
70539
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
wenzelm
parents:
70534
diff
changeset
|
402 |
{ |
70884 | 403 |
val body = provider.focus(id.theory_name).uncompressed_yxml(export_prefix_proofs + id.serial) |
404 |
if (body.isEmpty) error("Bad proof export " + id.serial) |
|
70843
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
405 |
|
70881 | 406 |
val (typargs, (args, (prop_body, proof_body))) = |
70843
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
407 |
{ |
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
408 |
import XML.Decode._ |
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
409 |
import Term_XML.Decode._ |
70881 | 410 |
pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(x => x, x => x)))(body) |
70843
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
411 |
} |
70881 | 412 |
val env = args.toMap |
70843
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
413 |
val prop = Term_XML.Decode.term_env(env)(prop_body) |
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
414 |
val proof = Term_XML.Decode.proof_env(env)(proof_body) |
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70790
diff
changeset
|
415 |
|
70884 | 416 |
val result = Proof(typargs, args, prop, proof) |
70881 | 417 |
cache.map(result.cache(_)) getOrElse result |
70539
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
wenzelm
parents:
70534
diff
changeset
|
418 |
} |
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
wenzelm
parents:
70534
diff
changeset
|
419 |
|
68264 | 420 |
|
421 |
/* type classes */ |
|
422 |
||
423 |
sealed case class Class( |
|
69023
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents:
69019
diff
changeset
|
424 |
entity: Entity, params: List[(String, Term.Typ)], axioms: List[Prop]) |
68267 | 425 |
{ |
426 |
def cache(cache: Term.Cache): Class = |
|
427 |
Class(entity.cache(cache), |
|
428 |
params.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), |
|
69023
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents:
69019
diff
changeset
|
429 |
axioms.map(_.cache(cache))) |
68267 | 430 |
} |
68264 | 431 |
|
68418 | 432 |
def read_classes(provider: Export.Provider): List[Class] = |
433 |
provider.uncompressed_yxml(export_prefix + "classes").map((tree: XML.Tree) => |
|
434 |
{ |
|
68714 | 435 |
val (entity, body) = decode_entity(Kind.CLASS, tree) |
68418 | 436 |
val (params, axioms) = |
68264 | 437 |
{ |
68418 | 438 |
import XML.Decode._ |
439 |
import Term_XML.Decode._ |
|
69023
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents:
69019
diff
changeset
|
440 |
pair(list(pair(string, typ)), list(decode_prop))(body) |
68418 | 441 |
} |
442 |
Class(entity, params, axioms) |
|
443 |
}) |
|
68264 | 444 |
|
445 |
||
68862 | 446 |
/* locales */ |
447 |
||
448 |
sealed case class Locale( |
|
69019 | 449 |
entity: Entity, |
450 |
typargs: List[(String, Term.Sort)], |
|
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
451 |
args: List[((String, Term.Typ), Syntax)], |
69023
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents:
69019
diff
changeset
|
452 |
axioms: List[Prop]) |
68862 | 453 |
{ |
454 |
def cache(cache: Term.Cache): Locale = |
|
455 |
Locale(entity.cache(cache), |
|
68864 | 456 |
typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), |
69076 | 457 |
args.map({ case ((name, typ), syntax) => ((cache.string(name), cache.typ(typ)), syntax) }), |
69023
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents:
69019
diff
changeset
|
458 |
axioms.map(_.cache(cache))) |
68862 | 459 |
} |
460 |
||
461 |
def read_locales(provider: Export.Provider): List[Locale] = |
|
462 |
provider.uncompressed_yxml(export_prefix + "locales").map((tree: XML.Tree) => |
|
463 |
{ |
|
464 |
val (entity, body) = decode_entity(Kind.LOCALE, tree) |
|
69019 | 465 |
val (typargs, args, axioms) = |
68862 | 466 |
{ |
467 |
import XML.Decode._ |
|
468 |
import Term_XML.Decode._ |
|
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69076
diff
changeset
|
469 |
triple(list(pair(string, sort)), list(pair(pair(string, typ), decode_syntax)), |
69076 | 470 |
list(decode_prop))(body) |
68862 | 471 |
} |
69019 | 472 |
Locale(entity, typargs, args, axioms) |
68862 | 473 |
}) |
474 |
||
475 |
||
69069
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
476 |
/* locale dependencies */ |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
477 |
|
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
478 |
sealed case class Locale_Dependency( |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
479 |
entity: Entity, |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
480 |
source: String, |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
481 |
target: String, |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
482 |
prefix: List[(String, Boolean)], |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
483 |
subst_types: List[((String, Term.Sort), Term.Typ)], |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
484 |
subst_terms: List[((String, Term.Typ), Term.Term)]) |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
485 |
{ |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
486 |
def cache(cache: Term.Cache): Locale_Dependency = |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
487 |
Locale_Dependency(entity.cache(cache), |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
488 |
cache.string(source), |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
489 |
cache.string(target), |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
490 |
prefix.map({ case (name, mandatory) => (cache.string(name), mandatory) }), |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
491 |
subst_types.map({ case ((a, s), ty) => ((cache.string(a), cache.sort(s)), cache.typ(ty)) }), |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
492 |
subst_terms.map({ case ((x, ty), t) => ((cache.string(x), cache.typ(ty)), cache.term(t)) })) |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
493 |
|
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
494 |
def is_inclusion: Boolean = |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
495 |
subst_types.isEmpty && subst_terms.isEmpty |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
496 |
} |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
497 |
|
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
498 |
def read_locale_dependencies(provider: Export.Provider): List[Locale_Dependency] = |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
499 |
provider.uncompressed_yxml(export_prefix + "locale_dependencies").map((tree: XML.Tree) => |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
500 |
{ |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
501 |
val (entity, body) = decode_entity(Kind.LOCALE_DEPENDENCY, tree) |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
502 |
val (source, (target, (prefix, (subst_types, subst_terms)))) = |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
503 |
{ |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
504 |
import XML.Decode._ |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
505 |
import Term_XML.Decode._ |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
506 |
pair(string, pair(string, pair(list(pair(string, bool)), |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
507 |
pair(list(pair(pair(string, sort), typ)), list(pair(pair(string, typ), term))))))(body) |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
508 |
} |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
509 |
Locale_Dependency(entity, source, target, prefix, subst_types, subst_terms) |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
510 |
}) |
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
511 |
|
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
wenzelm
parents:
69023
diff
changeset
|
512 |
|
68295 | 513 |
/* sort algebra */ |
514 |
||
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
515 |
sealed case class Classrel(class1: String, class2: String, prop: Prop) |
68295 | 516 |
{ |
517 |
def cache(cache: Term.Cache): Classrel = |
|
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
518 |
Classrel(cache.string(class1), cache.string(class2), prop.cache(cache)) |
68295 | 519 |
} |
520 |
||
68418 | 521 |
def read_classrel(provider: Export.Provider): List[Classrel] = |
522 |
{ |
|
523 |
val body = provider.uncompressed_yxml(export_prefix + "classrel") |
|
524 |
val classrel = |
|
525 |
{ |
|
526 |
import XML.Decode._ |
|
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
527 |
list(pair(decode_prop, pair(string, string)))(body) |
68418 | 528 |
} |
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
529 |
for ((prop, (c1, c2)) <- classrel) yield Classrel(c1, c2, prop) |
68418 | 530 |
} |
68295 | 531 |
|
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
532 |
sealed case class Arity(type_name: String, domain: List[Term.Sort], codomain: String, prop: Prop) |
68295 | 533 |
{ |
534 |
def cache(cache: Term.Cache): Arity = |
|
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
535 |
Arity(cache.string(type_name), domain.map(cache.sort(_)), cache.string(codomain), |
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
536 |
prop.cache(cache)) |
68295 | 537 |
} |
538 |
||
68418 | 539 |
def read_arities(provider: Export.Provider): List[Arity] = |
540 |
{ |
|
541 |
val body = provider.uncompressed_yxml(export_prefix + "arities") |
|
542 |
val arities = |
|
543 |
{ |
|
544 |
import XML.Decode._ |
|
545 |
import Term_XML.Decode._ |
|
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
546 |
list(pair(decode_prop, triple(string, list(sort), string)))(body) |
68418 | 547 |
} |
70384
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
wenzelm
parents:
69996
diff
changeset
|
548 |
for ((prop, (a, b, c)) <- arities) yield Arity(a, b, c, prop) |
68418 | 549 |
} |
68295 | 550 |
|
551 |
||
68264 | 552 |
/* HOL typedefs */ |
553 |
||
554 |
sealed case class Typedef(name: String, |
|
555 |
rep_type: Term.Typ, abs_type: Term.Typ, rep_name: String, abs_name: String, axiom_name: String) |
|
68267 | 556 |
{ |
557 |
def cache(cache: Term.Cache): Typedef = |
|
558 |
Typedef(cache.string(name), |
|
559 |
cache.typ(rep_type), |
|
560 |
cache.typ(abs_type), |
|
561 |
cache.string(rep_name), |
|
562 |
cache.string(abs_name), |
|
563 |
cache.string(axiom_name)) |
|
564 |
} |
|
68264 | 565 |
|
68418 | 566 |
def read_typedefs(provider: Export.Provider): List[Typedef] = |
567 |
{ |
|
568 |
val body = provider.uncompressed_yxml(export_prefix + "typedefs") |
|
569 |
val typedefs = |
|
570 |
{ |
|
571 |
import XML.Decode._ |
|
572 |
import Term_XML.Decode._ |
|
573 |
list(pair(string, pair(typ, pair(typ, pair(string, pair(string, string))))))(body) |
|
574 |
} |
|
575 |
for { (name, (rep_type, (abs_type, (rep_name, (abs_name, axiom_name))))) <- typedefs } |
|
576 |
yield Typedef(name, rep_type, abs_type, rep_name, abs_name, axiom_name) |
|
577 |
} |
|
68171 | 578 |
} |