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