| author | wenzelm | 
| Fri, 09 Oct 2020 13:12:56 +0200 | |
| changeset 72411 | b8cc129ece05 | 
| parent 71726 | a5fda30edae2 | 
| child 72691 | 2126cf946086 | 
| 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 | ||
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 10 | import scala.collection.immutable.SortedMap | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 11 | |
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 12 | |
| 68171 | 13 | object Export_Theory | 
| 14 | {
 | |
| 68206 | 15 | /** session content **/ | 
| 16 | ||
| 70789 | 17 | sealed case class Session(name: String, theory_graph: Graph[String, Option[Theory]]) | 
| 68206 | 18 |   {
 | 
| 19 | override def toString: String = name | |
| 20 | ||
| 70789 | 21 | def theory(theory_name: String): Option[Theory] = | 
| 68206 | 22 | if (theory_graph.defined(theory_name)) theory_graph.get_node(theory_name) | 
| 70789 | 23 | else None | 
| 68206 | 24 | |
| 25 | def theories: List[Theory] = | |
| 71601 | 26 | theory_graph.topological_order.flatMap(theory) | 
| 68206 | 27 | } | 
| 28 | ||
| 70790 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 29 | def read_session( | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 30 | store: Sessions.Store, | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 31 | sessions_structure: Sessions.Structure, | 
| 68209 | 32 | session_name: String, | 
| 71726 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 wenzelm parents: 
71601diff
changeset | 33 | progress: Progress = new Progress, | 
| 68267 | 34 | cache: Term.Cache = Term.make_cache()): Session = | 
| 68206 | 35 |   {
 | 
| 36 | val thys = | |
| 70790 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 37 | sessions_structure.build_requirements(List(session_name)).flatMap(session => | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 38 | using(store.open_database(session))(db => | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 39 |         {
 | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 40 |           db.transaction {
 | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 41 | for (theory <- Export.read_theory_names(db, session)) | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 42 |             yield {
 | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 43 |               progress.echo("Reading theory " + theory)
 | 
| 70891 | 44 | read_theory(Export.Provider.database(db, session, theory), | 
| 71251 | 45 | session, theory, cache = Some(cache)) | 
| 70790 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 46 | } | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 47 | } | 
| 
73514ccad7a6
clarified signature: read full session requirements;
 wenzelm parents: 
70789diff
changeset | 48 | })) | 
| 68206 | 49 | |
| 50 | val graph0 = | |
| 70890 | 51 |       (Graph.string[Option[Theory]] /: thys) {
 | 
| 52 | case (g, thy) => g.default_node(thy.name, Some(thy)) } | |
| 68206 | 53 | val graph1 = | 
| 54 |       (graph0 /: thys) { case (g0, thy) =>
 | |
| 55 |         (g0 /: thy.parents) { case (g1, parent) =>
 | |
| 70789 | 56 | g1.default_node(parent, None).add_edge_acyclic(parent, thy.name) } } | 
| 68206 | 57 | |
| 58 | Session(session_name, graph1) | |
| 59 | } | |
| 60 | ||
| 61 | ||
| 62 | ||
| 68203 | 63 | /** theory content **/ | 
| 64 | ||
| 68346 | 65 | val export_prefix: String = "theory/" | 
| 70539 
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
 wenzelm parents: 
70534diff
changeset | 66 | val export_prefix_proofs: String = "proofs/" | 
| 68346 | 67 | |
| 68208 | 68 | sealed case class Theory(name: String, parents: List[String], | 
| 69 | types: List[Type], | |
| 70 | consts: List[Const], | |
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 71 | axioms: List[Axiom], | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 72 | thms: List[Thm], | 
| 68264 | 73 | classes: List[Class], | 
| 68862 | 74 | locales: List[Locale], | 
| 69069 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 75 | locale_dependencies: List[Locale_Dependency], | 
| 68295 | 76 | classrel: List[Classrel], | 
| 68862 | 77 | arities: List[Arity], | 
| 70920 | 78 | constdefs: List[Constdef], | 
| 71202 | 79 | typedefs: List[Typedef], | 
| 71248 | 80 | datatypes: List[Datatype], | 
| 71202 | 81 | spec_rules: List[Spec_Rule]) | 
| 68206 | 82 |   {
 | 
| 83 | override def toString: String = name | |
| 68267 | 84 | |
| 70789 | 85 | def entity_iterator: Iterator[Entity] = | 
| 86 | types.iterator.map(_.entity) ++ | |
| 87 | consts.iterator.map(_.entity) ++ | |
| 88 | axioms.iterator.map(_.entity) ++ | |
| 89 | thms.iterator.map(_.entity) ++ | |
| 90 | classes.iterator.map(_.entity) ++ | |
| 91 | locales.iterator.map(_.entity) ++ | |
| 92 | locale_dependencies.iterator.map(_.entity) | |
| 68711 | 93 | |
| 68267 | 94 | def cache(cache: Term.Cache): Theory = | 
| 95 | Theory(cache.string(name), | |
| 96 | parents.map(cache.string(_)), | |
| 97 | types.map(_.cache(cache)), | |
| 98 | consts.map(_.cache(cache)), | |
| 99 | axioms.map(_.cache(cache)), | |
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 100 | thms.map(_.cache(cache)), | 
| 68267 | 101 | classes.map(_.cache(cache)), | 
| 68862 | 102 | locales.map(_.cache(cache)), | 
| 69069 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 103 | locale_dependencies.map(_.cache(cache)), | 
| 68295 | 104 | classrel.map(_.cache(cache)), | 
| 68862 | 105 | arities.map(_.cache(cache)), | 
| 70920 | 106 | constdefs.map(_.cache(cache)), | 
| 71202 | 107 | typedefs.map(_.cache(cache)), | 
| 71248 | 108 | datatypes.map(_.cache(cache)), | 
| 71202 | 109 | spec_rules.map(_.cache(cache))) | 
| 68206 | 110 | } | 
| 111 | ||
| 68418 | 112 | def read_theory(provider: Export.Provider, session_name: String, theory_name: String, | 
| 68267 | 113 | cache: Option[Term.Cache] = None): Theory = | 
| 68203 | 114 |   {
 | 
| 68206 | 115 | val parents = | 
| 70891 | 116 | if (theory_name == Thy_Header.PURE) Nil | 
| 117 |       else {
 | |
| 118 |         provider(export_prefix + "parents") match {
 | |
| 119 | case Some(entry) => split_lines(entry.uncompressed().text) | |
| 120 | case None => | |
| 121 |             error("Missing theory export in session " + quote(session_name) + ": " +
 | |
| 122 | quote(theory_name)) | |
| 123 | } | |
| 68206 | 124 | } | 
| 68267 | 125 | val theory = | 
| 126 | Theory(theory_name, parents, | |
| 71251 | 127 | read_types(provider), | 
| 128 | read_consts(provider), | |
| 129 | read_axioms(provider), | |
| 130 | read_thms(provider), | |
| 131 | read_classes(provider), | |
| 132 | read_locales(provider), | |
| 133 | read_locale_dependencies(provider), | |
| 134 | read_classrel(provider), | |
| 135 | read_arities(provider), | |
| 136 | read_constdefs(provider), | |
| 137 | read_typedefs(provider), | |
| 138 | read_datatypes(provider), | |
| 139 | read_spec_rules(provider)) | |
| 68267 | 140 | if (cache.isDefined) theory.cache(cache.get) else theory | 
| 68203 | 141 | } | 
| 142 | ||
| 70883 | 143 | def read_pure[A](store: Sessions.Store, read: (Export.Provider, String, String) => A): A = | 
| 68710 | 144 |   {
 | 
| 145 | val session_name = Thy_Header.PURE | |
| 146 | val theory_name = Thy_Header.PURE | |
| 147 | ||
| 148 | using(store.open_database(session_name))(db => | |
| 149 |     {
 | |
| 150 |       db.transaction {
 | |
| 70883 | 151 | read(Export.Provider.database(db, session_name, theory_name), session_name, theory_name) | 
| 68710 | 152 | } | 
| 153 | }) | |
| 154 | } | |
| 155 | ||
| 70883 | 156 | def read_pure_theory(store: Sessions.Store, cache: Option[Term.Cache] = None): Theory = | 
| 157 | read_pure(store, read_theory(_, _, _, cache = cache)) | |
| 158 | ||
| 70899 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 159 | def read_pure_proof( | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 160 | store: Sessions.Store, id: Thm_Id, cache: Option[Term.Cache] = None): Option[Proof] = | 
| 70884 | 161 | read_pure(store, (provider, _, _) => read_proof(provider, id, cache = cache)) | 
| 70883 | 162 | |
| 68203 | 163 | |
| 68171 | 164 | /* entities */ | 
| 165 | ||
| 68714 | 166 | object Kind extends Enumeration | 
| 68171 | 167 |   {
 | 
| 68714 | 168 |     val TYPE = Value("type")
 | 
| 169 |     val CONST = Value("const")
 | |
| 170 |     val AXIOM = Value("axiom")
 | |
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 171 |     val THM = Value("thm")
 | 
| 70882 | 172 |     val PROOF = Value("proof")
 | 
| 68714 | 173 |     val CLASS = Value("class")
 | 
| 68862 | 174 |     val LOCALE = Value("locale")
 | 
| 69069 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 175 |     val LOCALE_DEPENDENCY = Value("locale_dependency")
 | 
| 70913 | 176 |     val DOCUMENT_HEADING = Value("document_heading")
 | 
| 177 |     val DOCUMENT_TEXT = Value("document_text")
 | |
| 178 |     val PROOF_TEXT = Value("proof_text")
 | |
| 68714 | 179 | } | 
| 180 | ||
| 68835 | 181 | sealed case class Entity( | 
| 68997 | 182 | kind: Kind.Value, name: String, xname: String, pos: Position.T, id: Option[Long], serial: Long) | 
| 68714 | 183 |   {
 | 
| 68718 | 184 | override def toString: String = kind.toString + " " + quote(name) | 
| 68267 | 185 | |
| 186 | def cache(cache: Term.Cache): Entity = | |
| 68997 | 187 | Entity(kind, cache.string(name), cache.string(xname), cache.position(pos), id, serial) | 
| 68171 | 188 | } | 
| 189 | ||
| 68714 | 190 | def decode_entity(kind: Kind.Value, tree: XML.Tree): (Entity, XML.Body) = | 
| 68171 | 191 |   {
 | 
| 192 | def err(): Nothing = throw new XML.XML_Body(List(tree)) | |
| 193 | ||
| 194 |     tree match {
 | |
| 195 | case XML.Elem(Markup(Markup.ENTITY, props), body) => | |
| 196 | val name = Markup.Name.unapply(props) getOrElse err() | |
| 68997 | 197 | val xname = Markup.XName.unapply(props) getOrElse err() | 
| 68830 
44ec6fdaacf8
retain original id, which is command_id/exec_id for PIDE;
 wenzelm parents: 
68726diff
changeset | 198 |         val pos = props.filter({ case (a, _) => Markup.POSITION_PROPERTIES(a) && a != Markup.ID })
 | 
| 68835 | 199 | val id = Position.Id.unapply(props) | 
| 68171 | 200 | val serial = Markup.Serial.unapply(props) getOrElse err() | 
| 68997 | 201 | (Entity(kind, name, xname, pos, id, serial), body) | 
| 68171 | 202 | case _ => err() | 
| 203 | } | |
| 204 | } | |
| 205 | ||
| 206 | ||
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 207 | /* approximative syntax */ | 
| 69003 | 208 | |
| 209 | object Assoc extends Enumeration | |
| 210 |   {
 | |
| 211 | val NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC = Value | |
| 212 | } | |
| 213 | ||
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 214 | sealed abstract class Syntax | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 215 | case object No_Syntax extends Syntax | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 216 | case class Prefix(delim: String) extends Syntax | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 217 | case class Infix(assoc: Assoc.Value, delim: String, pri: Int) extends Syntax | 
| 69003 | 218 | |
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 219 | def decode_syntax: XML.Decode.T[Syntax] = | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 220 | XML.Decode.variant(List( | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 221 |       { case (Nil, Nil) => No_Syntax },
 | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 222 |       { case (List(delim), Nil) => Prefix(delim) },
 | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 223 |       { case (Nil, body) =>
 | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 224 | import XML.Decode._ | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 225 | val (ass, delim, pri) = triple(int, string, int)(body) | 
| 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 226 | Infix(Assoc(ass), delim, pri) })) | 
| 69003 | 227 | |
| 228 | ||
| 68171 | 229 | /* types */ | 
| 230 | ||
| 69003 | 231 | sealed case class Type( | 
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 232 | entity: Entity, syntax: Syntax, args: List[String], abbrev: Option[Term.Typ]) | 
| 68267 | 233 |   {
 | 
| 234 | def cache(cache: Term.Cache): Type = | |
| 235 | Type(entity.cache(cache), | |
| 69003 | 236 | syntax, | 
| 68267 | 237 | args.map(cache.string(_)), | 
| 238 | abbrev.map(cache.typ(_))) | |
| 239 | } | |
| 68171 | 240 | |
| 68418 | 241 | def read_types(provider: Export.Provider): List[Type] = | 
| 242 | provider.uncompressed_yxml(export_prefix + "types").map((tree: XML.Tree) => | |
| 243 |       {
 | |
| 68714 | 244 | val (entity, body) = decode_entity(Kind.TYPE, tree) | 
| 69003 | 245 | val (syntax, args, abbrev) = | 
| 68203 | 246 |         {
 | 
| 68418 | 247 | import XML.Decode._ | 
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 248 | triple(decode_syntax, list(string), option(Term_XML.Decode.typ))(body) | 
| 68418 | 249 | } | 
| 69003 | 250 | Type(entity, syntax, args, abbrev) | 
| 68418 | 251 | }) | 
| 68171 | 252 | |
| 253 | ||
| 254 | /* consts */ | |
| 255 | ||
| 68173 | 256 | sealed case class Const( | 
| 69003 | 257 | entity: Entity, | 
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 258 | syntax: Syntax, | 
| 69003 | 259 | typargs: List[String], | 
| 260 | typ: Term.Typ, | |
| 69988 | 261 | abbrev: Option[Term.Term], | 
| 71222 
2bc39c80a95d
clarified export of consts: recursion is accessible via spec_rules;
 wenzelm parents: 
71211diff
changeset | 262 | propositional: Boolean) | 
| 68267 | 263 |   {
 | 
| 264 | def cache(cache: Term.Cache): Const = | |
| 265 | Const(entity.cache(cache), | |
| 69003 | 266 | syntax, | 
| 68267 | 267 | typargs.map(cache.string(_)), | 
| 268 | cache.typ(typ), | |
| 69988 | 269 | abbrev.map(cache.term(_)), | 
| 71222 
2bc39c80a95d
clarified export of consts: recursion is accessible via spec_rules;
 wenzelm parents: 
71211diff
changeset | 270 | propositional) | 
| 68267 | 271 | } | 
| 68171 | 272 | |
| 68418 | 273 | def read_consts(provider: Export.Provider): List[Const] = | 
| 274 | provider.uncompressed_yxml(export_prefix + "consts").map((tree: XML.Tree) => | |
| 275 |       {
 | |
| 68714 | 276 | val (entity, body) = decode_entity(Kind.CONST, tree) | 
| 71222 
2bc39c80a95d
clarified export of consts: recursion is accessible via spec_rules;
 wenzelm parents: 
71211diff
changeset | 277 | val (syntax, (typargs, (typ, (abbrev, propositional)))) = | 
| 68203 | 278 |         {
 | 
| 68418 | 279 | import XML.Decode._ | 
| 69992 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 wenzelm parents: 
69988diff
changeset | 280 | pair(decode_syntax, | 
| 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 wenzelm parents: 
69988diff
changeset | 281 | pair(list(string), | 
| 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 wenzelm parents: 
69988diff
changeset | 282 | pair(Term_XML.Decode.typ, | 
| 71222 
2bc39c80a95d
clarified export of consts: recursion is accessible via spec_rules;
 wenzelm parents: 
71211diff
changeset | 283 | pair(option(Term_XML.Decode.term), bool))))(body) | 
| 68418 | 284 | } | 
| 71222 
2bc39c80a95d
clarified export of consts: recursion is accessible via spec_rules;
 wenzelm parents: 
71211diff
changeset | 285 | Const(entity, syntax, typargs, typ, abbrev, propositional) | 
| 68418 | 286 | }) | 
| 68208 | 287 | |
| 288 | ||
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 289 | /* axioms */ | 
| 68232 | 290 | |
| 68726 | 291 | sealed case class Prop( | 
| 292 | typargs: List[(String, Term.Sort)], | |
| 293 | args: List[(String, Term.Typ)], | |
| 294 | term: Term.Term) | |
| 68232 | 295 |   {
 | 
| 68726 | 296 | def cache(cache: Term.Cache): Prop = | 
| 297 | Prop( | |
| 298 |         typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
 | |
| 299 |         args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }),
 | |
| 300 | cache.term(term)) | |
| 68232 | 301 | } | 
| 68208 | 302 | |
| 68726 | 303 | def decode_prop(body: XML.Body): Prop = | 
| 68267 | 304 |   {
 | 
| 68726 | 305 | val (typargs, args, t) = | 
| 306 |     {
 | |
| 307 | import XML.Decode._ | |
| 308 | import Term_XML.Decode._ | |
| 309 | triple(list(pair(string, sort)), list(pair(string, typ)), term)(body) | |
| 310 | } | |
| 311 | Prop(typargs, args, t) | |
| 68267 | 312 | } | 
| 68208 | 313 | |
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 314 | 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: 
70384diff
changeset | 315 |   {
 | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 316 | def cache(cache: Term.Cache): Axiom = | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 317 | 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: 
70384diff
changeset | 318 | } | 
| 
fb876ebbf5a7
export facts with reconstructed proof term (if possible), but its PThm boxes need to be collected separately;
 wenzelm parents: 
