| author | wenzelm | 
| Thu, 08 Dec 2022 22:38:03 +0100 | |
| changeset 76610 | 6e2383488a55 | 
| parent 75972 | d574b55c4e83 | 
| child 76656 | a8f452f7c503 | 
| permissions | -rw-r--r-- | 
| 75907 | 1 | /* Title: Pure/PIDE/document_info.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Persistent document information --- for presentation purposes. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 10 | object Document_Info {
 | |
| 11 | sealed case class Session( | |
| 12 | name: String, | |
| 13 | used_theories: List[String], | |
| 75971 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 14 | loaded_theories: Map[String, Theory], | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 15 | build_uuid: String | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 16 |   ) {
 | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 17 |     if (build_uuid.isEmpty) error("Missing build_uuid for session " + quote(name))
 | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 18 | } | 
| 75907 | 19 | |
| 20 |   object Theory {
 | |
| 21 | def apply( | |
| 22 | name: String, | |
| 23 | files: List[String], | |
| 24 | static_session: String, | |
| 25 | dynamic_session: String, | |
| 26 | entities: List[Export_Theory.Entity0], | |
| 27 | others: List[String] | |
| 28 |     ): Theory = {
 | |
| 29 | val entities1 = | |
| 30 | entities.filter(e => e.file.nonEmpty && Position.Range.unapply(e.pos).isDefined) | |
| 31 | new Theory(name, files, static_session, dynamic_session, entities1, others) | |
| 32 | } | |
| 33 | } | |
| 34 | ||
| 35 | class Theory private( | |
| 36 | val name: String, | |
| 37 | val files: List[String], | |
| 38 | val static_session: String, | |
| 39 | val dynamic_session: String, | |
| 40 | entities: List[Export_Theory.Entity0], | |
| 41 | others: List[String] | |
| 42 |   ) {
 | |
| 43 | override def toString: String = name | |
| 44 | ||
| 45 | val (thy_file, blobs_files) = | |
| 46 |       files match {
 | |
| 47 |         case Nil => error("Unknown theory file for " + quote(name))
 | |
| 48 | case a :: bs => | |
| 49 | def for_theory: String = " for theory " + quote(name) | |
| 50 |           if (!File.is_thy(a)) error("Bad .thy file " + quote(a) + for_theory)
 | |
| 51 |           for (b <- bs if File.is_thy(b)) error("Bad auxiliary file " + quote(b) + for_theory)
 | |
| 52 | (a, bs) | |
| 53 | } | |
| 54 | ||
| 55 | def home_session: Boolean = static_session == dynamic_session | |
| 56 | ||
| 57 | def print_short: String = | |
| 58 | if (home_session) Long_Name.base_name(name) else name | |
| 59 | ||
| 60 | def print_long: String = | |
| 61 | "theory " + quote(name) + | |
| 62 | (if (home_session) "" else " (session " + quote(dynamic_session) + ")") | |
| 63 | ||
| 64 | private lazy val by_file_range: Map[(String, Symbol.Range), List[Export_Theory.Entity0]] = | |
| 65 | entities.groupBy(entity => (entity.file, entity.range)) | |
| 66 | ||
| 67 | private lazy val by_file_kname: Map[(String, String), Export_Theory.Entity0] = | |
| 68 |       (for {
 | |
| 69 | entity <- entities | |
| 70 | file <- Position.File.unapply(entity.pos) | |
| 71 | } yield (file, entity.kname) -> entity).toMap | |
| 72 | ||
| 73 | def get_defs(file: String, range: Symbol.Range): List[Export_Theory.Entity0] = | |
| 74 | by_file_range.getOrElse((file, range), Nil) | |
| 75 | ||
| 76 | def get_def(file: String, kind: String, name: String): Option[Export_Theory.Entity0] = | |
| 77 | by_file_kname.get((file, Export_Theory.export_kind_name(kind, name))) | |
| 78 | ||
| 75941 | 79 | def elements(elements: Browser_Info.Elements): Browser_Info.Elements = | 
| 75907 | 80 | elements.copy(entity = others.foldLeft(elements.entity)(_ + _)) | 
| 81 | } | |
| 82 | ||
| 83 | val empty: Document_Info = new Document_Info(Map.empty) | |
| 84 | ||
| 85 | def read( | |
| 86 | database_context: Export.Database_Context, | |
| 87 | deps: Sessions.Deps, | |
| 88 | sessions: List[String] | |
| 89 |   ): Document_Info = {
 | |
| 90 | val sessions_structure = deps.sessions_structure | |
| 91 | val sessions_requirements = sessions_structure.build_requirements(sessions) | |
| 92 | ||
| 75910 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 93 | def read_theory(theory_context: Export.Theory_Context): Option[Document_Info.Theory] = | 
| 75907 | 94 |     {
 | 
| 95 | val session_name = theory_context.session_context.session_name | |
| 96 | val theory_name = theory_context.theory | |
| 97 | ||
| 75910 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 98 |       theory_context.files0(permissive = true) match {
 | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 99 | case Nil => None | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 100 | case files => | 
| 75972 
d574b55c4e83
read full sessions_requirements, for more complete entity hyperlinks;
 wenzelm parents: 
75971diff
changeset | 101 | val theory_export = Export_Theory.read_theory(theory_context, permissive = true) | 
| 75910 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 102 | val theory = | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 103 | Theory(theory_name, | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 104 | static_session = sessions_structure.theory_qualifier(theory_name), | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 105 | dynamic_session = session_name, | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 106 | files = files, | 
| 75972 
d574b55c4e83
read full sessions_requirements, for more complete entity hyperlinks;
 wenzelm parents: 
75971diff
changeset | 107 | entities = theory_export.entity_iterator.toList, | 
| 
d574b55c4e83
read full sessions_requirements, for more complete entity hyperlinks;
 wenzelm parents: 
75971diff
changeset | 108 | others = theory_export.others.keySet.toList) | 
| 75910 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 109 | Some(theory) | 
| 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 110 | } | 
| 75907 | 111 | } | 
| 112 | ||
| 113 |     def read_session(session_name: String): Document_Info.Session = {
 | |
| 75910 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 114 | val static_theories = deps(session_name).used_theories.map(_._1.theory) | 
| 75971 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 115 |       val (thys, build_uuid) = {
 | 
| 75907 | 116 |         using(database_context.open_session(deps.base_info(session_name))) { session_context =>
 | 
| 75971 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 117 | val thys = | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 118 |             for {
 | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 119 | theory_name <- static_theories | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 120 | theory <- read_theory(session_context.theory(theory_name)) | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 121 | } yield theory_name -> theory | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 122 | val build_uuid = | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 123 |             (for {
 | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 124 | db <- session_context.session_db(session_name) | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 125 | build <- database_context.store.read_build(db, session_name) | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 126 |             } yield build.uuid).getOrElse("")
 | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 127 | (thys, build_uuid) | 
| 75907 | 128 | } | 
| 75971 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 129 | } | 
| 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 130 | val loaded_theories0 = thys.toMap | 
| 75910 
529e4ac56904
more robust: theories could have been suppressed via option "condition";
 wenzelm parents: 
75907diff
changeset | 131 | val used_theories = static_theories.filter(loaded_theories0.keySet) | 
| 75971 
cec22f403c25
more accurate management of dependencies: change of build_uuid causes output of HTML, but already existing/current HTML is not produced again;
 wenzelm parents: 
75941diff
changeset | 132 | Session(session_name, used_theories, loaded_theories0, build_uuid) | 
| 75907 | 133 | } | 
| 134 | ||
| 135 | val result0 = | |
| 136 | (for (session <- Par_List.map(read_session, sessions_requirements).iterator) | |
| 137 | yield session.name -> session).toMap | |
| 138 | ||
| 139 | val result1 = | |
| 140 |       sessions_requirements.foldLeft(Map.empty[String, Session]) {
 | |
| 141 | case (seen, session_name) => | |
| 142 | val session0 = result0(session_name) | |
| 143 | val loaded_theories1 = | |
| 144 |             sessions_structure(session_name).parent.map(seen) match {
 | |
| 145 | case None => session0.loaded_theories | |
| 146 | case Some(parent_session) => | |
| 147 | parent_session.loaded_theories ++ session0.loaded_theories | |
| 148 | } | |
| 149 | val session1 = session0.copy(loaded_theories = loaded_theories1) | |
| 150 | seen + (session_name -> session1) | |
| 151 | } | |
| 152 | ||
| 153 | new Document_Info(result1) | |
| 154 | } | |
| 155 | } | |
| 156 | ||
| 157 | class Document_Info private(sessions: Map[String, Document_Info.Session]) {
 | |
| 158 | override def toString: String = | |
| 159 |     sessions.keysIterator.toList.sorted.mkString("Document_Info(", ", ", ")")
 | |
| 160 | ||
| 161 | def the_session(session: String): Document_Info.Session = | |
| 162 | sessions.getOrElse(session, | |
| 163 |       error("Unknown document information for session: " + quote(session)))
 | |
| 164 | ||
| 165 | def theory_by_name(session: String, theory: String): Option[Document_Info.Theory] = | |
| 166 | by_session_and_theory_name.get((session, theory)) | |
| 167 | ||
| 168 | def theory_by_file(session: String, file: String): Option[Document_Info.Theory] = | |
| 169 | by_session_and_theory_file.get((session, file)) | |
| 170 | ||
| 171 | private lazy val by_session_and_theory_name: Map[(String, String), Document_Info.Theory] = | |
| 172 |     (for {
 | |
| 173 | session <- sessions.valuesIterator | |
| 174 | theory <- session.loaded_theories.valuesIterator | |
| 175 | } yield (session.name, theory.name) -> theory).toMap | |
| 176 | ||
| 177 |   private lazy val by_session_and_theory_file: Map[(String, String), Document_Info.Theory] = {
 | |
| 178 |     (for {
 | |
| 179 | session <- sessions.valuesIterator | |
| 180 | theory <- session.loaded_theories.valuesIterator | |
| 181 | file <- theory.files.iterator | |
| 182 | } yield (session.name, file) -> theory).toMap | |
| 183 | } | |
| 184 | } |