diff --git a/src/Pure/General/sql.scala b/src/Pure/General/sql.scala --- a/src/Pure/General/sql.scala +++ b/src/Pure/General/sql.scala @@ -1,627 +1,634 @@ /* Title: Pure/General/sql.scala Author: Makarius Support for SQL databases: SQLite and PostgreSQL. */ package isabelle import java.time.OffsetDateTime import java.sql.{DriverManager, Connection, PreparedStatement, ResultSet} import org.sqlite.jdbc4.JDBC4Connection import org.postgresql.{PGConnection, PGNotification} import scala.collection.mutable object SQL { /** SQL language **/ type Source = String /* concrete syntax */ def escape_char(c: Char): String = c match { case '\u0000' => "\\0" case '\'' => "\\'" case '\"' => "\\\"" case '\b' => "\\b" case '\n' => "\\n" case '\r' => "\\r" case '\t' => "\\t" case '\u001a' => "\\Z" case '\\' => "\\\\" case _ => c.toString } def string(s: String): Source = s.iterator.map(escape_char).mkString("'", "", "'") def ident(s: String): Source = Long_Name.implode(Long_Name.explode(s).map(a => quote(a.replace("\"", "\"\"")))) def enclose(s: Source): Source = "(" + s + ")" def enclosure(ss: Iterable[Source]): Source = ss.mkString("(", ", ", ")") def separate(sql: Source): Source = (if (sql.isEmpty || sql.startsWith(" ")) "" else " ") + sql def select(columns: List[Column] = Nil, distinct: Boolean = false, sql: Source = ""): Source = "SELECT " + (if (distinct) "DISTINCT " else "") + (if (columns.isEmpty) "*" else commas(columns.map(_.ident))) + " FROM " + sql val join_outer: Source = " LEFT OUTER JOIN " val join_inner: Source = " INNER JOIN " def join(outer: Boolean): Source = if (outer) join_outer else join_inner def infix(op: Source, args: Iterable[Source]): Source = { val body = args.iterator.filter(_.nonEmpty).mkString(" " + op + " ") if_proper(body, enclose(body)) } def AND(args: Iterable[Source]): Source = infix("AND", args) def OR(args: Iterable[Source]): Source = infix("OR", args) def and(args: Source*): Source = AND(args) def or(args: Source*): Source = OR(args) val TRUE: Source = "TRUE" val FALSE: Source = "FALSE" def equal(sql: Source, s: String): Source = sql + " = " + string(s) def member(sql: Source, set: Iterable[String]): Source = if (set.isEmpty) FALSE else OR(set.iterator.map(equal(sql, _)).toList) def where(sql: Source): Source = if_proper(sql, " WHERE " + sql) /* types */ object Type extends Enumeration { val Boolean = Value("BOOLEAN") val Int = Value("INTEGER") val Long = Value("BIGINT") val Double = Value("DOUBLE PRECISION") val String = Value("TEXT") val Bytes = Value("BLOB") val Date = Value("TIMESTAMP WITH TIME ZONE") } def sql_type_default(T: Type.Value): Source = T.toString def sql_type_sqlite(T: Type.Value): Source = if (T == Type.Boolean) "INTEGER" else if (T == Type.Date) "TEXT" else sql_type_default(T) def sql_type_postgresql(T: Type.Value): Source = if (T == Type.Bytes) "BYTEA" else sql_type_default(T) /* columns */ object Column { def bool(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Boolean, strict, primary_key) def int(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Int, strict, primary_key) def long(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Long, strict, primary_key) def double(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Double, strict, primary_key) def string(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.String, strict, primary_key) def bytes(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Bytes, strict, primary_key) def date(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Date, strict, primary_key) } sealed case class Column( name: String, T: Type.Value, strict: Boolean = false, primary_key: Boolean = false, expr: SQL.Source = "" ) { def make_primary_key: Column = copy(primary_key = true) def apply(table: Table): Column = Column(Long_Name.qualify(table.name, name), T, strict = strict, primary_key = primary_key) def ident: Source = if (expr == "") SQL.ident(name) else enclose(expr) + " AS " + SQL.ident(name) def decl(sql_type: Type.Value => Source): Source = ident + " " + sql_type(T) + (if (strict || primary_key) " NOT NULL" else "") def defined: String = ident + " IS NOT NULL" def undefined: String = ident + " IS NULL" def equal(s: String): Source = SQL.equal(ident, s) def member(set: Iterable[String]): Source = SQL.member(ident, set) def where_equal(s: String): Source = SQL.where(equal(s)) def where_member(set: Iterable[String]): Source = SQL.where(member(set)) override def toString: Source = ident } def order_by(columns: List[Column], descending: Boolean = false): Source = " ORDER BY " + columns.mkString(", ") + (if (descending) " DESC" else "") /* tables */ sealed case class Table(name: String, columns: List[Column], body: Source = "") { Library.duplicates(columns.map(_.name)) match { case Nil => case bad => error("Duplicate column names " + commas_quote(bad) + " for table " + quote(name)) } def ident: Source = SQL.ident(name) def query: Source = if (body == "") error("Missing SQL body for table " + quote(name)) else SQL.enclose(body) def query_named: Source = query + " AS " + SQL.ident(name) def create(strict: Boolean, sql_type: Type.Value => Source): Source = { val primary_key = columns.filter(_.primary_key).map(_.name) match { case Nil => Nil case keys => List("PRIMARY KEY " + enclosure(keys)) } "CREATE TABLE " + (if (strict) "" else "IF NOT EXISTS ") + ident + " " + enclosure(columns.map(_.decl(sql_type)) ::: primary_key) } def create_index(index_name: String, index_columns: List[Column], strict: Boolean = false, unique: Boolean = false): Source = "CREATE " + (if (unique) "UNIQUE " else "") + "INDEX " + (if (strict) "" else "IF NOT EXISTS ") + SQL.ident(index_name) + " ON " + ident + " " + enclosure(index_columns.map(_.name)) def insert_cmd(cmd: Source = "INSERT", sql: Source = ""): Source = cmd + " INTO " + ident + " VALUES " + enclosure(columns.map(_ => "?")) + SQL.separate(sql) def insert(sql: Source = ""): Source = insert_cmd(sql = sql) def delete(sql: Source = ""): Source = "DELETE FROM " + ident + SQL.separate(sql) def update(update_columns: List[Column] = Nil, sql: Source = ""): Source = "UPDATE " + ident + " SET " + commas(update_columns.map(c => c.ident + " = ?")) + SQL.separate(sql) def select( select_columns: List[Column] = Nil, distinct: Boolean = false, sql: Source = "" ): Source = SQL.select(select_columns, distinct = distinct, sql = ident + SQL.separate(sql)) override def toString: Source = ident } /* table groups */ object Tables { def list(list: List[Table]): Tables = new Tables(list) val empty: Tables = list(Nil) def apply(args: Table*): Tables = list(args.toList) } final class Tables private(val list: List[Table]) extends Iterable[Table] { override def toString: String = list.mkString("SQL.Tables(", ", ", ")") def ::: (other: Tables): Tables = new Tables(other.list ::: list) def iterator: Iterator[Table] = list.iterator // requires transaction def create_lock(db: Database): Unit = { foreach(db.create_table(_)) lock(db) } // requires transaction def lock(db: Database): Unit = { val sql = db.lock_tables(list) if (sql.nonEmpty) db.execute_statement(sql) } } /** SQL database operations **/ /* statements */ class Statement private[SQL](val db: Database, val rep: PreparedStatement) extends AutoCloseable { stmt => object bool { def update(i: Int, x: Boolean): Unit = rep.setBoolean(i, x) def update(i: Int, x: Option[Boolean]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BOOLEAN) } } object int { def update(i: Int, x: Int): Unit = rep.setInt(i, x) def update(i: Int, x: Option[Int]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.INTEGER) } } object long { def update(i: Int, x: Long): Unit = rep.setLong(i, x) def update(i: Int, x: Option[Long]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BIGINT) } } object double { def update(i: Int, x: Double): Unit = rep.setDouble(i, x) def update(i: Int, x: Option[Double]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.DOUBLE) } } object string { def update(i: Int, x: String): Unit = rep.setString(i, x) def update(i: Int, x: Option[String]): Unit = update(i, x.orNull) } object bytes { def update(i: Int, bytes: Bytes): Unit = { if (bytes == null) rep.setBytes(i, null) else rep.setBinaryStream(i, bytes.stream(), bytes.length) } def update(i: Int, bytes: Option[Bytes]): Unit = update(i, bytes.orNull) } object date { def update(i: Int, date: Date): Unit = db.update_date(stmt, i, date) def update(i: Int, date: Option[Date]): Unit = update(i, date.orNull) } def execute(): Boolean = rep.execute() def execute_query(): Result = new Result(this, rep.executeQuery()) def close(): Unit = rep.close() } /* results */ class Result private[SQL](val stmt: Statement, val rep: ResultSet) { res => def next(): Boolean = rep.next() def iterator[A](get: Result => A): Iterator[A] = new Iterator[A] { private var _next: Boolean = res.next() def hasNext: Boolean = _next def next(): A = { val x = get(res); _next = res.next(); x } } def bool(column: Column): Boolean = rep.getBoolean(column.name) def int(column: Column): Int = rep.getInt(column.name) def long(column: Column): Long = rep.getLong(column.name) def double(column: Column): Double = rep.getDouble(column.name) def string(column: Column): String = { val s = rep.getString(column.name) if (s == null) "" else s } def bytes(column: Column): Bytes = { val bs = rep.getBytes(column.name) if (bs == null) Bytes.empty else Bytes(bs) } def date(column: Column): Date = stmt.db.date(res, column) def timing(c1: Column, c2: Column, c3: Column): Timing = Timing(Time.ms(long(c1)), Time.ms(long(c2)), Time.ms(long(c3))) def get[A](column: Column, f: Column => A): Option[A] = { val x = f(column) if (rep.wasNull || x == null) None else Some(x) } def get_bool(column: Column): Option[Boolean] = get(column, bool) def get_int(column: Column): Option[Int] = get(column, int) def get_long(column: Column): Option[Long] = get(column, long) def get_double(column: Column): Option[Double] = get(column, double) def get_string(column: Column): Option[String] = get(column, string) def get_bytes(column: Column): Option[Bytes] = get(column, bytes) def get_date(column: Column): Option[Date] = get(column, date) } /* database */ trait Database extends AutoCloseable { db => def is_sqlite: Boolean = isInstanceOf[SQLite.Database] def is_postgresql: Boolean = isInstanceOf[PostgreSQL.Database] def vacuum(tables: SQL.Tables = SQL.Tables.empty): Unit def now(): Date /* types */ def sql_type(T: Type.Value): Source /* connection */ def connection: Connection def sqlite_connection: Option[JDBC4Connection] = connection match { case conn: JDBC4Connection => Some(conn) case _ => None } def postgresql_connection: Option[PGConnection] = connection match { case conn: PGConnection => Some(conn) case _ => None } def the_sqlite_connection: JDBC4Connection = sqlite_connection getOrElse error("SQLite connection expected, but found " + connection.getClass.getName) def the_postgresql_connection: PGConnection = postgresql_connection getOrElse error("PostgreSQL connection expected, but found " + connection.getClass.getName) def close(): Unit = connection.close() def transaction[A](body: => A): A = { val auto_commit = connection.getAutoCommit() try { connection.setAutoCommit(false) val savepoint = connection.setSavepoint() try { val result = body connection.commit() result } catch { case exn: Throwable => connection.rollback(savepoint); throw exn } } finally { connection.setAutoCommit(auto_commit) } } def transaction_lock[A](tables: Tables)(body: => A): A = transaction { tables.lock(db); body } def lock_tables(tables: List[Table]): Source = "" // PostgreSQL only /* statements and results */ def statement(sql: Source): Statement = new Statement(db, connection.prepareStatement(sql)) def using_statement[A](sql: Source)(f: Statement => A): A = using(statement(sql))(f) def execute_statement(sql: Source, body: Statement => Unit = _ => ()): Unit = using_statement(sql) { stmt => body(stmt); stmt.execute() } def execute_query_statement[A, B]( sql: Source, make_result: Iterator[A] => B, get: Result => A ): B = using_statement(sql)(stmt => make_result(stmt.execute_query().iterator(get))) def execute_query_statementO[A](sql: Source, get: Result => A): Option[A] = execute_query_statement[A, Option[A]](sql, _.nextOption, get) def execute_query_statementB(sql: Source): Boolean = using_statement(sql)(stmt => stmt.execute_query().next()) def update_date(stmt: Statement, i: Int, date: Date): Unit def date(res: Result, column: Column): Date def insert_permissive(table: Table, sql: Source = ""): Source /* tables and views */ def tables: List[String] = { val result = new mutable.ListBuffer[String] val rs = connection.getMetaData.getTables(null, null, "%", null) while (rs.next) { result += rs.getString(3) } result.toList } - def create_table(table: Table, strict: Boolean = false, sql: Source = ""): Unit = + def create_table(table: Table, strict: Boolean = false, sql: Source = ""): Unit = { execute_statement(table.create(strict, sql_type) + SQL.separate(sql)) + if (is_postgresql) { + for (column <- table.columns if column.T == SQL.Type.Bytes) { + execute_statement( + "ALTER TABLE " + table + " ALTER COLUMN " + column + " SET STORAGE EXTERNAL") + } + } + } def create_index(table: Table, name: String, columns: List[Column], strict: Boolean = false, unique: Boolean = false): Unit = execute_statement(table.create_index(name, columns, strict, unique)) def create_view(table: Table, strict: Boolean = false): Unit = { if (strict || !tables.contains(table.name)) { execute_statement("CREATE VIEW " + table + " AS " + { table.query; table.body }) } } } } /** SQLite **/ object SQLite { // see https://www.sqlite.org/lang_datefunc.html val date_format: Date.Format = Date.Format("uuuu-MM-dd HH:mm:ss.SSS x") lazy val init_jdbc: Unit = { val lib_path = Path.explode("$ISABELLE_SQLITE_HOME/" + Platform.jvm_platform) val lib_name = File.get_file(lib_path).file_name System.setProperty("org.sqlite.lib.path", File.platform_path(lib_path)) System.setProperty("org.sqlite.lib.name", lib_name) Class.forName("org.sqlite.JDBC") } def open_database(path: Path): Database = { init_jdbc val path0 = path.expand val s0 = File.platform_path(path0) val s1 = if (Platform.is_windows) s0.replace('\\', '/') else s0 val connection = DriverManager.getConnection("jdbc:sqlite:" + s1) new Database(path0.toString, connection) } class Database private[SQLite](name: String, val connection: Connection) extends SQL.Database { override def toString: String = name override def vacuum(tables: SQL.Tables = SQL.Tables.empty): Unit = execute_statement("VACUUM") // always FULL override def now(): Date = Date.now() def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_sqlite(T) def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit = if (date == null) stmt.string(i) = (null: String) else stmt.string(i) = date_format(date) def date(res: SQL.Result, column: SQL.Column): Date = proper_string(res.string(column)) match { case None => null case Some(s) => date_format.parse(s) } def insert_permissive(table: SQL.Table, sql: SQL.Source = ""): SQL.Source = table.insert_cmd(cmd = "INSERT OR IGNORE", sql = sql) } } /** PostgreSQL **/ object PostgreSQL { type Source = SQL.Source val default_port = 5432 lazy val init_jdbc: Unit = Class.forName("org.postgresql.Driver") def open_database( user: String, password: String, database: String = "", host: String = "", port: Int = 0, ssh: Option[SSH.Session] = None, ssh_close: Boolean = false, // see https://www.postgresql.org/docs/current/transaction-iso.html transaction_isolation: Int = Connection.TRANSACTION_SERIALIZABLE ): Database = { init_jdbc if (user == "") error("Undefined database user") val db_host = proper_string(host) getOrElse "localhost" val db_port = if (port > 0 && port != default_port) ":" + port else "" val db_name = "/" + (proper_string(database) getOrElse user) val (url, name, port_forwarding) = ssh match { case None => val spec = db_host + db_port + db_name val url = "jdbc:postgresql://" + spec val name = user + "@" + spec (url, name, None) case Some(ssh) => val fw = ssh.port_forwarding(remote_host = db_host, remote_port = if (port > 0) port else default_port, ssh_close = ssh_close) val url = "jdbc:postgresql://localhost:" + fw.port + db_name val name = user + "@" + fw + db_name + " via ssh " + ssh (url, name, Some(fw)) } try { val connection = DriverManager.getConnection(url, user, password) connection.setTransactionIsolation(transaction_isolation) new Database(name, connection, port_forwarding) } catch { case exn: Throwable => port_forwarding.foreach(_.close()); throw exn } } class Database private[PostgreSQL]( name: String, val connection: Connection, port_forwarding: Option[SSH.Port_Forwarding] ) extends SQL.Database { override def toString: String = name override def vacuum(tables: SQL.Tables = SQL.Tables.empty): Unit = execute_statement("VACUUM" + if_proper(tables.list, " " + commas(tables.list.map(_.ident)))) override def now(): Date = { val now = SQL.Column.date("now") execute_query_statementO[Date]("SELECT NOW() as " + now.ident, res => res.date(now)) .getOrElse(error("Failed to get current date/time from database server " + toString)) } def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_postgresql(T) // see https://jdbc.postgresql.org/documentation/head/8-date-time.html def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit = if (date == null) stmt.rep.setObject(i, null) else stmt.rep.setObject(i, OffsetDateTime.from(date.to(Date.timezone_utc).rep)) def date(res: SQL.Result, column: SQL.Column): Date = { val obj = res.rep.getObject(column.name, classOf[OffsetDateTime]) if (obj == null) null else Date.instant(obj.toInstant) } def insert_permissive(table: SQL.Table, sql: SQL.Source = ""): SQL.Source = table.insert_cmd(sql = if_proper(sql, sql + " ") + "ON CONFLICT DO NOTHING") /* explicit locking: only applicable to PostgreSQL within transaction context */ // see https://www.postgresql.org/docs/current/sql-lock.html // see https://www.postgresql.org/docs/current/explicit-locking.html override def lock_tables(tables: List[SQL.Table]): PostgreSQL.Source = "LOCK TABLE " + tables.mkString(", ") + " IN ACCESS EXCLUSIVE MODE" /* notifications: IPC via database server */ // see https://www.postgresql.org/docs/current/sql-notify.html def listen(name: String): Unit = execute_statement("LISTEN " + SQL.ident(name)) def unlisten(name: String = "*"): Unit = execute_statement("UNLISTEN " + (if (name == "*") name else SQL.ident(name))) def notify(name: String, payload: String = ""): Unit = execute_statement("NOTIFY " + SQL.ident(name) + if_proper(payload, ", " + SQL.string(payload))) def get_notifications(): List[PGNotification] = the_postgresql_connection.getNotifications() match { case null => Nil case array => array.toList } override def close(): Unit = { super.close(); port_forwarding.foreach(_.close()) } } } diff --git a/src/Pure/Thy/export.scala b/src/Pure/Thy/export.scala --- a/src/Pure/Thy/export.scala +++ b/src/Pure/Thy/export.scala @@ -1,617 +1,616 @@ /* Title: Pure/Thy/export.scala Author: Makarius Manage theory exports: compressed blobs. */ package isabelle import scala.annotation.tailrec import scala.util.matching.Regex object Export { /* artefact names */ val DOCUMENT_ID = "PIDE/document_id" val FILES = "PIDE/files" val MARKUP = "PIDE/markup" val MESSAGES = "PIDE/messages" val DOCUMENT_PREFIX = "document/" val DOCUMENT_LATEX = DOCUMENT_PREFIX + "latex" val THEORY_PREFIX: String = "theory/" val PROOFS_PREFIX: String = "proofs/" def explode_name(s: String): List[String] = space_explode('/', s) def implode_name(elems: Iterable[String]): String = elems.mkString("/") /* SQL data model */ object Data { val session_name = SQL.Column.string("session_name").make_primary_key val theory_name = SQL.Column.string("theory_name").make_primary_key val name = SQL.Column.string("name").make_primary_key val executable = SQL.Column.bool("executable") val compressed = SQL.Column.bool("compressed") val body = SQL.Column.bytes("body") val table = SQL.Table("isabelle_exports", List(session_name, theory_name, name, executable, compressed, body)) val tables = SQL.Tables(table) def where_equal(session_name: String, theory_name: String = "", name: String = ""): SQL.Source = SQL.where( SQL.and( Data.session_name.equal(session_name), if_proper(theory_name, Data.theory_name.equal(theory_name)), if_proper(name, Data.name.equal(name)))) } def compound_name(a: String, b: String): String = if (a.isEmpty) b else a + ":" + b sealed case class Entry_Name(session: String = "", theory: String = "", name: String = "") { val compound_name: String = Export.compound_name(theory, name) def make_path(prune: Int = 0): Path = { val elems = theory :: space_explode('/', name) if (elems.length < prune + 1) { error("Cannot prune path by " + prune + " element(s): " + Path.make(elems)) } else Path.make(elems.drop(prune)) } def readable(db: SQL.Database): Boolean = { db.execute_query_statementB( Data.table.select(List(Data.name), sql = Data.where_equal(session, theory, name))) } def read(db: SQL.Database, cache: XML.Cache): Option[Entry] = db.execute_query_statementO[Entry]( Data.table.select(List(Data.executable, Data.compressed, Data.body), sql = Data.where_equal(session, theory, name)), { res => val executable = res.bool(Data.executable) val compressed = res.bool(Data.compressed) val bytes = res.bytes(Data.body) val body = Future.value(compressed, bytes) Entry(this, executable, body, cache) } ) } def read_theory_names(db: SQL.Database, session_name: String): List[String] = db.execute_query_statement( Data.table.select(List(Data.theory_name), distinct = true, sql = Data.where_equal(session_name) + SQL.order_by(List(Data.theory_name))), List.from[String], res => res.string(Data.theory_name)) def read_entry_names(db: SQL.Database, session_name: String): List[Entry_Name] = db.execute_query_statement( Data.table.select(List(Data.theory_name, Data.name), sql = Data.where_equal(session_name)) + SQL.order_by(List(Data.theory_name, Data.name)), List.from[Entry_Name], { res => Entry_Name( session = session_name, theory = res.string(Data.theory_name), name = res.string(Data.name)) }) def message(msg: String, theory_name: String, name: String): String = msg + " " + quote(name) + " for theory " + quote(theory_name) object Entry { def apply( entry_name: Entry_Name, executable: Boolean, body: Future[(Boolean, Bytes)], cache: XML.Cache ): Entry = new Entry(entry_name, executable, body, cache) def empty(theory_name: String, name: String): Entry = Entry(Entry_Name(theory = theory_name, name = name), false, Future.value(false, Bytes.empty), XML.Cache.none) def make( session_name: String, args: Protocol.Export.Args, bytes: Bytes, cache: XML.Cache ): Entry = { val body = if (args.compress) Future.fork(bytes.maybe_compress(cache = cache.compress)) else Future.value((false, bytes)) val entry_name = Entry_Name(session = session_name, theory = args.theory_name, name = args.name) Entry(entry_name, args.executable, body, cache) } } final class Entry private( val entry_name: Entry_Name, val executable: Boolean, body: Future[(Boolean, Bytes)], cache: XML.Cache ) { def session_name: String = entry_name.session def theory_name: String = entry_name.theory def name: String = entry_name.name override def toString: String = name def is_finished: Boolean = body.is_finished def cancel(): Unit = body.cancel() def compound_name: String = entry_name.compound_name def name_has_prefix(s: String): Boolean = name.startsWith(s) val name_elems: List[String] = explode_name(name) def name_extends(elems: List[String]): Boolean = name_elems.startsWith(elems) && name_elems != elems def bytes: Bytes = { val (compressed, bs) = body.join if (compressed) bs.uncompress(cache = cache.compress) else bs } def text: String = bytes.text def yxml: XML.Body = YXML.parse_body(UTF8.decode_permissive(bytes), cache = cache) def write(db: SQL.Database): Unit = { val (compressed, bs) = body.join db.execute_statement(Data.table.insert(), body = { stmt => stmt.string(1) = session_name stmt.string(2) = theory_name stmt.string(3) = name stmt.bool(4) = executable stmt.bool(5) = compressed stmt.bytes(6) = bs }) } } def make_regex(pattern: String): Regex = { @tailrec def make(result: List[String], depth: Int, chs: List[Char]): Regex = chs match { case '*' :: '*' :: rest => make("[^:]*" :: result, depth, rest) case '*' :: rest => make("[^:/]*" :: result, depth, rest) case '?' :: rest => make("[^:/]" :: result, depth, rest) case '\\' :: c :: rest => make(("\\" + c) :: result, depth, rest) case '{' :: rest => make("(" :: result, depth + 1, rest) case ',' :: rest if depth > 0 => make("|" :: result, depth, rest) case '}' :: rest if depth > 0 => make(")" :: result, depth - 1, rest) case c :: rest if ".+()".contains(c) => make(("\\" + c) :: result, depth, rest) case c :: rest => make(c.toString :: result, depth, rest) case Nil => result.reverse.mkString.r } make(Nil, 0, pattern.toList) } def make_matcher(pats: List[String]): Entry_Name => Boolean = { val regs = pats.map(make_regex) (entry_name: Entry_Name) => regs.exists(_.matches(entry_name.compound_name)) } /* database consumer thread */ def consumer(db: SQL.Database, cache: XML.Cache, progress: Progress = new Progress): Consumer = new Consumer(db, cache, progress) class Consumer private[Export](db: SQL.Database, cache: XML.Cache, progress: Progress) { private val errors = Synchronized[List[String]](Nil) private val consumer = Consumer_Thread.fork_bulk[(Entry, Boolean)](name = "export")( bulk = { case (entry, _) => entry.is_finished }, consume = { (args: List[(Entry, Boolean)]) => val results = db.transaction { for ((entry, strict) <- args) yield { if (progress.stopped) { entry.cancel() Exn.Res(()) } else if (entry.entry_name.readable(db)) { if (strict) { val msg = message("Duplicate export", entry.theory_name, entry.name) errors.change(msg :: _) } Exn.Res(()) } else Exn.capture { entry.write(db) } } } (results, true) }) - def make_entry(session_name: String, args0: Protocol.Export.Args, body: Bytes): Unit = { + def make_entry(session_name: String, args: Protocol.Export.Args, body: Bytes): Unit = { if (!progress.stopped && !body.is_empty) { - val args = if (db.is_postgresql) args0.copy(compress = false) else args0 consumer.send(Entry.make(session_name, args, body, cache) -> args.strict) } } def shutdown(close: Boolean = false): List[String] = { consumer.shutdown() if (close) db.close() errors.value.reverse ::: (if (progress.stopped) List("Export stopped") else Nil) } } /* context for database access */ def open_database_context(store: Sessions.Store): Database_Context = { val database_server = if (store.database_server) Some(store.open_database_server()) else None new Database_Context(store, database_server) } def open_session_context0(store: Sessions.Store, session: String): Session_Context = open_database_context(store).open_session0(session, close_database_context = true) def open_session_context( store: Sessions.Store, session_background: Sessions.Background, document_snapshot: Option[Document.Snapshot] = None ): Session_Context = { open_database_context(store).open_session( session_background, document_snapshot = document_snapshot, close_database_context = true) } class Database_Context private[Export]( val store: Sessions.Store, val database_server: Option[SQL.Database] ) extends AutoCloseable { database_context => override def toString: String = { val s = database_server match { case Some(db) => db.toString case None => "input_dirs = " + store.input_dirs.map(_.absolute).mkString(", ") } "Database_Context(" + s + ")" } def cache: Term.Cache = store.cache def close(): Unit = database_server.foreach(_.close()) def open_database(session: String, output: Boolean = false): Session_Database = database_server match { case Some(db) => new Session_Database(session, db) case None => new Session_Database(session, store.open_database(session, output = output)) { override def close(): Unit = db.close() } } def open_session0(session: String, close_database_context: Boolean = false): Session_Context = open_session(Sessions.background0(session), close_database_context = close_database_context) def open_session( session_background: Sessions.Background, document_snapshot: Option[Document.Snapshot] = None, close_database_context: Boolean = false ): Session_Context = { val session_name = session_background.check_errors.session_name val session_hierarchy = session_background.sessions_structure.build_hierarchy(session_name) val session_databases = database_server match { case Some(db) => session_hierarchy.map(name => new Session_Database(name, db)) case None => val attempts = session_hierarchy.map(name => name -> store.try_open_database(name, server = false)) attempts.collectFirst({ case (name, None) => name }) match { case Some(bad) => for ((_, Some(db)) <- attempts) db.close() store.error_database(bad) case None => for ((name, Some(db)) <- attempts) yield { new Session_Database(name, db) { override def close(): Unit = this.db.close() } } } } new Session_Context( database_context, session_background, session_databases, document_snapshot) { override def close(): Unit = { session_databases.foreach(_.close()) if (close_database_context) database_context.close() } } } } class Session_Database private[Export](val session: String, val db: SQL.Database) extends AutoCloseable { def close(): Unit = () lazy private [Export] val theory_names: List[String] = read_theory_names(db, session) lazy private [Export] val entry_names: List[Entry_Name] = read_entry_names(db, session) } class Session_Context private[Export]( val database_context: Database_Context, session_background: Sessions.Background, db_hierarchy: List[Session_Database], val document_snapshot: Option[Document.Snapshot] ) extends AutoCloseable { session_context => def close(): Unit = () def cache: Term.Cache = database_context.cache def sessions_structure: Sessions.Structure = session_background.sessions_structure def session_base: Sessions.Base = session_background.base def session_name: String = if (document_snapshot.isDefined) Sessions.DRAFT else session_base.session_name def session_database(session: String = session_name): Option[Session_Database] = db_hierarchy.find(_.session == session) def session_db(session: String = session_name): Option[SQL.Database] = session_database(session = session).map(_.db) def session_stack: List[String] = ((if (document_snapshot.isDefined) List(session_name) else Nil) ::: db_hierarchy.map(_.session)).reverse private def select[A]( session: String, select: Session_Database => List[A], project: Entry_Name => A, sort_key: A => String ): List[A] = { def result(name: String): List[A] = if (name == Sessions.DRAFT) { (for { snapshot <- document_snapshot.iterator entry_name <- snapshot.all_exports.keysIterator } yield project(entry_name)).toSet.toList.sortBy(sort_key) } else session_database(name).map(select).getOrElse(Nil) if (session.nonEmpty) result(session) else session_stack.flatMap(result) } def entry_names(session: String = session_name): List[Entry_Name] = select(session, _.entry_names, identity, _.compound_name) def theory_names(session: String = session_name): List[String] = select(session, _.theory_names, _.theory, identity) def get(theory: String, name: String): Option[Entry] = { def snapshot_entry: Option[Entry] = for { snapshot <- document_snapshot entry_name = Entry_Name(session = Sessions.DRAFT, theory = theory, name = name) entry <- snapshot.all_exports.get(entry_name) } yield entry def db_entry: Option[Entry] = db_hierarchy.view.map(database => Export.Entry_Name(session = database.session, theory = theory, name = name) .read(database.db, cache)) .collectFirst({ case Some(entry) => entry }) snapshot_entry orElse db_entry } def apply(theory: String, name: String, permissive: Boolean = false): Entry = get(theory, name) match { case None if permissive => Entry.empty(theory, name) case None => error("Missing export entry " + quote(compound_name(theory, name))) case Some(entry) => entry } def theory(theory: String, other_cache: Option[Term.Cache] = None): Theory_Context = new Theory_Context(session_context, theory, other_cache) def get_source_file(name: String): Option[Sessions.Source_File] = { val store = database_context.store (for { database <- db_hierarchy.iterator file <- store.read_sources(database.db, database.session, name = name).iterator } yield file).nextOption() } def source_file(name: String): Sessions.Source_File = get_source_file(name).getOrElse(error("Missing session source file " + quote(name))) def theory_source(theory: String, unicode_symbols: Boolean = false): String = { def snapshot_source: Option[String] = for { snapshot <- document_snapshot text <- snapshot.version.nodes.iterator.collectFirst( { case (name, node) if name.theory == theory => node.source }) if text.nonEmpty } yield Symbol.output(unicode_symbols, text) def db_source: Option[String] = { val theory_context = session_context.theory(theory) for { name <- theory_context.files0(permissive = true).headOption file <- get_source_file(name) } yield Symbol.output(unicode_symbols, file.bytes.text) } snapshot_source orElse db_source getOrElse "" } def classpath(): List[File.Content] = { (for { session <- session_stack.iterator info <- sessions_structure.get(session).iterator if info.export_classpath.nonEmpty matcher = make_matcher(info.export_classpath) entry_name <- entry_names(session = session).iterator if matcher(entry_name) entry <- get(entry_name.theory, entry_name.name).iterator } yield File.content(entry.entry_name.make_path(), entry.bytes)).toList } override def toString: String = "Export.Session_Context(" + commas_quote(session_stack) + ")" } class Theory_Context private[Export]( val session_context: Session_Context, val theory: String, other_cache: Option[Term.Cache] ) { def cache: Term.Cache = other_cache getOrElse session_context.cache def get(name: String): Option[Entry] = session_context.get(theory, name) def apply(name: String, permissive: Boolean = false): Entry = session_context.apply(theory, name, permissive = permissive) def yxml(name: String): XML.Body = get(name) match { case Some(entry) => entry.yxml case None => Nil } def document_id(): Option[Long] = apply(DOCUMENT_ID, permissive = true).text match { case Value.Long(id) => Some(id) case _ => None } def files0(permissive: Boolean = false): List[String] = split_lines(apply(FILES, permissive = permissive).text) def files(permissive: Boolean = false): Option[(String, List[String])] = files0(permissive = permissive) match { case Nil => None case a :: bs => Some((a, bs)) } override def toString: String = "Export.Theory_Context(" + quote(theory) + ")" } /* export to file-system */ def export_files( store: Sessions.Store, session_name: String, export_dir: Path, progress: Progress = new Progress, export_prune: Int = 0, export_list: Boolean = false, export_patterns: List[String] = Nil ): Unit = { using(store.open_database(session_name)) { db => val entry_names = read_entry_names(db, session_name) // list if (export_list) { for (entry_name <- entry_names) progress.echo(entry_name.compound_name) } // export if (export_patterns.nonEmpty) { val matcher = make_matcher(export_patterns) for { entry_name <- entry_names if matcher(entry_name) entry <- entry_name.read(db, store.cache) } { val path = export_dir + entry_name.make_path(prune = export_prune) progress.echo("export " + path + (if (entry.executable) " (executable)" else "")) Isabelle_System.make_directory(path.dir) val bytes = entry.bytes if (!path.is_file || Bytes.read(path) != bytes) Bytes.write(path, bytes) File.set_executable(path, entry.executable) } } } } /* Isabelle tool wrapper */ val default_export_dir: Path = Path.explode("export") val isabelle_tool = Isabelle_Tool("export", "retrieve theory exports", Scala_Project.here, { args => /* arguments */ var export_dir = default_export_dir var dirs: List[Path] = Nil var export_list = false var no_build = false var options = Options.init() var export_prune = 0 var export_patterns: List[String] = Nil val getopts = Getopts(""" Usage: isabelle export [OPTIONS] SESSION Options are: -O DIR output directory for exported files (default: """ + default_export_dir + """) -d DIR include session directory -l list exports -n no build of session -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NUM prune path of exported files by NUM elements -x PATTERN extract files matching pattern (e.g. "*:**" for all) List or export theory exports for SESSION: named blobs produced by isabelle build. Option -l or -x is required; option -x may be repeated. The PATTERN language resembles glob patterns in the shell, with ? and * (both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc], and variants {pattern1,pattern2,pattern3}. """, "O:" -> (arg => export_dir = Path.explode(arg)), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "l" -> (_ => export_list = true), "n" -> (_ => no_build = true), "o:" -> (arg => options = options + arg), "p:" -> (arg => export_prune = Value.Int.parse(arg)), "x:" -> (arg => export_patterns ::= arg)) val more_args = getopts(args) val session_name = more_args match { case List(session_name) if export_list || export_patterns.nonEmpty => session_name case _ => getopts.usage() } val progress = new Console_Progress() /* build */ if (!no_build) { val rc = progress.interrupt_handler { Build.build_logic(options, session_name, progress = progress, dirs = dirs) } if (rc != Process_Result.RC.ok) sys.exit(rc) } /* export files */ val store = Sessions.store(options) export_files(store, session_name, export_dir, progress = progress, export_prune = export_prune, export_list = export_list, export_patterns = export_patterns) }) }