70384diff
changeset | 319 | |
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 320 | def read_axioms(provider: Export.Provider): List[Axiom] = | 
| 68418 | 321 | provider.uncompressed_yxml(export_prefix + "axioms").map((tree: XML.Tree) => | 
| 322 |       {
 | |
| 68714 | 323 | val (entity, body) = decode_entity(Kind.AXIOM, tree) | 
| 68726 | 324 | val prop = decode_prop(body) | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 325 | Axiom(entity, prop) | 
| 68418 | 326 | }) | 
| 68232 | 327 | |
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 328 | |
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 329 | /* theorems */ | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 330 | |
| 70884 | 331 | sealed case class Thm_Id(serial: Long, theory_name: String) | 
| 332 |   {
 | |
| 333 | def pure: Boolean = theory_name == Thy_Header.PURE | |
| 334 | } | |
| 335 | ||
| 336 | sealed case class Thm( | |
| 337 | entity: Entity, | |
| 338 | prop: Prop, | |
| 339 | deps: List[String], | |
| 70896 | 340 | proof: Term.Proof) | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 341 |   {
 | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 342 | def cache(cache: Term.Cache): Thm = | 
| 70884 | 343 | Thm( | 
| 344 | entity.cache(cache), | |
| 345 | prop.cache(cache), | |
| 346 | deps.map(cache.string _), | |
| 70896 | 347 | cache.proof(proof)) | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 348 | } | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 349 | |
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 350 | def read_thms(provider: Export.Provider): List[Thm] = | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 351 | provider.uncompressed_yxml(export_prefix + "thms").map((tree: XML.Tree) => | 
| 68418 | 352 |       {
 | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 353 | val (entity, body) = decode_entity(Kind.THM, tree) | 
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 354 | val (prop, deps, prf) = | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 355 |         {
 | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 356 | import XML.Decode._ | 
| 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 357 | import Term_XML.Decode._ | 
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 358 | triple(decode_prop, list(string), proof)(body) | 
| 70579 
5a8e3e4b3760
clarified export of axioms and theorems (identified derivations instead of projected facts);
 wenzelm parents: 
70539diff
changeset | 359 | } | 
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 360 | Thm(entity, prop, deps, prf) | 
| 68418 | 361 | }) | 
| 68264 | 362 | |
| 70881 | 363 | sealed case class Proof( | 
| 364 | typargs: List[(String, Term.Sort)], | |
| 365 | args: List[(String, Term.Typ)], | |
| 366 | term: Term.Term, | |
| 367 | proof: Term.Proof) | |
| 368 |   {
 | |
| 70883 | 369 | def prop: Prop = Prop(typargs, args, term) | 
| 370 | ||
| 70881 | 371 | def cache(cache: Term.Cache): Proof = | 
| 372 | Proof( | |
| 373 |         typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
 | |
| 374 |         args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }),
 | |
