--- a/Admin/component_repository/components.sha1 Fri Nov 30 17:12:01 2012 +0100
+++ b/Admin/component_repository/components.sha1 Fri Nov 30 23:06:11 2012 +0100
@@ -3,10 +3,12 @@
b98a98025d1f7e560ca6864a53296137dae736b4 e-1.6.tar.gz
6d34b18ca0aa1e10bab6413045d079188c0e2dfb exec_process-1.0.1.tar.gz
8b9bffd10e396d965e815418295f2ee2849bea75 exec_process-1.0.2.tar.gz
+e6aada354da11e533af2dee3dcdd96c06479b053 exec_process-1.0.3.tar.gz
ae7ee5becb26512f18c609e83b34612918bae5f0 exec_process-1.0.tar.gz
683acd94761ef460cca1a628f650355370de5afb hol-light-bundle-0.5-126.tar.gz
8d83e433c1419e0c0cc5fd1762903d11b4a5752c jdk-6u31.tar.gz
ec740ee9ffd43551ddf1e5b91641405116af6291 jdk-7u6.tar.gz
+7d5b152ac70f720bb9e783fa45ecadcf95069584 jdk-7u9.tar.gz
44775a22f42a9d665696bfb49e53c79371c394b0 jedit_build-20111217.tar.gz
a242a688810f2bccf24587b0062ce8027bf77fa2 jedit_build-20120304.tar.gz
4c948dee53f74361c097c08f49a1a5ff9b17bd1d jedit_build-20120307.tar.gz
--- a/Admin/components/main Fri Nov 30 17:12:01 2012 +0100
+++ b/Admin/components/main Fri Nov 30 23:06:11 2012 +0100
@@ -1,8 +1,8 @@
#main components for everyday use, without big impact on overall build time
cvc3-2.4.1
e-1.6
-exec_process-1.0.2
-jdk-7u6
+exec_process-1.0.3
+jdk-7u9
jedit_build-20120903
kodkodi-1.5.2
polyml-5.5.0
--- a/Admin/exec_process/exec_process.c Fri Nov 30 17:12:01 2012 +0100
+++ b/Admin/exec_process/exec_process.c Fri Nov 30 23:06:11 2012 +0100
@@ -28,15 +28,6 @@
char *script = argv[2];
- /* report pid */
-
- FILE *pid_file;
- pid_file = fopen(pid_name, "w");
- if (pid_file == NULL) fail("Cannot open pid file");
- fprintf(pid_file, "%d", getpid());
- fclose(pid_file);
-
-
/* setsid */
if (setsid() == -1) {
@@ -47,6 +38,15 @@
}
+ /* report pid */
+
+ FILE *pid_file;
+ pid_file = fopen(pid_name, "w");
+ if (pid_file == NULL) fail("Cannot open pid file");
+ fprintf(pid_file, "%d", getpid());
+ fclose(pid_file);
+
+
/* exec */
char *cmd_line[4];
--- a/Admin/java/build Fri Nov 30 17:12:01 2012 +0100
+++ b/Admin/java/build Fri Nov 30 23:06:11 2012 +0100
@@ -11,12 +11,12 @@
## parameters
-ARCHIVE_LINUX32="jdk-7u6-linux-i586.tar.gz"
-ARCHIVE_LINUX64="jdk-7u6-linux-x64.tar.gz"
-ARCHIVE_DARWIN="jdk1.7.0_06.jdk.tar.gz"
-ARCHIVE_WINDOWS="jdk1.7.0_06.tar.gz"
+ARCHIVE_LINUX32="jdk-7u9-linux-i586.tar.gz"
+ARCHIVE_LINUX64="jdk-7u9-linux-x64.tar.gz"
+ARCHIVE_DARWIN="jdk1.7.0_09.jdk.tar.gz"
+ARCHIVE_WINDOWS="jdk1.7.0_09.tar.gz"
-VERSION="7u6"
+VERSION="7u9"
## variations on version
@@ -47,13 +47,16 @@
# README
cat >> "$DIR/README" << EOF
-This is JDK $FULL_VERSION for Linux, Mac OS X, Windows.
+This is JDK $FULL_VERSION as required for Isabelle.
See http://www.oracle.com/technetwork/java/javase/downloads/index.html
for the original downloads, which are covered by the Oracle Binary
Code License Agreement for Java SE.
-Note that Java 1.7 requires 64bit hardware on Mac OS X.
+Linux, Mac OS X, Windows work uniformly, depending on certain
+platform-specific subdirectories.
+
+Note that Java 1.7 on Mac OS X requires 64bit hardware!
EOF
@@ -76,7 +79,12 @@
esac
if [ -n "\$ISABELLE_JDK_HOME" ]; then
- ISABELLE_JAVA_EXT="\${ISABELLE_JDK_HOME}/jre/lib/ext"
+ if [ -d "\$ISABELLE_JDK_HOME" ]; then
+ ISABELLE_JAVA_EXT="\${ISABELLE_JDK_HOME}/jre/lib/ext"
+ else
+ echo "### Missing Java 1.7 platform directory: \"\$ISABELLE_JDK_HOME\"" >&2
+ unset ISABELLE_JDK_HOME
+ fi
fi
EOF
@@ -94,6 +102,8 @@
chmod -R a+r "$DIR"
chmod -R a+X "$DIR"
+find "$DIR/x86_64-darwin" -name "._*" -exec rm -f {} ";"
+
(
cd "$DIR/x86-linux/jdk${FULL_VERSION}"
for FILE in $(find . -type f)
--- a/etc/isar-keywords.el Fri Nov 30 17:12:01 2012 +0100
+++ b/etc/isar-keywords.el Fri Nov 30 23:06:11 2012 +0100
@@ -190,6 +190,7 @@
"print_drafts"
"print_facts"
"print_induct_rules"
+ "print_inductives"
"print_interps"
"print_locale"
"print_locales"
@@ -413,6 +414,7 @@
"print_drafts"
"print_facts"
"print_induct_rules"
+ "print_inductives"
"print_interps"
"print_locale"
"print_locales"
--- a/src/Doc/IsarRef/HOL_Specific.thy Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Doc/IsarRef/HOL_Specific.thy Fri Nov 30 23:06:11 2012 +0100
@@ -66,6 +66,7 @@
@{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
@{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
@{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "print_inductives"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
@{attribute_def (HOL) mono} & : & @{text attribute} \\
\end{matharray}
@@ -135,6 +136,9 @@
native HOL predicates. This allows to define (co)inductive sets,
where multiple arguments are simulated via tuples.
+ \item @{command "print_inductives"} prints (co)inductive definitions and
+ monotonicity rules.
+
\item @{attribute (HOL) mono} declares monotonicity rules in the
context. These rule are involved in the automated monotonicity
proof of the above inductive and coinductive definitions.
--- a/src/HOL/Inductive.thy Fri Nov 30 17:12:01 2012 +0100
+++ b/src/HOL/Inductive.thy Fri Nov 30 23:06:11 2012 +0100
@@ -9,6 +9,7 @@
keywords
"inductive" "coinductive" :: thy_decl and
"inductive_cases" "inductive_simps" :: thy_script and "monos" and
+ "print_inductives" :: diag and
"rep_datatype" :: thy_goal and
"primrec" :: thy_decl
begin
--- a/src/HOL/Tools/inductive.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/HOL/Tools/inductive.ML Fri Nov 30 23:06:11 2012 +0100
@@ -220,10 +220,12 @@
val {infos, monos, ...} = rep_data ctxt;
val space = Consts.space_of (Proof_Context.consts_of ctxt);
in
- [Pretty.strs ("(co)inductives:" :: map #1 (Name_Space.extern_table ctxt (space, infos))),
+ [Pretty.block
+ (Pretty.breaks
+ (Pretty.str "(co)inductives:" ::
+ map (Pretty.mark_str o #1) (Name_Space.extern_table ctxt (space, infos)))),
Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm ctxt) monos)]
- |> Pretty.chunks |> Pretty.writeln
- end;
+ end |> Pretty.chunks |> Pretty.writeln;
(* inductive info *)
@@ -1165,4 +1167,11 @@
"create simplification rules for inductive predicates"
(Parse.and_list1 Parse_Spec.specs >> (snd oo inductive_simps));
+val _ =
+ Outer_Syntax.improper_command @{command_spec "print_inductives"}
+ "print (co)inductive definitions and monotonicity rules"
+ (Scan.succeed
+ (Toplevel.no_timing o Toplevel.unknown_context o
+ Toplevel.keep (print_inductives o Toplevel.context_of)));
+
end;
--- a/src/Pure/General/name_space.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/General/name_space.ML Fri Nov 30 23:06:11 2012 +0100
@@ -63,7 +63,7 @@
val join_tables: (string -> 'a * 'a -> 'a) (*Symtab.SAME*) ->
'a table * 'a table -> 'a table
val dest_table: Proof.context -> 'a table -> (string * 'a) list
- val extern_table: Proof.context -> 'a table -> (xstring * 'a) list
+ val extern_table: Proof.context -> 'a table -> ((Markup.T * xstring) * 'a) list
end;
structure Name_Space: NAME_SPACE =
@@ -450,8 +450,10 @@
Symtab.fold (fn (name, x) => cons ((name, extern ctxt space name), x)) tab []
|> Library.sort_wrt (#2 o #1);
-fun dest_table ctxt tab = map (apfst #1) (ext_table ctxt tab);
-fun extern_table ctxt tab = map (apfst #2) (ext_table ctxt tab);
+fun dest_table ctxt table = map (apfst #1) (ext_table ctxt table);
+
+fun extern_table ctxt (space, tab) =
+ map (fn ((name, xname), x) => ((markup space name, xname), x)) (ext_table ctxt (space, tab));
end;
--- a/src/Pure/General/symbol.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/General/symbol.scala Fri Nov 30 23:06:11 2012 +0100
@@ -380,6 +380,18 @@
def decode(text: String): String = symbols.decode(text)
def encode(text: String): String = symbols.encode(text)
+ def decode_strict(text: String): String =
+ {
+ val decoded = decode(text)
+ if (encode(decoded) == text) decoded
+ else {
+ val bad = new mutable.ListBuffer[Symbol]
+ for (s <- iterator(text) if encode(decode(s)) != s && !bad.contains(s))
+ bad += s
+ error("Bad Unicode symbols in text: " + commas_quote(bad))
+ }
+ }
+
def fonts: Map[Symbol, String] = symbols.fonts
def font_names: List[String] = symbols.font_names
def font_index: Map[String, Int] = symbols.font_index
--- a/src/Pure/General/symbol_pos.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/General/symbol_pos.ML Fri Nov 30 23:06:11 2012 +0100
@@ -251,11 +251,11 @@
val scan_ident =
Scan.one (Symbol.is_letter o symbol) ::: Scan.many (Symbol.is_letdig o symbol);
-fun is_ident [] = false
- | is_ident (s :: ss) = Symbol.is_letter (symbol s) andalso forall (Symbol.is_letdig o symbol) ss;
-
fun is_identifier s =
- Symbol.is_ascii_identifier s orelse is_ident (explode (s, Position.none));
+ Symbol.is_ascii_identifier s orelse
+ (case try (Scan.finite stopper scan_ident) (explode (s, Position.none)) of
+ SOME (_, []) => true
+ | _ => false);
end;
--- a/src/Pure/General/time.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/General/time.scala Fri Nov 30 23:06:11 2012 +0100
@@ -8,6 +8,9 @@
package isabelle
+import java.util.Locale
+
+
object Time
{
def seconds(s: Double): Time = new Time((s * 1000.0).round)
@@ -24,7 +27,7 @@
def is_relevant: Boolean = ms >= 1
override def toString =
- String.format(java.util.Locale.ROOT, "%.3f", seconds.asInstanceOf[AnyRef])
+ String.format(Locale.ROOT, "%.3f", seconds.asInstanceOf[AnyRef])
def message: String = toString + "s"
}
--- a/src/Pure/Isar/attrib.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/Isar/attrib.ML Fri Nov 30 23:06:11 2012 +0100
@@ -91,9 +91,10 @@
let
val ctxt = Proof_Context.init_global thy;
val attribs = Attributes.get thy;
- fun prt_attr (name, (_, "")) = Pretty.str name
+ fun prt_attr (name, (_, "")) = Pretty.mark_str name
| prt_attr (name, (_, comment)) =
- Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
+ Pretty.block
+ (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);
in
[Pretty.big_list "attributes:" (map prt_attr (Name_Space.extern_table ctxt attribs))]
|> Pretty.chunks |> Pretty.writeln
@@ -440,8 +441,8 @@
val thy = Proof_Context.theory_of ctxt;
fun prt (name, config) =
let val value = Config.get ctxt config in
- Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,
- Pretty.str (Config.print_value value)]
+ Pretty.block [Pretty.mark_str name, Pretty.str (": " ^ Config.print_type value ^ " ="),
+ Pretty.brk 1, Pretty.str (Config.print_value value)]
end;
val configs = Name_Space.extern_table ctxt (#1 (Attributes.get thy), Configs.get thy);
in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;
--- a/src/Pure/Isar/bundle.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/Isar/bundle.ML Fri Nov 30 23:06:11 2012 +0100
@@ -138,7 +138,7 @@
(Pretty.breaks (map prt_thm ths)) :: Attrib.pretty_attribs ctxt atts;
fun prt_bundle (name, bundle) =
- Pretty.block (Pretty.command "bundle" :: Pretty.str " " :: Pretty.str name ::
+ Pretty.block (Pretty.command "bundle" :: Pretty.str " " :: Pretty.mark_str name ::
Pretty.breaks (Pretty.str " =" :: maps prt_fact bundle));
in
map prt_bundle (Name_Space.extern_table ctxt (get_bundles ctxt))
--- a/src/Pure/Isar/locale.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/Isar/locale.ML Fri Nov 30 23:06:11 2012 +0100
@@ -608,8 +608,11 @@
(*** diagnostic commands and interfaces ***)
fun print_locales thy =
- Pretty.strs ("locales:" ::
- map #1 (Name_Space.extern_table (Proof_Context.init_global thy) (Locales.get thy)))
+ Pretty.block
+ (Pretty.breaks
+ (Pretty.str "locales:" ::
+ map (Pretty.mark_str o #1)
+ (Name_Space.extern_table (Proof_Context.init_global thy) (Locales.get thy))))
|> Pretty.writeln;
fun pretty_locale thy show_facts name =
--- a/src/Pure/Isar/method.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/Isar/method.ML Fri Nov 30 23:06:11 2012 +0100
@@ -316,9 +316,10 @@
let
val ctxt = Proof_Context.init_global thy;
val meths = Methods.get thy;
- fun prt_meth (name, (_, "")) = Pretty.str name
+ fun prt_meth (name, (_, "")) = Pretty.mark_str name
| prt_meth (name, (_, comment)) =
- Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
+ Pretty.block
+ (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment);
in
[Pretty.big_list "methods:" (map prt_meth (Name_Space.extern_table ctxt meths))]
|> Pretty.chunks |> Pretty.writeln
--- a/src/Pure/System/isabelle_system.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/System/isabelle_system.scala Fri Nov 30 23:06:11 2012 +0100
@@ -9,7 +9,6 @@
import java.lang.System
import java.util.regex.Pattern
-import java.util.Locale
import java.io.{InputStream, OutputStream, File => JFile, BufferedReader, InputStreamReader,
BufferedWriter, OutputStreamWriter, IOException, FileInputStream, BufferedInputStream}
import java.awt.{GraphicsEnvironment, Font}
--- a/src/Pure/System/options.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/System/options.scala Fri Nov 30 23:06:11 2012 +0100
@@ -7,7 +7,6 @@
package isabelle
-import java.util.Locale
import java.util.Calendar
@@ -22,7 +21,7 @@
sealed abstract class Type
{
- def print: String = toString.toLowerCase(Locale.ENGLISH)
+ def print: String = Library.lowercase(toString)
}
case object Bool extends Type
case object Int extends Type
--- a/src/Pure/System/standard_system.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/System/standard_system.scala Fri Nov 30 23:06:11 2012 +0100
@@ -9,7 +9,6 @@
import java.lang.System
import java.util.regex.Pattern
-import java.util.Locale
import java.net.URL
import java.io.{File => JFile}
@@ -94,7 +93,7 @@
val rest =
posix_path match {
case Cygdrive(drive, rest) =>
- result_path ++= (drive.toUpperCase(Locale.ENGLISH) + ":" + JFile.separator)
+ result_path ++= (Library.uppercase(drive) + ":" + JFile.separator)
rest
case Named_Root(root, rest) =>
result_path ++= JFile.separator
@@ -129,7 +128,7 @@
jvm_path.replace('/', '\\') match {
case Platform_Root(rest) => "/" + rest.replace('\\', '/')
case Drive(letter, rest) =>
- "/cygdrive/" + letter.toLowerCase(Locale.ENGLISH) +
+ "/cygdrive/" + Library.lowercase(letter) +
(if (rest == "") "" else "/" + rest.replace('\\', '/'))
case path => path.replace('\\', '/')
}
--- a/src/Pure/Thy/thy_load.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/Thy/thy_load.scala Fri Nov 30 23:06:11 2012 +0100
@@ -47,7 +47,10 @@
{
val path = Path.explode(name.node)
if (!path.is_file) error("No such file: " + path.toString)
- f(File.read(path))
+
+ val text = File.read(path)
+ Symbol.decode_strict(text)
+ f(text)
}
--- a/src/Pure/Thy/thy_output.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/Thy/thy_output.ML Fri Nov 30 23:06:11 2012 +0100
@@ -127,8 +127,8 @@
val command_names = map #1 (Name_Space.extern_table ctxt commands);
val option_names = map #1 (Name_Space.extern_table ctxt options);
in
- [Pretty.big_list "document antiquotations:" (map Pretty.str command_names),
- Pretty.big_list "document antiquotation options:" (map Pretty.str option_names)]
+ [Pretty.big_list "document antiquotations:" (map Pretty.mark_str command_names),
+ Pretty.big_list "document antiquotation options:" (map Pretty.mark_str option_names)]
|> Pretty.chunks |> Pretty.writeln
end;
--- a/src/Pure/display.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/display.ML Fri Nov 30 23:06:11 2012 +0100
@@ -117,7 +117,7 @@
val prt_typ_no_tvars = prt_typ o Logic.unvarifyT_global;
fun prt_term t = Pretty.quote (Syntax.pretty_term ctxt t);
val prt_term_no_vars = prt_term o Logic.unvarify_global;
- fun prt_const (c, ty) = [Pretty.str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty];
+ fun prt_const (c, ty) = [Pretty.mark_str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty];
val prt_const' = Defs.pretty_const ctxt;
fun pretty_classrel (c, []) = prt_cls c
@@ -145,7 +145,8 @@
fun pretty_abbrev (c, (ty, t)) = Pretty.block
(prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]);
- fun pretty_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, prt_term_no_vars t];
+ fun pretty_axm (a, t) =
+ Pretty.block [Pretty.mark_str a, Pretty.str ":", Pretty.brk 1, prt_term_no_vars t];
fun pretty_finals reds = Pretty.block
(Pretty.str "final:" :: Pretty.brk 1 :: Pretty.commas (map (prt_const' o fst) reds));
@@ -202,7 +203,8 @@
Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs),
Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs),
Pretty.big_list "axioms:" (map pretty_axm axms),
- Pretty.strs ("oracles:" :: Thm.extern_oracles ctxt),
+ Pretty.block
+ (Pretty.breaks (Pretty.str "oracles:" :: map Pretty.mark_str (Thm.extern_oracles ctxt))),
Pretty.big_list "definitions:"
[pretty_finals reds0,
Pretty.big_list "non-overloaded:" (map pretty_reduct reds1),
--- a/src/Pure/library.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/library.scala Fri Nov 30 23:06:11 2012 +0100
@@ -85,9 +85,12 @@
if (str.endsWith("\n")) str.substring(0, str.length - 1)
else str
+ def lowercase(str: String): String = str.toLowerCase(Locale.ENGLISH)
+ def uppercase(str: String): String = str.toUpperCase(Locale.ENGLISH)
+
def capitalize(str: String): String =
if (str.length == 0) str
- else str.substring(0, 1).toUpperCase(Locale.ENGLISH) + str.substring(1)
+ else uppercase(str.substring(0, 1)) + str.substring(1)
/* quote */
--- a/src/Pure/thm.ML Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Pure/thm.ML Fri Nov 30 23:06:11 2012 +0100
@@ -147,7 +147,7 @@
val compose_no_flatten: bool -> thm * int -> int -> thm -> thm Seq.seq
val bicompose: bool -> bool * thm * int -> int -> thm -> thm Seq.seq
val biresolution: bool -> (bool * thm) list -> int -> thm -> thm Seq.seq
- val extern_oracles: Proof.context -> xstring list
+ val extern_oracles: Proof.context -> (Markup.T * xstring) list
val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory
end;
--- a/src/Tools/jEdit/lib/Tools/jedit Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/lib/Tools/jedit Fri Nov 30 23:06:11 2012 +0100
@@ -35,10 +35,10 @@
"src/rich_text_area.scala"
"src/scala_console.scala"
"src/sendback.scala"
- "src/session_dockable.scala"
"src/symbols_dockable.scala"
"src/syslog_dockable.scala"
"src/text_overview.scala"
+ "src/theories_dockable.scala"
"src/token_markup.scala"
)
--- a/src/Tools/jEdit/src/Isabelle.props Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/src/Isabelle.props Fri Nov 30 23:06:11 2012 +0100
@@ -27,12 +27,11 @@
options.isabelle-rendering.code=new isabelle.jedit.Isabelle_Options2();
#actions
-isabelle.check-buffer.label=Commence full proof checking of current buffer
isabelle.check-buffer.shortcut=C+e SPACE
-isabelle.cancel-execution.label=Cancel current proof checking process
isabelle.cancel-execution.shortcut=C+e BACK_SPACE
isabelle.increase-font-size.label=Increase font size
isabelle.increase-font-size.shortcut=C+PLUS
+isabelle.increase-font-size.shortcut2=C+EQUALS
isabelle.decrease-font-size.label=Decrease font size
isabelle.decrease-font-size.shortcut=C+MINUS
isabelle.control-isub.label=Control subscript
@@ -46,8 +45,8 @@
#menu actions
plugin.isabelle.jedit.Plugin.menu.label=Isabelle
-plugin.isabelle.jedit.Plugin.menu=isabelle.session-panel isabelle.output-panel isabelle.graphview-panel isabelle.raw-output-panel isabelle.protocol-panel isabelle.readme-panel isabelle.symbols-panel isabelle.syslog-panel
-isabelle.session-panel.label=Prover Session panel
+plugin.isabelle.jedit.Plugin.menu=isabelle.theories-panel isabelle.output-panel isabelle.graphview-panel isabelle.raw-output-panel isabelle.protocol-panel isabelle.readme-panel isabelle.symbols-panel isabelle.syslog-panel
+isabelle.theories-panel.label=Theories panel
isabelle.output-panel.label=Output panel
isabelle.graphview-panel.label=Graphview panel
isabelle.raw-output-panel.label=Raw Output panel
@@ -57,7 +56,7 @@
isabelle.syslog-panel.label=Syslog panel
#dockables
-isabelle-session.title=Prover Session
+isabelle-theories.title=Theories
isabelle-output.title=Output
isabelle-info.title=Info
isabelle-graphview.title=Graphview
--- a/src/Tools/jEdit/src/actions.xml Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/src/actions.xml Fri Nov 30 23:06:11 2012 +0100
@@ -2,9 +2,9 @@
<!DOCTYPE ACTIONS SYSTEM "actions.dtd">
<ACTIONS>
- <ACTION NAME="isabelle.session-panel">
+ <ACTION NAME="isabelle.theories-panel">
<CODE>
- wm.addDockableWindow("isabelle-session");
+ wm.addDockableWindow("isabelle-theories");
</CODE>
</ACTION>
<ACTION NAME="isabelle.syslog-panel">
--- a/src/Tools/jEdit/src/dockables.xml Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/src/dockables.xml Fri Nov 30 23:06:11 2012 +0100
@@ -2,8 +2,8 @@
<!DOCTYPE DOCKABLES SYSTEM "dockables.dtd">
<DOCKABLES>
- <DOCKABLE NAME="isabelle-session" MOVABLE="TRUE">
- new isabelle.jedit.Session_Dockable(view, position);
+ <DOCKABLE NAME="isabelle-theories" MOVABLE="TRUE">
+ new isabelle.jedit.Theories_Dockable(view, position);
</DOCKABLE>
<DOCKABLE NAME="isabelle-syslog" MOVABLE="TRUE">
new isabelle.jedit.Syslog_Dockable(view, position);
--- a/src/Tools/jEdit/src/isabelle.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/src/isabelle.scala Fri Nov 30 23:06:11 2012 +0100
@@ -20,9 +20,9 @@
private def wm(view: View): DockableWindowManager = view.getDockableWindowManager
- def docked_session(view: View): Option[Session_Dockable] =
- wm(view).getDockableWindow("isabelle-session") match {
- case dockable: Session_Dockable => Some(dockable)
+ def docked_theories(view: View): Option[Theories_Dockable] =
+ wm(view).getDockableWindow("isabelle-theories") match {
+ case dockable: Theories_Dockable => Some(dockable)
case _ => None
}
--- a/src/Tools/jEdit/src/jEdit.props Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/src/jEdit.props Fri Nov 30 23:06:11 2012 +0100
@@ -181,7 +181,7 @@
isabelle-output.height=174
isabelle-output.width=412
isabelle-readme.dock-position=bottom
-isabelle-session.dock-position=bottom
+isabelle-theories.dock-position=bottom
isabelle-symbols.dock-position=bottom
line-end.shortcut=END
line-home.shortcut=HOME
--- a/src/Tools/jEdit/src/session_dockable.scala Fri Nov 30 17:12:01 2012 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,176 +0,0 @@
-/* Title: Tools/jEdit/src/session_dockable.scala
- Author: Makarius
-
-Dockable window for prover session management.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.actors.Actor._
-import scala.swing.{FlowPanel, Button, TextArea, Label, ListView, Alignment, ScrollPane, Component}
-import scala.swing.event.{ButtonClicked, MouseClicked}
-
-import java.lang.System
-import java.awt.{BorderLayout, Graphics2D, Insets}
-import javax.swing.{JList, BorderFactory}
-import javax.swing.border.{BevelBorder, SoftBevelBorder}
-
-import org.gjt.sp.jedit.{View, jEdit}
-
-
-class Session_Dockable(view: View, position: String) extends Dockable(view, position)
-{
- /* status */
-
- private val status = new ListView(Nil: List[Document.Node.Name]) {
- listenTo(mouse.clicks)
- reactions += {
- case MouseClicked(_, point, _, clicks, _) if clicks == 2 =>
- val index = peer.locationToIndex(point)
- if (index >= 0) jEdit.openFile(view, listData(index).node)
- }
- }
- status.peer.setLayoutOrientation(JList.HORIZONTAL_WRAP)
- status.peer.setVisibleRowCount(0)
- status.selection.intervalMode = ListView.IntervalMode.Single
-
- set_content(new ScrollPane(status))
-
-
- /* controls */
-
- private val session_phase = new Label(PIDE.session.phase.toString)
- session_phase.border = new SoftBevelBorder(BevelBorder.LOWERED)
- session_phase.tooltip = "Prover status"
-
- private def handle_phase(phase: Session.Phase)
- {
- Swing_Thread.later { session_phase.text = " " + phase.toString + " " }
- }
-
- private val cancel = new Button("Cancel") {
- reactions += { case ButtonClicked(_) => PIDE.cancel_execution() }
- }
- cancel.tooltip = jEdit.getProperty("isabelle.cancel-execution.label")
-
- private val check = new Button("Check") {
- reactions += { case ButtonClicked(_) => PIDE.check_buffer(view.getBuffer) }
- }
- check.tooltip = jEdit.getProperty("isabelle.check-buffer.label")
-
- private val logic = Isabelle_Logic.logic_selector(true)
-
- private val controls =
- new FlowPanel(FlowPanel.Alignment.Right)(check, cancel, session_phase, logic)
- add(controls.peer, BorderLayout.NORTH)
-
-
- /* component state -- owned by Swing thread */
-
- private var nodes_status: Map[Document.Node.Name, Protocol.Node_Status] = Map.empty
-
- private object Node_Renderer_Component extends Label
- {
- opaque = false
- xAlignment = Alignment.Leading
- border = BorderFactory.createEmptyBorder(2, 2, 2, 2)
-
- var node_name = Document.Node.Name.empty
- override def paintComponent(gfx: Graphics2D)
- {
- nodes_status.get(node_name) match {
- case Some(st) if st.total > 0 =>
- val size = peer.getSize()
- val insets = border.getBorderInsets(this.peer)
- val w = size.width - insets.left - insets.right
- val h = size.height - insets.top - insets.bottom
- var end = size.width - insets.right
- for {
- (n, color) <- List(
- (st.unprocessed, PIDE.options.color_value("unprocessed1_color")),
- (st.running, PIDE.options.color_value("running_color")),
- (st.warned, PIDE.options.color_value("warning_color")),
- (st.failed, PIDE.options.color_value("error_color"))) }
- {
- gfx.setColor(color)
- val v = (n * w / st.total) max (if (n > 0) 2 else 0)
- gfx.fillRect(end - v, insets.top, v, h)
- end -= v
- }
- case _ =>
- }
- super.paintComponent(gfx)
- }
- }
-
- private class Node_Renderer extends ListView.Renderer[Document.Node.Name]
- {
- def componentFor(list: ListView[_], isSelected: Boolean, focused: Boolean,
- name: Document.Node.Name, index: Int): Component =
- {
- val component = Node_Renderer_Component
- component.node_name = name
- component.text = name.theory
- component
- }
- }
- status.renderer = new Node_Renderer
-
- private def handle_update(restriction: Option[Set[Document.Node.Name]] = None)
- {
- Swing_Thread.now {
- val snapshot = PIDE.session.snapshot()
-
- val iterator =
- restriction match {
- case Some(names) => names.iterator.map(name => (name, snapshot.version.nodes(name)))
- case None => snapshot.version.nodes.entries
- }
- val nodes_status1 =
- (nodes_status /: iterator)({ case (status, (name, node)) =>
- if (PIDE.thy_load.loaded_theories(name.theory)) status
- else status + (name -> Protocol.node_status(snapshot.state, snapshot.version, node)) })
-
- if (nodes_status != nodes_status1) {
- nodes_status = nodes_status1
- status.listData =
- snapshot.version.nodes.topological_order.filter(
- (name: Document.Node.Name) => nodes_status.isDefinedAt(name))
- }
- }
- }
-
-
- /* main actor */
-
- private val main_actor = actor {
- loop {
- react {
- case phase: Session.Phase => handle_phase(phase)
-
- case _: Session.Global_Options => Swing_Thread.later { logic.load () }
-
- case changed: Session.Commands_Changed => handle_update(Some(changed.nodes))
-
- case bad => System.err.println("Session_Dockable: ignoring bad message " + bad)
- }
- }
- }
-
- override def init()
- {
- PIDE.session.phase_changed += main_actor; handle_phase(PIDE.session.phase)
- PIDE.session.global_options += main_actor
- PIDE.session.commands_changed += main_actor; handle_update()
- }
-
- override def exit()
- {
- PIDE.session.phase_changed -= main_actor
- PIDE.session.global_options -= main_actor
- PIDE.session.commands_changed -= main_actor
- }
-}
--- a/src/Tools/jEdit/src/symbols_dockable.scala Fri Nov 30 17:12:01 2012 +0100
+++ b/src/Tools/jEdit/src/symbols_dockable.scala Fri Nov 30 23:06:11 2012 +0100
@@ -10,16 +10,18 @@
import isabelle._
import java.awt.Font
-import org.gjt.sp.jedit.View
import scala.swing.event.ValueChanged
import scala.swing.{Action, Button, FlowPanel, TabbedPane, TextField, BorderPanel, Label}
+import org.gjt.sp.jedit.View
+
+
class Symbols_Dockable(view: View, position: String) extends Dockable(view, position)
{
val searchspace =
for ((group, symbols) <- Symbol.groups; sym <- symbols)
- yield (sym, sym.toLowerCase)
+ yield (sym, Library.lowercase(sym))
private class Symbol_Component(val symbol: String) extends Button
{
@@ -75,7 +77,8 @@
val max_results = PIDE.options.int("jedit_symbols_search_limit") max 0
results_panel.contents.clear
val results =
- (searchspace filter (search.text.toLowerCase.split("\\s+") forall _._2.contains)
+ (searchspace filter
+ (Library.lowercase(search.text).split("\\s+") forall _._2.contains)
take (max_results + 1)).toList
for ((sym, _) <- results)
results_panel.contents += new Symbol_Component(sym)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/theories_dockable.scala Fri Nov 30 23:06:11 2012 +0100
@@ -0,0 +1,179 @@
+/* Title: Tools/jEdit/src/theories_dockable.scala
+ Author: Makarius
+
+Dockable window for theories managed by prover.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.actors.Actor._
+import scala.swing.{FlowPanel, Button, TextArea, Label, ListView, Alignment, ScrollPane, Component}
+import scala.swing.event.{ButtonClicked, MouseClicked}
+
+import java.lang.System
+import java.awt.{BorderLayout, Graphics2D, Insets}
+import javax.swing.{JList, BorderFactory}
+import javax.swing.border.{BevelBorder, SoftBevelBorder}
+
+import org.gjt.sp.jedit.{View, jEdit}
+
+
+class Theories_Dockable(view: View, position: String) extends Dockable(view, position)
+{
+ /* status */
+
+ private val status = new ListView(Nil: List[Document.Node.Name]) {
+ listenTo(mouse.clicks)
+ reactions += {
+ case MouseClicked(_, point, _, clicks, _) if clicks == 2 =>
+ val index = peer.locationToIndex(point)
+ if (index >= 0) jEdit.openFile(view, listData(index).node)
+ }
+ }
+ status.peer.setLayoutOrientation(JList.HORIZONTAL_WRAP)
+ status.peer.setVisibleRowCount(0)
+ status.selection.intervalMode = ListView.IntervalMode.Single
+
+ set_content(new ScrollPane(status))
+
+
+ /* controls */
+
+ def phase_text(phase: Session.Phase): String =
+ "Prover: " + Library.lowercase(phase.toString)
+
+ private val session_phase = new Label(phase_text(PIDE.session.phase))
+ session_phase.border = new SoftBevelBorder(BevelBorder.LOWERED)
+ session_phase.tooltip = "Status of prover session"
+
+ private def handle_phase(phase: Session.Phase)
+ {
+ Swing_Thread.later { session_phase.text = " " + phase_text(phase) + " " }
+ }
+
+ private val cancel = new Button("Cancel") {
+ reactions += { case ButtonClicked(_) => PIDE.cancel_execution() }
+ }
+ cancel.tooltip = "Cancel current checking process"
+
+ private val check = new Button("Check") {
+ reactions += { case ButtonClicked(_) => PIDE.check_buffer(view.getBuffer) }
+ }
+ check.tooltip = "Commence full checking of current buffer"
+
+ private val logic = Isabelle_Logic.logic_selector(true)
+
+ private val controls =
+ new FlowPanel(FlowPanel.Alignment.Right)(check, cancel, session_phase, logic)
+ add(controls.peer, BorderLayout.NORTH)
+
+
+ /* component state -- owned by Swing thread */
+
+ private var nodes_status: Map[Document.Node.Name, Protocol.Node_Status] = Map.empty
+
+ private object Node_Renderer_Component extends Label
+ {
+ opaque = false
+ xAlignment = Alignment.Leading
+ border = BorderFactory.createEmptyBorder(2, 2, 2, 2)
+
+ var node_name = Document.Node.Name.empty
+ override def paintComponent(gfx: Graphics2D)
+ {
+ nodes_status.get(node_name) match {
+ case Some(st) if st.total > 0 =>
+ val size = peer.getSize()
+ val insets = border.getBorderInsets(this.peer)
+ val w = size.width - insets.left - insets.right
+ val h = size.height - insets.top - insets.bottom
+ var end = size.width - insets.right
+ for {
+ (n, color) <- List(
+ (st.unprocessed, PIDE.options.color_value("unprocessed1_color")),
+ (st.running, PIDE.options.color_value("running_color")),
+ (st.warned, PIDE.options.color_value("warning_color")),
+ (st.failed, PIDE.options.color_value("error_color"))) }
+ {
+ gfx.setColor(color)
+ val v = (n * w / st.total) max (if (n > 0) 2 else 0)
+ gfx.fillRect(end - v, insets.top, v, h)
+ end -= v
+ }
+ case _ =>
+ }
+ super.paintComponent(gfx)
+ }
+ }
+
+ private class Node_Renderer extends ListView.Renderer[Document.Node.Name]
+ {
+ def componentFor(list: ListView[_], isSelected: Boolean, focused: Boolean,
+ name: Document.Node.Name, index: Int): Component =
+ {
+ val component = Node_Renderer_Component
+ component.node_name = name
+ component.text = name.theory
+ component
+ }
+ }
+ status.renderer = new Node_Renderer
+
+ private def handle_update(restriction: Option[Set[Document.Node.Name]] = None)
+ {
+ Swing_Thread.now {
+ val snapshot = PIDE.session.snapshot()
+
+ val iterator =
+ restriction match {
+ case Some(names) => names.iterator.map(name => (name, snapshot.version.nodes(name)))
+ case None => snapshot.version.nodes.entries
+ }
+ val nodes_status1 =
+ (nodes_status /: iterator)({ case (status, (name, node)) =>
+ if (PIDE.thy_load.loaded_theories(name.theory)) status
+ else status + (name -> Protocol.node_status(snapshot.state, snapshot.version, node)) })
+
+ if (nodes_status != nodes_status1) {
+ nodes_status = nodes_status1
+ status.listData =
+ snapshot.version.nodes.topological_order.filter(
+ (name: Document.Node.Name) => nodes_status.isDefinedAt(name))
+ }
+ }
+ }
+
+
+ /* main actor */
+
+ private val main_actor = actor {
+ loop {
+ react {
+ case phase: Session.Phase => handle_phase(phase)
+
+ case _: Session.Global_Options => Swing_Thread.later { logic.load () }
+
+ case changed: Session.Commands_Changed => handle_update(Some(changed.nodes))
+
+ case bad => System.err.println("Theories_Dockable: ignoring bad message " + bad)
+ }
+ }
+ }
+
+ override def init()
+ {
+ PIDE.session.phase_changed += main_actor; handle_phase(PIDE.session.phase)
+ PIDE.session.global_options += main_actor
+ PIDE.session.commands_changed += main_actor; handle_update()
+ }
+
+ override def exit()
+ {
+ PIDE.session.phase_changed -= main_actor
+ PIDE.session.global_options -= main_actor
+ PIDE.session.commands_changed -= main_actor
+ }
+}