| 375 | cache.term(term), | |
| 376 | cache.proof(proof)) | |
| 377 | } | |
| 378 | ||
| 70899 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 379 | def read_proof( | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 380 | provider: Export.Provider, id: Thm_Id, cache: Option[Term.Cache] = None): Option[Proof] = | 
| 70539 
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
 wenzelm parents: 
70534diff
changeset | 381 |   {
 | 
| 70899 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 382 |     for { entry <- provider.focus(id.theory_name)(export_prefix_proofs + id.serial) }
 | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 383 |     yield {
 | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 384 | val body = entry.uncompressed_yxml() | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 385 | val (typargs, (args, (prop_body, proof_body))) = | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 386 |       {
 | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 387 | import XML.Decode._ | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 388 | import Term_XML.Decode._ | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 389 | pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(x => x, x => x)))(body) | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 390 | } | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 391 | val env = args.toMap | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 392 | val prop = Term_XML.Decode.term_env(env)(prop_body) | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 393 | val proof = Term_XML.Decode.proof_env(env)(proof_body) | 
| 70843 
cc987440d776
more compact XML: separate environment for free variables;
 wenzelm parents: 
70790diff
changeset | 394 | |
| 70899 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 395 | val result = Proof(typargs, args, prop, proof) | 
| 
5f6dea6a7a4c
clarified signature: support partial read_proof to accommodate proof term normalization vs. approximative proof_boxes as upper bound;
 wenzelm parents: 
70896diff
changeset | 396 | cache.map(result.cache(_)) getOrElse result | 
| 70843 
cc987440d776
more compact XML: separate environment for free variables;
 wenzelm parents: 
70790diff
changeset | 397 | } | 
| 70539 
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
 wenzelm parents: 
70534diff
changeset | 398 | } | 
| 
30b3c58a1933
support Export_Theory.read_proof, based on theory_name and serial;
 wenzelm parents: 
70534diff
changeset | 399 | |
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 400 | def read_proof_boxes( | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 401 | store: Sessions.Store, | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 402 | provider: Export.Provider, | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 403 | proof: Term.Proof, | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 404 | suppress: Thm_Id => Boolean = _ => false, | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 405 | cache: Option[Term.Cache] = None): List[(Thm_Id, Proof)] = | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 406 |   {
 | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 407 | var seen = Set.empty[Long] | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 408 | var result = SortedMap.empty[Long, (Thm_Id, Proof)] | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 409 | |
| 71016 | 410 | def boxes(context: Option[(Long, Term.Proof)], prf: Term.Proof) | 
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 411 |     {
 | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 412 |       prf match {
 | 
| 71016 | 413 | case Term.Abst(_, _, p) => boxes(context, p) | 
| 414 | case Term.AbsP(_, _, p) => boxes(context, p) | |
| 415 | case Term.Appt(p, _) => boxes(context, p) | |
| 416 | case Term.AppP(p, q) => boxes(context, p); boxes(context, q) | |
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 417 | case thm: Term.PThm if !seen(thm.serial) => | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 418 | seen += thm.serial | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 419 | val id = Thm_Id(thm.serial, thm.theory_name) | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 420 |           if (!suppress(id)) {
 | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 421 | val read = | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 422 | if (id.pure) Export_Theory.read_pure_proof(store, id, cache = cache) | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 423 | else Export_Theory.read_proof(provider, id, cache = cache) | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 424 |             read match {
 | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 425 | case Some(p) => | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 426 | result += (thm.serial -> (id -> p)) | 
| 71016 | 427 | boxes(Some((thm.serial, p.proof)), p.proof) | 
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 428 | case None => | 
| 71016 | 429 |                 error("Missing proof " + thm.serial + " (theory " + quote (thm.theory_name) + ")" +
 | 
| 430 |                   (context match {
 | |
| 431 | case None => "" | |
| 432 | case Some((i, p)) => " in proof " + i + ":\n" + p | |
| 433 | })) | |
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 434 | } | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 435 | } | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 436 | case _ => | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 437 | } | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 438 | } | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 439 | |
| 71016 | 440 | boxes(None, proof) | 
| 71015 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 441 | result.iterator.map(_._2).toList | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 442 | } | 
| 
bb49abc2ecbb
determine proof boxes from exported proof (NB: thm_boxes is not sufficient due to OfClass proofs);
 wenzelm parents: 
70920diff
changeset | 443 | |
| 68264 | 444 | |
| 445 | /* type classes */ | |
| 446 | ||
| 447 | sealed case class Class( | |
| 69023 
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
 wenzelm parents: 
69019diff
changeset | 448 | entity: Entity, params: List[(String, Term.Typ)], axioms: List[Prop]) | 
| 68267 | 449 |   {
 | 
| 450 | def cache(cache: Term.Cache): Class = | |
| 451 | Class(entity.cache(cache), | |
| 452 |         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: 
69019diff
changeset | 453 | axioms.map(_.cache(cache))) | 
| 68267 | 454 | } | 
| 68264 | 455 | |
| 68418 | 456 | def read_classes(provider: Export.Provider): List[Class] = | 
| 457 | provider.uncompressed_yxml(export_prefix + "classes").map((tree: XML.Tree) => | |
| 458 |       {
 | |
| 68714 | 459 | val (entity, body) = decode_entity(Kind.CLASS, tree) | 
| 68418 | 460 | val (params, axioms) = | 
| 68264 | 461 |         {
 | 
| 68418 | 462 | import XML.Decode._ | 
| 463 | import Term_XML.Decode._ | |
| 69023 
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
 wenzelm parents: 
69019diff
changeset | 464 | pair(list(pair(string, typ)), list(decode_prop))(body) | 
| 68418 | 465 | } | 
| 466 | Class(entity, params, axioms) | |
| 467 | }) | |
| 68264 | 468 | |
| 469 | ||
| 68862 | 470 | /* locales */ | 
| 471 | ||
| 472 | sealed case class Locale( | |
| 69019 | 473 | entity: Entity, | 
| 474 | typargs: List[(String, Term.Sort)], | |
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 475 | args: List[((String, Term.Typ), Syntax)], | 
| 69023 
cef000855cf4
clarified standardization of variables, with proper treatment of local variables;
 wenzelm parents: 
69019diff
changeset | 476 | axioms: List[Prop]) | 
| 68862 | 477 |   {
 | 
| 478 | def cache(cache: Term.Cache): Locale = | |
| 479 | Locale(entity.cache(cache), | |
| 68864 | 480 |         typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
 | 
| 69076 | 481 |         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: 
69019diff
changeset | 482 | axioms.map(_.cache(cache))) | 
| 68862 | 483 | } | 
| 484 | ||
| 485 | def read_locales(provider: Export.Provider): List[Locale] = | |
| 486 | provider.uncompressed_yxml(export_prefix + "locales").map((tree: XML.Tree) => | |
| 487 |       {
 | |
| 488 | val (entity, body) = decode_entity(Kind.LOCALE, tree) | |
| 69019 | 489 | val (typargs, args, axioms) = | 
| 68862 | 490 |         {
 | 
| 491 | import XML.Decode._ | |
| 492 | import Term_XML.Decode._ | |
| 69077 
11529ae45786
more approximative prefix syntax, including binder;
 wenzelm parents: 
69076diff
changeset | 493 | triple(list(pair(string, sort)), list(pair(pair(string, typ), decode_syntax)), | 
| 69076 | 494 | list(decode_prop))(body) | 
| 68862 | 495 | } | 
| 69019 | 496 | Locale(entity, typargs, args, axioms) | 
| 68862 | 497 | }) | 
| 498 | ||
| 499 | ||
| 69069 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 500 | /* locale dependencies */ | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 501 | |
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 502 | sealed case class Locale_Dependency( | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 503 | entity: Entity, | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 504 | source: String, | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 505 | target: String, | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 506 | prefix: List[(String, Boolean)], | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 507 | subst_types: List[((String, Term.Sort), Term.Typ)], | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 508 | subst_terms: List[((String, Term.Typ), Term.Term)]) | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 509 |   {
 | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 510 | def cache(cache: Term.Cache): Locale_Dependency = | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 511 | Locale_Dependency(entity.cache(cache), | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 512 | cache.string(source), | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 513 | cache.string(target), | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 514 |         prefix.map({ case (name, mandatory) => (cache.string(name), mandatory) }),
 | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 515 |         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: 
69023diff
changeset | 516 |         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: 
69023diff
changeset | 517 | |
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 518 | def is_inclusion: Boolean = | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 519 | subst_types.isEmpty && subst_terms.isEmpty | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 520 | } | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 521 | |
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 522 | def read_locale_dependencies(provider: Export.Provider): List[Locale_Dependency] = | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 523 | provider.uncompressed_yxml(export_prefix + "locale_dependencies").map((tree: XML.Tree) => | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 524 |       {
 | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 525 | val (entity, body) = decode_entity(Kind.LOCALE_DEPENDENCY, tree) | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 526 | val (source, (target, (prefix, (subst_types, subst_terms)))) = | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 527 |         {
 | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 528 | import XML.Decode._ | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 529 | import Term_XML.Decode._ | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 530 | pair(string, pair(string, pair(list(pair(string, bool)), | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 531 | 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: 
69023diff
changeset | 532 | } | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 533 | Locale_Dependency(entity, source, target, prefix, subst_types, subst_terms) | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 534 | }) | 
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 535 | |
| 
b9aca3b9619f
export locale dependencies, with approx. morphism as type/term substitution;
 wenzelm parents: 
69023diff
changeset | 536 | |
| 68295 | 537 | /* sort algebra */ | 
| 538 | ||
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 539 | sealed case class Classrel(class1: String, class2: String, prop: Prop) | 
| 68295 | 540 |   {
 | 
| 541 | def cache(cache: Term.Cache): Classrel = | |
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 542 | Classrel(cache.string(class1), cache.string(class2), prop.cache(cache)) | 
| 68295 | 543 | } | 
| 544 | ||
| 68418 | 545 | def read_classrel(provider: Export.Provider): List[Classrel] = | 
| 546 |   {
 | |
| 547 | val body = provider.uncompressed_yxml(export_prefix + "classrel") | |
| 548 | val classrel = | |
| 549 |     {
 | |
| 550 | import XML.Decode._ | |
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 551 | list(pair(decode_prop, pair(string, string)))(body) | 
| 68418 | 552 | } | 
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 553 | for ((prop, (c1, c2)) <- classrel) yield Classrel(c1, c2, prop) | 
| 68418 | 554 | } | 
| 68295 | 555 | |
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 556 | sealed case class Arity(type_name: String, domain: List[Term.Sort], codomain: String, prop: Prop) | 
| 68295 | 557 |   {
 | 
| 558 | def cache(cache: Term.Cache): Arity = | |
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 559 | 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: 
69996diff
changeset | 560 | prop.cache(cache)) | 
| 68295 | 561 | } | 
| 562 | ||
| 68418 | 563 | def read_arities(provider: Export.Provider): List[Arity] = | 
| 564 |   {
 | |
| 565 | val body = provider.uncompressed_yxml(export_prefix + "arities") | |
| 566 | val arities = | |
| 567 |     {
 | |
| 568 | import XML.Decode._ | |
| 569 | import Term_XML.Decode._ | |
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 570 | list(pair(decode_prop, triple(string, list(sort), string)))(body) | 
| 68418 | 571 | } | 
| 70384 
8ce08b154aa1
clarified export of sort algebra: avoid logical operations in Isabelle/Scala;
 wenzelm parents: 
69996diff
changeset | 572 | for ((prop, (a, b, c)) <- arities) yield Arity(a, b, c, prop) | 
| 68418 | 573 | } | 
| 68295 | 574 | |
| 575 | ||
| 70920 | 576 | /* Pure constdefs */ | 
| 577 | ||
| 578 | sealed case class Constdef(name: String, axiom_name: String) | |
| 579 |   {
 | |
| 580 | def cache(cache: Term.Cache): Constdef = | |
| 581 | Constdef(cache.string(name), cache.string(axiom_name)) | |
| 582 | } | |
| 583 | ||
| 584 | def read_constdefs(provider: Export.Provider): List[Constdef] = | |
| 585 |   {
 | |
| 586 | val body = provider.uncompressed_yxml(export_prefix + "constdefs") | |
| 587 | val constdefs = | |
| 588 |     {
 | |
| 589 | import XML.Decode._ | |
| 590 | list(pair(string, string))(body) | |
| 591 | } | |
| 592 | for ((name, axiom_name) <- constdefs) yield Constdef(name, axiom_name) | |
| 593 | } | |
| 594 | ||
| 595 | ||
| 68264 | 596 | /* HOL typedefs */ | 
| 597 | ||
| 598 | sealed case class Typedef(name: String, | |
| 599 | rep_type: Term.Typ, abs_type: Term.Typ, rep_name: String, abs_name: String, axiom_name: String) | |
| 68267 | 600 |   {
 | 
| 601 | def cache(cache: Term.Cache): Typedef = | |
| 602 | Typedef(cache.string(name), | |
| 603 | cache.typ(rep_type), | |
| 604 | cache.typ(abs_type), | |
| 605 | cache.string(rep_name), | |
| 606 | cache.string(abs_name), | |
| 607 | cache.string(axiom_name)) | |
| 608 | } | |
| 68264 | 609 | |
| 68418 | 610 | def read_typedefs(provider: Export.Provider): List[Typedef] = | 
| 611 |   {
 | |
| 612 | val body = provider.uncompressed_yxml(export_prefix + "typedefs") | |
| 613 | val typedefs = | |
| 614 |     {
 | |
| 615 | import XML.Decode._ | |
| 616 | import Term_XML.Decode._ | |
| 617 | list(pair(string, pair(typ, pair(typ, pair(string, pair(string, string))))))(body) | |
| 618 | } | |
| 619 |     for { (name, (rep_type, (abs_type, (rep_name, (abs_name, axiom_name))))) <- typedefs }
 | |
| 620 | yield Typedef(name, rep_type, abs_type, rep_name, abs_name, axiom_name) | |
| 621 | } | |
| 71202 | 622 | |
| 623 | ||
| 71248 | 624 | /* HOL datatypes */ | 
| 625 | ||
| 626 | sealed case class Datatype( | |
| 627 | pos: Position.T, | |
| 628 | name: String, | |
| 629 | co: Boolean, | |
| 630 | typargs: List[(String, Term.Sort)], | |
| 631 | typ: Term.Typ, | |
| 632 | constructors: List[(Term.Term, Term.Typ)]) | |
| 633 |   {
 | |
| 634 | def id: Option[Long] = Position.Id.unapply(pos) | |
| 635 | ||
| 636 | def cache(cache: Term.Cache): Datatype = | |
| 637 | Datatype( | |
| 638 | cache.position(pos), | |
| 639 | cache.string(name), | |
| 640 | co, | |
| 641 |         typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
 | |
| 642 | cache.typ(typ), | |
| 643 |         constructors.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) }))
 | |
| 644 | } | |
| 645 | ||
| 646 | def read_datatypes(provider: Export.Provider): List[Datatype] = | |
| 647 |   {
 | |
| 648 | val body = provider.uncompressed_yxml(export_prefix + "datatypes") | |
| 649 | val datatypes = | |
| 650 |     {
 | |
| 651 | import XML.Decode._ | |
| 652 | import Term_XML.Decode._ | |
| 653 | list(pair(properties, pair(string, pair(bool, pair(list(pair(string, sort)), | |
| 654 | pair(typ, list(pair(term, typ))))))))(body) | |
| 655 | } | |
| 656 | for ((pos, (name, (co, (typargs, (typ, constructors))))) <- datatypes) | |
| 657 | yield Datatype(pos, name, co, typargs, typ, constructors) | |
| 658 | } | |
| 659 | ||
| 660 | ||
| 71202 | 661 | /* Pure spec rules */ | 
| 662 | ||
| 663 | sealed abstract class Recursion | |
| 664 |   {
 | |
| 665 | def cache(cache: Term.Cache): Recursion = | |
| 666 |       this match {
 | |
| 667 | case Primrec(types) => Primrec(types.map(cache.string)) | |
| 668 | case Primcorec(types) => Primcorec(types.map(cache.string)) | |
| 669 | case _ => this | |
| 670 | } | |
| 671 | } | |
| 672 | case class Primrec(types: List[String]) extends Recursion | |
| 673 | case object Recdef extends Recursion | |
| 674 | case class Primcorec(types: List[String]) extends Recursion | |
| 675 | case object Corec extends Recursion | |
| 676 | case object Unknown_Recursion extends Recursion | |
| 677 | ||
| 678 | val decode_recursion: XML.Decode.T[Recursion] = | |
| 679 |   {
 | |
| 680 | import XML.Decode._ | |
| 681 | variant(List( | |
| 682 |       { case (Nil, a) => Primrec(list(string)(a)) },
 | |
| 683 |       { case (Nil, Nil) => Recdef },
 | |
| 684 |       { case (Nil, a) => Primcorec(list(string)(a)) },
 | |
| 685 |       { case (Nil, Nil) => Corec },
 | |
| 686 |       { case (Nil, Nil) => Unknown_Recursion }))
 | |
| 687 | } | |
| 688 | ||
| 689 | ||
| 690 | sealed abstract class Rough_Classification | |
| 691 |   {
 | |
| 692 | def is_equational: Boolean = this.isInstanceOf[Equational] | |
| 693 | def is_inductive: Boolean = this == Inductive | |
| 694 | def is_co_inductive: Boolean = this == Co_Inductive | |
| 695 | def is_relational: Boolean = is_inductive || is_co_inductive | |
| 696 | def is_unknown: Boolean = this == Unknown | |
| 697 | ||
| 698 | def cache(cache: Term.Cache): Rough_Classification = | |
| 699 |       this match {
 | |
| 700 | case Equational(recursion) => Equational(recursion.cache(cache)) | |
| 701 | case _ => this | |
| 702 | } | |
| 703 | } | |
| 704 | case class Equational(recursion: Recursion) extends Rough_Classification | |
| 705 | case object Inductive extends Rough_Classification | |
| 706 | case object Co_Inductive extends Rough_Classification | |
| 707 | case object Unknown extends Rough_Classification | |
| 708 | ||
| 709 | val decode_rough_classification: XML.Decode.T[Rough_Classification] = | |
| 710 |   {
 | |
| 711 | import XML.Decode._ | |
| 712 | variant(List( | |
| 713 |       { case (Nil, a) => Equational(decode_recursion(a)) },
 | |
| 714 |       { case (Nil, Nil) => Inductive },
 | |
| 715 |       { case (Nil, Nil) => Co_Inductive },
 | |
| 716 |       { case (Nil, Nil) => Unknown }))
 | |
| 717 | } | |
| 718 | ||
| 719 | ||
| 720 | sealed case class Spec_Rule( | |
| 71207 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 wenzelm parents: 
71202diff
changeset | 721 | pos: Position.T, | 
| 71202 | 722 | name: String, | 
| 723 | rough_classification: Rough_Classification, | |
| 71208 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 724 | typargs: List[(String, Term.Sort)], | 
| 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 725 | args: List[(String, Term.Typ)], | 
| 71211 | 726 | terms: List[(Term.Term, Term.Typ)], | 
| 71208 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 727 | rules: List[Term.Term]) | 
| 71202 | 728 |   {
 | 
| 71207 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 wenzelm parents: 
71202diff
changeset | 729 | def id: Option[Long] = Position.Id.unapply(pos) | 
| 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 wenzelm parents: 
71202diff
changeset | 730 | |
| 71202 | 731 | def cache(cache: Term.Cache): Spec_Rule = | 
| 732 | Spec_Rule( | |
| 71207 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 wenzelm parents: 
71202diff
changeset | 733 | cache.position(pos), | 
| 71202 | 734 | cache.string(name), | 
| 735 | rough_classification.cache(cache), | |
| 71208 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 736 |         typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }),
 | 
| 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 737 |         args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }),
 | 
| 71211 | 738 |         terms.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) }),
 | 
| 71208 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 739 | rules.map(cache.term(_))) | 
| 71202 | 740 | } | 
| 741 | ||
| 742 | def read_spec_rules(provider: Export.Provider): List[Spec_Rule] = | |
| 743 |   {
 | |
| 744 | val body = provider.uncompressed_yxml(export_prefix + "spec_rules") | |
| 745 | val spec_rules = | |
| 746 |     {
 | |
| 747 | import XML.Decode._ | |
| 748 | import Term_XML.Decode._ | |
| 71207 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 wenzelm parents: 
71202diff
changeset | 749 | list( | 
| 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 wenzelm parents: 
71202diff
changeset | 750 | pair(properties, pair(string, pair(decode_rough_classification, | 
| 71208 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 751 | pair(list(pair(string, sort)), pair(list(pair(string, typ)), | 
| 71211 | 752 | pair(list(pair(term, typ)), list(term))))))))(body) | 
| 71202 | 753 | } | 
| 71208 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 754 | for ((pos, (name, (rough_classification, (typargs, (args, (terms, rules)))))) <- spec_rules) | 
| 
5e0050eb64f2
clarified export of spec rules: more like locale;
 wenzelm parents: 
71207diff
changeset | 755 | yield Spec_Rule(pos, name, rough_classification, typargs, args, terms, rules) | 
| 71202 | 756 | } | 
| 68171 | 757 | } |