--- a/.hgignore Wed Jun 08 17:01:07 2011 +0200
+++ b/.hgignore Wed Jun 08 22:13:49 2011 +0200
@@ -21,8 +21,4 @@
^doc-src/.*\.log
^doc-src/.*\.out
^doc-src/.*\.toc
-
-^src/Tools/jEdit/nbproject/private/
-^src/Tools/jEdit/build/
^src/Tools/jEdit/dist/
-^src/Tools/jEdit/contrib/
--- a/Admin/build Wed Jun 08 17:01:07 2011 +0200
+++ b/Admin/build Wed Jun 08 22:13:49 2011 +0200
@@ -27,6 +27,7 @@
browser graph browser (requires jdk)
doc documentation (requires latex and rail)
jars Isabelle/Scala layer (requires Scala in \$SCALA_HOME)
+ jars_fresh fresh build of jars
EOF
exit 1
@@ -84,7 +85,7 @@
function build_jars ()
{
pushd "$ISABELLE_HOME/src/Pure" >/dev/null
- "$ISABELLE_TOOL" env ./build-jars || exit $?
+ "$ISABELLE_TOOL" env ./build-jars "$@" || exit $?
popd >/dev/null
}
@@ -98,6 +99,7 @@
browser) build_browser;;
doc) build_doc;;
jars) build_jars;;
+ jars_fresh) build_jars -f;;
*) fail "Bad module $MODULE"
esac
done
--- a/etc/components Wed Jun 08 17:01:07 2011 +0200
+++ b/etc/components Wed Jun 08 22:13:49 2011 +0200
@@ -11,6 +11,7 @@
src/Sequents
#misc components
src/Tools/Code
+src/Tools/jEdit
src/Tools/WWW_Find
src/HOL/Tools/ATP
src/HOL/Mirabelle
--- a/src/HOL/Mutabelle/mutabelle_extra.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML Wed Jun 08 22:13:49 2011 +0200
@@ -389,7 +389,7 @@
fun mutate_theorem create_entry thy mtds thm =
let
val exec = is_executable_thm thy thm
- val _ = Output.tracing (if exec then "EXEC" else "NOEXEC")
+ val _ = tracing (if exec then "EXEC" else "NOEXEC")
val mutants =
(if num_mutations = 0 then
[Thm.prop_of thm]
--- a/src/HOL/Nominal/nominal_fresh_fun.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML Wed Jun 08 22:13:49 2011 +0200
@@ -5,8 +5,13 @@
a tactic to analyse instances of the fresh_fun.
*)
-(* First some functions that should be in the library *)
+(* First some functions that should be in the library *) (* FIXME really?? *)
+
+(* FIXME proper ML structure *)
+(* FIXME res_inst_tac mostly obsolete, cf. Subgoal.FOCUS *)
+
+(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
fun gen_res_inst_tac_term instf tyinst tinst elim th i st =
let
val thy = theory_of_thm st;
@@ -25,7 +30,8 @@
(Thm.lift_rule cgoal th)
in
compose_tac (elim, th', nprems_of th) i st
- end handle Subscript => Seq.empty;
+ end handle General.Subscript => Seq.empty;
+(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
val res_inst_tac_term =
gen_res_inst_tac_term (curry Thm.instantiate);
--- a/src/HOL/Nominal/nominal_permeq.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Nominal/nominal_permeq.ML Wed Jun 08 22:13:49 2011 +0200
@@ -287,6 +287,7 @@
| collect_vars i (Abs (_, _, t)) vs = collect_vars (i+1) t vs
| collect_vars i (t $ u) vs = collect_vars i u (collect_vars i t vs);
+(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
fun finite_guess_tac_i tactical ss i st =
let val goal = nth (cprems_of st) (i - 1)
in
@@ -318,12 +319,14 @@
end
| _ => Seq.empty
end
- handle Subscript => Seq.empty
+ handle General.Subscript => Seq.empty
+(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
(* tactic that guesses whether an atom is fresh for an expression *)
(* it first collects all free variables and tries to show that the *)
(* support of these free variables (op supports) the goal *)
+(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
fun fresh_guess_tac_i tactical ss i st =
let
val goal = nth (cprems_of st) (i - 1)
@@ -361,7 +364,8 @@
| _ => (tactical ("if it is not of the form _\<sharp>_, then try the simplifier",
(asm_full_simp_tac (HOL_ss addsimps [fresh_prod]@fresh_atm) i))) st
end
- handle Subscript => Seq.empty;
+ handle General.Subscript => Seq.empty;
+(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
val eqvt_simp_tac = eqvt_asm_full_simp_tac_i NO_DEBUG_tac;
--- a/src/HOL/Statespace/state_space.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Statespace/state_space.ML Wed Jun 08 22:13:49 2011 +0200
@@ -360,7 +360,7 @@
val all_comps = rename renaming (parent_comps @ map (apsnd subst) components);
in all_comps end;
-fun take_upto i xs = List.take(xs,i) handle Subscript => xs;
+fun take_upto i xs = List.take(xs,i) handle General.Subscript => xs;
fun statespace_definition state_type args name parents parent_comps components thy =
let
--- a/src/HOL/Tools/ATP/atp_translate.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/ATP/atp_translate.ML Wed Jun 08 22:13:49 2011 +0200
@@ -1,4 +1,4 @@
-(* Title: HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
+(* Title: HOL/Tools/ATP/atp_translate.ML
Author: Fabian Immler, TU Muenchen
Author: Makarius
Author: Jasmin Blanchette, TU Muenchen
@@ -239,7 +239,7 @@
(* translation of #" " to #"/" *)
un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs
else
- let val digits = List.take (c::cs, 3) handle Subscript => [] in
+ let val digits = List.take (c::cs, 3) handle General.Subscript => [] in
case Int.fromString (String.implode digits) of
SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))
| NONE => un (c:: #"_"::rcs) cs (* ERROR *)
--- a/src/HOL/Tools/Function/fun.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/Function/fun.ML Wed Jun 08 22:13:49 2011 +0200
@@ -86,9 +86,9 @@
fun warnings ctxt origs tss =
let
fun warn_redundant t =
- Output.warning ("Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t))
+ warning ("Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t))
fun warn_missing strs =
- Output.warning (cat_lines ("Missing patterns in function definition:" :: strs))
+ warning (cat_lines ("Missing patterns in function definition:" :: strs))
val (tss', added) = chop (length origs) tss
--- a/src/HOL/Tools/Function/function.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/Function/function.ML Wed Jun 08 22:13:49 2011 +0200
@@ -87,8 +87,7 @@
val defname = mk_defname fixes
val FunctionConfig {partials, default, ...} = config
val _ =
- if is_some default then Output.legacy_feature
- "'function (default)'. Use 'partial_function'."
+ if is_some default then legacy_feature "'function (default)'. Use 'partial_function'."
else ()
val ((goal_state, cont), lthy') =
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML Wed Jun 08 22:13:49 2011 +0200
@@ -377,7 +377,7 @@
val p' = if adjustment > p then p else p - adjustment
val tm_p =
nth args p'
- handle Subscript => path_finder_fail tm (p :: ps) (SOME t)
+ handle General.Subscript => path_finder_fail tm (p :: ps) (SOME t)
val _ = trace_msg ctxt (fn () =>
"path_finder: " ^ string_of_int p ^ " " ^
Syntax.string_of_term ctxt tm_p)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Wed Jun 08 22:13:49 2011 +0200
@@ -211,7 +211,7 @@
| ths => rev ths
val _ =
if show_intermediate_results options then
- Output.tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^
+ tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^
commas (map (Display.string_of_thm_global thy) spec))
else ()
in
--- a/src/HOL/Tools/inductive_set.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/inductive_set.ML Wed Jun 08 22:13:49 2011 +0200
@@ -206,7 +206,7 @@
| (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of
SOME (SOME (_, (arity, _))) =>
(fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs
- handle Subscript => error "infer_arities: bad term")
+ handle General.Subscript => error "infer_arities: bad term")
| _ => fold (infer_arities thy arities) (map (pair NONE) ts)
(case optf of
NONE => fs
--- a/src/HOL/Tools/recdef.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/Tools/recdef.ML Wed Jun 08 22:13:49 2011 +0200
@@ -261,7 +261,7 @@
NONE => error ("No recdef definition of constant: " ^ quote name)
| SOME {tcs, ...} => tcs);
val i = the_default 1 opt_i;
- val tc = nth tcs (i - 1) handle Subscript =>
+ val tc = nth tcs (i - 1) handle General.Subscript =>
error ("No termination condition #" ^ string_of_int i ^
" in recdef definition of " ^ quote name);
in
--- a/src/HOL/ex/atp_export.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/HOL/ex/atp_export.ML Wed Jun 08 22:13:49 2011 +0200
@@ -22,7 +22,7 @@
val fact_name_of = prefix ATP_Translate.fact_prefix o ascii_of
fun facts_of thy =
- Sledgehammer_Filter.all_facts (ProofContext.init_global thy) Symtab.empty
+ Sledgehammer_Filter.all_facts (Proof_Context.init_global thy) Symtab.empty
true (K true) [] []
fun fold_body_thms f =
@@ -44,7 +44,7 @@
in names end
fun interesting_const_names ctxt =
- let val thy = ProofContext.theory_of ctxt in
+ let val thy = Proof_Context.theory_of ctxt in
Sledgehammer_Filter.const_names_in_fact thy
(Sledgehammer_Provers.is_built_in_const_for_prover ctxt ATP_Systems.eN)
end
--- a/src/Provers/order.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Provers/order.ML Wed Jun 08 22:13:49 2011 +0200
@@ -1247,9 +1247,9 @@
end
handle Contr p =>
(Subgoal.FOCUS (fn {prems = asms, ...} => rtac (prove asms p) 1) ctxt n st
- handle Subscript => Seq.empty)
+ handle General.Subscript => Seq.empty)
| Cannot => Seq.empty
- | Subscript => Seq.empty)
+ | General.Subscript => Seq.empty)
end;
(* partial_tac - solves partial orders *)
--- a/src/Provers/quasi.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Provers/quasi.ML Wed Jun 08 22:13:49 2011 +0200
@@ -589,8 +589,8 @@
end
handle Contr p =>
(Subgoal.FOCUS (fn {prems, ...} => rtac (prove prems p) 1) ctxt n st
- handle Subscript => Seq.empty)
+ handle General.Subscript => Seq.empty)
| Cannot => Seq.empty
- | Subscript => Seq.empty);
+ | General.Subscript => Seq.empty);
end;
--- a/src/Provers/trancl.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Provers/trancl.ML Wed Jun 08 22:13:49 2011 +0200
@@ -567,6 +567,6 @@
val thms = map (prove thy rel' prems) prfs
in rtac (prove thy rel' thms prf) 1 end) ctxt n st
end
- handle Cannot => Seq.empty | Subscript => Seq.empty);
+ handle Cannot => Seq.empty | General.Subscript => Seq.empty);
end;
--- a/src/Pure/Isar/parse.scala Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/Isar/parse.scala Wed Jun 08 22:13:49 2011 +0200
@@ -1,4 +1,4 @@
-/* Title: Pure/Isar/outer_parse.scala
+/* Title: Pure/Isar/parse.scala
Author: Makarius
Generic parsers for Isabelle/Isar outer syntax.
--- a/src/Pure/Proof/reconstruct.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/Proof/reconstruct.ML Wed Jun 08 22:13:49 2011 +0200
@@ -88,7 +88,7 @@
in (t' $ u', V, vTs2, unifyT thy env3 T (U --> V)) end)
end
| infer_type thy env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env)
- handle Subscript => error ("infer_type: bad variable index " ^ string_of_int i));
+ handle General.Subscript => error ("infer_type: bad variable index " ^ string_of_int i));
fun cantunify thy (t, u) = error ("Non-unifiable terms:\n" ^
Syntax.string_of_term_global thy t ^ "\n\n" ^ Syntax.string_of_term_global thy u);
@@ -145,7 +145,7 @@
(Envir.head_norm env prop, prf, cnstrts, env, vTs);
fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs)
- handle Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i))
+ handle General.Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i))
| mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) =
let
val (T, env') =
--- a/src/Pure/System/isabelle_process.scala Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/System/isabelle_process.scala Wed Jun 08 22:13:49 2011 +0200
@@ -1,4 +1,4 @@
-/* Title: Pure/System/isabelle_process.ML
+/* Title: Pure/System/isabelle_process.scala
Author: Makarius
Options: :folding=explicit:collapseFolds=1:
--- a/src/Pure/Thy/term_style.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/Thy/term_style.ML Wed Jun 08 22:13:49 2011 +0200
@@ -53,7 +53,7 @@
>> fold I
|| Scan.succeed I));
-val parse_bare = Args.context :|-- (fn ctxt => (Output.legacy_feature "Old-style antiquotation style.";
+val parse_bare = Args.context :|-- (fn ctxt => (legacy_feature "Old-style antiquotation style.";
Scan.lift Args.liberal_name
>> (fn name => fst (Args.context_syntax "style"
(Scan.lift (the_style (Proof_Context.theory_of ctxt) name))
@@ -84,7 +84,7 @@
fun style_parm_premise i = Scan.succeed (fn ctxt => fn t =>
let
val i_str = string_of_int i;
- val _ = Output.legacy_feature (quote ("prem" ^ i_str)
+ val _ = legacy_feature (quote ("prem" ^ i_str)
^ " term style encountered; use explicit argument syntax "
^ quote ("prem " ^ i_str) ^ " instead.");
val prems = Logic.strip_imp_prems t;
--- a/src/Pure/build-jars Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/build-jars Wed Jun 08 22:13:49 2011 +0200
@@ -9,6 +9,19 @@
## diagnostics
+PRG="$(basename "$0")"
+
+function usage()
+{
+ echo
+ echo "Usage: isabelle $PRG [OPTIONS]"
+ echo
+ echo " Options are:"
+ echo " -f fresh build"
+ echo
+ exit 1
+}
+
function fail()
{
echo "$1" >&2
@@ -19,6 +32,33 @@
[ -z "$SCALA_HOME" ] && fail "Unknown SCALA_HOME -- Scala unavailable"
+## process command line
+
+# options
+
+FRESH=""
+
+while getopts "f" OPT
+do
+ case "$OPT" in
+ f)
+ FRESH=true
+ ;;
+ \?)
+ usage
+ ;;
+ esac
+done
+
+shift $(($OPTIND - 1))
+
+
+# args
+
+[ "$#" -ne 0 ] && usage
+
+
+
## dependencies
declare -a SOURCES=(
@@ -75,16 +115,19 @@
## main
-OUTDATED=false
-
-for SOURCE in "${SOURCES[@]}"
-do
- [ ! -e "$SOURCE" ] && fail "Missing source file: $SOURCE"
- for TARGET in "${TARGETS[@]}"
+if [ -n "$FRESH" ]; then
+ OUTDATED=true
+else
+ OUTDATED=false
+ for SOURCE in "${SOURCES[@]}"
do
- [ ! -e "$TARGET" -o "$SOURCE" -nt "$TARGET" ] && OUTDATED=true
+ [ ! -e "$SOURCE" ] && fail "Missing source file: $SOURCE"
+ for TARGET in "${TARGETS[@]}"
+ do
+ [ ! -e "$TARGET" -o "$SOURCE" -nt "$TARGET" ] && OUTDATED=true
+ done
done
-done
+fi
if [ "$OUTDATED" = true ]
then
@@ -96,17 +139,20 @@
"$SCALA_HOME/bin/scalac" -unchecked -deprecation -d classes -target:jvm-1.5 "${SOURCES[@]}" || \
fail "Failed to compile sources"
mkdir -p "$TARGET_DIR" || fail "Failed to create directory $TARGET_DIR"
- (
- cd classes
- jar cfe "$(jvmpath "$PURE_JAR")" isabelle.GUI_Setup isabelle || \
- fail "Failed to produce $PURE_JAR"
+
+ pushd classes >/dev/null
+
+ jar cfe "$(jvmpath "$PURE_JAR")" isabelle.GUI_Setup isabelle || \
+ fail "Failed to produce $PURE_JAR"
- cp "$SCALA_HOME/lib/scala-swing.jar" .
- jar xf scala-swing.jar
+ cp "$SCALA_HOME/lib/scala-swing.jar" .
+ jar xf scala-swing.jar
- cp "$SCALA_HOME/lib/scala-library.jar" "$FULL_JAR"
- jar ufe "$(jvmpath "$FULL_JAR")" isabelle.GUI_Setup isabelle scala || \
- fail "Failed to produce $FULL_JAR"
- )
+ cp "$SCALA_HOME/lib/scala-library.jar" "$FULL_JAR"
+ jar ufe "$(jvmpath "$FULL_JAR")" isabelle.GUI_Setup isabelle scala || \
+ fail "Failed to produce $FULL_JAR"
+
+ popd >/dev/null
+
rm -rf classes
fi
--- a/src/Pure/envir.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/envir.ML Wed Jun 08 22:13:49 2011 +0200
@@ -289,7 +289,7 @@
| fast Ts (Const (_, T)) = T
| fast Ts (Free (_, T)) = T
| fast Ts (Bound i) =
- (nth Ts i handle Subscript => raise TERM ("fastype: Bound", [Bound i]))
+ (nth Ts i handle General.Subscript => raise TERM ("fastype: Bound", [Bound i]))
| fast Ts (Var (_, T)) = T
| fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u;
in fast end;
--- a/src/Pure/library.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/library.ML Wed Jun 08 22:13:49 2011 +0200
@@ -429,7 +429,7 @@
raise Subscript if list too short*)
fun nth xs i = List.nth (xs, i);
-fun nth_list xss i = nth xss i handle Subscript => [];
+fun nth_list xss i = nth xss i handle General.Subscript => [];
fun nth_map 0 f (x :: xs) = f x :: xs
| nth_map n f (x :: xs) = x :: nth_map (n - 1) f xs
--- a/src/Pure/proofterm.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/proofterm.ML Wed Jun 08 22:13:49 2011 +0200
@@ -529,7 +529,7 @@
fun subst' lev (Bound i) =
(if i<lev then raise Same.SAME (*var is locally bound*)
else incr_boundvars lev (nth args (i-lev))
- handle Subscript => Bound (i-n)) (*loose: change it*)
+ handle General.Subscript => Bound (i-n)) (*loose: change it*)
| subst' lev (Abs (a, T, body)) = Abs (a, T, subst' (lev+1) body)
| subst' lev (f $ t) = (subst' lev f $ substh' lev t
handle Same.SAME => f $ subst' lev t)
@@ -554,7 +554,7 @@
fun subst (PBound i) Plev tlev =
(if i < Plev then raise Same.SAME (*var is locally bound*)
else incr_pboundvars Plev tlev (nth args (i-Plev))
- handle Subscript => PBound (i-n) (*loose: change it*))
+ handle General.Subscript => PBound (i-n) (*loose: change it*))
| subst (AbsP (a, t, body)) Plev tlev = AbsP (a, t, subst body (Plev+1) tlev)
| subst (Abst (a, T, body)) Plev tlev = Abst (a, T, subst body Plev (tlev+1))
| subst (prf %% prf') Plev tlev = (subst prf Plev tlev %% substh prf' Plev tlev
--- a/src/Pure/sign.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/sign.ML Wed Jun 08 22:13:49 2011 +0200
@@ -269,7 +269,7 @@
fun typ_of (_, Const (_, T)) = T
| typ_of (_, Free (_, T)) = T
| typ_of (_, Var (_, T)) = T
- | typ_of (bs, Bound i) = snd (nth bs i handle Subscript =>
+ | typ_of (bs, Bound i) = snd (nth bs i handle General.Subscript =>
raise TYPE ("Loose bound variable: B." ^ string_of_int i, [], [Bound i]))
| typ_of (bs, Abs (x, T, body)) = T --> typ_of ((x, T) :: bs, body)
| typ_of (bs, t $ u) =
--- a/src/Pure/tactical.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/tactical.ML Wed Jun 08 22:13:49 2011 +0200
@@ -354,7 +354,7 @@
orelse
not (Pattern.aeconv (t, Thm.term_of (Thm.cprem_of st' (Thm.nprems_of st' - d))))
in Seq.filter diff (tac i st) end
- handle Subscript => Seq.empty (*no subgoal i*);
+ handle General.Subscript => Seq.empty (*no subgoal i*);
(*Returns all states where some subgoals have been solved. For
subgoal-based tactics this means subgoal i has been solved
--- a/src/Pure/term.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/term.ML Wed Jun 08 22:13:49 2011 +0200
@@ -311,7 +311,7 @@
fun type_of1 (Ts, Const (_,T)) = T
| type_of1 (Ts, Free (_,T)) = T
| type_of1 (Ts, Bound i) = (nth Ts i
- handle Subscript => raise TYPE("type_of: bound variable", [], [Bound i]))
+ handle General.Subscript => raise TYPE("type_of: bound variable", [], [Bound i]))
| type_of1 (Ts, Var (_,T)) = T
| type_of1 (Ts, Abs (_,T,body)) = T --> type_of1(T::Ts, body)
| type_of1 (Ts, f$u) =
@@ -336,7 +336,7 @@
| fastype_of1 (_, Const (_,T)) = T
| fastype_of1 (_, Free (_,T)) = T
| fastype_of1 (Ts, Bound i) = (nth Ts i
- handle Subscript => raise TERM("fastype_of: Bound", [Bound i]))
+ handle General.Subscript => raise TERM("fastype_of: Bound", [Bound i]))
| fastype_of1 (_, Var (_,T)) = T
| fastype_of1 (Ts, Abs (_,T,u)) = T --> fastype_of1 (T::Ts, u);
@@ -668,7 +668,7 @@
fun subst (t as Bound i, lev) =
(if i < lev then raise Same.SAME (*var is locally bound*)
else incr_boundvars lev (nth args (i - lev))
- handle Subscript => Bound (i - n)) (*loose: change it*)
+ handle General.Subscript => Bound (i - n)) (*loose: change it*)
| subst (Abs (a, T, body), lev) = Abs (a, T, subst (body, lev + 1))
| subst (f $ t, lev) =
(subst (f, lev) $ (subst (t, lev) handle Same.SAME => t)
--- a/src/Pure/thm.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/thm.ML Wed Jun 08 22:13:49 2011 +0200
@@ -1423,7 +1423,7 @@
and concl = Logic.strip_imp_concl prop;
val moved_prems = List.drop (prems, j)
and fixed_prems = List.take (prems, j)
- handle Subscript => raise THM ("permute_prems: j", j, [rl]);
+ handle General.Subscript => raise THM ("permute_prems: j", j, [rl]);
val n_j = length moved_prems;
val m = if k < 0 then n_j + k else k;
val prop' =
--- a/src/Pure/type_infer_context.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Pure/type_infer_context.ML Wed Jun 08 22:13:49 2011 +0200
@@ -220,7 +220,7 @@
| inf _ (Free (_, T)) tye_idx = (T, tye_idx)
| inf _ (Var (_, T)) tye_idx = (T, tye_idx)
| inf bs (Bound i) tye_idx =
- (snd (nth bs i handle Subscript => err_loose i), tye_idx)
+ (snd (nth bs i handle General.Subscript => err_loose i), tye_idx)
| inf bs (Abs (x, T, t)) tye_idx =
let val (U, tye_idx') = inf ((x, T) :: bs) t tye_idx
in (T --> U, tye_idx') end
--- a/src/Tools/WWW_Find/http_util.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Tools/WWW_Find/http_util.ML Wed Jun 08 22:13:49 2011 +0200
@@ -60,7 +60,7 @@
then f (Substring.full " "::pre::done, Substring.triml 1 post)
else let
val (c, rest) = Substring.splitAt (post, 3)
- handle Subscript =>
+ handle General.Subscript =>
(Substring.full "%25", Substring.triml 1 post);
in f (to_char c::pre::done, rest) end
end;
--- a/src/Tools/WWW_Find/yxml_find_theorems.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Tools/WWW_Find/yxml_find_theorems.ML Wed Jun 08 22:13:49 2011 +0200
@@ -1,7 +1,7 @@
-(* Title: src/Pure/Tools/yxml_find_theorems.ML
+(* Title: Tools/WWW_Find/yxml_find_theorems.ML
Author: Sree Harsha Totakura, TUM
- Lars Noschinski, TUM
- Alexander Krauss, TUM
+ Author: Lars Noschinski, TUM
+ Author: Alexander Krauss, TUM
Simple find theorems web service with yxml interface for programmatic
invocation.
@@ -32,7 +32,7 @@
Facts.dest_static [] facts
|> filter_out (Facts.is_concealed facts o #1);
-fun init () =
+fun init () =
let
val all_facts =
maps Facts.selections
--- a/src/Tools/induct.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Tools/induct.ML Wed Jun 08 22:13:49 2011 +0200
@@ -590,7 +590,7 @@
Unify.smash_unifiers thy [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule'))
|> Seq.map (fn env => Drule.instantiate (dest_env thy env) rule')
end
- end handle Subscript => Seq.empty;
+ end handle General.Subscript => Seq.empty;
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/README.html Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,137 @@
+<?xml version="1.0" encoding="UTF-8" ?>
+<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN" "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd">
+<html xmlns="http://www.w3.org/1999/xhtml">
+
+<head>
+<meta http-equiv="Content-Type" content="text/html; charset=utf-8"/>
+<title>Notes on the Isabelle/jEdit Prover IDE</title>
+</head>
+
+<body>
+
+<h2>The Isabelle/jEdit Prover IDE</h2>
+
+<ul>
+
+<li>The original jEdit look-and-feel is generally preserved, although
+ some default properties have been changed to accommodate Isabelle
+ (e.g. the text area font).</li>
+
+<li>Formal Isabelle/Isar text is checked asynchronously while editing.
+ The user is in full command of the editor, and the prover refrains
+ from locking portions of the buffer etc.</li>
+
+<li>Prover feedback works via tooltips, syntax highlighting, colors,
+ boxes etc. based on semantic markup provided by Isabelle in the
+ background.</li>
+
+<li>Using the mouse together with the modifier key <tt>C</tt>
+(<tt>CONTROL</tt> on Linux or Windows,
+ <tt>COMMAND</tt> on Mac OS) exposes additional information.</li>
+
+<li>Dockable panels (e.g. <em>Output</em>) are managed as independent
+ windows by jEdit, which also allows multiple instances.</li>
+
+</ul>
+
+
+<h2>Isabelle symbols and fonts</h2>
+
+<ul>
+
+ <li>Isabelle supports infinitely many symbols:<br/>
+ α, β, γ, …<br/>
+ ∀, ∃, ∨, ∧, ⟶, ⟷, …<br/>
+ ≤, ≥, ⊓, ⊔, …<br/>
+ ℵ, △, ∇, …<br/>
+ <tt>\<foo></tt>, <tt>\<bar></tt>, <tt>\<baz></tt>, …<br/>
+ </li>
+
+ <li>A default mapping relates some Isabelle symbols to Unicode points
+ (see <tt>$ISABELLE_HOME/etc/symbols</tt> and <tt>$ISABELLE_HOME_USER/etc/symbols</tt>).
+ </li>
+
+ <li>The <em>IsabelleText</em> font ensures that Unicode points are actually
+ seen on the screen (or printer).
+ </li>
+
+ <li>Input methods:
+ <ul>
+ <li>copy/paste from decoded source files</li>
+ <li>copy/paste from prover output</li>
+ <li>completion provided by Isabelle plugin, e.g.<br/>
+
+ <table border="1">
+ <tr><th><b>name</b></th> <th><b>abbreviation</b></th> <th><b>symbol</b></th></tr>
+
+ <tr><td>lambda</td> <td></td> <td>λ</td></tr>
+ <tr><td>Rightarrow</td> <td><tt>=></tt></td> <td>⇒</td></tr>
+ <tr><td>Longrightarrow</td> <td><tt>==></tt></td> <td>⟹</td></tr>
+
+ <tr><td>And</td> <td><tt>!!</tt></td> <td>⋀</td></tr>
+ <tr><td>equiv</td> <td><tt>==</tt></td> <td>≡</td></tr>
+
+ <tr><td>forall</td> <td><tt>!</tt></td> <td>∀</td></tr>
+ <tr><td>exists</td> <td><tt>?</tt></td> <td>∃</td></tr>
+ <tr><td>longrightarrow</td> <td><tt>--></tt></td> <td>⟶</td></tr>
+ <tr><td>and</td> <td><tt>/\</tt></td> <td>∧</td></tr>
+ <tr><td>or</td> <td><tt>\/</tt></td> <td>∨</td></tr>
+ <tr><td>not</td> <td><tt>~ </tt></td> <td>¬</td></tr>
+ <tr><td>noteq</td> <td><tt>~=</tt></td> <td>≠</td></tr>
+ <tr><td>in</td> <td><tt>:</tt></td> <td>∈</td></tr>
+ <tr><td>notin</td> <td><tt>~:</tt></td> <td>∉</td></tr>
+ </table>
+ </li>
+ </ul>
+ </li>
+
+ <li><b>NOTE:</b> The above abbreviations refer to the input method.
+ The logical notation provides ASCII alternatives that often
+ coincide, but deviate occasionally.
+ </li>
+
+ <li><b>NOTE:</b> Generic jEdit abbreviations or plugins perform similar
+ source replacement operations; this works for Isabelle as long
+ as the Unicode sequences coincide with the symbol mapping.
+ </li>
+
+</ul>
+
+
+<h2>Limitations and workrounds (January 2011)</h2>
+
+<ul>
+ <li>No way to start/stop prover or switch to a different logic.<br/>
+ <em>Workaround:</em> Change options and restart editor.</li>
+
+ <li>Multiple theory buffers cannot depend on each other, imports are
+ resolved via the file-system.<br/>
+ <em>Workaround:</em> Save/reload files manually.</li>
+
+ <li>No reclaiming of old/unused document versions in prover or
+ editor.<br/>
+ <em>Workaround:</em> Avoid large files; restart after a few hours of use.</li>
+
+ <li>Incremental reparsing sometimes produces unexpected command
+ spans.<br/>
+ <em>Workaround:</em> Cut/paste larger parts or reload buffer.</li>
+
+ <li>Command execution sometimes gets stuck (purple background).<br/>
+ <em>Workaround:</em> Force reparsing as above.</li>
+
+ <li>Odd behavior of some diagnostic commands, notably those
+ starting external processes asynchronously
+ (e.g. <tt>thy_deps</tt>, <tt>sledgehammer</tt>).<br/>
+ <em>Workaround:</em> Avoid such commands.</li>
+
+ <li>No support for non-local markup, e.g. commands reporting on
+ previous commands (proof end on proof head), or markup produced by
+ loading external files.</li>
+
+ <li>General lack of various conveniences known from Proof
+ General.</li>
+</ul>
+
+</body>
+</html>
+
--- a/src/Tools/jEdit/README_BUILD Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Tools/jEdit/README_BUILD Wed Jun 08 22:13:49 2011 +0200
@@ -1,86 +1,29 @@
+Requirements for instantaneous build from sources
+=================================================
-Requirements to build from sources
-==================================
-
-* Proper Java JRE/JDK from Sun, e.g. 1.6.0_24 or 1.6.0_25
+* Proper Java JDK from Sun/Oracle/Apple, e.g. 1.6.0_24 or 1.6.0_25
http://java.sun.com/javase/downloads/index.jsp
-* Netbeans 6.9
- http://www.netbeans.org/downloads/index.html
-
-* Scala for Netbeans: version 6.9v1.1.0
- http://wiki.netbeans.org/Scala
- http://wiki.netbeans.org/Scala68v1
- http://sourceforge.net/projects/erlybird/files/nb-scala/6.9v1.1.0
-
-* jEdit 4.3.2
- http://www.jedit.org/
- Netbeans Project "jEdit": install official sources as ./contrib/jEdit/.
+* Scala Compiler 2.8.1.final
+ http://www.scala-lang.org
-* jEdit plugins:
- Netbeans Library "Console" = $HOME/.jedit/jars/Console.jar
- Netbeans Library "SideKick" = $HOME/.jedit/jars/SideKick.jar
- Netbeans Library "ErrorList" = $HOME/.jedit/jars/ErrorList.jar
- Netbeans Library "Hyperlinks" = $HOME/.jedit/jars/Hyperlinks.jar
-
-* Cobra Renderer 0.98.4
- http://lobobrowser.org/cobra.jsp
- Netbeans Library "Cobra-Renderer" = .../cobra.jar
- Netbenas Library "Rhino-JavaScript" = .../js.jar
-
-* Isabelle/Pure Scala components
- Netbeans Library "Isabelle-Pure" = ~~/lib/classes/Pure.jar
-
-* Scala Compiler 2.8.1.final or 2.9.0.final
- http://www.scala-lang.org
- Netbeans Library "Scala-compiler" = $SCALA_HOME/lib/scala-compiler.jar
+* Auxiliary jedit_build component
+ http://www4.in.tum.de/~wenzelm/test/jedit_build-20110521.tar.gz
-Building and running from command line (recommended)
-====================================================
-
-* Manual build:
+Important settings within Isabelle environment
+==============================================
- ( cd "$HOME/isabelle/repos/src/Pure" && isabelle env ./build-jars )
- ant jar
- makedist -j "$HOME/lib/jedit-4.3.2" #location of original jedit distribution
-
-* Component setup, e.g. in $ISABELLE_HOME_USER/.isabelle/etc/settings:
-
- [ -d "$HOME/isabelle/jedit" ] && init_component "$HOME/isabelle/jedit" #build location
-
-* Run:
-
- isabelle jedit -l HOL Test.thy &
+- JAVA_HOME
+- SCALA_HOME
+- JEDIT_BUILD_HOME (via "init_component .../jedit_build...")
-Running the application within Netbeans (not recommended)
-=========================================================
-
-* Project properties: add "Run" argument like
- -noserver -nobackground -settings=/home/makarius/isabelle/isabelle-jedit/dist
-
-* The Isabelle environment is obtained automatically via
- "$ISABELLE_HOME/bin/isabelle getenv", where ISABELLE_HOME is determined as follows:
-
- (1) via regular Isabelle settings,
- e.g. "isabelle env netbeans"
+Build and run
+=============
- (2) or via ISABELLE_HOME from raw process environment,
- e.g. "env ISABELLE_HOME=.../Isabelle netbeans"
-
- (3) or via JVM system properties (cf. "Run / VM Options")
- e.g. -Disabelle.home=.../Isabelle
-
+isabelle jedit -l HOL Test.thy
-Misc notes
-==========
-
-- Netbeans config/Editors/Preferences/...-CustomPreferences.xml
-
- <entry javaType="java.lang.Integer" name="caret-blink-rate" xml:space="preserve">
- <value><![CDATA[0]]></value>
- </entry>
-----------------------------------------------------------------------
To run jedit with remote debugging enabled, I use the following
--- a/src/Tools/jEdit/build.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,102 +0,0 @@
-<?xml version="1.0" encoding="UTF-8"?>
-<!-- You may freely edit this file. See commented blocks below for -->
-<!-- some examples of how to customize the build. -->
-<!-- (If you delete it and reopen the project it will be recreated.) -->
-<project name="Isabelle-jEdit" default="default" basedir=".">
- <description>Builds, tests, and runs the project Isabelle-jEdit.</description>
- <import file="nbproject/build-impl.xml"/>
- <!--
-
- There exist several targets which are by default empty and which can be
- used for execution of your tasks. These targets are usually executed
- before and after some main targets. They are:
-
- -pre-init: called before initialization of project properties
- -post-init: called after initialization of project properties
- -pre-compile: called before javac compilation
- -post-compile: called after javac compilation
- -pre-compile-single: called before javac compilation of single file
- -post-compile-single: called after javac compilation of single file
- -pre-compile-test: called before javac compilation of JUnit tests
- -post-compile-test: called after javac compilation of JUnit tests
- -pre-compile-test-single: called before javac compilation of single JUnit test
- -post-compile-test-single: called after javac compilation of single JUunit test
- -pre-jar: called before JAR building
- -post-jar: called after JAR building
- -post-clean: called after cleaning build products
-
- (Targets beginning with '-' are not intended to be called on their own.)
-
- Example of inserting an obfuscator after compilation could look like this:
-
- <target name="-post-compile">
- <obfuscate>
- <fileset dir="${build.classes.dir}"/>
- </obfuscate>
- </target>
-
- For list of available properties check the imported
- nbproject/build-impl.xml file.
-
-
- Another way to customize the build is by overriding existing main targets.
- The targets of interest are:
-
- -init-macrodef-javac: defines macro for javac compilation
- -init-macrodef-junit: defines macro for junit execution
- -init-macrodef-debug: defines macro for class debugging
- -init-macrodef-java: defines macro for class execution
- -do-jar-with-manifest: JAR building (if you are using a manifest)
- -do-jar-without-manifest: JAR building (if you are not using a manifest)
- run: execution of project
- -javadoc-build: Javadoc generation
- test-report: JUnit report generation
-
- An example of overriding the target for project execution could look like this:
-
- <target name="run" depends="Isabelle-jEdit-impl.jar">
- <exec dir="bin" executable="launcher.exe">
- <arg file="${dist.jar}"/>
- </exec>
- </target>
-
- Notice that the overridden target depends on the jar target and not only on
- the compile target as the regular run target does. Again, for a list of available
- properties which you can use, check the target you are overriding in the
- nbproject/build-impl.xml file.
-
- -->
- <target name="run" depends="Isabelle-jEdit-impl.jar,Isabelle-jEdit-impl.run">
- </target>
- <target name="debug" depends="Isabelle-jEdit-impl.jar,Isabelle-jEdit-impl.debug">
- </target>
- <target name="-pre-jar">
- <copy file="plugin/services.xml" todir="${build.classes.dir}" />
- <copy file="plugin/dockables.xml" todir="${build.classes.dir}" />
- <copy file="plugin/actions.xml" todir="${build.classes.dir}" />
- <copy file="plugin/Isabelle.props" todir="${build.classes.dir}" />
- </target>
- <target name="-post-jar">
- <!-- jars -->
- <delete file="${dist.dir}/jars/lib/jEdit.jar" />
- <move todir="${dist.dir}/jars">
- <fileset dir="${dist.dir}/jars/lib" />
- </move>
- <copy file="${scala.library}" todir="${dist.dir}/jars" />
- <copy file="${scala.lib}/scala-swing.jar" todir="${dist.dir}/jars" />
- <!-- clean up -->
- <delete dir="{dist.dir}/jars/lib" />
- <!-- dist-template -->
- <copy file="dist-template/properties/jedit.props" tofile="${dist.dir}/properties" />
- <copy todir="${dist.dir}/modes">
- <fileset dir="dist-template/modes" />
- </copy>
- <copy todir="${dist.dir}/modes" overwrite="true">
- <fileset dir="${project.jEdit}/modes" />
- </copy>
- <replaceregexp byline="true" file="${dist.dir}/modes/catalog">
- <regexp pattern='(^.*NAME="javacc".*$)'/>
- <substitution expression="<MODE NAME="isabelle" FILE="isabelle.xml" FILE_NAME_GLOB="*.thy"/>${line.separator}${line.separator}\1"/>
- </replaceregexp>
- </target>
-</project>
--- a/src/Tools/jEdit/contrib/jEdit/build-nb.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-<?xml version="1.0" encoding="UTF-8"?>
-<project name="jEdit">
- <import file="./build.xml"/>
- <property environment="env"/>
-
- <target name="run" depends="compile">
- <java classname="org.gjt.sp.jedit.jEdit" classpath="./build/jEdit.jar" dir="./" fork="true">
- <arg value="-noserver"/>
- <arg value="-nobackground"/>
- <arg value="-settings=${env.ISABELLE_HOME_USER}/jedit"/>
- </java>
- </target>
-
- <target name="debug-nb" depends="compile">
- <path id="cp" location="./build/jEdit.jar" />
-
- <nbjpdastart addressproperty="jpda.address" name="jEdit" transport="dt_socket">
- <classpath refid="cp"/>
- </nbjpdastart>
-
- <java classname="org.gjt.sp.jedit.jEdit" classpathref="cp" fork="true" dir="./">
- <jvmarg value="-Xdebug"/>
- <jvmarg value="-Xrunjdwp:transport=dt_socket,address=${jpda.address}"/>
- </java>
- </target>
-
- <target name="profile-nb" depends="compile">
- <fail unless="netbeans.home">This target can only run inside the NetBeans IDE.</fail>
-
- <path id="cp" location="./build/jEdit.jar" />
-
- <nbprofiledirect>
- <classpath refid="cp"/>
- </nbprofiledirect>
-
- <java classname="org.gjt.sp.jedit.jEdit" fork="true" logError="yes" dir="." classpathref="cp">
- <classpath refid="cp"/>
- <jvmarg value="${profiler.info.jvmargs.agent}"/>
- </java>
- </target>
-</project>
\ No newline at end of file
--- a/src/Tools/jEdit/contrib/jEdit/nbproject/project.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,118 +0,0 @@
-<?xml version="1.0" encoding="UTF-8"?>
-<project xmlns="http://www.netbeans.org/ns/project/1">
- <type>org.netbeans.modules.ant.freeform</type>
- <configuration>
- <general-data xmlns="http://www.netbeans.org/ns/freeform-project/1">
- <name>jEdit</name>
- </general-data>
- <general-data xmlns="http://www.netbeans.org/ns/freeform-project/2">
- <!-- Do not use Project Properties customizer when editing this file manually. -->
- <name>jEdit</name>
- <properties>
- <property name="ant.script">build-nb.xml</property>
- </properties>
- <folders>
- <source-folder>
- <label>jEdit</label>
- <location>.</location>
- <encoding>UTF-8</encoding>
- </source-folder>
- <source-folder>
- <label>Source Packages</label>
- <type>java</type>
- <location>.</location>
- <excludes>build/** doc/** icons/** macros/** modes/** package-files/**</excludes>
- <encoding>UTF-8</encoding>
- </source-folder>
- </folders>
- <ide-actions>
- <action name="build">
- <script>${ant.script}</script>
- <target>build</target>
- </action>
- <action name="clean">
- <script>${ant.script}</script>
- <target>clean</target>
- </action>
- <action name="javadoc">
- <script>${ant.script}</script>
- <target>docs-javadoc</target>
- </action>
- <action name="run">
- <script>${ant.script}</script>
- <target>run</target>
- </action>
- <action name="rebuild">
- <script>${ant.script}</script>
- <target>clean</target>
- <target>build</target>
- </action>
- <action name="debug">
- <script>build-nb.xml</script>
- <target>debug-nb</target>
- </action>
- </ide-actions>
- <export>
- <type>jar</type>
- <location>build/jEdit.jar</location>
- <script>${ant.script}</script>
- <build-target>build</build-target>
- </export>
- <view>
- <items>
- <source-folder style="packages">
- <label>Source Packages</label>
- <location>.</location>
- <excludes>build/** doc/** icons/** macros/** modes/** package-files/**</excludes>
- </source-folder>
- <source-file>
- <location>${ant.script}</location>
- </source-file>
- </items>
- <context-menu>
- <ide-action name="build"/>
- <ide-action name="rebuild"/>
- <ide-action name="clean"/>
- <ide-action name="javadoc"/>
- <ide-action name="run"/>
- <ide-action name="debug"/>
- </context-menu>
- </view>
- <subprojects/>
- </general-data>
- <java-data xmlns="http://www.netbeans.org/ns/freeform-project-java/1">
- <compilation-unit>
- <package-root>.</package-root>
- <classpath mode="compile">.</classpath>
- <built-to>build/jEdit.jar</built-to>
- <source-level>1.5</source-level>
- </compilation-unit>
- </java-data>
- <preferences xmlns="http://www.netbeans.org/ns/auxiliary-configuration-preferences/1">
- <module name="org-netbeans-modules-editor-indent">
- <node name="CodeStyle">
- <property name="usedProfile" value="default"/>
- <node name="project">
- <property name="spaces-per-tab" value="2"/>
- <property name="tab-size" value="2"/>
- <property name="indent-shift-width" value="2"/>
- <property name="text-limit-width" value="100"/>
- <property name="expand-tabs" value="true"/>
- </node>
- </node>
- <node name="text">
- <node name="x-java">
- <node name="CodeStyle">
- <node name="project">
- <property name="tab-size" value="4"/>
- <property name="text-limit-width" value="100"/>
- <property name="spaces-per-tab" value="2"/>
- <property name="indent-shift-width" value="2"/>
- </node>
- </node>
- </node>
- </node>
- </module>
- </preferences>
- </configuration>
-</project>
--- a/src/Tools/jEdit/dist-template/README.html Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,137 +0,0 @@
-<?xml version="1.0" encoding="UTF-8" ?>
-<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN" "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd">
-<html xmlns="http://www.w3.org/1999/xhtml">
-
-<head>
-<meta http-equiv="Content-Type" content="text/html; charset=utf-8"/>
-<title>Notes on the Isabelle/jEdit Prover IDE</title>
-</head>
-
-<body>
-
-<h2>The Isabelle/jEdit Prover IDE</h2>
-
-<ul>
-
-<li>The original jEdit look-and-feel is generally preserved, although
- some default properties have been changed to accommodate Isabelle
- (e.g. the text area font).</li>
-
-<li>Formal Isabelle/Isar text is checked asynchronously while editing.
- The user is in full command of the editor, and the prover refrains
- from locking portions of the buffer etc.</li>
-
-<li>Prover feedback works via tooltips, syntax highlighting, colors,
- boxes etc. based on semantic markup provided by Isabelle in the
- background.</li>
-
-<li>Using the mouse together with the modifier key <tt>C</tt>
-(<tt>CONTROL</tt> on Linux or Windows,
- <tt>COMMAND</tt> on Mac OS) exposes additional information.</li>
-
-<li>Dockable panels (e.g. <em>Output</em>) are managed as independent
- windows by jEdit, which also allows multiple instances.</li>
-
-</ul>
-
-
-<h2>Isabelle symbols and fonts</h2>
-
-<ul>
-
- <li>Isabelle supports infinitely many symbols:<br/>
- α, β, γ, …<br/>
- ∀, ∃, ∨, ∧, ⟶, ⟷, …<br/>
- ≤, ≥, ⊓, ⊔, …<br/>
- ℵ, △, ∇, …<br/>
- <tt>\<foo></tt>, <tt>\<bar></tt>, <tt>\<baz></tt>, …<br/>
- </li>
-
- <li>A default mapping relates some Isabelle symbols to Unicode points
- (see <tt>$ISABELLE_HOME/etc/symbols</tt> and <tt>$ISABELLE_HOME_USER/etc/symbols</tt>).
- </li>
-
- <li>The <em>IsabelleText</em> font ensures that Unicode points are actually
- seen on the screen (or printer).
- </li>
-
- <li>Input methods:
- <ul>
- <li>copy/paste from decoded source files</li>
- <li>copy/paste from prover output</li>
- <li>completion provided by Isabelle plugin, e.g.<br/>
-
- <table border="1">
- <tr><th><b>name</b></th> <th><b>abbreviation</b></th> <th><b>symbol</b></th></tr>
-
- <tr><td>lambda</td> <td></td> <td>λ</td></tr>
- <tr><td>Rightarrow</td> <td><tt>=></tt></td> <td>⇒</td></tr>
- <tr><td>Longrightarrow</td> <td><tt>==></tt></td> <td>⟹</td></tr>
-
- <tr><td>And</td> <td><tt>!!</tt></td> <td>⋀</td></tr>
- <tr><td>equiv</td> <td><tt>==</tt></td> <td>≡</td></tr>
-
- <tr><td>forall</td> <td><tt>!</tt></td> <td>∀</td></tr>
- <tr><td>exists</td> <td><tt>?</tt></td> <td>∃</td></tr>
- <tr><td>longrightarrow</td> <td><tt>--></tt></td> <td>⟶</td></tr>
- <tr><td>and</td> <td><tt>/\</tt></td> <td>∧</td></tr>
- <tr><td>or</td> <td><tt>\/</tt></td> <td>∨</td></tr>
- <tr><td>not</td> <td><tt>~ </tt></td> <td>¬</td></tr>
- <tr><td>noteq</td> <td><tt>~=</tt></td> <td>≠</td></tr>
- <tr><td>in</td> <td><tt>:</tt></td> <td>∈</td></tr>
- <tr><td>notin</td> <td><tt>~:</tt></td> <td>∉</td></tr>
- </table>
- </li>
- </ul>
- </li>
-
- <li><b>NOTE:</b> The above abbreviations refer to the input method.
- The logical notation provides ASCII alternatives that often
- coincide, but deviate occasionally.
- </li>
-
- <li><b>NOTE:</b> Generic jEdit abbreviations or plugins perform similar
- source replacement operations; this works for Isabelle as long
- as the Unicode sequences coincide with the symbol mapping.
- </li>
-
-</ul>
-
-
-<h2>Limitations and workrounds (January 2011)</h2>
-
-<ul>
- <li>No way to start/stop prover or switch to a different logic.<br/>
- <em>Workaround:</em> Change options and restart editor.</li>
-
- <li>Multiple theory buffers cannot depend on each other, imports are
- resolved via the file-system.<br/>
- <em>Workaround:</em> Save/reload files manually.</li>
-
- <li>No reclaiming of old/unused document versions in prover or
- editor.<br/>
- <em>Workaround:</em> Avoid large files; restart after a few hours of use.</li>
-
- <li>Incremental reparsing sometimes produces unexpected command
- spans.<br/>
- <em>Workaround:</em> Cut/paste larger parts or reload buffer.</li>
-
- <li>Command execution sometimes gets stuck (purple background).<br/>
- <em>Workaround:</em> Force reparsing as above.</li>
-
- <li>Odd behavior of some diagnostic commands, notably those
- starting external processes asynchronously
- (e.g. <tt>thy_deps</tt>, <tt>sledgehammer</tt>).<br/>
- <em>Workaround:</em> Avoid such commands.</li>
-
- <li>No support for non-local markup, e.g. commands reporting on
- previous commands (proof end on proof head), or markup produced by
- loading external files.</li>
-
- <li>General lack of various conveniences known from Proof
- General.</li>
-</ul>
-
-</body>
-</html>
-
--- a/src/Tools/jEdit/dist-template/etc/isabelle-jedit.css Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-/* additional style file for Isabelle/jEdit output */
-
-.message { margin-top: 0.3ex; background-color: #F0F0F0; }
-
-.writeln { }
-.tracing { background-color: #F0F8FF; }
-.warning { background-color: #EEE8AA; }
-.error { background-color: #FFC1C1; }
-
-.report { display: none; }
-
-.hilite { background-color: #FFCC66; }
-
-.keyword { font-weight: bold; color: #009966; }
-.operator { font-weight: bold; }
-.command { font-weight: bold; color: #006699; }
-
-.sendback { text-decoration: underline; }
-.sendback:hover { background-color: #FFCC66; }
--- a/src/Tools/jEdit/dist-template/etc/settings Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-# -*- shell-script -*- :mode=shellscript:
-
-JEDIT_HOME="$COMPONENT"
-JEDIT_SETTINGS="$ISABELLE_HOME_USER/jedit"
-
-JEDIT_OPTIONS="-reuseview -noserver -nobackground"
-JEDIT_JAVA_OPTIONS="-Xms128m -Xmx1024m -Xss4m -Dactors.corePoolSize=4 -Dactors.enableForkJoin=false"
-#JEDIT_JAVA_OPTIONS="-Xms128m -Xmx512m -Xss2m -Dactors.corePoolSize=4 -Dactors.enableForkJoin=false"
-JEDIT_SYSTEM_OPTIONS="-Dapple.laf.useScreenMenuBar=true -Dcom.apple.mrj.application.apple.menu.about.name=Isabelle/jEdit"
-
-JEDIT_STYLE_SHEETS="$ISABELLE_HOME/lib/html/isabelle.css:$JEDIT_HOME/etc/isabelle-jedit.css:$ISABELLE_HOME_USER/etc/isabelle.css:$ISABELLE_HOME_USER/etc/isabelle-jedit.css"
-
-ISABELLE_JEDIT_OPTIONS="-m xsymbols -m no_brackets -m no_type_brackets"
-
-ISABELLE_TOOLS="$ISABELLE_TOOLS:$JEDIT_HOME/lib/Tools"
-
--- a/src/Tools/jEdit/dist-template/lib/Tools/jedit Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,132 +0,0 @@
-#!/usr/bin/env bash
-#
-# Author: Makarius
-#
-# DESCRIPTION: Isabelle/jEdit interface wrapper
-
-
-## diagnostics
-
-PRG="$(basename "$0")"
-
-usage()
-{
- echo
- echo "Usage: isabelle $PRG [OPTIONS] [FILES ...]"
- echo
- echo " Options are:"
- echo " -J OPTION add JVM runtime option"
- echo " (default JEDIT_JAVA_OPTIONS=$JEDIT_JAVA_OPTIONS)"
- echo " -d enable debugger"
- echo " -j OPTION add jEdit runtime option"
- echo " (default JEDIT_OPTIONS=$JEDIT_OPTIONS)"
- echo " -l NAME logic image name (default ISABELLE_LOGIC=$ISABELLE_LOGIC)"
- echo " -m MODE add print mode for output"
- echo
- echo "Start jEdit with Isabelle plugin setup and opens theory FILES"
- echo
- exit 1
-}
-
-fail()
-{
- echo "$1" >&2
- exit 2
-}
-
-
-## process command line
-
-# options
-
-JEDIT_LOGIC="$ISABELLE_LOGIC"
-JEDIT_PRINT_MODE=""
-
-getoptions()
-{
- OPTIND=1
- while getopts "J:dj:l:m:" OPT
- do
- case "$OPT" in
- J)
- JAVA_ARGS["${#JAVA_ARGS[@]}"]="$OPTARG"
- ;;
- d)
- JAVA_ARGS["${#JAVA_ARGS[@]}"]="-Xdebug"
- JAVA_ARGS["${#JAVA_ARGS[@]}"]="-Xrunjdwp:transport=dt_socket,server=y,suspend=n"
- ;;
- j)
- ARGS["${#ARGS[@]}"]="$OPTARG"
- ;;
- l)
- JEDIT_LOGIC="$OPTARG"
- ;;
- m)
- if [ -z "$JEDIT_PRINT_MODE" ]; then
- JEDIT_PRINT_MODE="$OPTARG"
- else
- JEDIT_PRINT_MODE="$JEDIT_PRINT_MODE,$OPTARG"
- fi
- ;;
- \?)
- usage
- ;;
- esac
- done
-}
-
-declare -a JAVA_ARGS; eval "JAVA_ARGS=($JEDIT_JAVA_OPTIONS $JEDIT_SYSTEM_OPTIONS)"
-[ -n "$SCALA_HOME" ] && JAVA_ARGS["${#JAVA_ARGS[@]}"]="-Dscala.home=$SCALA_HOME"
-
-declare -a ARGS; eval "ARGS=($JEDIT_OPTIONS)"
-
-declare -a OPTIONS; eval "OPTIONS=($ISABELLE_JEDIT_OPTIONS)"
-getoptions "${OPTIONS[@]}"
-
-getoptions "$@"
-shift $(($OPTIND - 1))
-
-
-# args
-
-while [ "$#" -gt 0 ]; do
- ARGS["${#ARGS[@]}"]="$(jvmpath "$1")"
- shift
-done
-
-
-## default perspective
-
-mkdir -p "$JEDIT_SETTINGS/DockableWindowManager"
-
-if [ ! -e "$JEDIT_SETTINGS/perspective.xml" ]; then
- cat > "$JEDIT_SETTINGS/DockableWindowManager/perspective-view0.xml" <<EOF
-<DOCKING LEFT="" TOP="" RIGHT="" BOTTOM="isabelle-session" LEFT_POS="0" TOP_POS="0" RIGHT_POS="250" BOTTOM_POS="250" />
-EOF
- cat > "$JEDIT_SETTINGS/perspective.xml" <<EOF
-<?xml version="1.0" encoding="UTF-8" ?>
-<!DOCTYPE PERSPECTIVE SYSTEM "perspective.dtd">
-<PERSPECTIVE>
-<VIEW PLAIN="FALSE">
-<GEOMETRY X="0" Y="35" WIDTH="1072" HEIGHT="787" EXT_STATE="0" />
-</VIEW>
-</PERSPECTIVE>
-EOF
-fi
-
-
-## main
-
-case "$JEDIT_LOGIC" in
- /*)
- ;;
- */*)
- JEDIT_LOGIC="$(pwd -P)/$JEDIT_LOGIC"
- ;;
-esac
-
-export JEDIT_LOGIC JEDIT_PRINT_MODE
-
-exec "$ISABELLE_TOOL" java "${JAVA_ARGS[@]}" \
- -jar "$(jvmpath "$JEDIT_HOME/jedit.jar")" \
- "-settings=$(jvmpath "$JEDIT_SETTINGS")" "${ARGS[@]}"
--- a/src/Tools/jEdit/dist-template/modes/isabelle-session.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-<?xml version="1.0"?>
-<!DOCTYPE MODE SYSTEM "xmode.dtd">
-
-<!-- Isabelle session mode -->
-<MODE>
- <PROPS>
- <PROPERTY NAME="commentStart" VALUE="(*"/>
- <PROPERTY NAME="commentEnd" VALUE="*)"/>
- <PROPERTY NAME="noWordSep" VALUE="_'.?"/>
- <PROPERTY NAME="unalignedOpenBrackets" VALUE="(" />
- <PROPERTY NAME="unalignedCloseBrackets" VALUE=")" />
- <PROPERTY NAME="tabSize" VALUE="2" />
- <PROPERTY NAME="indentSize" VALUE="2" />
- </PROPS>
- <RULES IGNORE_CASE="FALSE" HIGHLIGHT_DIGITS="FALSE" ESCAPE="\">
- <SPAN TYPE="COMMENT1">
- <BEGIN>(*</BEGIN>
- <END>*)</END>
- </SPAN>
- <SPAN TYPE="COMMENT3">
- <BEGIN>{*</BEGIN>
- <END>*}</END>
- </SPAN>
- <SPAN TYPE="LITERAL1">
- <BEGIN>`</BEGIN>
- <END>`</END>
- </SPAN>
- <SPAN TYPE="LITERAL3">
- <BEGIN>"</BEGIN>
- <END>"</END>
- </SPAN>
- <KEYWORDS>
- <KEYWORD1>session</KEYWORD1>
- <KEYWORD2>parent</KEYWORD2>
- <KEYWORD2>imports</KEYWORD2>
- <KEYWORD2>uses</KEYWORD2>
- <KEYWORD2>options</KEYWORD2>
- <KEYWORD2>dependencies</KEYWORD2>
- </KEYWORDS>
- </RULES>
-</MODE>
--- a/src/Tools/jEdit/dist-template/modes/isabelle.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-<?xml version="1.0"?>
-<!DOCTYPE MODE SYSTEM "xmode.dtd">
-
-<!-- Isabelle theory mode -->
-<MODE>
- <PROPS>
- <PROPERTY NAME="commentStart" VALUE="(*"/>
- <PROPERTY NAME="commentEnd" VALUE="*)"/>
- <PROPERTY NAME="noWordSep" VALUE="_'.?"/>
- <PROPERTY NAME="unalignedOpenBrackets" VALUE="{[(" />
- <PROPERTY NAME="unalignedCloseBrackets" VALUE=")]}" />
- <PROPERTY NAME="tabSize" VALUE="2" />
- <PROPERTY NAME="indentSize" VALUE="2" />
- </PROPS>
- <RULES IGNORE_CASE="FALSE" HIGHLIGHT_DIGITS="FALSE" ESCAPE="\">
- <SPAN TYPE="COMMENT1">
- <BEGIN>(*</BEGIN>
- <END>*)</END>
- </SPAN>
- <SPAN TYPE="COMMENT3">
- <BEGIN>{*</BEGIN>
- <END>*}</END>
- </SPAN>
- <SPAN TYPE="LITERAL1">
- <BEGIN>`</BEGIN>
- <END>`</END>
- </SPAN>
- <SPAN TYPE="LITERAL3">
- <BEGIN>"</BEGIN>
- <END>"</END>
- </SPAN>
- <KEYWORDS>
- <KEYWORD2>header</KEYWORD2>
- <KEYWORD1>theory</KEYWORD1>
- <KEYWORD2>imports</KEYWORD2>
- <KEYWORD2>uses</KEYWORD2>
- <KEYWORD2>begin</KEYWORD2>
- <KEYWORD2>end</KEYWORD2>
- </KEYWORDS>
- </RULES>
-</MODE>
--- a/src/Tools/jEdit/dist-template/properties/jedit.props Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,213 +0,0 @@
-#jEdit properties
-buffer.deepIndent=false
-buffer.encoding=UTF-8-Isabelle
-buffer.indentSize=2
-buffer.lineSeparator=\n
-buffer.maxLineLen=100
-buffer.noTabs=true
-buffer.sidekick.keystroke-parse=false
-buffer.tabSize=2
-console.encoding=UTF-8
-console.font=IsabelleText
-console.fontsize=14
-delete-line.shortcut=A+d
-delete.shortcut2=C+d
-encoding.opt-out.Big5-HKSCS=true
-encoding.opt-out.Big5=true
-encoding.opt-out.COMPOUND_TEXT=true
-encoding.opt-out.EUC-JP=true
-encoding.opt-out.EUC-KR=true
-encoding.opt-out.GB18030=true
-encoding.opt-out.GB2312=true
-encoding.opt-out.GBK=true
-encoding.opt-out.IBM-Thai=true
-encoding.opt-out.IBM00858=true
-encoding.opt-out.IBM01140=true
-encoding.opt-out.IBM01141=true
-encoding.opt-out.IBM01142=true
-encoding.opt-out.IBM01143=true
-encoding.opt-out.IBM01144=true
-encoding.opt-out.IBM01145=true
-encoding.opt-out.IBM01146=true
-encoding.opt-out.IBM01147=true
-encoding.opt-out.IBM01148=true
-encoding.opt-out.IBM01149=true
-encoding.opt-out.IBM037=true
-encoding.opt-out.IBM1026=true
-encoding.opt-out.IBM1047=true
-encoding.opt-out.IBM273=true
-encoding.opt-out.IBM277=true
-encoding.opt-out.IBM278=true
-encoding.opt-out.IBM280=true
-encoding.opt-out.IBM284=true
-encoding.opt-out.IBM285=true
-encoding.opt-out.IBM297=true
-encoding.opt-out.IBM420=true
-encoding.opt-out.IBM424=true
-encoding.opt-out.IBM437=true
-encoding.opt-out.IBM500=true
-encoding.opt-out.IBM775=true
-encoding.opt-out.IBM850=true
-encoding.opt-out.IBM852=true
-encoding.opt-out.IBM855=true
-encoding.opt-out.IBM857=true
-encoding.opt-out.IBM860=true
-encoding.opt-out.IBM861=true
-encoding.opt-out.IBM862=true
-encoding.opt-out.IBM863=true
-encoding.opt-out.IBM864=true
-encoding.opt-out.IBM865=true
-encoding.opt-out.IBM866=true
-encoding.opt-out.IBM868=true
-encoding.opt-out.IBM869=true
-encoding.opt-out.IBM870=true
-encoding.opt-out.IBM871=true
-encoding.opt-out.IBM918=true
-encoding.opt-out.ISO-2022-CN=true
-encoding.opt-out.ISO-2022-JP-2=true
-encoding.opt-out.ISO-2022-JP=true
-encoding.opt-out.ISO-2022-KR=true
-encoding.opt-out.ISO-8859-13=true
-encoding.opt-out.ISO-8859-2=true
-encoding.opt-out.ISO-8859-3=true
-encoding.opt-out.ISO-8859-4=true
-encoding.opt-out.ISO-8859-5=true
-encoding.opt-out.ISO-8859-6=true
-encoding.opt-out.ISO-8859-7=true
-encoding.opt-out.ISO-8859-8=true
-encoding.opt-out.ISO-8859-9=true
-encoding.opt-out.JIS_X0201=true
-encoding.opt-out.JIS_X0212-1990=true
-encoding.opt-out.KOI8-R=true
-encoding.opt-out.KOI8-U=true
-encoding.opt-out.Shift_JIS=true
-encoding.opt-out.TIS-620=true
-encoding.opt-out.UTF-16=true
-encoding.opt-out.UTF-16BE=true
-encoding.opt-out.UTF-16LE=true
-encoding.opt-out.UTF-32=true
-encoding.opt-out.UTF-32BE=true
-encoding.opt-out.UTF-32LE=true
-encoding.opt-out.X-UTF-32BE-BOM=true
-encoding.opt-out.X-UTF-32LE-BOM=true
-encoding.opt-out.windows-1250=true
-encoding.opt-out.windows-1251=true
-encoding.opt-out.windows-1253=true
-encoding.opt-out.windows-1254=true
-encoding.opt-out.windows-1255=true
-encoding.opt-out.windows-1256=true
-encoding.opt-out.windows-1257=true
-encoding.opt-out.windows-1258=true
-encoding.opt-out.windows-31j=true
-encoding.opt-out.x-Big5-Solaris=true
-encoding.opt-out.x-EUC-TW=true
-encoding.opt-out.x-IBM1006=true
-encoding.opt-out.x-IBM1025=true
-encoding.opt-out.x-IBM1046=true
-encoding.opt-out.x-IBM1097=true
-encoding.opt-out.x-IBM1098=true
-encoding.opt-out.x-IBM1112=true
-encoding.opt-out.x-IBM1122=true
-encoding.opt-out.x-IBM1123=true
-encoding.opt-out.x-IBM1124=true
-encoding.opt-out.x-IBM1381=true
-encoding.opt-out.x-IBM1383=true
-encoding.opt-out.x-IBM33722=true
-encoding.opt-out.x-IBM737=true
-encoding.opt-out.x-IBM834=true
-encoding.opt-out.x-IBM856=true
-encoding.opt-out.x-IBM874=true
-encoding.opt-out.x-IBM875=true
-encoding.opt-out.x-IBM921=true
-encoding.opt-out.x-IBM922=true
-encoding.opt-out.x-IBM930=true
-encoding.opt-out.x-IBM933=true
-encoding.opt-out.x-IBM935=true
-encoding.opt-out.x-IBM937=true
-encoding.opt-out.x-IBM939=true
-encoding.opt-out.x-IBM942=true
-encoding.opt-out.x-IBM942C=true
-encoding.opt-out.x-IBM943=true
-encoding.opt-out.x-IBM943C=true
-encoding.opt-out.x-IBM948=true
-encoding.opt-out.x-IBM949=true
-encoding.opt-out.x-IBM949C=true
-encoding.opt-out.x-IBM950=true
-encoding.opt-out.x-IBM964=true
-encoding.opt-out.x-IBM970=true
-encoding.opt-out.x-ISCII91=true
-encoding.opt-out.x-ISO-2022-CN-CNS=true
-encoding.opt-out.x-ISO-2022-CN-GB=true
-encoding.opt-out.x-JIS0208=true
-encoding.opt-out.x-JISAutoDetect=true
-encoding.opt-out.x-Johab=true
-encoding.opt-out.x-MS932_0213=true
-encoding.opt-out.x-MS950-HKSCS=true
-encoding.opt-out.x-MacArabic=true
-encoding.opt-out.x-MacCentralEurope=true
-encoding.opt-out.x-MacCroatian=true
-encoding.opt-out.x-MacCyrillic=true
-encoding.opt-out.x-MacDingbat=true
-encoding.opt-out.x-MacGreek=true
-encoding.opt-out.x-MacHebrew=true
-encoding.opt-out.x-MacIceland=true
-encoding.opt-out.x-MacRoman=true
-encoding.opt-out.x-MacRomania=true
-encoding.opt-out.x-MacSymbol=true
-encoding.opt-out.x-MacThai=true
-encoding.opt-out.x-MacTurkish=true
-encoding.opt-out.x-MacUkraine=true
-encoding.opt-out.x-PCK=true
-encoding.opt-out.x-SJIS_0213=true
-encoding.opt-out.x-UTF-16LE-BOM=true
-encoding.opt-out.x-euc-jp-linux=true
-encoding.opt-out.x-eucJP-Open=true
-encoding.opt-out.x-iso-8859-11=true
-encoding.opt-out.x-mswin-936=true
-encoding.opt-out.x-windows-50220=true
-encoding.opt-out.x-windows-50221=true
-encoding.opt-out.x-windows-874=true
-encoding.opt-out.x-windows-949=true
-encoding.opt-out.x-windows-950=true
-encoding.opt-out.x-windows-iso2022jp=true
-encodingDetectors=BOM XML-PI buffer-local-property
-end.shortcut=
-fallbackEncodings=UTF-8 ISO-8859-15 US-ASCII
-firstTime=false
-home.shortcut=
-insert-newline-indent.shortcut=
-insert-newline.shortcut=ENTER
-isabelle-output.dock-position=bottom
-isabelle-output.height=174
-isabelle-output.width=412
-isabelle-session.dock-position=bottom
-line-end.shortcut=END
-line-home.shortcut=HOME
-lookAndFeel=com.sun.java.swing.plaf.nimbus.NimbusLookAndFeel
-mode.isabelle.customSettings=true
-mode.isabelle.folding=sidekick
-mode.isabelle.sidekick.showStatusWindow.label=true
-print.font=IsabelleText
-restore.remote=false
-restore=false
-sidekick-tree.dock-position=right
-sidekick.buffer-save-parse=true
-sidekick.complete-delay=300
-sidekick.splitter.location=721
-tip.show=false
-twoStageSave=false
-view.antiAlias=standard
-view.blockCaret=true
-view.caretBlink=false
-view.eolMarkers=false
-view.extendedState=0
-view.font=IsabelleText
-view.fontsize=18
-view.fracFontMetrics=false
-view.gutter.fontsize=12
-view.gutter.selectionAreaWidth=18
-view.height=787
-view.middleMousePaste=true
-view.showToolbar=false
-view.thickCaret=true
-view.width=1072
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/etc/isabelle-jedit.css Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,19 @@
+/* additional style file for Isabelle/jEdit output */
+
+.message { margin-top: 0.3ex; background-color: #F0F0F0; }
+
+.writeln { }
+.tracing { background-color: #F0F8FF; }
+.warning { background-color: #EEE8AA; }
+.error { background-color: #FFC1C1; }
+
+.report { display: none; }
+
+.hilite { background-color: #FFCC66; }
+
+.keyword { font-weight: bold; color: #009966; }
+.operator { font-weight: bold; }
+.command { font-weight: bold; color: #006699; }
+
+.sendback { text-decoration: underline; }
+.sendback:hover { background-color: #FFCC66; }
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/etc/settings Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,16 @@
+# -*- shell-script -*- :mode=shellscript:
+
+JEDIT_HOME="$COMPONENT"
+JEDIT_SETTINGS="$ISABELLE_HOME_USER/jedit"
+
+JEDIT_OPTIONS="-reuseview -noserver -nobackground"
+JEDIT_JAVA_OPTIONS="-Xms128m -Xmx1024m -Xss4m -Dactors.corePoolSize=4 -Dactors.enableForkJoin=false"
+#JEDIT_JAVA_OPTIONS="-Xms128m -Xmx512m -Xss2m -Dactors.corePoolSize=4 -Dactors.enableForkJoin=false"
+JEDIT_SYSTEM_OPTIONS="-Dapple.laf.useScreenMenuBar=true -Dcom.apple.mrj.application.apple.menu.about.name=Isabelle/jEdit"
+
+JEDIT_STYLE_SHEETS="$ISABELLE_HOME/lib/html/isabelle.css:$JEDIT_HOME/etc/isabelle-jedit.css:$ISABELLE_HOME_USER/etc/isabelle.css:$ISABELLE_HOME_USER/etc/isabelle-jedit.css"
+
+ISABELLE_JEDIT_OPTIONS="-m xsymbols -m no_brackets -m no_type_brackets"
+
+ISABELLE_TOOLS="$ISABELLE_TOOLS:$JEDIT_HOME/lib/Tools"
+
--- a/src/Tools/jEdit/jedit_build/Tools/jedit Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,141 +0,0 @@
-#!/usr/bin/env bash
-#
-# Author: Makarius
-#
-# DESCRIPTION: build and run Isabelle/jEdit on the spot
-
-## diagnostics
-
-function fail()
-{
- echo "$1" >&2
- exit 2
-}
-
-function failed()
-{
- fail "Failed!"
-}
-
-
-## dependencies
-
-[ -z "$ISABELLE_HOME" ] && fail "Missing Isabelle settings environment"
-[ -z "$SCALA_HOME" ] && fail "Unknown SCALA_HOME -- Scala unavailable"
-
-[ -z "$ISABELLE_JEDIT_BUILD_HOME" ] && \
- fail "Unknown ISABELLE_JEDIT_BUILD_HOME -- missing auxiliary component"
-
-[ -e "$ISABELLE_HOME/Admin/build" ] && \
- { "$ISABELLE_HOME/Admin/build" jars || exit $?; }
-
-
-# sources
-
-SRC_DIR="$ISABELLE_HOME/src/Tools/jEdit"
-
-declare -a SOURCES=(
- "$SRC_DIR/src/jedit/dockable.scala"
- "$SRC_DIR/src/jedit/document_model.scala"
- "$SRC_DIR/src/jedit/document_view.scala"
- "$SRC_DIR/src/jedit/html_panel.scala"
- "$SRC_DIR/src/jedit/isabelle_encoding.scala"
- "$SRC_DIR/src/jedit/isabelle_hyperlinks.scala"
- "$SRC_DIR/src/jedit/isabelle_markup.scala"
- "$SRC_DIR/src/jedit/isabelle_options.scala"
- "$SRC_DIR/src/jedit/isabelle_sidekick.scala"
- "$SRC_DIR/src/jedit/output_dockable.scala"
- "$SRC_DIR/src/jedit/plugin.scala"
- "$SRC_DIR/src/jedit/protocol_dockable.scala"
- "$SRC_DIR/src/jedit/raw_output_dockable.scala"
- "$SRC_DIR/src/jedit/scala_console.scala"
- "$SRC_DIR/src/jedit/session_dockable.scala"
-)
-
-declare -a PLUGIN_FILES=(
- "$SRC_DIR/plugin/actions.xml"
- "$SRC_DIR/plugin/dockables.xml"
- "$SRC_DIR/plugin/Isabelle.props"
- "$SRC_DIR/plugin/services.xml"
-)
-
-JEDIT_JAR="$ISABELLE_JEDIT_BUILD_HOME/contrib/$ISABELLE_JEDIT_BUILD_VERSION/jedit.jar"
-
-JEDIT_JARS=(
- "$ISABELLE_JEDIT_BUILD_HOME/contrib/Console.jar"
- "$ISABELLE_JEDIT_BUILD_HOME/contrib/ErrorList.jar"
- "$ISABELLE_JEDIT_BUILD_HOME/contrib/Hyperlinks.jar"
- "$ISABELLE_JEDIT_BUILD_HOME/contrib/SideKick.jar"
- "$ISABELLE_JEDIT_BUILD_HOME/contrib/cobra.jar"
- "$ISABELLE_JEDIT_BUILD_HOME/contrib/js.jar"
-)
-
-SCALA_JARS=(
- "$SCALA_HOME/lib/scala-compiler.jar"
- "$SCALA_HOME/lib/scala-library.jar"
- "$SCALA_HOME/lib/scala-swing.jar"
-)
-
-
-# target
-
-TARGET_DIR="$ISABELLE_JEDIT_BUILD_HOME/${ISABELLE_JEDIT_BUILD_VERSION}_Isabelle-$("$ISABELLE_TOOL" version -i)"
-TARGET="$TARGET_DIR/jars/Isabelle-jEdit.jar"
-
-OUTDATED=false
-for SOURCE in "${SOURCES[@]}" "${PLUGIN_FILES[@]}" "$JEDIT_JAR" "${JEDIT_JARS[@]}"
-do
- [ ! -e "$SOURCE" ] && fail "Missing file: $SOURCE"
- [ ! -e "$TARGET" -o "$SOURCE" -nt "$TARGET" ] && OUTDATED=true
-done
-
-
-## main
-
-if [ "$OUTDATED" = true ]
-then
- echo "###"
- echo "### Building Isabelle/jEdit ..."
- echo "###"
-
- rm -rf "$TARGET_DIR" || failed
- mkdir -p "$TARGET_DIR" "$TARGET_DIR/classes" || failed
- cp -a "$ISABELLE_JEDIT_BUILD_HOME/contrib/$ISABELLE_JEDIT_BUILD_VERSION/." "$TARGET_DIR/."
-
- cp -a "${PLUGIN_FILES[@]}" "$TARGET_DIR/classes/."
- cp -a "$SRC_DIR/dist-template/." "$TARGET_DIR/."
- cp -a "$SRC_DIR/README" "$TARGET_DIR/."
-
- perl -i -e 'while (<>) { if (m/NAME="javacc"/) {
- print qq,<MODE NAME="isabelle" FILE="isabelle.xml" FILE_NAME_GLOB="*.thy"/>\n\n,;
- print qq,<MODE NAME="isabelle-session" FILE="isabelle-session.xml" FILE_NAME_GLOB="session.root"/>\n\n,; }
- print; }' "$TARGET_DIR/modes/catalog"
-
- (
- for JAR in "$JEDIT_JAR" "${JEDIT_JARS[@]}" "$SCALA_HOME/lib/scala-compiler.jar"
- do
- CLASSPATH="$CLASSPATH:$JAR"
- done
-
- cp -a "${JEDIT_JARS[@]}" "${SCALA_JARS[@]}" "$ISABELLE_HOME/lib/classes/Pure.jar" \
- "$TARGET_DIR/jars/." || failed
-
- CLASSPATH="$(jvmpath "$CLASSPATH")"
- "$SCALA_HOME/bin/scalac" -unchecked -deprecation \
- -d "$TARGET_DIR/classes" -target:jvm-1.5 "${SOURCES[@]}" || \
- fail "Failed to compile sources"
-
- cd "$TARGET_DIR/classes"
- jar cf "../jars/Isabelle-jEdit.jar" * || failed
- cd ..
- rm -rf classes
- )
-fi
-
-set -o allexport
-init_component "$TARGET_DIR"
-[ -f "$ISABELLE_JEDIT_BUILD_HOME/etc/user-settings" ] && \
- source "$ISABELLE_JEDIT_BUILD_HOME/etc/user-settings"
-set +o allexport
-
-exec "$TARGET_DIR/lib/Tools/jedit" "$@"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/lib/Tools/jedit Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,271 @@
+#!/usr/bin/env bash
+#
+# Author: Makarius
+#
+# DESCRIPTION: Isabelle/jEdit interface wrapper
+
+
+## sources
+
+declare -a SOURCE_FILES=(
+ "src/dockable.scala"
+ "src/document_model.scala"
+ "src/document_view.scala"
+ "src/html_panel.scala"
+ "src/isabelle_encoding.scala"
+ "src/isabelle_hyperlinks.scala"
+ "src/isabelle_markup.scala"
+ "src/isabelle_options.scala"
+ "src/isabelle_sidekick.scala"
+ "src/output_dockable.scala"
+ "src/plugin.scala"
+ "src/protocol_dockable.scala"
+ "src/raw_output_dockable.scala"
+ "src/scala_console.scala"
+ "src/session_dockable.scala"
+)
+
+declare -a MORE_FILES=(
+ "src/actions.xml"
+ "src/dockables.xml"
+ "src/Isabelle.props"
+ "src/services.xml"
+)
+
+
+## diagnostics
+
+PRG="$(basename "$0")"
+
+function usage()
+{
+ echo
+ echo "Usage: isabelle $PRG [OPTIONS] [FILES ...]"
+ echo
+ echo " Options are:"
+ echo " -J OPTION add JVM runtime option"
+ echo " (default JEDIT_JAVA_OPTIONS=$JEDIT_JAVA_OPTIONS)"
+ echo " -b build only"
+ echo " -d enable debugger"
+ echo " -f fresh build"
+ echo " -j OPTION add jEdit runtime option"
+ echo " (default JEDIT_OPTIONS=$JEDIT_OPTIONS)"
+ echo " -l NAME logic image name (default ISABELLE_LOGIC=$ISABELLE_LOGIC)"
+ echo " -m MODE add print mode for output"
+ echo
+ echo "Start jEdit with Isabelle plugin setup and opens theory FILES"
+ echo
+ exit 1
+}
+
+function fail()
+{
+ echo "$1" >&2
+ exit 2
+}
+
+function failed()
+{
+ fail "Failed!"
+}
+
+
+## process command line
+
+# options
+
+BUILD_ONLY=false
+BUILD_JARS="jars"
+JEDIT_LOGIC="$ISABELLE_LOGIC"
+JEDIT_PRINT_MODE=""
+
+function getoptions()
+{
+ OPTIND=1
+ while getopts "J:bdfj:l:m:" OPT
+ do
+ case "$OPT" in
+ J)
+ JAVA_ARGS["${#JAVA_ARGS[@]}"]="$OPTARG"
+ ;;
+ b)
+ BUILD_ONLY=true
+ ;;
+ d)
+ JAVA_ARGS["${#JAVA_ARGS[@]}"]="-Xdebug"
+ JAVA_ARGS["${#JAVA_ARGS[@]}"]="-Xrunjdwp:transport=dt_socket,server=y,suspend=n"
+ ;;
+ f)
+ BUILD_JARS="jars_fresh"
+ ;;
+ j)
+ ARGS["${#ARGS[@]}"]="$OPTARG"
+ ;;
+ l)
+ JEDIT_LOGIC="$OPTARG"
+ ;;
+ m)
+ if [ -z "$JEDIT_PRINT_MODE" ]; then
+ JEDIT_PRINT_MODE="$OPTARG"
+ else
+ JEDIT_PRINT_MODE="$JEDIT_PRINT_MODE,$OPTARG"
+ fi
+ ;;
+ \?)
+ usage
+ ;;
+ esac
+ done
+}
+
+declare -a JAVA_ARGS; eval "JAVA_ARGS=($JEDIT_JAVA_OPTIONS $JEDIT_SYSTEM_OPTIONS)"
+[ -n "$SCALA_HOME" ] && JAVA_ARGS["${#JAVA_ARGS[@]}"]="-Dscala.home=$SCALA_HOME"
+
+declare -a ARGS; eval "ARGS=($JEDIT_OPTIONS)"
+
+declare -a OPTIONS; eval "OPTIONS=($ISABELLE_JEDIT_OPTIONS)"
+getoptions "${OPTIONS[@]}"
+
+getoptions "$@"
+shift $(($OPTIND - 1))
+
+
+# args
+
+while [ "$#" -gt 0 ]
+do
+ ARGS["${#ARGS[@]}"]="$(jvmpath "$1")"
+ shift
+done
+
+
+## dependencies
+
+pushd "$JEDIT_HOME" >/dev/null || failed
+
+JEDIT_JAR="$ISABELLE_JEDIT_BUILD_HOME/contrib/$ISABELLE_JEDIT_BUILD_VERSION/jedit.jar"
+
+JEDIT_JARS=(
+ "$ISABELLE_JEDIT_BUILD_HOME/contrib/Console.jar"
+ "$ISABELLE_JEDIT_BUILD_HOME/contrib/ErrorList.jar"
+ "$ISABELLE_JEDIT_BUILD_HOME/contrib/Hyperlinks.jar"
+ "$ISABELLE_JEDIT_BUILD_HOME/contrib/SideKick.jar"
+ "$ISABELLE_JEDIT_BUILD_HOME/contrib/cobra.jar"
+ "$ISABELLE_JEDIT_BUILD_HOME/contrib/js.jar"
+)
+
+SCALA_JARS=(
+ "$SCALA_HOME/lib/scala-compiler.jar"
+ "$SCALA_HOME/lib/scala-library.jar"
+ "$SCALA_HOME/lib/scala-swing.jar"
+)
+
+
+# target
+
+TARGET="dist/jars/Isabelle-jEdit.jar"
+
+if [ "$BUILD_JARS" = jars_fresh ]; then
+ OUTDATED=true
+else
+ OUTDATED=false
+ for SOURCE in "${SOURCE_FILES[@]}" "${MORE_FILES[@]}" "$JEDIT_JAR" "${JEDIT_JARS[@]}"
+ do
+ [ ! -e "$SOURCE" ] && fail "Missing file: $SOURCE"
+ [ ! -e "$TARGET" -o "$SOURCE" -nt "$TARGET" ] && OUTDATED=true
+ done
+fi
+
+
+# build
+
+if [ "$OUTDATED" = true ]
+then
+ [ -z "$SCALA_HOME" ] && fail "Unknown SCALA_HOME -- Scala unavailable"
+
+ [ -z "$ISABELLE_JEDIT_BUILD_HOME" ] && \
+ fail "Unknown ISABELLE_JEDIT_BUILD_HOME -- missing auxiliary component"
+
+ [ -e "$ISABELLE_HOME/Admin/build" ] && \
+ { "$ISABELLE_HOME/Admin/build" "$BUILD_JARS" || exit $?; }
+
+
+ echo "###"
+ echo "### Building Isabelle/jEdit ..."
+ echo "###"
+
+ rm -rf dist || failed
+ mkdir -p dist dist/classes || failed
+
+ cp -a "$ISABELLE_JEDIT_BUILD_HOME/contrib/$ISABELLE_JEDIT_BUILD_VERSION/." dist/.
+ cp -a "${MORE_FILES[@]}" dist/classes/.
+ cp src/jEdit.props dist/properties/.
+ cp -a src/modes/. dist/modes/.
+
+ perl -i -e 'while (<>) { if (m/NAME="javacc"/) {
+ print qq,<MODE NAME="isabelle" FILE="isabelle.xml" FILE_NAME_GLOB="*.thy"/>\n\n,;
+ print qq,<MODE NAME="isabelle-session" FILE="isabelle-session.xml" FILE_NAME_GLOB="session.root"/>\n\n,; }
+ print; }' dist/modes/catalog
+
+ cp -a "${JEDIT_JARS[@]}" "${SCALA_JARS[@]}" "$ISABELLE_HOME/lib/classes/Pure.jar" \
+ dist/jars/. || failed
+
+ (
+ for JAR in "$JEDIT_JAR" "${JEDIT_JARS[@]}" "$SCALA_HOME/lib/scala-compiler.jar"
+ do
+ CLASSPATH="$CLASSPATH:$JAR"
+ done
+ CLASSPATH="$(jvmpath "$CLASSPATH")"
+
+ exec "$SCALA_HOME/bin/scalac" -unchecked -deprecation \
+ -d dist/classes -target:jvm-1.5 "${SOURCE_FILES[@]}"
+ ) || fail "Failed to compile sources"
+
+ cd dist/classes
+ jar cf "../jars/Isabelle-jEdit.jar" * || failed
+ cd ../..
+ rm -rf dist/classes
+fi
+
+popd >/dev/null
+
+
+## main
+
+# perspective
+
+mkdir -p "$JEDIT_SETTINGS/DockableWindowManager"
+
+if [ ! -e "$JEDIT_SETTINGS/perspective.xml" ]; then
+ cat > "$JEDIT_SETTINGS/DockableWindowManager/perspective-view0.xml" <<EOF
+<DOCKING LEFT="" TOP="" RIGHT="" BOTTOM="isabelle-session" LEFT_POS="0" TOP_POS="0" RIGHT_POS="250" BOTTOM_POS="250" />
+EOF
+ cat > "$JEDIT_SETTINGS/perspective.xml" <<EOF
+<?xml version="1.0" encoding="UTF-8" ?>
+<!DOCTYPE PERSPECTIVE SYSTEM "perspective.dtd">
+<PERSPECTIVE>
+<VIEW PLAIN="FALSE">
+<GEOMETRY X="0" Y="35" WIDTH="1072" HEIGHT="787" EXT_STATE="0" />
+</VIEW>
+</PERSPECTIVE>
+EOF
+fi
+
+
+# run
+
+[ "$BUILD_ONLY" = true ] || {
+ case "$JEDIT_LOGIC" in
+ /*)
+ ;;
+ */*)
+ JEDIT_LOGIC="$(pwd -P)/$JEDIT_LOGIC"
+ ;;
+ esac
+
+ export JEDIT_LOGIC JEDIT_PRINT_MODE
+
+ exec "$ISABELLE_TOOL" java "${JAVA_ARGS[@]}" \
+ -jar "$(jvmpath "$JEDIT_HOME/dist/jedit.jar")" \
+ "-settings=$(jvmpath "$JEDIT_SETTINGS")" "${ARGS[@]}"
+}
--- a/src/Tools/jEdit/makedist Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,101 +0,0 @@
-#!/usr/bin/env bash
-#
-# makedist -- make Isabelle/jEdit distribution
-
-## self references
-
-PRG="$(basename "$0")"
-THIS="$(cd "$(dirname "$0")"; pwd)"
-SUPER="$(cd "$THIS/.."; pwd)"
-
-
-## diagnostics
-
-JEDIT_HOME=""
-
-function usage()
-{
- echo
- echo "Usage: $PRG [OPTIONS]"
- echo
- echo " Options are:"
- echo " -j DIR specify original jEdit distribution"
- echo
- echo " Produce Isabelle/jEdit distribution from Netbeans build"
- echo " in $THIS/dist"
- echo
- exit 1
-}
-
-fail()
-{
- echo "$1" >&2
- exit 2
-}
-
-
-## process command line
-
-# options
-
-while getopts "j:" OPT
-do
- case "$OPT" in
- j)
- JEDIT_HOME="$OPTARG"
- ;;
- \?)
- usage
- ;;
- esac
-done
-
-shift $(($OPTIND - 1))
-
-
-# args
-
-[ "$#" -ne 0 ] && usage
-
-
-## main
-
-cd "$THIS/dist" || fail "Bad directory: $THIS/dist"
-
-
-# target name
-
-[ -z "$JEDIT_HOME" ] && fail "Unknown JEDIT_HOME"
-[ -n "$ISABELLE_HOME" ] || fail "Missing Isabelle settings environment"
-
-VERSION="$(basename "$JEDIT_HOME")_Isabelle-$("$ISABELLE_TOOL" version -i)"
-JEDIT="jedit-${VERSION}"
-
-rm -rf "$JEDIT" jedit
-mkdir "$JEDIT"
-
-rm -f jedit && ln -s "$JEDIT" jedit
-
-
-# copy stuff
-
-[ "$JEDIT_HOME/jedit.jar" ] || fail "Bad original jEdit directory: $JEDIT_HOME"
-cp -R "$JEDIT_HOME/." "$JEDIT/."
-rm -rf "$JEDIT/jEdit" "$JEDIT/build-support"
-
-mkdir -p "$JEDIT/jars"
-cp -R jars/. "$JEDIT/jars/."
-
-cp -R "$THIS/dist-template/." "$JEDIT/."
-cp "$THIS/README" "$JEDIT/."
-
-perl -i -e 'while (<>) { if (m/NAME="javacc"/) {
- print qq,<MODE NAME="isabelle" FILE="isabelle.xml" FILE_NAME_GLOB="*.thy"/>\n\n,;
- print qq,<MODE NAME="isabelle-session" FILE="isabelle-session.xml" FILE_NAME_GLOB="session.root"/>\n\n,; }
- print; }' "$JEDIT/modes/catalog"
-
-
-# build archive
-
-echo "${JEDIT}.tar.gz"
-tar czf "${JEDIT}.tar.gz" "$JEDIT"
--- a/src/Tools/jEdit/manifest.mf Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-Manifest-Version: 1.0
-X-COMMENT: Main-Class will be added automatically by build
-
--- a/src/Tools/jEdit/nbproject/build-impl.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,707 +0,0 @@
-<?xml version="1.0" encoding="UTF-8"?>
-<!--
-*** GENERATED FROM project.xml - DO NOT EDIT ***
-*** EDIT ../build.xml INSTEAD ***
-
-For the purpose of easier reading the script
-is divided into following sections:
-
- - initialization
- - compilation
- - jar
- - execution
- - debugging
- - javadoc
- - junit compilation
- - junit execution
- - junit debugging
- - applet
- - cleanup
-
-
- -->
-<project xmlns:jaxrpc="http://www.netbeans.org/ns/scala-project/jax-rpc" xmlns:scalaProject1="http://www.netbeans.org/ns/scala-project/1" basedir=".." default="default" name="Isabelle-jEdit-impl">
- <target depends="test,jar,javadoc" description="Build and test whole project." name="default"/>
- <!--
- ======================
- INITIALIZATION SECTION
- ======================
- -->
- <target name="-pre-init">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="-pre-init" name="-init-private">
- <property file="nbproject/private/config.properties"/>
- <property file="nbproject/private/configs/${config}.properties"/>
- <property file="nbproject/private/private.properties"/>
- <property environment="env"/>
- <condition property="scala.home" value="${env.SCALA_HOME}">
- <isset property="env.SCALA_HOME"/>
- </condition>
- <fail unless="scala.home">
- You must set SCALA_HOME or environment property and append "-J-Dscala.home=scalahomepath"
- property to the end of "netbeans_default_options" in NetBeansInstallationPath/etc/netbeans.conf to point to
- Scala installation directory.
- </fail>
- <property name="scala.compiler" value="${scala.home}/lib/scala-compiler.jar"/>
- <property name="scala.library" value="${scala.home}/lib/scala-library.jar"/>
- <property name="scala.lib" value="${scala.home}/lib"/>
- <taskdef resource="scala/tools/ant/antlib.xml">
- <classpath>
- <pathelement location="${scala.compiler}"/>
- <pathelement location="${scala.library}"/>
- </classpath>
- </taskdef>
- </target>
- <target depends="-pre-init,-init-private" name="-init-user">
- <property file="${user.properties.file}"/>
- <!-- The two properties below are usually overridden -->
- <!-- by the active platform. Just a fallback. -->
- <property name="default.javac.source" value="1.5"/>
- <property name="default.javac.target" value="1.5"/>
- </target>
- <target depends="-pre-init,-init-private,-init-user" name="-init-project">
- <property file="nbproject/configs/${config}.properties"/>
- <property file="nbproject/project.properties"/>
- </target>
- <target depends="-pre-init,-init-private,-init-user,-init-project,-init-macrodef-property" name="-do-init">
- <available file="${manifest.file}" property="manifest.available"/>
- <condition property="manifest.available+main.class">
- <and>
- <isset property="manifest.available"/>
- <isset property="main.class"/>
- <not>
- <equals arg1="${main.class}" arg2="" trim="true"/>
- </not>
- </and>
- </condition>
- <condition property="manifest.available+main.class+mkdist.available">
- <and>
- <istrue value="${manifest.available+main.class}"/>
- <isset property="libs.CopyLibs.classpath"/>
- </and>
- </condition>
- <condition property="have.tests">
- <or/>
- </condition>
- <condition property="have.sources">
- <or>
- <available file="${src.dir}"/>
- </or>
- </condition>
- <condition property="netbeans.home+have.tests">
- <and>
- <isset property="netbeans.home"/>
- <isset property="have.tests"/>
- </and>
- </condition>
- <condition property="no.javadoc.preview">
- <and>
- <isset property="javadoc.preview"/>
- <isfalse value="${javadoc.preview}"/>
- </and>
- </condition>
- <property name="run.jvmargs" value=""/>
- <property name="javac.compilerargs" value=""/>
- <property name="work.dir" value="${basedir}"/>
- <condition property="no.deps">
- <and>
- <istrue value="${no.dependencies}"/>
- </and>
- </condition>
- <property name="javac.debug" value="true"/>
- <property name="javadoc.preview" value="true"/>
- <property name="application.args" value=""/>
- <property name="source.encoding" value="${file.encoding}"/>
- <condition property="javadoc.encoding.used" value="${javadoc.encoding}">
- <and>
- <isset property="javadoc.encoding"/>
- <not>
- <equals arg1="${javadoc.encoding}" arg2=""/>
- </not>
- </and>
- </condition>
- <property name="javadoc.encoding.used" value="${source.encoding}"/>
- <property name="includes" value="**"/>
- <property name="excludes" value=""/>
- <property name="extdirs" value=" "/>
- <property name="do.depend" value="false"/>
- <condition property="do.depend.true">
- <istrue value="${do.depend}"/>
- </condition>
- <condition else="" property="javac.compilerargs.jaxws" value="-Djava.endorsed.dirs='${jaxws.endorsed.dir}'">
- <and>
- <isset property="jaxws.endorsed.dir"/>
- <available file="nbproject/jaxws-build.xml"/>
- </and>
- </condition>
- </target>
- <target name="-post-init">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="-pre-init,-init-private,-init-user,-init-project,-do-init" name="-init-check">
- <fail unless="src.dir">Must set src.dir</fail>
- <fail unless="build.dir">Must set build.dir</fail>
- <fail unless="dist.dir">Must set dist.dir</fail>
- <fail unless="build.classes.dir">Must set build.classes.dir</fail>
- <fail unless="dist.javadoc.dir">Must set dist.javadoc.dir</fail>
- <fail unless="build.test.classes.dir">Must set build.test.classes.dir</fail>
- <fail unless="build.test.results.dir">Must set build.test.results.dir</fail>
- <fail unless="build.classes.excludes">Must set build.classes.excludes</fail>
- <fail unless="dist.jar">Must set dist.jar</fail>
- </target>
- <target name="-init-macrodef-property">
- <macrodef name="property" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute name="name"/>
- <attribute name="value"/>
- <sequential>
- <property name="@{name}" value="${@{value}}"/>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-macrodef-javac">
- <macrodef name="javac" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${src.dir}" name="srcdir"/>
- <attribute default="${build.classes.dir}" name="destdir"/>
- <attribute default="${javac.classpath}" name="classpath"/>
- <attribute default="${includes}" name="includes"/>
- <attribute default="${excludes}" name="excludes"/>
- <attribute default="${javac.debug}" name="debug"/>
- <attribute default="" name="sourcepath"/>
- <element name="customize" optional="true"/>
- <sequential>
- <javac debug="@{debug}" deprecation="${javac.deprecation}" destdir="@{destdir}" encoding="${source.encoding}" excludes="@{excludes}" includeantruntime="false" includes="@{includes}" source="${javac.source}" sourcepath="@{sourcepath}" srcdir="@{srcdir}" target="${javac.target}">
- <classpath>
- <path path="@{classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- </classpath>
- <compilerarg line="${javac.compilerargs} ${javac.compilerargs.jaxws}"/>
- <customize/>
- </javac>
- </sequential>
- </macrodef>
- <macrodef name="depend" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${src.dir}" name="srcdir"/>
- <attribute default="${build.classes.dir}" name="destdir"/>
- <attribute default="${javac.classpath}" name="classpath"/>
- <sequential>
- <depend cache="${build.dir}/depcache" destdir="@{destdir}" excludes="${excludes}" includes="${includes}" srcdir="@{srcdir}">
- <classpath>
- <path>
- <pathelement path="@{classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- <pathelement location="${build.classes.dir}"/>
- </path>
- </classpath>
- </depend>
- </sequential>
- </macrodef>
- <macrodef name="force-recompile" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${build.classes.dir}" name="destdir"/>
- <sequential>
- <fail unless="javac.includes">Must set javac.includes</fail>
- <pathconvert pathsep="," property="javac.includes.binary">
- <path>
- <filelist dir="@{destdir}" files="${javac.includes}"/>
- </path>
- <globmapper from="*.java" to="*.class"/>
- </pathconvert>
- <delete>
- <files includes="${javac.includes.binary}"/>
- </delete>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-macrodef-scalac">
- <macrodef name="scalac" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${src.dir}" name="srcdir"/>
- <attribute default="${build.classes.dir}" name="destdir"/>
- <attribute default="${javac.classpath}" name="classpath"/>
- <attribute default="${extdirs}" name="extdirs"/>
- <attribute default="${includes}" name="includes"/>
- <attribute default="${excludes}" name="excludes"/>
- <attribute default="${scalac.compilerargs}" name="addparams"/>
- <attribute default="" name="sourcepath"/>
- <element name="customize" optional="true"/>
- <sequential>
- <scalac addparams="-make:transitive -dependencyfile "${basedir}/${build.dir}/.scala_dependencies" @{addparams}" deprecation="${scalac.deprecation}" destdir="@{destdir}" encoding="${source.encoding}" excludes="@{excludes}" extdirs="@{extdirs}" force="yes" fork="true" includes="@{includes}" sourcepath="@{sourcepath}" srcdir="@{srcdir}" target="jvm-${javac.target}" unchecked="${scalac.unchecked}">
- <classpath>
- <path>
- <pathelement path="@{classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- <pathelement location="${build.classes.dir}"/>
- </path>
- </classpath>
- <customize/>
- </scalac>
- </sequential>
- </macrodef>
- <macrodef name="force-recompile" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${build.classes.dir}" name="destdir"/>
- <sequential>
- <fail unless="javac.includes">Must set javac.includes</fail>
- <pathconvert pathsep="," property="javac.includes.binary">
- <path>
- <filelist dir="@{destdir}" files="${javac.includes}"/>
- </path>
- <globmapper from="*.scala" to="*.class"/>
- </pathconvert>
- <delete>
- <files includes="${javac.includes.binary}"/>
- </delete>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-macrodef-junit">
- <macrodef name="junit" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${includes}" name="includes"/>
- <attribute default="${excludes}" name="excludes"/>
- <attribute default="**" name="testincludes"/>
- <sequential>
- <junit dir="${work.dir}" errorproperty="tests.failed" failureproperty="tests.failed" fork="true" showoutput="true">
- <batchtest todir="${build.test.results.dir}">
- <fileset dir="${build.test.classes.dir}" excludes="@{excludes},${excludes}" includes="@{includes}">
- <filename name="@{testincludes}"/>
- </fileset>
- </batchtest>
- <classpath>
- <path path="${run.test.classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- </classpath>
- <syspropertyset>
- <propertyref prefix="test-sys-prop."/>
- <mapper from="test-sys-prop.*" to="*" type="glob"/>
- </syspropertyset>
- <formatter type="brief" usefile="false"/>
- <formatter type="xml"/>
- <jvmarg line="${run.jvmargs}"/>
- </junit>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-macrodef-nbjpda">
- <macrodef name="nbjpdastart" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${main.class}" name="name"/>
- <attribute default="${debug.classpath}" name="classpath"/>
- <attribute default="" name="stopclassname"/>
- <sequential>
- <nbjpdastart addressproperty="jpda.address" name="@{name}" stopclassname="@{stopclassname}" transport="dt_socket">
- <classpath>
- <path path="@{classpath}"/>
- </classpath>
- </nbjpdastart>
- </sequential>
- </macrodef>
- <macrodef name="nbjpdareload" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${build.classes.dir}" name="dir"/>
- <sequential>
- <nbjpdareload>
- <fileset dir="@{dir}" includes="${fix.includes}*.class"/>
- </nbjpdareload>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-debug-args">
- <property name="version-output" value="java version "${ant.java.version}"/>
- <condition property="have-jdk-older-than-1.4">
- <or>
- <contains string="${version-output}" substring="java version "1.0"/>
- <contains string="${version-output}" substring="java version "1.1"/>
- <contains string="${version-output}" substring="java version "1.2"/>
- <contains string="${version-output}" substring="java version "1.3"/>
- </or>
- </condition>
- <condition else="-Xdebug" property="debug-args-line" value="-Xdebug -Xnoagent -Djava.compiler=none">
- <istrue value="${have-jdk-older-than-1.4}"/>
- </condition>
- </target>
- <target depends="-init-debug-args" name="-init-macrodef-debug">
- <macrodef name="debug" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${main.class}" name="classname"/>
- <attribute default="${debug.classpath}" name="classpath"/>
- <element name="customize" optional="true"/>
- <sequential>
- <java classname="@{classname}" dir="${work.dir}" fork="true">
- <jvmarg line="${debug-args-line}"/>
- <jvmarg value="-Xrunjdwp:transport=dt_socket,address=${jpda.address}"/>
- <jvmarg line="${run.jvmargs}"/>
- <classpath>
- <path path="@{classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- </classpath>
- <syspropertyset>
- <propertyref prefix="run-sys-prop."/>
- <mapper from="run-sys-prop.*" to="*" type="glob"/>
- </syspropertyset>
- <customize/>
- </java>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-macrodef-java">
- <macrodef name="java" uri="http://www.netbeans.org/ns/scala-project/1">
- <attribute default="${main.class}" name="classname"/>
- <element name="customize" optional="true"/>
- <sequential>
- <java classname="@{classname}" dir="${work.dir}" fork="true">
- <jvmarg line="${run.jvmargs}"/>
- <classpath>
- <path path="${run.classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- </classpath>
- <syspropertyset>
- <propertyref prefix="run-sys-prop."/>
- <mapper from="run-sys-prop.*" to="*" type="glob"/>
- </syspropertyset>
- <customize/>
- </java>
- </sequential>
- </macrodef>
- </target>
- <target name="-init-presetdef-jar">
- <presetdef name="jar" uri="http://www.netbeans.org/ns/scala-project/1">
- <jar compress="${jar.compress}" jarfile="${dist.jar}">
- <scalaProject1:fileset dir="${build.classes.dir}"/>
- </jar>
- </presetdef>
- </target>
- <target depends="-pre-init,-init-private ,-init-user,-init-project,-do-init,-post-init,-init-check,-init-macrodef-property,-init-macrodef-javac,-init-macrodef-scalac,-init-macrodef-junit,-init-macrodef-nbjpda,-init-macrodef-debug,-init-macrodef-java,-init-presetdef-jar" name="init"/>
- <!--
- ===================
- COMPILATION SECTION
- ===================
- -->
- <target depends="init" name="deps-jar" unless="no.deps">
- <ant antfile="${project.jEdit}/build-nb.xml" inheritall="false" target="build"/>
- </target>
- <target depends="init,deps-jar" name="-pre-pre-compile">
- <mkdir dir="${build.classes.dir}"/>
- </target>
- <target name="-pre-compile">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target if="do.depend.true" name="-compile-depend">
- <scalaProject1:depend/>
- </target>
- <target depends="init,deps-jar,-pre-pre-compile,-pre-compile,-compile-depend" if="have.sources" name="-do-compile">
- <scalaProject1:scalac/>
- <scalaProject1:javac/>
- <copy todir="${build.classes.dir}">
- <fileset dir="${src.dir}" excludes="${build.classes.excludes},${excludes} " includes="${includes}"/>
- </copy>
- </target>
- <target name="-post-compile">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,deps-jar,-pre-pre-compile,-pre-compile,-do-compile,-post-compile" description="Compile project." name="compile"/>
- <target name="-pre-compile-single">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,deps-jar,-pre-pre-compile" name="-do-compile-single">
- <fail unless="javac.includes">Must select some files in the IDE or set javac.includes</fail>
- <scalaProject1:force-recompile/>
- <scalaProject1:scalac excludes="" includes="${javac.includes}" sourcepath="${src.dir}"/>
- </target>
- <target name="-post-compile-single">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,deps-jar,-pre-pre-compile,-pre-compile-single,-do-compile-single,-post-compile-single" name="compile-single"/>
- <!--
- ====================
- JAR BUILDING SECTION
- ====================
- -->
- <target depends="init" name="-pre-pre-jar">
- <dirname file="${dist.jar}" property="dist.jar.dir"/>
- <mkdir dir="${dist.jar.dir}"/>
- </target>
- <target name="-pre-jar">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,compile,-pre-pre-jar,-pre-jar" name="-do-jar-without-manifest" unless="manifest.available">
- <scalaProject1:jar/>
- </target>
- <target depends="init,compile,-pre-pre-jar,-pre-jar" if="manifest.available" name="-do-jar-with-manifest" unless="manifest.available+main.class">
- <scalaProject1:jar manifest="${manifest.file}"/>
- </target>
- <target depends="init,compile,-pre-pre-jar,-pre-jar" if="manifest.available+main.class" name="-do-jar-with-mainclass" unless="manifest.available+main.class+mkdist.available">
- <scalaProject1:jar manifest="${manifest.file}">
- <scalaProject1:manifest>
- <scalaProject1:attribute name="Main-Class" value="${main.class}"/>
- </scalaProject1:manifest>
- </scalaProject1:jar>
- <echo>To run this application from the command line without Ant, try:</echo>
- <property location="${build.classes.dir}" name="build.classes.dir.resolved"/>
- <property location="${dist.jar}" name="dist.jar.resolved"/>
- <pathconvert property="run.classpath.with.dist.jar">
- <path path="${run.classpath}"/>
- <map from="${build.classes.dir.resolved}" to="${dist.jar.resolved}"/>
- </pathconvert>
- <echo>java -cp "${run.classpath.with.dist.jar}" ${main.class}
- </echo>
- </target>
- <target depends="init,compile,-pre-pre-jar,-pre-jar" if="manifest.available+main.class+mkdist.available" name="-do-jar-with-libraries">
- <property location="${build.classes.dir}" name="build.classes.dir.resolved"/>
- <pathconvert property="run.classpath.without.build.classes.dir">
- <path path="${run.classpath}"/>
- <map from="${build.classes.dir.resolved}" to=""/>
- </pathconvert>
- <pathconvert pathsep=" " property="jar.classpath">
- <path path="${run.classpath.without.build.classes.dir}"/>
- <chainedmapper>
- <flattenmapper/>
- <globmapper from="*" to="lib/*"/>
- </chainedmapper>
- </pathconvert>
- <taskdef classname="org.netbeans.modules.java.j2seproject.copylibstask.CopyLibs" classpath="${libs.CopyLibs.classpath}" name="copylibs"/>
- <copylibs compress="${jar.compress}" jarfile="${dist.jar}" manifest="${manifest.file}" runtimeclasspath="${run.classpath.without.build.classes.dir}">
- <fileset dir="${build.classes.dir}"/>
- <manifest>
- <attribute name="Main-Class" value="${main.class}"/>
- <attribute name="Class-Path" value="${jar.classpath}"/>
- </manifest>
- </copylibs>
- <echo>To run this application from the command line without Ant, try:</echo>
- <property location="${dist.jar}" name="dist.jar.resolved"/>
- <echo>java -jar "${dist.jar.resolved}"
- </echo>
- </target>
- <target name="-post-jar">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,compile,-pre-jar,-do-jar-with-manifest,-do-jar-without-manifest,-do-jar-with-mainclass,-do-jar-with-libraries,-post-jar" description="Build JAR." name="jar"/>
- <!--
- =================
- EXECUTION SECTION
- =================
- -->
- <target depends="init,compile" description="Run a main class." name="run">
- <scalaProject1:java>
- <customize>
- <arg line="${application.args}"/>
- </customize>
- </scalaProject1:java>
- </target>
- <target name="-do-not-recompile">
- <property name="javac.includes.binary" value=""/>
- </target>
- <target depends="init,-do-not-recompile,compile-single" name="run-single">
- <fail unless="run.class">Must select one file in the IDE or set run.class</fail>
- <scalaProject1:java classname="${run.class}"/>
- </target>
- <!--
- =================
- DEBUGGING SECTION
- =================
- -->
- <target depends="init" if="netbeans.home" name="-debug-start-debugger">
- <scalaProject1:nbjpdastart name="${debug.class}"/>
- </target>
- <target depends="init,compile" name="-debug-start-debuggee">
- <scalaProject1:debug>
- <customize>
- <arg line="${application.args}"/>
- </customize>
- </scalaProject1:debug>
- </target>
- <target depends="init,compile,-debug-start-debugger,-debug-start-debuggee" description="Debug project in IDE." if="netbeans.home" name="debug"/>
- <target depends="init" if="netbeans.home" name="-debug-start-debugger-stepinto">
- <scalaProject1:nbjpdastart stopclassname="${main.class}"/>
- </target>
- <target depends="init,compile,-debug-start-debugger-stepinto,-debug-start-debuggee" if="netbeans.home" name="debug-stepinto"/>
- <target depends="init,compile-single" if="netbeans.home" name="-debug-start-debuggee-single">
- <fail unless="debug.class">Must select one file in the IDE or set debug.class</fail>
- <scalaProject1:debug classname="${debug.class}"/>
- </target>
- <target depends="init,-do-not-recompile,compile-single,-debug-start-debugger,-debug-start-debuggee-single" if="netbeans.home" name="debug-single"/>
- <target depends="init" name="-pre-debug-fix">
- <fail unless="fix.includes">Must set fix.includes</fail>
- <property name="javac.includes" value="${fix.includes}.java"/>
- </target>
- <target depends="init,-pre-debug-fix,compile-single" if="netbeans.home" name="-do-debug-fix">
- <scalaProject1:nbjpdareload/>
- </target>
- <target depends="init,-pre-debug-fix,-do-debug-fix" if="netbeans.home" name="debug-fix"/>
- <!--
- ===============
- JAVADOC SECTION
- ===============
- -->
- <target depends="init" name="-javadoc-build">
- <mkdir dir="${dist.javadoc.dir}"/>
- <scaladoc addparams="${javadoc.additionalparam}" deprecation="yes" destdir="${dist.javadoc.dir}" doctitle="${javadoc.windowtitle}" encoding="${javadoc.encoding.used}" srcdir="${src.dir}" unchecked="yes">
- <classpath>
- <path path="${javac.classpath}"/>
- <fileset dir="${scala.lib}">
- <include name="**/*.jar"/>
- </fileset>
- </classpath>
- </scaladoc>
- </target>
- <target depends="init,-javadoc-build" if="netbeans.home" name="-javadoc-browse" unless="no.javadoc.preview">
- <nbbrowse file="${dist.javadoc.dir}/index.html"/>
- </target>
- <target depends="init,-javadoc-build,-javadoc-browse" description="Build Javadoc." name="javadoc"/>
- <!--
- =========================
- JUNIT COMPILATION SECTION
- =========================
- -->
- <target depends="init,compile" if="have.tests" name="-pre-pre-compile-test">
- <mkdir dir="${build.test.classes.dir}"/>
- </target>
- <target name="-pre-compile-test">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target if="do.depend.true" name="-compile-test-depend">
- <scalaProject1:depend classpath="${javac.test.classpath}" destdir="${build.test.classes.dir}" srcdir=""/>
- </target>
- <target depends="init,compile,-pre-pre-compile-test,-pre-compile-test,-compile-test-depend" if="have.tests" name="-do-compile-test">
- <scalaProject1:scalac classpath="${javac.test.classpath}" destdir="${build.test.classes.dir}" srcdir=""/>
- <copy todir="${build.test.classes.dir}"/>
- </target>
- <target name="-post-compile-test">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,compile,-pre-pre-compile-test,-pre-compile-test,-do-compile-test,-post-compile-test" name="compile-test"/>
- <target name="-pre-compile-test-single">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,compile,-pre-pre-compile-test,-pre-compile-test-single" if="have.tests" name="-do-compile-test-single">
- <fail unless="javac.includes">Must select some files in the IDE or set javac.includes</fail>
- <scalaProject1:force-recompile destdir="${build.test.classes.dir}"/>
- <scalaProject1:scalac classpath="${javac.test.classpath}" destdir="${build.test.classes.dir}" excludes="" includes="${javac.includes}" sourcepath="" srcdir=""/>
- <copy todir="${build.test.classes.dir}"/>
- </target>
- <target name="-post-compile-test-single">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,compile,-pre-pre-compile-test,-pre-compile-test-single,-do-compile-test-single,-post-compile-test-single" name="compile-test-single"/>
- <!--
- =======================
- JUNIT EXECUTION SECTION
- =======================
- -->
- <target depends="init" if="have.tests" name="-pre-test-run">
- <mkdir dir="${build.test.results.dir}"/>
- </target>
- <target depends="init,compile-test,-pre-test-run" if="have.tests" name="-do-test-run">
- <scalaProject1:junit testincludes="**/*Test.class"/>
- </target>
- <target depends="init,compile-test,-pre-test-run,-do-test-run" if="have.tests" name="-post-test-run">
- <fail if="tests.failed">Some tests failed; see details above.</fail>
- </target>
- <target depends="init" if="have.tests" name="test-report"/>
- <target depends="init" if="netbeans.home+have.tests" name="-test-browse"/>
- <target depends="init,compile-test,-pre-test-run,-do-test-run,test-report,-post-test-run,-test-browse" description="Run unit tests." name="test"/>
- <target depends="init" if="have.tests" name="-pre-test-run-single">
- <mkdir dir="${build.test.results.dir}"/>
- </target>
- <target depends="init,compile-test-single,-pre-test-run-single" if="have.tests" name="-do-test-run-single">
- <fail unless="test.includes">Must select some files in the IDE or set test.includes</fail>
- <scalaProject1:junit excludes="" includes="${test.includes}"/>
- </target>
- <target depends="init,compile-test-single,-pre-test-run-single,-do-test-run-single" if="have.tests" name="-post-test-run-single">
- <fail if="tests.failed">Some tests failed; see details above.</fail>
- </target>
- <target depends="init,-do-not-recompile,compile-test-single,-pre-test-run-single,-do-test-run-single,-post-test-run-single" description="Run single unit test." name="test-single"/>
- <!--
- =======================
- JUNIT DEBUGGING SECTION
- =======================
- -->
- <target depends="init,compile-test" if="have.tests" name="-debug-start-debuggee-test">
- <fail unless="test.class">Must select one file in the IDE or set test.class</fail>
- <property location="${build.test.results.dir}/TEST-${test.class}.xml" name="test.report.file"/>
- <delete file="${test.report.file}"/>
- <mkdir dir="${build.test.results.dir}"/>
- <scalaProject1:debug classname="org.apache.tools.ant.taskdefs.optional.junit.JUnitTestRunner" classpath="${ant.home}/lib/ant.jar:${ant.home}/lib/ant-junit.jar:${debug.test.classpath}">
- <customize>
- <syspropertyset>
- <propertyref prefix="test-sys-prop."/>
- <mapper from="test-sys-prop.*" to="*" type="glob"/>
- </syspropertyset>
- <arg value="${test.class}"/>
- <arg value="showoutput=true"/>
- <arg value="formatter=org.apache.tools.ant.taskdefs.optional.junit.BriefJUnitResultFormatter"/>
- <arg value="formatter=org.apache.tools.ant.taskdefs.optional.junit.XMLJUnitResultFormatter,${test.report.file}"/>
- </customize>
- </scalaProject1:debug>
- </target>
- <target depends="init,compile-test" if="netbeans.home+have.tests" name="-debug-start-debugger-test">
- <scalaProject1:nbjpdastart classpath="${debug.test.classpath}" name="${test.class}"/>
- </target>
- <target depends="init,-do-not-recompile,compile-test-single,-debug-start-debugger-test,-debug-start-debuggee-test" name="debug-test"/>
- <target depends="init,-pre-debug-fix,compile-test-single" if="netbeans.home" name="-do-debug-fix-test">
- <scalaProject1:nbjpdareload dir="${build.test.classes.dir}"/>
- </target>
- <target depends="init,-pre-debug-fix,-do-debug-fix-test" if="netbeans.home" name="debug-fix-test"/>
- <!--
- =========================
- APPLET EXECUTION SECTION
- =========================
- -->
- <target depends="init,compile-single" name="run-applet">
- <fail unless="applet.url">Must select one file in the IDE or set applet.url</fail>
- <scalaProject1:java classname="sun.applet.AppletViewer">
- <customize>
- <arg value="${applet.url}"/>
- </customize>
- </scalaProject1:java>
- </target>
- <!--
- =========================
- APPLET DEBUGGING SECTION
- =========================
- -->
- <target depends="init,compile-single" if="netbeans.home" name="-debug-start-debuggee-applet">
- <fail unless="applet.url">Must select one file in the IDE or set applet.url</fail>
- <scalaProject1:debug classname="sun.applet.AppletViewer">
- <customize>
- <arg value="${applet.url}"/>
- </customize>
- </scalaProject1:debug>
- </target>
- <target depends="init,compile-single,-debug-start-debugger,-debug-start-debuggee-applet" if="netbeans.home" name="debug-applet"/>
- <!--
- ===============
- CLEANUP SECTION
- ===============
- -->
- <target depends="init" name="deps-clean" unless="no.deps">
- <ant antfile="${project.jEdit}/build-nb.xml" inheritall="false" target="clean"/>
- </target>
- <target depends="init" name="-do-clean">
- <delete dir="${build.dir}"/>
- <delete dir="${dist.dir}"/>
- </target>
- <target name="-post-clean">
- <!-- Empty placeholder for easier customization. -->
- <!-- You can override this target in the ../build.xml file. -->
- </target>
- <target depends="init,deps-clean,-do-clean,-post-clean" description="Clean build products." name="clean"/>
-</project>
--- a/src/Tools/jEdit/nbproject/genfiles.properties Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-build.xml.data.CRC32=d2379ac2
-build.xml.script.CRC32=2db9d955
-build.xml.stylesheet.CRC32=ca9d572e
-# This file is used by a NetBeans-based IDE to track changes in generated files such as build-impl.xml.
-# Do not edit this file. You may delete it but then the IDE will never regenerate such files for you.
-nbproject/build-impl.xml.data.CRC32=8f41dcce
-nbproject/build-impl.xml.script.CRC32=e3e2a5d5
-nbproject/build-impl.xml.stylesheet.CRC32=5220179f@1.3.5
--- a/src/Tools/jEdit/nbproject/project.properties Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,81 +0,0 @@
-application.title=Isabelle-jEdit
-application.vendor=makarius
-application.args=-noserver -nobackground
-build.classes.dir=${build.dir}/classes
-build.classes.excludes=**/*.java,**/*.form,**/*.scala
-# This directory is removed when the project is cleaned:
-build.dir=build
-build.generated.dir=${build.dir}/generated
-# Only compile against the classpath explicitly listed here:
-build.sysclasspath=ignore
-build.test.classes.dir=${build.dir}/test/classes
-build.test.results.dir=${build.dir}/test/results
-debug.classpath=\
- ${run.classpath}
-debug.test.classpath=\
- ${run.test.classpath}
-# This directory is removed when the project is cleaned:
-dist.dir=dist
-# dist can be used as jEdits settings-directory;
-# jEdit searches for plugins in the 'jars' subdirectory
-# must include something like this to private.properties:
-# application.args=-noserver -nobackground -settings=/absolute/path/to/project/dist
-#
-dist.jar=${dist.dir}/jars/Isabelle-jEdit.jar
-dist.javadoc.dir=${dist.dir}/javadoc
-excludes=
-file.reference.isabelle-jedit-src=src
-file.reference.jedit.jar=/home/makarius/lib/jedit/current/jedit.jar
-includes=**
-jar.compress=false
-java.platform.active=java_default_platform
-javac.classpath=\
- ${reference.jEdit.build}:\
- ${libs.Isabelle-Pure.classpath}:\
- ${libs.Cobra-Renderer.classpath}:\
- ${libs.Rhino-JavaScript.classpath}:\
- ${libs.ErrorList.classpath}:\
- ${libs.Hyperlinks.classpath}:\
- ${libs.SideKick.classpath}:\
- ${libs.Console.classpath}:\
- ${libs.Scala-compiler.classpath}
-# Space-separated list of extra javac options
-javac.compilerargs=
-javac.deprecation=false
-javac.source=1.5
-javac.target=1.5
-javac.test.classpath=\
- ${javac.classpath}:\
- ${build.classes.dir}:\
- ${libs.junit.classpath}:\
- ${libs.junit_4.classpath}
-javadoc.additionalparam=
-javadoc.author=false
-javadoc.encoding=${source.encoding}
-javadoc.noindex=false
-javadoc.nonavbar=false
-javadoc.notree=false
-javadoc.private=false
-javadoc.splitindex=true
-javadoc.use=true
-javadoc.version=false
-javadoc.windowtitle=
-main.class=org.gjt.sp.jedit.jEdit
-manifest.file=manifest.mf
-meta.inf.dir=${src.dir}/META-INF
-platform.active=default_platform
-project.jEdit=contrib/jEdit
-reference.jEdit.build=${project.jEdit}/build/jEdit.jar
-run.classpath=\
- ${javac.classpath}:\
- ${build.classes.dir}
-run.jvmargs=-Xms128m -Xmx512m
-run.test.classpath=\
- ${javac.test.classpath}:\
- ${build.test.classes.dir}
-source.encoding=UTF-8
-src.dir=${file.reference.isabelle-jedit-src}
-scalac.compilerargs=
-scalac.deprecation=yes
-scalac.unchecked=yes
-
--- a/src/Tools/jEdit/nbproject/project.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-<?xml version="1.0" encoding="UTF-8"?>
-<project xmlns="http://www.netbeans.org/ns/project/1">
- <type>org.netbeans.modules.scala.project</type>
- <configuration>
- <data xmlns="http://www.netbeans.org/ns/scala-project/1">
- <name>Isabelle-jEdit</name>
- <minimum-ant-version>1.6.5</minimum-ant-version>
- <source-roots>
- <root id="src.dir"/>
- </source-roots>
- <test-roots/>
- </data>
- <references xmlns="http://www.netbeans.org/ns/ant-project-references/1">
- <reference>
- <foreign-project>jEdit</foreign-project>
- <artifact-type>jar</artifact-type>
- <script>build-nb.xml</script>
- <target>build</target>
- <clean-target>clean</clean-target>
- <id>build</id>
- </reference>
- </references>
- </configuration>
-</project>
--- a/src/Tools/jEdit/plugin/Isabelle.props Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-## Isabelle plugin properties
-##
-##:encoding=ISO-8859-1:wrap=soft:maxLineLen=100:
-
-#identification
-plugin.isabelle.jedit.Plugin.name=Isabelle
-plugin.isabelle.jedit.Plugin.author=Johannes Hölzl, Fabian Immler, Makarius Wenzel
-plugin.isabelle.jedit.Plugin.version=0.1.0
-plugin.isabelle.jedit.Plugin.description=Isabelle/Isar asynchronous proof document editing
-
-#system parameters
-plugin.isabelle.jedit.Plugin.activate=startup
-plugin.isabelle.jedit.Plugin.usePluginHome=false
-
-#dependencies
-plugin.isabelle.jedit.Plugin.depend.0=jdk 1.6
-plugin.isabelle.jedit.Plugin.depend.1=jedit 04.03.99.00
-plugin.isabelle.jedit.Plugin.depend.2=plugin console.ConsolePlugin 4.4.1
-plugin.isabelle.jedit.Plugin.depend.3=plugin errorlist.ErrorListPlugin 1.8
-plugin.isabelle.jedit.Plugin.depend.4=plugin sidekick.SideKickPlugin 0.8
-plugin.isabelle.jedit.Plugin.depend.5=plugin gatchan.jedit.hyperlinks.HyperlinksPlugin 1.0.1
-
-#options
-plugin.isabelle.jedit.Plugin.option-pane=isabelle
-options.isabelle.label=Isabelle
-options.isabelle.code=new isabelle.jedit.Isabelle_Options();
-options.isabelle.logic.title=Logic
-options.isabelle.relative-font-size.title=Relative Font Size
-options.isabelle.relative-font-size=100
-options.isabelle.tooltip-font-size.title=Tooltip Font Size
-options.isabelle.tooltip-font-size=10
-options.isabelle.tooltip-margin.title=Tooltip Margin
-options.isabelle.tooltip-margin=40
-options.isabelle.tooltip-dismiss-delay.title=Tooltip Dismiss Delay (global)
-options.isabelle.tooltip-dismiss-delay=8.0
-options.isabelle.startup-timeout=25.0
-options.isabelle.auto-start.title=Auto Start
-options.isabelle.auto-start=true
-
-#menu actions
-plugin.isabelle.jedit.Plugin.menu.label=Isabelle
-plugin.isabelle.jedit.Plugin.menu=isabelle.session-panel isabelle.output-panel isabelle.raw-output-panel isabelle.protocol-panel
-isabelle.session-panel.label=Prover Session panel
-isabelle.output-panel.label=Output panel
-isabelle.raw-output-panel.label=Raw Output panel
-isabelle.protocol-panel.label=Protocol panel
-
-#dockables
-isabelle-session.title=Prover Session
-isabelle-output.title=Output
-isabelle-raw-output.title=Raw Output
-isabelle-protocol.title=Protocol
-
-#SideKick
-sidekick.parser.isabelle.label=Isabelle
-mode.isabelle.sidekick.parser=isabelle
-mode.ml.sidekick.parser=isabelle
-
-#Hyperlinks
-mode.isabelle.hyperlink.source=isabelle
--- a/src/Tools/jEdit/plugin/actions.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-<?xml version="1.0"?>
-<!DOCTYPE ACTIONS SYSTEM "actions.dtd">
-
-<ACTIONS>
- <ACTION NAME="isabelle.session-panel">
- <CODE>
- wm.addDockableWindow("isabelle-session");
- </CODE>
- </ACTION>
- <ACTION NAME="isabelle.output-panel">
- <CODE>
- wm.addDockableWindow("isabelle-output");
- </CODE>
- </ACTION>
- <ACTION NAME="isabelle.raw-output-panel">
- <CODE>
- wm.addDockableWindow("isabelle-raw-output");
- </CODE>
- </ACTION>
- <ACTION NAME="isabelle.protocol-panel">
- <CODE>
- wm.addDockableWindow("isabelle-protocol");
- </CODE>
- </ACTION>
-</ACTIONS>
\ No newline at end of file
--- a/src/Tools/jEdit/plugin/dockables.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-<?xml version="1.0"?>
-<!DOCTYPE DOCKABLES SYSTEM "dockables.dtd">
-
-<DOCKABLES>
- <DOCKABLE NAME="isabelle-session" MOVABLE="TRUE">
- new isabelle.jedit.Session_Dockable(view, position);
- </DOCKABLE>
- <DOCKABLE NAME="isabelle-output" MOVABLE="TRUE">
- new isabelle.jedit.Output_Dockable(view, position);
- </DOCKABLE>
- <DOCKABLE NAME="isabelle-raw-output" MOVABLE="TRUE">
- new isabelle.jedit.Raw_Output_Dockable(view, position);
- </DOCKABLE>
- <DOCKABLE NAME="isabelle-protocol" MOVABLE="TRUE">
- new isabelle.jedit.Protocol_Dockable(view, position);
- </DOCKABLE>
-</DOCKABLES>
\ No newline at end of file
--- a/src/Tools/jEdit/plugin/services.xml Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,20 +0,0 @@
-<?xml version="1.0"?>
-<!DOCTYPE SERVICES SYSTEM "services.dtd">
-
-<SERVICES>
- <SERVICE NAME="UTF-8-Isabelle" CLASS="org.gjt.sp.jedit.io.Encoding">
- new isabelle.jedit.Isabelle_Encoding();
- </SERVICE>
- <SERVICE NAME="isabelle" CLASS="sidekick.SideKickParser">
- new isabelle.jedit.Isabelle_Sidekick_Default();
- </SERVICE>
- <SERVICE NAME="isabelle-raw" CLASS="sidekick.SideKickParser">
- new isabelle.jedit.Isabelle_Sidekick_Raw();
- </SERVICE>
- <SERVICE NAME="isabelle" CLASS="gatchan.jedit.hyperlinks.HyperlinkSource">
- new isabelle.jedit.Isabelle_Hyperlinks();
- </SERVICE>
- <SERVICE CLASS="console.Shell" NAME="Scala">
- new isabelle.jedit.Scala_Console();
- </SERVICE>
-</SERVICES>
--- a/src/Tools/jEdit/src/Dummy.java Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-/* dummy class to make ant javadoc work */
-package isabelle;
-public class Dummy { }
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/Isabelle.props Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,60 @@
+## Isabelle plugin properties
+##
+##:encoding=ISO-8859-1:wrap=soft:maxLineLen=100:
+
+#identification
+plugin.isabelle.jedit.Plugin.name=Isabelle
+plugin.isabelle.jedit.Plugin.author=Johannes Hölzl, Fabian Immler, Makarius Wenzel
+plugin.isabelle.jedit.Plugin.version=0.1.0
+plugin.isabelle.jedit.Plugin.description=Isabelle/Isar asynchronous proof document editing
+
+#system parameters
+plugin.isabelle.jedit.Plugin.activate=startup
+plugin.isabelle.jedit.Plugin.usePluginHome=false
+
+#dependencies
+plugin.isabelle.jedit.Plugin.depend.0=jdk 1.6
+plugin.isabelle.jedit.Plugin.depend.1=jedit 04.03.99.00
+plugin.isabelle.jedit.Plugin.depend.2=plugin console.ConsolePlugin 4.4.1
+plugin.isabelle.jedit.Plugin.depend.3=plugin errorlist.ErrorListPlugin 1.8
+plugin.isabelle.jedit.Plugin.depend.4=plugin sidekick.SideKickPlugin 0.8
+plugin.isabelle.jedit.Plugin.depend.5=plugin gatchan.jedit.hyperlinks.HyperlinksPlugin 1.0.1
+
+#options
+plugin.isabelle.jedit.Plugin.option-pane=isabelle
+options.isabelle.label=Isabelle
+options.isabelle.code=new isabelle.jedit.Isabelle_Options();
+options.isabelle.logic.title=Logic
+options.isabelle.relative-font-size.title=Relative Font Size
+options.isabelle.relative-font-size=100
+options.isabelle.tooltip-font-size.title=Tooltip Font Size
+options.isabelle.tooltip-font-size=10
+options.isabelle.tooltip-margin.title=Tooltip Margin
+options.isabelle.tooltip-margin=40
+options.isabelle.tooltip-dismiss-delay.title=Tooltip Dismiss Delay (global)
+options.isabelle.tooltip-dismiss-delay=8.0
+options.isabelle.startup-timeout=25.0
+options.isabelle.auto-start.title=Auto Start
+options.isabelle.auto-start=true
+
+#menu actions
+plugin.isabelle.jedit.Plugin.menu.label=Isabelle
+plugin.isabelle.jedit.Plugin.menu=isabelle.session-panel isabelle.output-panel isabelle.raw-output-panel isabelle.protocol-panel
+isabelle.session-panel.label=Prover Session panel
+isabelle.output-panel.label=Output panel
+isabelle.raw-output-panel.label=Raw Output panel
+isabelle.protocol-panel.label=Protocol panel
+
+#dockables
+isabelle-session.title=Prover Session
+isabelle-output.title=Output
+isabelle-raw-output.title=Raw Output
+isabelle-protocol.title=Protocol
+
+#SideKick
+sidekick.parser.isabelle.label=Isabelle
+mode.isabelle.sidekick.parser=isabelle
+mode.ml.sidekick.parser=isabelle
+
+#Hyperlinks
+mode.isabelle.hyperlink.source=isabelle
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/actions.xml Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,25 @@
+<?xml version="1.0"?>
+<!DOCTYPE ACTIONS SYSTEM "actions.dtd">
+
+<ACTIONS>
+ <ACTION NAME="isabelle.session-panel">
+ <CODE>
+ wm.addDockableWindow("isabelle-session");
+ </CODE>
+ </ACTION>
+ <ACTION NAME="isabelle.output-panel">
+ <CODE>
+ wm.addDockableWindow("isabelle-output");
+ </CODE>
+ </ACTION>
+ <ACTION NAME="isabelle.raw-output-panel">
+ <CODE>
+ wm.addDockableWindow("isabelle-raw-output");
+ </CODE>
+ </ACTION>
+ <ACTION NAME="isabelle.protocol-panel">
+ <CODE>
+ wm.addDockableWindow("isabelle-protocol");
+ </CODE>
+ </ACTION>
+</ACTIONS>
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/dockable.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,41 @@
+/* Title: Tools/jEdit/src/dockable.scala
+ Author: Makarius
+
+Generic dockable window.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.awt.{Dimension, Component, BorderLayout}
+import javax.swing.JPanel
+
+import org.gjt.sp.jedit.View
+import org.gjt.sp.jedit.gui.DockableWindowManager
+
+
+class Dockable(view: View, position: String) extends JPanel(new BorderLayout)
+{
+ if (position == DockableWindowManager.FLOATING)
+ setPreferredSize(new Dimension(500, 250))
+
+ def set_content(c: Component) { add(c, BorderLayout.CENTER) }
+ def set_content(c: scala.swing.Component) { add(c.peer, BorderLayout.CENTER) }
+
+ protected def init() { }
+ protected def exit() { }
+
+ override def addNotify()
+ {
+ super.addNotify()
+ init()
+ }
+
+ override def removeNotify()
+ {
+ exit()
+ super.removeNotify()
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/dockables.xml Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,17 @@
+<?xml version="1.0"?>
+<!DOCTYPE DOCKABLES SYSTEM "dockables.dtd">
+
+<DOCKABLES>
+ <DOCKABLE NAME="isabelle-session" MOVABLE="TRUE">
+ new isabelle.jedit.Session_Dockable(view, position);
+ </DOCKABLE>
+ <DOCKABLE NAME="isabelle-output" MOVABLE="TRUE">
+ new isabelle.jedit.Output_Dockable(view, position);
+ </DOCKABLE>
+ <DOCKABLE NAME="isabelle-raw-output" MOVABLE="TRUE">
+ new isabelle.jedit.Raw_Output_Dockable(view, position);
+ </DOCKABLE>
+ <DOCKABLE NAME="isabelle-protocol" MOVABLE="TRUE">
+ new isabelle.jedit.Protocol_Dockable(view, position);
+ </DOCKABLE>
+</DOCKABLES>
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/document_model.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,252 @@
+/* Title: Tools/jEdit/src/document_model.scala
+ Author: Fabian Immler, TU Munich
+ Author: Makarius
+
+Document model connected to jEdit buffer -- single node in theory graph.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.collection.mutable
+
+import org.gjt.sp.jedit.Buffer
+import org.gjt.sp.jedit.buffer.{BufferAdapter, BufferListener, JEditBuffer}
+import org.gjt.sp.jedit.syntax.{SyntaxStyle, Token, TokenMarker, TokenHandler, ParserRuleSet}
+import org.gjt.sp.jedit.textarea.TextArea
+
+import java.awt.font.TextAttribute
+import javax.swing.text.Segment;
+
+
+object Document_Model
+{
+ object Token_Markup
+ {
+ /* extended token styles */
+
+ private val plain_range: Int = Token.ID_COUNT
+ private val full_range: Int = 3 * plain_range
+ private def check_range(i: Int) { require(0 <= i && i < plain_range) }
+
+ def subscript(i: Byte): Byte = { check_range(i); (i + plain_range).toByte }
+ def superscript(i: Byte): Byte = { check_range(i); (i + 2 * plain_range).toByte }
+
+ private def script_style(style: SyntaxStyle, i: Int): SyntaxStyle =
+ {
+ import scala.collection.JavaConversions._
+ val script_font =
+ style.getFont.deriveFont(Map(TextAttribute.SUPERSCRIPT -> new java.lang.Integer(i)))
+ new SyntaxStyle(style.getForegroundColor, style.getBackgroundColor, script_font)
+ }
+
+ def extend_styles(styles: Array[SyntaxStyle]): Array[SyntaxStyle] =
+ {
+ val new_styles = new Array[SyntaxStyle](full_range)
+ for (i <- 0 until plain_range) {
+ val style = styles(i)
+ new_styles(i) = style
+ new_styles(subscript(i.toByte)) = script_style(style, -1)
+ new_styles(superscript(i.toByte)) = script_style(style, 1)
+ }
+ new_styles
+ }
+
+
+ /* line context */
+
+ private val dummy_rules = new ParserRuleSet("isabelle", "MAIN")
+
+ class Line_Context(val line: Int, prev: Line_Context)
+ extends TokenMarker.LineContext(dummy_rules, prev)
+ {
+ override def hashCode: Int = line
+ override def equals(that: Any): Boolean =
+ that match {
+ case other: Line_Context => line == other.line
+ case _ => false
+ }
+ }
+ }
+
+
+ /* document model of buffer */
+
+ private val key = "isabelle.document_model"
+
+ def init(session: Session, buffer: Buffer, thy_name: String): Document_Model =
+ {
+ Swing_Thread.require()
+ val model = new Document_Model(session, buffer, thy_name)
+ buffer.setProperty(key, model)
+ model.activate()
+ model
+ }
+
+ def apply(buffer: Buffer): Option[Document_Model] =
+ {
+ Swing_Thread.require()
+ buffer.getProperty(key) match {
+ case model: Document_Model => Some(model)
+ case _ => None
+ }
+ }
+
+ def exit(buffer: Buffer)
+ {
+ Swing_Thread.require()
+ apply(buffer) match {
+ case None =>
+ case Some(model) =>
+ model.deactivate()
+ buffer.unsetProperty(key)
+ }
+ }
+}
+
+
+class Document_Model(val session: Session, val buffer: Buffer, val thy_name: String)
+{
+ /* pending text edits */
+
+ object pending_edits // owned by Swing thread
+ {
+ private val pending = new mutable.ListBuffer[Text.Edit]
+ def snapshot(): List[Text.Edit] = pending.toList
+
+ def flush(more_edits: Option[List[Text.Edit]]*)
+ {
+ Swing_Thread.require()
+ val edits = snapshot()
+ pending.clear
+
+ val all_edits =
+ if (edits.isEmpty) more_edits.toList
+ else Some(edits) :: more_edits.toList
+ if (!all_edits.isEmpty) session.edit_version(all_edits.map((thy_name, _)))
+ }
+
+ def init()
+ {
+ Swing_Thread.require()
+ flush(None, Some(List(Text.Edit.insert(0, Isabelle.buffer_text(buffer)))))
+ }
+
+ private val delay_flush =
+ Swing_Thread.delay_last(session.input_delay) { flush() }
+
+ def +=(edit: Text.Edit)
+ {
+ Swing_Thread.require()
+ pending += edit
+ delay_flush()
+ }
+ }
+
+
+ /* snapshot */
+
+ def snapshot(): Document.Snapshot =
+ {
+ Swing_Thread.require()
+ session.snapshot(thy_name, pending_edits.snapshot())
+ }
+
+
+ /* buffer listener */
+
+ private val buffer_listener: BufferListener = new BufferAdapter
+ {
+ override def bufferLoaded(buffer: JEditBuffer)
+ {
+ pending_edits.init()
+ }
+
+ override def contentInserted(buffer: JEditBuffer,
+ start_line: Int, offset: Int, num_lines: Int, length: Int)
+ {
+ if (!buffer.isLoading)
+ pending_edits += Text.Edit.insert(offset, buffer.getText(offset, length))
+ }
+
+ override def preContentRemoved(buffer: JEditBuffer,
+ start_line: Int, offset: Int, num_lines: Int, removed_length: Int)
+ {
+ if (!buffer.isLoading)
+ pending_edits += Text.Edit.remove(offset, buffer.getText(offset, removed_length))
+ }
+ }
+
+
+ /* semantic token marker */
+
+ private val token_marker = new TokenMarker
+ {
+ override def markTokens(prev: TokenMarker.LineContext,
+ handler: TokenHandler, line_segment: Segment): TokenMarker.LineContext =
+ {
+ Isabelle.swing_buffer_lock(buffer) {
+ val snapshot = Document_Model.this.snapshot()
+
+ val previous = prev.asInstanceOf[Document_Model.Token_Markup.Line_Context]
+ val line = if (prev == null) 0 else previous.line + 1
+ val context = new Document_Model.Token_Markup.Line_Context(line, previous)
+
+ val start = buffer.getLineStartOffset(line)
+ val stop = start + line_segment.count
+
+ /* FIXME
+ for (text_area <- Isabelle.jedit_text_areas(buffer)
+ if Document_View(text_area).isDefined)
+ Document_View(text_area).get.set_styles()
+ */
+
+ def handle_token(style: Byte, offset: Text.Offset, length: Int) =
+ handler.handleToken(line_segment, style, offset, length, context)
+
+ val syntax = session.current_syntax()
+ val tokens = snapshot.select_markup(Text.Range(start, stop))(Isabelle_Markup.tokens(syntax))
+
+ var last = start
+ for (token <- tokens.iterator) {
+ val Text.Range(token_start, token_stop) = token.range
+ if (last < token_start)
+ handle_token(Token.COMMENT1, last - start, token_start - last)
+ handle_token(token.info getOrElse Token.NULL,
+ token_start - start, token_stop - token_start)
+ last = token_stop
+ }
+ if (last < stop) handle_token(Token.COMMENT1, last - start, stop - last)
+
+ handle_token(Token.END, line_segment.count, 0)
+ handler.setLineContext(context)
+ context
+ }
+ }
+ }
+
+
+ /* activation */
+
+ def activate()
+ {
+ buffer.setTokenMarker(token_marker)
+ buffer.addBufferListener(buffer_listener)
+ buffer.propertiesChanged()
+ pending_edits.init()
+ }
+
+ def refresh()
+ {
+ buffer.setTokenMarker(token_marker)
+ }
+
+ def deactivate()
+ {
+ pending_edits.flush()
+ buffer.setTokenMarker(buffer.getMode.getTokenMarker)
+ buffer.removeBufferListener(buffer_listener)
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/document_view.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,587 @@
+/* Title: Tools/jEdit/src/document_view.scala
+ Author: Fabian Immler, TU Munich
+ Author: Makarius
+
+Document view connected to jEdit text area.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.actors.Actor._
+
+import java.awt.{BorderLayout, Graphics, Color, Dimension, Graphics2D, Point}
+import java.awt.event.{MouseAdapter, MouseMotionAdapter, MouseEvent,
+ FocusAdapter, FocusEvent, WindowEvent, WindowAdapter}
+import javax.swing.{JPanel, ToolTipManager, Popup, PopupFactory, SwingUtilities, BorderFactory}
+import javax.swing.event.{CaretListener, CaretEvent}
+import java.util.ArrayList
+
+import org.gjt.sp.jedit.{jEdit, OperatingSystem, Debug}
+import org.gjt.sp.jedit.gui.RolloverButton
+import org.gjt.sp.jedit.options.GutterOptionPane
+import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, TextAreaExtension, TextAreaPainter}
+import org.gjt.sp.jedit.syntax.{SyntaxStyle, DisplayTokenHandler, Chunk, Token}
+
+
+object Document_View
+{
+ /* document view of text area */
+
+ private val key = new Object
+
+ def init(model: Document_Model, text_area: JEditTextArea): Document_View =
+ {
+ Swing_Thread.require()
+ val doc_view = new Document_View(model, text_area)
+ text_area.putClientProperty(key, doc_view)
+ doc_view.activate()
+ doc_view
+ }
+
+ def apply(text_area: JEditTextArea): Option[Document_View] =
+ {
+ Swing_Thread.require()
+ text_area.getClientProperty(key) match {
+ case doc_view: Document_View => Some(doc_view)
+ case _ => None
+ }
+ }
+
+ def exit(text_area: JEditTextArea)
+ {
+ Swing_Thread.require()
+ apply(text_area) match {
+ case None =>
+ case Some(doc_view) =>
+ doc_view.deactivate()
+ text_area.putClientProperty(key, null)
+ }
+ }
+}
+
+
+class Document_View(val model: Document_Model, text_area: JEditTextArea)
+{
+ private val session = model.session
+
+
+ /** token handling **/
+
+ /* extended token styles */
+
+ private var styles: Array[SyntaxStyle] = null // owned by Swing thread
+
+ def extend_styles()
+ {
+ Swing_Thread.require()
+ styles = Document_Model.Token_Markup.extend_styles(text_area.getPainter.getStyles)
+ }
+ extend_styles()
+
+ def set_styles()
+ {
+ Swing_Thread.require()
+ text_area.getPainter.setStyles(styles)
+ }
+
+
+ /* wrap_margin -- cf. org.gjt.sp.jedit.textarea.TextArea.propertiesChanged */
+
+ def wrap_margin(): Int =
+ {
+ val buffer = text_area.getBuffer
+ val painter = text_area.getPainter
+ val font = painter.getFont
+ val font_context = painter.getFontRenderContext
+
+ val soft_wrap = buffer.getStringProperty("wrap") == "soft"
+ val max = buffer.getIntegerProperty("maxLineLen", 0)
+
+ if (max > 0) font.getStringBounds(" " * max, font_context).getWidth.toInt
+ else if (soft_wrap)
+ painter.getWidth - (font.getStringBounds(" ", font_context).getWidth.round.toInt) * 3
+ else 0
+ }
+
+
+ /* chunks */
+
+ def line_chunks(physical_lines: Set[Int]): Map[Text.Offset, Chunk] =
+ {
+ import scala.collection.JavaConversions._
+
+ val buffer = text_area.getBuffer
+ val painter = text_area.getPainter
+ val margin = wrap_margin().toFloat
+
+ val out = new ArrayList[Chunk]
+ val handler = new DisplayTokenHandler
+
+ var result = Map[Text.Offset, Chunk]()
+ for (line <- physical_lines) {
+ out.clear
+ handler.init(painter.getStyles, painter.getFontRenderContext, painter, out, margin)
+ buffer.markTokens(line, handler)
+
+ val line_start = buffer.getLineStartOffset(line)
+ for (chunk <- handler.getChunkList.iterator)
+ result += (line_start + chunk.offset -> chunk)
+ }
+ result
+ }
+
+
+ /* visible line ranges */
+
+ // simplify slightly odd result of TextArea.getScreenLineEndOffset etc.
+ // NB: jEdit already normalizes \r\n and \r to \n
+ def proper_line_range(start: Text.Offset, end: Text.Offset): Text.Range =
+ {
+ val stop = if (start < end) end - 1 else end min model.buffer.getLength
+ Text.Range(start, stop)
+ }
+
+ def screen_lines_range(): Text.Range =
+ {
+ val start = text_area.getScreenLineStartOffset(0)
+ val raw_end = text_area.getScreenLineEndOffset(text_area.getVisibleLines - 1 max 0)
+ proper_line_range(start, if (raw_end >= 0) raw_end else model.buffer.getLength)
+ }
+
+ def invalidate_line_range(range: Text.Range)
+ {
+ text_area.invalidateLineRange(
+ model.buffer.getLineOfOffset(range.start),
+ model.buffer.getLineOfOffset(range.stop))
+ }
+
+
+ /* HTML popups */
+
+ private var html_popup: Option[Popup] = None
+
+ private def exit_popup() { html_popup.map(_.hide) }
+
+ private val html_panel =
+ new HTML_Panel(Isabelle.system, Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
+ html_panel.setBorder(BorderFactory.createLineBorder(Color.black))
+
+ private def html_panel_resize()
+ {
+ Swing_Thread.now {
+ html_panel.resize(Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
+ }
+ }
+
+ private def init_popup(snapshot: Document.Snapshot, x: Int, y: Int)
+ {
+ exit_popup()
+/* FIXME broken
+ val offset = text_area.xyToOffset(x, y)
+ val p = new Point(x, y); SwingUtilities.convertPointToScreen(p, text_area.getPainter)
+
+ // FIXME snapshot.cumulate
+ snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.popup) match {
+ case Text.Info(_, Some(msg)) #:: _ =>
+ val popup = PopupFactory.getSharedInstance().getPopup(text_area, html_panel, p.x, p.y + 60)
+ html_panel.render_sync(List(msg))
+ Thread.sleep(10) // FIXME !?
+ popup.show
+ html_popup = Some(popup)
+ case _ =>
+ }
+*/
+ }
+
+
+ /* subexpression highlighting */
+
+ private def subexp_range(snapshot: Document.Snapshot, x: Int, y: Int)
+ : Option[(Text.Range, Color)] =
+ {
+ val offset = text_area.xyToOffset(x, y)
+ snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.subexp) match {
+ case Text.Info(_, Some((range, color))) #:: _ => Some((snapshot.convert(range), color))
+ case _ => None
+ }
+ }
+
+ private var highlight_range: Option[(Text.Range, Color)] = None
+
+
+ /* CONTROL-mouse management */
+
+ private var control: Boolean = false
+
+ private def exit_control()
+ {
+ exit_popup()
+ highlight_range = None
+ }
+
+ private val focus_listener = new FocusAdapter {
+ override def focusLost(e: FocusEvent) {
+ highlight_range = None // FIXME exit_control !?
+ }
+ }
+
+ private val window_listener = new WindowAdapter {
+ override def windowIconified(e: WindowEvent) { exit_control() }
+ override def windowDeactivated(e: WindowEvent) { exit_control() }
+ }
+
+ private val mouse_motion_listener = new MouseMotionAdapter {
+ override def mouseMoved(e: MouseEvent) {
+ control = if (OperatingSystem.isMacOS()) e.isMetaDown else e.isControlDown
+ val x = e.getX()
+ val y = e.getY()
+
+ if (!model.buffer.isLoaded) exit_control()
+ else
+ Isabelle.swing_buffer_lock(model.buffer) {
+ val snapshot = model.snapshot
+
+ if (control) init_popup(snapshot, x, y)
+
+ highlight_range map { case (range, _) => invalidate_line_range(range) }
+ highlight_range = if (control) subexp_range(snapshot, x, y) else None
+ highlight_range map { case (range, _) => invalidate_line_range(range) }
+ }
+ }
+ }
+
+
+ /* TextArea painters */
+
+ private val background_painter = new TextAreaExtension
+ {
+ override def paintScreenLineRange(gfx: Graphics2D,
+ first_line: Int, last_line: Int, physical_lines: Array[Int],
+ start: Array[Int], end: Array[Int], y: Int, line_height: Int)
+ {
+ Isabelle.swing_buffer_lock(model.buffer) {
+ val snapshot = model.snapshot()
+ val ascent = text_area.getPainter.getFontMetrics.getAscent
+
+ for (i <- 0 until physical_lines.length) {
+ if (physical_lines(i) != -1) {
+ val line_range = proper_line_range(start(i), end(i))
+
+ // background color: status
+ val cmds = snapshot.node.command_range(snapshot.revert(line_range))
+ for {
+ (command, command_start) <- cmds if !command.is_ignored
+ val range = line_range.restrict(snapshot.convert(command.range + command_start))
+ r <- Isabelle.gfx_range(text_area, range)
+ color <- Isabelle_Markup.status_color(snapshot, command)
+ } {
+ gfx.setColor(color)
+ gfx.fillRect(r.x, y + i * line_height, r.length, line_height)
+ }
+
+ // background color (1): markup
+ for {
+ Text.Info(range, Some(color)) <-
+ snapshot.select_markup(line_range)(Isabelle_Markup.background1).iterator
+ r <- Isabelle.gfx_range(text_area, range)
+ } {
+ gfx.setColor(color)
+ gfx.fillRect(r.x, y + i * line_height, r.length, line_height)
+ }
+
+ // background color (2): markup
+ for {
+ Text.Info(range, Some(color)) <-
+ snapshot.select_markup(line_range)(Isabelle_Markup.background2).iterator
+ r <- Isabelle.gfx_range(text_area, range)
+ } {
+ gfx.setColor(color)
+ gfx.fillRect(r.x + 2, y + i * line_height + 2, r.length - 4, line_height - 4)
+ }
+
+ // sub-expression highlighting -- potentially from other snapshot
+ highlight_range match {
+ case Some((range, color)) if line_range.overlaps(range) =>
+ Isabelle.gfx_range(text_area, line_range.restrict(range)) match {
+ case None =>
+ case Some(r) =>
+ gfx.setColor(color)
+ gfx.drawRect(r.x, y + i * line_height, r.length - 1, line_height - 1)
+ }
+ case _ =>
+ }
+
+ // squiggly underline
+ for {
+ Text.Info(range, Some(color)) <-
+ snapshot.select_markup(line_range)(Isabelle_Markup.message).iterator
+ r <- Isabelle.gfx_range(text_area, range)
+ } {
+ gfx.setColor(color)
+ val x0 = (r.x / 2) * 2
+ val y0 = r.y + ascent + 1
+ for (x1 <- Range(x0, x0 + r.length, 2)) {
+ val y1 = if (x1 % 4 < 2) y0 else y0 + 1
+ gfx.drawLine(x1, y1, x1 + 1, y1)
+ }
+ }
+ }
+ }
+ }
+ }
+
+ override def getToolTipText(x: Int, y: Int): String =
+ {
+ Isabelle.swing_buffer_lock(model.buffer) {
+ val snapshot = model.snapshot()
+ val offset = text_area.xyToOffset(x, y)
+ if (control) {
+ snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.tooltip) match
+ {
+ case Text.Info(_, Some(text)) #:: _ => Isabelle.tooltip(text)
+ case _ => null
+ }
+ }
+ else {
+ // FIXME snapshot.cumulate
+ snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.popup) match
+ {
+ case Text.Info(_, Some(text)) #:: _ => Isabelle.tooltip(text)
+ case _ => null
+ }
+ }
+ }
+ }
+ }
+
+ var use_text_painter = false
+
+ private val text_painter = new TextAreaExtension
+ {
+ override def paintScreenLineRange(gfx: Graphics2D,
+ first_line: Int, last_line: Int, physical_lines: Array[Int],
+ start: Array[Int], end: Array[Int], y: Int, line_height: Int)
+ {
+ if (use_text_painter) {
+ Isabelle.swing_buffer_lock(model.buffer) {
+ val painter = text_area.getPainter
+ val fm = painter.getFontMetrics
+
+ val all_chunks = line_chunks(Set[Int]() ++ physical_lines.iterator.filter(i => i != -1))
+
+ val x0 = text_area.getHorizontalOffset
+ var y0 = y + fm.getHeight - (fm.getLeading + 1) - fm.getDescent
+ for (i <- 0 until physical_lines.length) {
+ if (physical_lines(i) != -1) {
+ all_chunks.get(start(i)) match {
+ case Some(chunk) =>
+ Chunk.paintChunkList(chunk, gfx, x0, y0, !Debug.DISABLE_GLYPH_VECTOR)
+ case None =>
+ }
+ }
+ y0 += line_height
+ }
+ }
+ }
+ }
+ }
+
+ private val gutter_painter = new TextAreaExtension
+ {
+ override def paintScreenLineRange(gfx: Graphics2D,
+ first_line: Int, last_line: Int, physical_lines: Array[Int],
+ start: Array[Int], end: Array[Int], y: Int, line_height: Int)
+ {
+ val gutter = text_area.getGutter
+ val width = GutterOptionPane.getSelectionAreaWidth
+ val border_width = jEdit.getIntegerProperty("view.gutter.borderWidth", 3)
+ val FOLD_MARKER_SIZE = 12
+
+ if (gutter.isSelectionAreaEnabled && !gutter.isExpanded && width >= 12 && line_height >= 12) {
+ Isabelle.swing_buffer_lock(model.buffer) {
+ val snapshot = model.snapshot()
+ for (i <- 0 until physical_lines.length) {
+ if (physical_lines(i) != -1) {
+ val line_range = proper_line_range(start(i), end(i))
+
+ // gutter icons
+ val icons =
+ (for (Text.Info(_, Some(icon)) <- // FIXME snapshot.cumulate
+ snapshot.select_markup(line_range)(Isabelle_Markup.gutter_message).iterator)
+ yield icon).toList.sortWith(_ >= _)
+ icons match {
+ case icon :: _ =>
+ val icn = icon.icon
+ val x0 = (FOLD_MARKER_SIZE + width - border_width - icn.getIconWidth) max 10
+ val y0 = y + i * line_height + (((line_height - icn.getIconHeight) / 2) max 0)
+ icn.paintIcon(gutter, gfx, x0, y0)
+ case Nil =>
+ }
+ }
+ }
+ }
+ }
+ }
+ }
+
+
+ /* caret handling */
+
+ def selected_command(): Option[Command] =
+ {
+ Swing_Thread.require()
+ model.snapshot().node.proper_command_at(text_area.getCaretPosition)
+ }
+
+ private val caret_listener = new CaretListener {
+ private val delay = Swing_Thread.delay_last(session.input_delay) {
+ session.perspective.event(Session.Perspective)
+ }
+ override def caretUpdate(e: CaretEvent) { delay() }
+ }
+
+
+ /* overview of command status left of scrollbar */
+
+ private val overview = new JPanel(new BorderLayout)
+ {
+ private val WIDTH = 10
+ private val HEIGHT = 2
+
+ setPreferredSize(new Dimension(WIDTH, 0))
+
+ setRequestFocusEnabled(false)
+
+ addMouseListener(new MouseAdapter {
+ override def mousePressed(event: MouseEvent) {
+ val line = y_to_line(event.getY)
+ if (line >= 0 && line < text_area.getLineCount)
+ text_area.setCaretPosition(text_area.getLineStartOffset(line))
+ }
+ })
+
+ override def addNotify() {
+ super.addNotify()
+ ToolTipManager.sharedInstance.registerComponent(this)
+ }
+
+ override def removeNotify() {
+ ToolTipManager.sharedInstance.unregisterComponent(this)
+ super.removeNotify
+ }
+
+ override def paintComponent(gfx: Graphics)
+ {
+ super.paintComponent(gfx)
+ Swing_Thread.assert()
+
+ val buffer = model.buffer
+ Isabelle.buffer_lock(buffer) {
+ val snapshot = model.snapshot()
+
+ def line_range(command: Command, start: Text.Offset): Option[(Int, Int)] =
+ {
+ try {
+ val line1 = buffer.getLineOfOffset(snapshot.convert(start))
+ val line2 = buffer.getLineOfOffset(snapshot.convert(start + command.length)) + 1
+ Some((line1, line2))
+ }
+ catch { case e: ArrayIndexOutOfBoundsException => None }
+ }
+ for {
+ (command, start) <- snapshot.node.command_starts
+ if !command.is_ignored
+ (line1, line2) <- line_range(command, start)
+ val y = line_to_y(line1)
+ val height = HEIGHT * (line2 - line1)
+ color <- Isabelle_Markup.overview_color(snapshot, command)
+ } {
+ gfx.setColor(color)
+ gfx.fillRect(0, y, getWidth - 1, height)
+ }
+ }
+ }
+
+ private def line_to_y(line: Int): Int =
+ (line * getHeight) / (text_area.getBuffer.getLineCount max text_area.getVisibleLines)
+
+ private def y_to_line(y: Int): Int =
+ (y * (text_area.getBuffer.getLineCount max text_area.getVisibleLines)) / getHeight
+ }
+
+
+ /* main actor */
+
+ private val main_actor = actor {
+ loop {
+ react {
+ case Session.Commands_Changed(changed) =>
+ val buffer = model.buffer
+ Isabelle.swing_buffer_lock(buffer) {
+ val snapshot = model.snapshot()
+
+ if (changed.exists(snapshot.node.commands.contains))
+ overview.repaint()
+
+ val visible_range = screen_lines_range()
+ val visible_cmds = snapshot.node.command_range(snapshot.revert(visible_range)).map(_._1)
+ if (visible_cmds.exists(changed)) {
+ for {
+ line <- 0 until text_area.getVisibleLines
+ val start = text_area.getScreenLineStartOffset(line) if start >= 0
+ val end = text_area.getScreenLineEndOffset(line) if end >= 0
+ val range = proper_line_range(start, end)
+ val line_cmds = snapshot.node.command_range(snapshot.revert(range)).map(_._1)
+ if line_cmds.exists(changed)
+ } text_area.invalidateScreenLineRange(line, line)
+
+ // FIXME danger of deadlock!?
+ // FIXME potentially slow!?
+ model.buffer.propertiesChanged()
+ }
+ }
+
+ case Session.Global_Settings => html_panel_resize()
+
+ case bad => System.err.println("command_change_actor: ignoring bad message " + bad)
+ }
+ }
+ }
+
+
+ /* activation */
+
+ private def activate()
+ {
+ val painter = text_area.getPainter
+ painter.addExtension(TextAreaPainter.LINE_BACKGROUND_LAYER + 1, background_painter)
+ painter.addExtension(TextAreaPainter.TEXT_LAYER + 1, text_painter)
+ text_area.getGutter.addExtension(gutter_painter)
+ text_area.addFocusListener(focus_listener)
+ text_area.getView.addWindowListener(window_listener)
+ painter.addMouseMotionListener(mouse_motion_listener)
+ text_area.addCaretListener(caret_listener)
+ text_area.addLeftOfScrollBar(overview)
+ session.commands_changed += main_actor
+ session.global_settings += main_actor
+ }
+
+ private def deactivate()
+ {
+ val painter = text_area.getPainter
+ session.commands_changed -= main_actor
+ session.global_settings -= main_actor
+ text_area.removeFocusListener(focus_listener)
+ text_area.getView.removeWindowListener(window_listener)
+ painter.removeMouseMotionListener(mouse_motion_listener)
+ text_area.removeCaretListener(caret_listener)
+ text_area.removeLeftOfScrollBar(overview)
+ text_area.getGutter.removeExtension(gutter_painter)
+ painter.removeExtension(text_painter)
+ painter.removeExtension(background_painter)
+ exit_popup()
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/html_panel.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,206 @@
+/* Title: Tools/jEdit/src/html_panel.scala
+ Author: Makarius
+
+HTML panel based on Lobo/Cobra.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.io.StringReader
+import java.awt.{Font, BorderLayout, Dimension, GraphicsEnvironment, Toolkit, FontMetrics}
+import java.awt.event.MouseEvent
+
+import java.util.logging.{Logger, Level}
+
+import org.w3c.dom.html2.HTMLElement
+
+import org.lobobrowser.html.parser.{DocumentBuilderImpl, InputSourceImpl}
+import org.lobobrowser.html.gui.HtmlPanel
+import org.lobobrowser.html.domimpl.{HTMLDocumentImpl, HTMLStyleElementImpl, NodeImpl}
+import org.lobobrowser.html.test.{SimpleHtmlRendererContext, SimpleUserAgentContext}
+
+import scala.actors.Actor._
+
+
+object HTML_Panel
+{
+ sealed abstract class Event { val element: HTMLElement; val mouse: MouseEvent }
+ case class Context_Menu(val element: HTMLElement, mouse: MouseEvent) extends Event
+ case class Mouse_Click(val element: HTMLElement, mouse: MouseEvent) extends Event
+ case class Double_Click(val element: HTMLElement, mouse: MouseEvent) extends Event
+ case class Mouse_Over(val element: HTMLElement, mouse: MouseEvent) extends Event
+ case class Mouse_Out(val element: HTMLElement, mouse: MouseEvent) extends Event
+}
+
+
+class HTML_Panel(
+ system: Isabelle_System,
+ initial_font_family: String,
+ initial_font_size: Int)
+ extends HtmlPanel
+{
+ /** Lobo setup **/
+
+ /* global logging */
+
+ Logger.getLogger("org.lobobrowser").setLevel(Level.WARNING)
+
+
+ /* pixel size -- cf. org.lobobrowser.html.style.HtmlValues.getFontSize */
+
+ val screen_resolution =
+ if (GraphicsEnvironment.isHeadless()) 72
+ else Toolkit.getDefaultToolkit().getScreenResolution()
+
+ def lobo_px(raw_px: Int): Int = raw_px * 96 / screen_resolution
+ def raw_px(lobo_px: Int): Int = (lobo_px * screen_resolution + 95) / 96
+
+
+ /* contexts and event handling */
+
+ protected val handler: PartialFunction[HTML_Panel.Event, Unit] = Library.undefined
+
+ private val ucontext = new SimpleUserAgentContext
+ private val rcontext = new SimpleHtmlRendererContext(this, ucontext)
+ {
+ private def handle(event: HTML_Panel.Event): Boolean =
+ if (handler.isDefinedAt(event)) { handler(event); false }
+ else true
+
+ override def onContextMenu(elem: HTMLElement, event: MouseEvent): Boolean =
+ handle(HTML_Panel.Context_Menu(elem, event))
+ override def onMouseClick(elem: HTMLElement, event: MouseEvent): Boolean =
+ handle(HTML_Panel.Mouse_Click(elem, event))
+ override def onDoubleClick(elem: HTMLElement, event: MouseEvent): Boolean =
+ handle(HTML_Panel.Double_Click(elem, event))
+ override def onMouseOver(elem: HTMLElement, event: MouseEvent)
+ { handle(HTML_Panel.Mouse_Over(elem, event)) }
+ override def onMouseOut(elem: HTMLElement, event: MouseEvent)
+ { handle(HTML_Panel.Mouse_Out(elem, event)) }
+ }
+
+ private val builder = new DocumentBuilderImpl(ucontext, rcontext)
+
+
+ /* document template with style sheets */
+
+ private val template_head =
+ """<?xml version="1.0" encoding="utf-8"?>
+<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN"
+ "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd">
+<html xmlns="http://www.w3.org/1999/xhtml">
+<head>
+<style media="all" type="text/css">
+""" +
+ system.try_read(system.getenv_strict("JEDIT_STYLE_SHEETS").split(":"))
+
+ private val template_tail =
+"""
+</style>
+</head>
+<body/>
+</html>
+"""
+
+ private def template(font_family: String, font_size: Int): String =
+ template_head +
+ "body { font-family: " + font_family + "; font-size: " + raw_px(font_size) + "px; }" +
+ template_tail
+
+
+ /** main actor **/
+
+ /* internal messages */
+
+ private case class Resize(font_family: String, font_size: Int)
+ private case class Render_Document(text: String)
+ private case class Render(body: XML.Body)
+ private case class Render_Sync(body: XML.Body)
+ private case object Refresh
+
+ private val main_actor = actor {
+
+ /* internal state */
+
+ var current_font_metrics: FontMetrics = null
+ var current_font_family = ""
+ var current_font_size: Int = 0
+ var current_margin: Int = 0
+ var current_body: XML.Body = Nil
+
+ def resize(font_family: String, font_size: Int)
+ {
+ val font = new Font(font_family, Font.PLAIN, lobo_px(raw_px(font_size)))
+ val (font_metrics, margin) =
+ Swing_Thread.now {
+ val metrics = getFontMetrics(font)
+ (metrics, (getWidth() / (metrics.charWidth(Symbol.spc) max 1) - 4) max 20)
+ }
+ if (current_font_metrics == null ||
+ current_font_family != font_family ||
+ current_font_size != font_size ||
+ current_margin != margin)
+ {
+ current_font_metrics = font_metrics
+ current_font_family = font_family
+ current_font_size = font_size
+ current_margin = margin
+ refresh()
+ }
+ }
+
+ def refresh() { render(current_body) }
+
+ def render_document(text: String)
+ {
+ val doc = builder.parse(new InputSourceImpl(new StringReader(text), "http://localhost"))
+ Swing_Thread.later { setDocument(doc, rcontext) }
+ }
+
+ def render(body: XML.Body)
+ {
+ current_body = body
+ val html_body =
+ current_body.flatMap(div =>
+ Pretty.formatted(List(div), current_margin, Pretty.font_metric(current_font_metrics))
+ .map(t =>
+ XML.Elem(Markup(HTML.PRE, List((Markup.CLASS, Markup.MESSAGE))),
+ HTML.spans(t, true))))
+ val doc =
+ builder.parse(
+ new InputSourceImpl(
+ new StringReader(template(current_font_family, current_font_size)), "http://localhost"))
+ doc.removeChild(doc.getLastChild())
+ doc.appendChild(XML.document_node(doc, XML.elem(HTML.BODY, html_body)))
+ Swing_Thread.later { setDocument(doc, rcontext) }
+ }
+
+
+ /* main loop */
+
+ resize(initial_font_family, initial_font_size)
+
+ loop {
+ react {
+ case Resize(font_family, font_size) => resize(font_family, font_size)
+ case Refresh => refresh()
+ case Render_Document(text) => render_document(text)
+ case Render(body) => render(body)
+ case Render_Sync(body) => render(body); reply(())
+ case bad => System.err.println("main_actor: ignoring bad message " + bad)
+ }
+ }
+ }
+
+
+ /* external methods */
+
+ def resize(font_family: String, font_size: Int) { main_actor ! Resize(font_family, font_size) }
+ def refresh() { main_actor ! Refresh }
+ def render_document(text: String) { main_actor ! Render_Document(text) }
+ def render(body: XML.Body) { main_actor ! Render(body) }
+ def render_sync(body: XML.Body) { main_actor !? Render_Sync(body) }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/isabelle_encoding.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,65 @@
+/* Title: Tools/jEdit/src/isabelle_encoding.scala
+ Author: Makarius
+
+Isabelle encoding -- based on UTF-8.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import org.gjt.sp.jedit.io.Encoding
+import org.gjt.sp.jedit.buffer.JEditBuffer
+
+import java.nio.charset.{Charset, CodingErrorAction}
+import java.io.{InputStream, OutputStream, Reader, Writer, InputStreamReader, OutputStreamWriter,
+ CharArrayReader, ByteArrayOutputStream}
+
+import scala.io.{Codec, Source, BufferedSource}
+
+
+object Isabelle_Encoding
+{
+ val NAME = "UTF-8-Isabelle"
+
+ def is_active(buffer: JEditBuffer): Boolean =
+ buffer.getStringProperty(JEditBuffer.ENCODING).asInstanceOf[String] == NAME
+}
+
+class Isabelle_Encoding extends Encoding
+{
+ private val charset = Charset.forName(Standard_System.charset)
+ val BUFSIZE = 32768
+
+ private def text_reader(in: InputStream, codec: Codec): Reader =
+ {
+ val source = new BufferedSource(in)(codec)
+ new CharArrayReader(Isabelle.system.symbols.decode(source.mkString).toArray)
+ }
+
+ override def getTextReader(in: InputStream): Reader =
+ text_reader(in, Standard_System.codec())
+
+ override def getPermissiveTextReader(in: InputStream): Reader =
+ {
+ val codec = Standard_System.codec()
+ codec.onMalformedInput(CodingErrorAction.REPLACE)
+ codec.onUnmappableCharacter(CodingErrorAction.REPLACE)
+ text_reader(in, codec)
+ }
+
+ override def getTextWriter(out: OutputStream): Writer =
+ {
+ val buffer = new ByteArrayOutputStream(BUFSIZE) {
+ override def flush()
+ {
+ val text = Isabelle.system.symbols.encode(toString(Standard_System.charset))
+ out.write(text.getBytes(Standard_System.charset))
+ out.flush()
+ }
+ override def close() { out.close() }
+ }
+ new OutputStreamWriter(buffer, charset.newEncoder())
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/isabelle_hyperlinks.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,88 @@
+/* Title: Tools/jEdit/src/isabelle_hyperlinks.scala
+ Author: Fabian Immler, TU Munich
+
+Hyperlink setup for Isabelle proof documents.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.io.File
+
+import gatchan.jedit.hyperlinks.{Hyperlink, HyperlinkSource, AbstractHyperlink}
+
+import org.gjt.sp.jedit.{View, jEdit, Buffer, TextUtilities}
+
+
+private class Internal_Hyperlink(start: Int, end: Int, line: Int, def_offset: Int)
+ extends AbstractHyperlink(start, end, line, "")
+{
+ override def click(view: View) {
+ view.getTextArea.moveCaretPosition(def_offset)
+ }
+}
+
+class External_Hyperlink(start: Int, end: Int, line: Int, def_file: String, def_line: Int)
+ extends AbstractHyperlink(start, end, line, "")
+{
+ override def click(view: View) = {
+ Isabelle.system.source_file(def_file) match {
+ case None =>
+ Library.error_dialog(view, "File not found", "Could not find source file " + def_file)
+ case Some(file) =>
+ jEdit.openFiles(view, file.getParent, Array(file.getName, "+line:" + def_line))
+ }
+ }
+}
+
+class Isabelle_Hyperlinks extends HyperlinkSource
+{
+ def getHyperlink(buffer: Buffer, buffer_offset: Int): Hyperlink =
+ {
+ Swing_Thread.assert()
+ Isabelle.buffer_lock(buffer) {
+ Document_Model(buffer) match {
+ case Some(model) =>
+ val snapshot = model.snapshot()
+ val markup =
+ snapshot.select_markup(Text.Range(buffer_offset, buffer_offset + 1)) {
+ // FIXME Isar_Document.Hyperlink extractor
+ case Text.Info(info_range,
+ XML.Elem(Markup(Markup.ENTITY, props), _))
+ if (props.find(
+ { case (Markup.KIND, Markup.ML_OPEN) => true
+ case (Markup.KIND, Markup.ML_STRUCT) => true
+ case _ => false }).isEmpty) =>
+ val Text.Range(begin, end) = info_range
+ val line = buffer.getLineOfOffset(begin)
+ (Position.File.unapply(props), Position.Line.unapply(props)) match {
+ case (Some(def_file), Some(def_line)) =>
+ new External_Hyperlink(begin, end, line, def_file, def_line)
+ case _ =>
+ (props, props) match {
+ case (Position.Id(def_id), Position.Offset(def_offset)) =>
+ snapshot.lookup_command(def_id) match {
+ case Some(def_cmd) =>
+ snapshot.node.command_start(def_cmd) match {
+ case Some(def_cmd_start) =>
+ new Internal_Hyperlink(begin, end, line,
+ snapshot.convert(def_cmd_start + def_cmd.decode(def_offset)))
+ case None => null
+ }
+ case None => null
+ }
+ case _ => null
+ }
+ }
+ }
+ markup match {
+ case Text.Info(_, Some(link)) #:: _ => link
+ case _ => null
+ }
+ case None => null
+ }
+ }
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/isabelle_markup.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,216 @@
+/* Title: Tools/jEdit/src/isabelle_markup.scala
+ Author: Makarius
+
+Isabelle specific physical rendering and markup selection.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.awt.Color
+
+import org.gjt.sp.jedit.syntax.Token
+
+
+object Isabelle_Markup
+{
+ /* physical rendering */
+
+ val outdated_color = new Color(240, 240, 240)
+ val unfinished_color = new Color(255, 228, 225)
+
+ val light_color = new Color(240, 240, 240)
+ val regular_color = new Color(192, 192, 192)
+ val warning_color = new Color(255, 140, 0)
+ val error_color = new Color(178, 34, 34)
+ val bad_color = new Color(255, 106, 106, 100)
+ val hilite_color = new Color(255, 204, 102, 100)
+
+ class Icon(val priority: Int, val icon: javax.swing.Icon)
+ {
+ def >= (that: Icon): Boolean = this.priority >= that.priority
+ }
+ val warning_icon = new Icon(1, Isabelle.load_icon("16x16/status/dialog-warning.png"))
+ val error_icon = new Icon(2, Isabelle.load_icon("16x16/status/dialog-error.png"))
+
+
+ /* command status */
+
+ def status_color(snapshot: Document.Snapshot, command: Command): Option[Color] =
+ {
+ val state = snapshot.state(command)
+ if (snapshot.is_outdated) Some(outdated_color)
+ else
+ Isar_Document.command_status(state.status) match {
+ case Isar_Document.Forked(i) if i > 0 => Some(unfinished_color)
+ case Isar_Document.Unprocessed => Some(unfinished_color)
+ case _ => None
+ }
+ }
+
+ def overview_color(snapshot: Document.Snapshot, command: Command): Option[Color] =
+ {
+ val state = snapshot.state(command)
+ if (snapshot.is_outdated) None
+ else
+ Isar_Document.command_status(state.status) match {
+ case Isar_Document.Forked(i) => if (i > 0) Some(unfinished_color) else None
+ case Isar_Document.Unprocessed => Some(unfinished_color)
+ case Isar_Document.Failed => Some(error_color)
+ case Isar_Document.Finished =>
+ if (state.results.exists(r => Isar_Document.is_error(r._2))) Some(error_color)
+ else if (state.results.exists(r => Isar_Document.is_warning(r._2))) Some(warning_color)
+ else None
+ }
+ }
+
+
+ /* markup selectors */
+
+ val message: Markup_Tree.Select[Color] =
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.WRITELN, _), _)) => regular_color
+ case Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), _)) => warning_color
+ case Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _)) => error_color
+ }
+
+ val popup: Markup_Tree.Select[String] =
+ {
+ case Text.Info(_, msg @ XML.Elem(Markup(markup, _), _))
+ if markup == Markup.WRITELN || markup == Markup.WARNING || markup == Markup.ERROR =>
+ Pretty.string_of(List(msg), margin = Isabelle.Int_Property("tooltip-margin"))
+ }
+
+ val gutter_message: Markup_Tree.Select[Icon] =
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), _)) => warning_icon
+ case Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _)) => error_icon
+ }
+
+ val background1: Markup_Tree.Select[Color] =
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _)) => bad_color
+ case Text.Info(_, XML.Elem(Markup(Markup.HILITE, _), _)) => hilite_color
+ }
+
+ val background2: Markup_Tree.Select[Color] =
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.TOKEN_RANGE, _), _)) => light_color
+ }
+
+ val tooltip: Markup_Tree.Select[String] =
+ {
+ case Text.Info(_, XML.Elem(Markup.Entity(kind, name), _)) => kind + " \"" + name + "\""
+ case Text.Info(_, XML.Elem(Markup(Markup.ML_TYPING, _), body)) =>
+ Pretty.string_of(List(Pretty.block(XML.Text("ML:") :: Pretty.Break(1) :: body)),
+ margin = Isabelle.Int_Property("tooltip-margin"))
+ case Text.Info(_, XML.Elem(Markup(Markup.SORT, _), _)) => "sort"
+ case Text.Info(_, XML.Elem(Markup(Markup.TYP, _), _)) => "type"
+ case Text.Info(_, XML.Elem(Markup(Markup.TERM, _), _)) => "term"
+ case Text.Info(_, XML.Elem(Markup(Markup.PROP, _), _)) => "proposition"
+ case Text.Info(_, XML.Elem(Markup(Markup.TOKEN_RANGE, _), _)) => "inner syntax token"
+ case Text.Info(_, XML.Elem(Markup(Markup.FREE, _), _)) => "free variable (globally fixed)"
+ case Text.Info(_, XML.Elem(Markup(Markup.SKOLEM, _), _)) => "skolem variable (locally fixed)"
+ case Text.Info(_, XML.Elem(Markup(Markup.BOUND, _), _)) => "bound variable (internally fixed)"
+ case Text.Info(_, XML.Elem(Markup(Markup.VAR, _), _)) => "schematic variable"
+ case Text.Info(_, XML.Elem(Markup(Markup.TFREE, _), _)) => "free type variable"
+ case Text.Info(_, XML.Elem(Markup(Markup.TVAR, _), _)) => "schematic type variable"
+ }
+
+ private val subexp_include =
+ Set(Markup.SORT, Markup.TYP, Markup.TERM, Markup.PROP, Markup.ML_TYPING, Markup.TOKEN_RANGE,
+ Markup.ENTITY, Markup.FREE, Markup.SKOLEM, Markup.BOUND, Markup.VAR,
+ Markup.TFREE, Markup.TVAR)
+
+ val subexp: Markup_Tree.Select[(Text.Range, Color)] =
+ {
+ case Text.Info(range, XML.Elem(Markup(name, _), _)) if subexp_include(name) =>
+ (range, Color.black)
+ }
+
+
+ /* token markup -- text styles */
+
+ private val command_style: Map[String, Byte] =
+ {
+ import Token._
+ Map[String, Byte](
+ Keyword.THY_END -> KEYWORD2,
+ Keyword.THY_SCRIPT -> LABEL,
+ Keyword.PRF_SCRIPT -> LABEL,
+ Keyword.PRF_ASM -> KEYWORD3,
+ Keyword.PRF_ASM_GOAL -> KEYWORD3
+ ).withDefaultValue(KEYWORD1)
+ }
+
+ private val token_style: Map[String, Byte] =
+ {
+ import Token._
+ Map[String, Byte](
+ // logical entities
+ Markup.TCLASS -> NULL,
+ Markup.TYCON -> NULL,
+ Markup.FIXED -> NULL,
+ Markup.CONST -> LITERAL2,
+ Markup.DYNAMIC_FACT -> LABEL,
+ // inner syntax
+ Markup.TFREE -> NULL,
+ Markup.FREE -> MARKUP,
+ Markup.TVAR -> NULL,
+ Markup.SKOLEM -> COMMENT2,
+ Markup.BOUND -> LABEL,
+ Markup.VAR -> NULL,
+ Markup.NUM -> DIGIT,
+ Markup.FLOAT -> DIGIT,
+ Markup.XNUM -> DIGIT,
+ Markup.XSTR -> LITERAL4,
+ Markup.LITERAL -> OPERATOR,
+ Markup.INNER_COMMENT -> COMMENT1,
+ Markup.SORT -> NULL,
+ Markup.TYP -> NULL,
+ Markup.TERM -> NULL,
+ Markup.PROP -> NULL,
+ // ML syntax
+ Markup.ML_KEYWORD -> KEYWORD1,
+ Markup.ML_DELIMITER -> OPERATOR,
+ Markup.ML_IDENT -> NULL,
+ Markup.ML_TVAR -> NULL,
+ Markup.ML_NUMERAL -> DIGIT,
+ Markup.ML_CHAR -> LITERAL1,
+ Markup.ML_STRING -> LITERAL1,
+ Markup.ML_COMMENT -> COMMENT1,
+ Markup.ML_MALFORMED -> INVALID,
+ // embedded source text
+ Markup.ML_SOURCE -> COMMENT3,
+ Markup.DOC_SOURCE -> COMMENT3,
+ Markup.ANTIQ -> COMMENT4,
+ Markup.ML_ANTIQ -> COMMENT4,
+ Markup.DOC_ANTIQ -> COMMENT4,
+ // outer syntax
+ Markup.KEYWORD -> KEYWORD2,
+ Markup.OPERATOR -> OPERATOR,
+ Markup.COMMAND -> KEYWORD1,
+ Markup.IDENT -> NULL,
+ Markup.VERBATIM -> COMMENT3,
+ Markup.COMMENT -> COMMENT1,
+ Markup.CONTROL -> COMMENT3,
+ Markup.MALFORMED -> INVALID,
+ Markup.STRING -> LITERAL3,
+ Markup.ALTSTRING -> LITERAL1
+ ).withDefaultValue(NULL)
+ }
+
+ def tokens(syntax: Outer_Syntax): Markup_Tree.Select[Byte] =
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.COMMAND, List((Markup.NAME, name))), _))
+ if syntax.keyword_kind(name).isDefined => command_style(syntax.keyword_kind(name).get)
+
+ case Text.Info(_, XML.Elem(Markup(Markup.ENTITY, Markup.Kind(kind)), _))
+ if token_style(kind) != Token.NULL => token_style(kind)
+
+ case Text.Info(_, XML.Elem(Markup(name, _), _))
+ if token_style(name) != Token.NULL => token_style(name)
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/isabelle_options.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,67 @@
+/* Title: Tools/jEdit/src/isabelle_options.scala
+ Author: Johannes Hölzl, TU Munich
+
+Editor pane for plugin options.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import javax.swing.JSpinner
+
+import scala.swing.CheckBox
+
+import org.gjt.sp.jedit.AbstractOptionPane
+
+
+class Isabelle_Options extends AbstractOptionPane("isabelle")
+{
+ private val logic_selector = Isabelle.logic_selector(Isabelle.Property("logic"))
+ private val auto_start = new CheckBox()
+ private val relative_font_size = new JSpinner()
+ private val tooltip_font_size = new JSpinner()
+ private val tooltip_margin = new JSpinner()
+ private val tooltip_dismiss_delay = new JSpinner()
+
+ override def _init()
+ {
+ addComponent(Isabelle.Property("logic.title"), logic_selector.peer)
+
+ addComponent(Isabelle.Property("auto-start.title"), auto_start.peer)
+ auto_start.selected = Isabelle.Boolean_Property("auto-start")
+
+ relative_font_size.setValue(Isabelle.Int_Property("relative-font-size", 100))
+ addComponent(Isabelle.Property("relative-font-size.title"), relative_font_size)
+
+ tooltip_font_size.setValue(Isabelle.Int_Property("tooltip-font-size", 10))
+ addComponent(Isabelle.Property("tooltip-font-size.title"), tooltip_font_size)
+
+ tooltip_margin.setValue(Isabelle.Int_Property("tooltip-margin", 40))
+ addComponent(Isabelle.Property("tooltip-margin.title"), tooltip_margin)
+
+ tooltip_dismiss_delay.setValue(
+ Isabelle.Time_Property("tooltip-dismiss-delay", Time.seconds(8.0)).ms.toInt)
+ addComponent(Isabelle.Property("tooltip-dismiss-delay.title"), tooltip_dismiss_delay)
+ }
+
+ override def _save()
+ {
+ Isabelle.Property("logic") = logic_selector.selection.item.name
+
+ Isabelle.Boolean_Property("auto-start") = auto_start.selected
+
+ Isabelle.Int_Property("relative-font-size") =
+ relative_font_size.getValue().asInstanceOf[Int]
+
+ Isabelle.Int_Property("tooltip-font-size") =
+ tooltip_font_size.getValue().asInstanceOf[Int]
+
+ Isabelle.Int_Property("tooltip-margin") =
+ tooltip_margin.getValue().asInstanceOf[Int]
+
+ Isabelle.Time_Property("tooltip-dismiss-delay") =
+ Time.seconds(tooltip_dismiss_delay.getValue().asInstanceOf[Int])
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/isabelle_sidekick.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,176 @@
+/* Title: Tools/jEdit/src/isabelle_sidekick.scala
+ Author: Fabian Immler, TU Munich
+ Author: Makarius
+
+SideKick parsers for Isabelle proof documents.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.collection.Set
+import scala.collection.immutable.TreeSet
+
+import javax.swing.tree.DefaultMutableTreeNode
+import javax.swing.text.Position
+import javax.swing.Icon
+
+import org.gjt.sp.jedit.{Buffer, EditPane, TextUtilities, View}
+import errorlist.DefaultErrorSource
+import sidekick.{SideKickParser, SideKickParsedData, SideKickCompletion, IAsset}
+
+
+object Isabelle_Sidekick
+{
+ def int_to_pos(offset: Text.Offset): Position =
+ new Position { def getOffset = offset; override def toString = offset.toString }
+
+ class Asset(name: String, start: Text.Offset, end: Text.Offset) extends IAsset
+ {
+ protected var _name = name
+ protected var _start = int_to_pos(start)
+ protected var _end = int_to_pos(end)
+ override def getIcon: Icon = null
+ override def getShortString: String = _name
+ override def getLongString: String = _name
+ override def getName: String = _name
+ override def setName(name: String) = _name = name
+ override def getStart: Position = _start
+ override def setStart(start: Position) = _start = start
+ override def getEnd: Position = _end
+ override def setEnd(end: Position) = _end = end
+ override def toString = _name
+ }
+}
+
+
+abstract class Isabelle_Sidekick(name: String) extends SideKickParser(name)
+{
+ /* parsing */
+
+ @volatile protected var stopped = false
+ override def stop() = { stopped = true }
+
+ def parser(data: SideKickParsedData, model: Document_Model): Unit
+
+ def parse(buffer: Buffer, error_source: DefaultErrorSource): SideKickParsedData =
+ {
+ stopped = false
+
+ // FIXME lock buffer (!??)
+ val data = new SideKickParsedData(buffer.getName)
+ val root = data.root
+ data.getAsset(root).setEnd(Isabelle_Sidekick.int_to_pos(buffer.getLength))
+
+ Swing_Thread.now { Document_Model(buffer) } match {
+ case Some(model) =>
+ parser(data, model)
+ if (stopped) root.add(new DefaultMutableTreeNode("<parser stopped>"))
+ case None => root.add(new DefaultMutableTreeNode("<buffer inactive>"))
+ }
+ data
+ }
+
+
+ /* completion */
+
+ override def supportsCompletion = true
+ override def canCompleteAnywhere = true
+
+ override def complete(pane: EditPane, caret: Text.Offset): SideKickCompletion =
+ {
+ val buffer = pane.getBuffer
+ Isabelle.swing_buffer_lock(buffer) {
+ Document_Model(buffer) match {
+ case None => null
+ case Some(model) =>
+ val line = buffer.getLineOfOffset(caret)
+ val start = buffer.getLineStartOffset(line)
+ val text = buffer.getSegment(start, caret - start)
+
+ val completion = model.session.current_syntax().completion
+ completion.complete(text) match {
+ case None => null
+ case Some((word, cs)) =>
+ val ds =
+ (if (Isabelle_Encoding.is_active(buffer))
+ cs.map(Isabelle.system.symbols.decode(_)).sortWith(_ < _)
+ else cs).filter(_ != word)
+ if (ds.isEmpty) null
+ else new SideKickCompletion(
+ pane.getView, word, ds.toArray.asInstanceOf[Array[Object]]) { }
+ }
+ }
+ }
+ }
+}
+
+
+class Isabelle_Sidekick_Default extends Isabelle_Sidekick("isabelle")
+{
+ import Thy_Syntax.Structure
+
+ def parser(data: SideKickParsedData, model: Document_Model)
+ {
+ val syntax = model.session.current_syntax()
+
+ def make_tree(offset: Text.Offset, entry: Structure.Entry): List[DefaultMutableTreeNode] =
+ entry match {
+ case Structure.Block(name, body) =>
+ val node =
+ new DefaultMutableTreeNode(
+ new Isabelle_Sidekick.Asset(Library.first_line(name), offset, offset + entry.length))
+ (offset /: body)((i, e) =>
+ {
+ make_tree(i, e) foreach (nd => node.add(nd))
+ i + e.length
+ })
+ List(node)
+ case Structure.Atom(command)
+ if command.is_command && syntax.heading_level(command).isEmpty =>
+ val node =
+ new DefaultMutableTreeNode(
+ new Isabelle_Sidekick.Asset(command.name, offset, offset + entry.length))
+ List(node)
+ case _ => Nil
+ }
+
+ val text = Isabelle.buffer_text(model.buffer)
+ val structure = Structure.parse(syntax, "theory " + model.thy_name, text)
+
+ make_tree(0, structure) foreach (node => data.root.add(node))
+ }
+}
+
+
+class Isabelle_Sidekick_Raw extends Isabelle_Sidekick("isabelle-raw")
+{
+ def parser(data: SideKickParsedData, model: Document_Model)
+ {
+ val root = data.root
+ val snapshot = Swing_Thread.now { model.snapshot() } // FIXME cover all nodes (!??)
+ for ((command, command_start) <- snapshot.node.command_range() if !stopped) {
+ snapshot.state(command).root_markup.swing_tree(root)((info: Text.Info[Any]) =>
+ {
+ val range = info.range + command_start
+ val content = command.source(info.range).replace('\n', ' ')
+ val info_text =
+ info.info match {
+ case elem @ XML.Elem(_, _) =>
+ Pretty.formatted(List(elem), margin = 40).mkString("\n")
+ case x => x.toString
+ }
+
+ new DefaultMutableTreeNode(
+ new Isabelle_Sidekick.Asset(command.toString, range.start, range.stop) {
+ override def getShortString: String = content
+ override def getLongString: String = info_text
+ override def toString = "\"" + content + "\" " + range.toString
+ })
+ })
+ }
+ }
+}
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/jEdit.props Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,213 @@
+#jEdit properties
+buffer.deepIndent=false
+buffer.encoding=UTF-8-Isabelle
+buffer.indentSize=2
+buffer.lineSeparator=\n
+buffer.maxLineLen=100
+buffer.noTabs=true
+buffer.sidekick.keystroke-parse=false
+buffer.tabSize=2
+console.encoding=UTF-8
+console.font=IsabelleText
+console.fontsize=14
+delete-line.shortcut=A+d
+delete.shortcut2=C+d
+encoding.opt-out.Big5-HKSCS=true
+encoding.opt-out.Big5=true
+encoding.opt-out.COMPOUND_TEXT=true
+encoding.opt-out.EUC-JP=true
+encoding.opt-out.EUC-KR=true
+encoding.opt-out.GB18030=true
+encoding.opt-out.GB2312=true
+encoding.opt-out.GBK=true
+encoding.opt-out.IBM-Thai=true
+encoding.opt-out.IBM00858=true
+encoding.opt-out.IBM01140=true
+encoding.opt-out.IBM01141=true
+encoding.opt-out.IBM01142=true
+encoding.opt-out.IBM01143=true
+encoding.opt-out.IBM01144=true
+encoding.opt-out.IBM01145=true
+encoding.opt-out.IBM01146=true
+encoding.opt-out.IBM01147=true
+encoding.opt-out.IBM01148=true
+encoding.opt-out.IBM01149=true
+encoding.opt-out.IBM037=true
+encoding.opt-out.IBM1026=true
+encoding.opt-out.IBM1047=true
+encoding.opt-out.IBM273=true
+encoding.opt-out.IBM277=true
+encoding.opt-out.IBM278=true
+encoding.opt-out.IBM280=true
+encoding.opt-out.IBM284=true
+encoding.opt-out.IBM285=true
+encoding.opt-out.IBM297=true
+encoding.opt-out.IBM420=true
+encoding.opt-out.IBM424=true
+encoding.opt-out.IBM437=true
+encoding.opt-out.IBM500=true
+encoding.opt-out.IBM775=true
+encoding.opt-out.IBM850=true
+encoding.opt-out.IBM852=true
+encoding.opt-out.IBM855=true
+encoding.opt-out.IBM857=true
+encoding.opt-out.IBM860=true
+encoding.opt-out.IBM861=true
+encoding.opt-out.IBM862=true
+encoding.opt-out.IBM863=true
+encoding.opt-out.IBM864=true
+encoding.opt-out.IBM865=true
+encoding.opt-out.IBM866=true
+encoding.opt-out.IBM868=true
+encoding.opt-out.IBM869=true
+encoding.opt-out.IBM870=true
+encoding.opt-out.IBM871=true
+encoding.opt-out.IBM918=true
+encoding.opt-out.ISO-2022-CN=true
+encoding.opt-out.ISO-2022-JP-2=true
+encoding.opt-out.ISO-2022-JP=true
+encoding.opt-out.ISO-2022-KR=true
+encoding.opt-out.ISO-8859-13=true
+encoding.opt-out.ISO-8859-2=true
+encoding.opt-out.ISO-8859-3=true
+encoding.opt-out.ISO-8859-4=true
+encoding.opt-out.ISO-8859-5=true
+encoding.opt-out.ISO-8859-6=true
+encoding.opt-out.ISO-8859-7=true
+encoding.opt-out.ISO-8859-8=true
+encoding.opt-out.ISO-8859-9=true
+encoding.opt-out.JIS_X0201=true
+encoding.opt-out.JIS_X0212-1990=true
+encoding.opt-out.KOI8-R=true
+encoding.opt-out.KOI8-U=true
+encoding.opt-out.Shift_JIS=true
+encoding.opt-out.TIS-620=true
+encoding.opt-out.UTF-16=true
+encoding.opt-out.UTF-16BE=true
+encoding.opt-out.UTF-16LE=true
+encoding.opt-out.UTF-32=true
+encoding.opt-out.UTF-32BE=true
+encoding.opt-out.UTF-32LE=true
+encoding.opt-out.X-UTF-32BE-BOM=true
+encoding.opt-out.X-UTF-32LE-BOM=true
+encoding.opt-out.windows-1250=true
+encoding.opt-out.windows-1251=true
+encoding.opt-out.windows-1253=true
+encoding.opt-out.windows-1254=true
+encoding.opt-out.windows-1255=true
+encoding.opt-out.windows-1256=true
+encoding.opt-out.windows-1257=true
+encoding.opt-out.windows-1258=true
+encoding.opt-out.windows-31j=true
+encoding.opt-out.x-Big5-Solaris=true
+encoding.opt-out.x-EUC-TW=true
+encoding.opt-out.x-IBM1006=true
+encoding.opt-out.x-IBM1025=true
+encoding.opt-out.x-IBM1046=true
+encoding.opt-out.x-IBM1097=true
+encoding.opt-out.x-IBM1098=true
+encoding.opt-out.x-IBM1112=true
+encoding.opt-out.x-IBM1122=true
+encoding.opt-out.x-IBM1123=true
+encoding.opt-out.x-IBM1124=true
+encoding.opt-out.x-IBM1381=true
+encoding.opt-out.x-IBM1383=true
+encoding.opt-out.x-IBM33722=true
+encoding.opt-out.x-IBM737=true
+encoding.opt-out.x-IBM834=true
+encoding.opt-out.x-IBM856=true
+encoding.opt-out.x-IBM874=true
+encoding.opt-out.x-IBM875=true
+encoding.opt-out.x-IBM921=true
+encoding.opt-out.x-IBM922=true
+encoding.opt-out.x-IBM930=true
+encoding.opt-out.x-IBM933=true
+encoding.opt-out.x-IBM935=true
+encoding.opt-out.x-IBM937=true
+encoding.opt-out.x-IBM939=true
+encoding.opt-out.x-IBM942=true
+encoding.opt-out.x-IBM942C=true
+encoding.opt-out.x-IBM943=true
+encoding.opt-out.x-IBM943C=true
+encoding.opt-out.x-IBM948=true
+encoding.opt-out.x-IBM949=true
+encoding.opt-out.x-IBM949C=true
+encoding.opt-out.x-IBM950=true
+encoding.opt-out.x-IBM964=true
+encoding.opt-out.x-IBM970=true
+encoding.opt-out.x-ISCII91=true
+encoding.opt-out.x-ISO-2022-CN-CNS=true
+encoding.opt-out.x-ISO-2022-CN-GB=true
+encoding.opt-out.x-JIS0208=true
+encoding.opt-out.x-JISAutoDetect=true
+encoding.opt-out.x-Johab=true
+encoding.opt-out.x-MS932_0213=true
+encoding.opt-out.x-MS950-HKSCS=true
+encoding.opt-out.x-MacArabic=true
+encoding.opt-out.x-MacCentralEurope=true
+encoding.opt-out.x-MacCroatian=true
+encoding.opt-out.x-MacCyrillic=true
+encoding.opt-out.x-MacDingbat=true
+encoding.opt-out.x-MacGreek=true
+encoding.opt-out.x-MacHebrew=true
+encoding.opt-out.x-MacIceland=true
+encoding.opt-out.x-MacRoman=true
+encoding.opt-out.x-MacRomania=true
+encoding.opt-out.x-MacSymbol=true
+encoding.opt-out.x-MacThai=true
+encoding.opt-out.x-MacTurkish=true
+encoding.opt-out.x-MacUkraine=true
+encoding.opt-out.x-PCK=true
+encoding.opt-out.x-SJIS_0213=true
+encoding.opt-out.x-UTF-16LE-BOM=true
+encoding.opt-out.x-euc-jp-linux=true
+encoding.opt-out.x-eucJP-Open=true
+encoding.opt-out.x-iso-8859-11=true
+encoding.opt-out.x-mswin-936=true
+encoding.opt-out.x-windows-50220=true
+encoding.opt-out.x-windows-50221=true
+encoding.opt-out.x-windows-874=true
+encoding.opt-out.x-windows-949=true
+encoding.opt-out.x-windows-950=true
+encoding.opt-out.x-windows-iso2022jp=true
+encodingDetectors=BOM XML-PI buffer-local-property
+end.shortcut=
+fallbackEncodings=UTF-8 ISO-8859-15 US-ASCII
+firstTime=false
+home.shortcut=
+insert-newline-indent.shortcut=
+insert-newline.shortcut=ENTER
+isabelle-output.dock-position=bottom
+isabelle-output.height=174
+isabelle-output.width=412
+isabelle-session.dock-position=bottom
+line-end.shortcut=END
+line-home.shortcut=HOME
+lookAndFeel=com.sun.java.swing.plaf.nimbus.NimbusLookAndFeel
+mode.isabelle.customSettings=true
+mode.isabelle.folding=sidekick
+mode.isabelle.sidekick.showStatusWindow.label=true
+print.font=IsabelleText
+restore.remote=false
+restore=false
+sidekick-tree.dock-position=right
+sidekick.buffer-save-parse=true
+sidekick.complete-delay=300
+sidekick.splitter.location=721
+tip.show=false
+twoStageSave=false
+view.antiAlias=standard
+view.blockCaret=true
+view.caretBlink=false
+view.eolMarkers=false
+view.extendedState=0
+view.font=IsabelleText
+view.fontsize=18
+view.fracFontMetrics=false
+view.gutter.fontsize=12
+view.gutter.selectionAreaWidth=18
+view.height=787
+view.middleMousePaste=true
+view.showToolbar=false
+view.thickCaret=true
+view.width=1072
--- a/src/Tools/jEdit/src/jedit/dockable.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/dockable.scala
- Author: Makarius
-
-Generic dockable window.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import java.awt.{Dimension, Component, BorderLayout}
-import javax.swing.JPanel
-
-import org.gjt.sp.jedit.View
-import org.gjt.sp.jedit.gui.DockableWindowManager
-
-
-class Dockable(view: View, position: String) extends JPanel(new BorderLayout)
-{
- if (position == DockableWindowManager.FLOATING)
- setPreferredSize(new Dimension(500, 250))
-
- def set_content(c: Component) { add(c, BorderLayout.CENTER) }
- def set_content(c: scala.swing.Component) { add(c.peer, BorderLayout.CENTER) }
-
- protected def init() { }
- protected def exit() { }
-
- override def addNotify()
- {
- super.addNotify()
- init()
- }
-
- override def removeNotify()
- {
- exit()
- super.removeNotify()
- }
-}
--- a/src/Tools/jEdit/src/jedit/document_model.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,252 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/document_model.scala
- Author: Fabian Immler, TU Munich
- Author: Makarius
-
-Document model connected to jEdit buffer -- single node in theory graph.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.collection.mutable
-
-import org.gjt.sp.jedit.Buffer
-import org.gjt.sp.jedit.buffer.{BufferAdapter, BufferListener, JEditBuffer}
-import org.gjt.sp.jedit.syntax.{SyntaxStyle, Token, TokenMarker, TokenHandler, ParserRuleSet}
-import org.gjt.sp.jedit.textarea.TextArea
-
-import java.awt.font.TextAttribute
-import javax.swing.text.Segment;
-
-
-object Document_Model
-{
- object Token_Markup
- {
- /* extended token styles */
-
- private val plain_range: Int = Token.ID_COUNT
- private val full_range: Int = 3 * plain_range
- private def check_range(i: Int) { require(0 <= i && i < plain_range) }
-
- def subscript(i: Byte): Byte = { check_range(i); (i + plain_range).toByte }
- def superscript(i: Byte): Byte = { check_range(i); (i + 2 * plain_range).toByte }
-
- private def script_style(style: SyntaxStyle, i: Int): SyntaxStyle =
- {
- import scala.collection.JavaConversions._
- val script_font =
- style.getFont.deriveFont(Map(TextAttribute.SUPERSCRIPT -> new java.lang.Integer(i)))
- new SyntaxStyle(style.getForegroundColor, style.getBackgroundColor, script_font)
- }
-
- def extend_styles(styles: Array[SyntaxStyle]): Array[SyntaxStyle] =
- {
- val new_styles = new Array[SyntaxStyle](full_range)
- for (i <- 0 until plain_range) {
- val style = styles(i)
- new_styles(i) = style
- new_styles(subscript(i.toByte)) = script_style(style, -1)
- new_styles(superscript(i.toByte)) = script_style(style, 1)
- }
- new_styles
- }
-
-
- /* line context */
-
- private val dummy_rules = new ParserRuleSet("isabelle", "MAIN")
-
- class Line_Context(val line: Int, prev: Line_Context)
- extends TokenMarker.LineContext(dummy_rules, prev)
- {
- override def hashCode: Int = line
- override def equals(that: Any): Boolean =
- that match {
- case other: Line_Context => line == other.line
- case _ => false
- }
- }
- }
-
-
- /* document model of buffer */
-
- private val key = "isabelle.document_model"
-
- def init(session: Session, buffer: Buffer, thy_name: String): Document_Model =
- {
- Swing_Thread.require()
- val model = new Document_Model(session, buffer, thy_name)
- buffer.setProperty(key, model)
- model.activate()
- model
- }
-
- def apply(buffer: Buffer): Option[Document_Model] =
- {
- Swing_Thread.require()
- buffer.getProperty(key) match {
- case model: Document_Model => Some(model)
- case _ => None
- }
- }
-
- def exit(buffer: Buffer)
- {
- Swing_Thread.require()
- apply(buffer) match {
- case None =>
- case Some(model) =>
- model.deactivate()
- buffer.unsetProperty(key)
- }
- }
-}
-
-
-class Document_Model(val session: Session, val buffer: Buffer, val thy_name: String)
-{
- /* pending text edits */
-
- object pending_edits // owned by Swing thread
- {
- private val pending = new mutable.ListBuffer[Text.Edit]
- def snapshot(): List[Text.Edit] = pending.toList
-
- def flush(more_edits: Option[List[Text.Edit]]*)
- {
- Swing_Thread.require()
- val edits = snapshot()
- pending.clear
-
- val all_edits =
- if (edits.isEmpty) more_edits.toList
- else Some(edits) :: more_edits.toList
- if (!all_edits.isEmpty) session.edit_version(all_edits.map((thy_name, _)))
- }
-
- def init()
- {
- Swing_Thread.require()
- flush(None, Some(List(Text.Edit.insert(0, Isabelle.buffer_text(buffer)))))
- }
-
- private val delay_flush =
- Swing_Thread.delay_last(session.input_delay) { flush() }
-
- def +=(edit: Text.Edit)
- {
- Swing_Thread.require()
- pending += edit
- delay_flush()
- }
- }
-
-
- /* snapshot */
-
- def snapshot(): Document.Snapshot =
- {
- Swing_Thread.require()
- session.snapshot(thy_name, pending_edits.snapshot())
- }
-
-
- /* buffer listener */
-
- private val buffer_listener: BufferListener = new BufferAdapter
- {
- override def bufferLoaded(buffer: JEditBuffer)
- {
- pending_edits.init()
- }
-
- override def contentInserted(buffer: JEditBuffer,
- start_line: Int, offset: Int, num_lines: Int, length: Int)
- {
- if (!buffer.isLoading)
- pending_edits += Text.Edit.insert(offset, buffer.getText(offset, length))
- }
-
- override def preContentRemoved(buffer: JEditBuffer,
- start_line: Int, offset: Int, num_lines: Int, removed_length: Int)
- {
- if (!buffer.isLoading)
- pending_edits += Text.Edit.remove(offset, buffer.getText(offset, removed_length))
- }
- }
-
-
- /* semantic token marker */
-
- private val token_marker = new TokenMarker
- {
- override def markTokens(prev: TokenMarker.LineContext,
- handler: TokenHandler, line_segment: Segment): TokenMarker.LineContext =
- {
- Isabelle.swing_buffer_lock(buffer) {
- val snapshot = Document_Model.this.snapshot()
-
- val previous = prev.asInstanceOf[Document_Model.Token_Markup.Line_Context]
- val line = if (prev == null) 0 else previous.line + 1
- val context = new Document_Model.Token_Markup.Line_Context(line, previous)
-
- val start = buffer.getLineStartOffset(line)
- val stop = start + line_segment.count
-
- /* FIXME
- for (text_area <- Isabelle.jedit_text_areas(buffer)
- if Document_View(text_area).isDefined)
- Document_View(text_area).get.set_styles()
- */
-
- def handle_token(style: Byte, offset: Text.Offset, length: Int) =
- handler.handleToken(line_segment, style, offset, length, context)
-
- val syntax = session.current_syntax()
- val tokens = snapshot.select_markup(Text.Range(start, stop))(Isabelle_Markup.tokens(syntax))
-
- var last = start
- for (token <- tokens.iterator) {
- val Text.Range(token_start, token_stop) = token.range
- if (last < token_start)
- handle_token(Token.COMMENT1, last - start, token_start - last)
- handle_token(token.info getOrElse Token.NULL,
- token_start - start, token_stop - token_start)
- last = token_stop
- }
- if (last < stop) handle_token(Token.COMMENT1, last - start, stop - last)
-
- handle_token(Token.END, line_segment.count, 0)
- handler.setLineContext(context)
- context
- }
- }
- }
-
-
- /* activation */
-
- def activate()
- {
- buffer.setTokenMarker(token_marker)
- buffer.addBufferListener(buffer_listener)
- buffer.propertiesChanged()
- pending_edits.init()
- }
-
- def refresh()
- {
- buffer.setTokenMarker(token_marker)
- }
-
- def deactivate()
- {
- pending_edits.flush()
- buffer.setTokenMarker(buffer.getMode.getTokenMarker)
- buffer.removeBufferListener(buffer_listener)
- }
-}
--- a/src/Tools/jEdit/src/jedit/document_view.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,587 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/document_view.scala
- Author: Fabian Immler, TU Munich
- Author: Makarius
-
-Document view connected to jEdit text area.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.actors.Actor._
-
-import java.awt.{BorderLayout, Graphics, Color, Dimension, Graphics2D, Point}
-import java.awt.event.{MouseAdapter, MouseMotionAdapter, MouseEvent,
- FocusAdapter, FocusEvent, WindowEvent, WindowAdapter}
-import javax.swing.{JPanel, ToolTipManager, Popup, PopupFactory, SwingUtilities, BorderFactory}
-import javax.swing.event.{CaretListener, CaretEvent}
-import java.util.ArrayList
-
-import org.gjt.sp.jedit.{jEdit, OperatingSystem, Debug}
-import org.gjt.sp.jedit.gui.RolloverButton
-import org.gjt.sp.jedit.options.GutterOptionPane
-import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, TextAreaExtension, TextAreaPainter}
-import org.gjt.sp.jedit.syntax.{SyntaxStyle, DisplayTokenHandler, Chunk, Token}
-
-
-object Document_View
-{
- /* document view of text area */
-
- private val key = new Object
-
- def init(model: Document_Model, text_area: JEditTextArea): Document_View =
- {
- Swing_Thread.require()
- val doc_view = new Document_View(model, text_area)
- text_area.putClientProperty(key, doc_view)
- doc_view.activate()
- doc_view
- }
-
- def apply(text_area: JEditTextArea): Option[Document_View] =
- {
- Swing_Thread.require()
- text_area.getClientProperty(key) match {
- case doc_view: Document_View => Some(doc_view)
- case _ => None
- }
- }
-
- def exit(text_area: JEditTextArea)
- {
- Swing_Thread.require()
- apply(text_area) match {
- case None =>
- case Some(doc_view) =>
- doc_view.deactivate()
- text_area.putClientProperty(key, null)
- }
- }
-}
-
-
-class Document_View(val model: Document_Model, text_area: JEditTextArea)
-{
- private val session = model.session
-
-
- /** token handling **/
-
- /* extended token styles */
-
- private var styles: Array[SyntaxStyle] = null // owned by Swing thread
-
- def extend_styles()
- {
- Swing_Thread.require()
- styles = Document_Model.Token_Markup.extend_styles(text_area.getPainter.getStyles)
- }
- extend_styles()
-
- def set_styles()
- {
- Swing_Thread.require()
- text_area.getPainter.setStyles(styles)
- }
-
-
- /* wrap_margin -- cf. org.gjt.sp.jedit.textarea.TextArea.propertiesChanged */
-
- def wrap_margin(): Int =
- {
- val buffer = text_area.getBuffer
- val painter = text_area.getPainter
- val font = painter.getFont
- val font_context = painter.getFontRenderContext
-
- val soft_wrap = buffer.getStringProperty("wrap") == "soft"
- val max = buffer.getIntegerProperty("maxLineLen", 0)
-
- if (max > 0) font.getStringBounds(" " * max, font_context).getWidth.toInt
- else if (soft_wrap)
- painter.getWidth - (font.getStringBounds(" ", font_context).getWidth.round.toInt) * 3
- else 0
- }
-
-
- /* chunks */
-
- def line_chunks(physical_lines: Set[Int]): Map[Text.Offset, Chunk] =
- {
- import scala.collection.JavaConversions._
-
- val buffer = text_area.getBuffer
- val painter = text_area.getPainter
- val margin = wrap_margin().toFloat
-
- val out = new ArrayList[Chunk]
- val handler = new DisplayTokenHandler
-
- var result = Map[Text.Offset, Chunk]()
- for (line <- physical_lines) {
- out.clear
- handler.init(painter.getStyles, painter.getFontRenderContext, painter, out, margin)
- buffer.markTokens(line, handler)
-
- val line_start = buffer.getLineStartOffset(line)
- for (chunk <- handler.getChunkList.iterator)
- result += (line_start + chunk.offset -> chunk)
- }
- result
- }
-
-
- /* visible line ranges */
-
- // simplify slightly odd result of TextArea.getScreenLineEndOffset etc.
- // NB: jEdit already normalizes \r\n and \r to \n
- def proper_line_range(start: Text.Offset, end: Text.Offset): Text.Range =
- {
- val stop = if (start < end) end - 1 else end min model.buffer.getLength
- Text.Range(start, stop)
- }
-
- def screen_lines_range(): Text.Range =
- {
- val start = text_area.getScreenLineStartOffset(0)
- val raw_end = text_area.getScreenLineEndOffset(text_area.getVisibleLines - 1 max 0)
- proper_line_range(start, if (raw_end >= 0) raw_end else model.buffer.getLength)
- }
-
- def invalidate_line_range(range: Text.Range)
- {
- text_area.invalidateLineRange(
- model.buffer.getLineOfOffset(range.start),
- model.buffer.getLineOfOffset(range.stop))
- }
-
-
- /* HTML popups */
-
- private var html_popup: Option[Popup] = None
-
- private def exit_popup() { html_popup.map(_.hide) }
-
- private val html_panel =
- new HTML_Panel(Isabelle.system, Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
- html_panel.setBorder(BorderFactory.createLineBorder(Color.black))
-
- private def html_panel_resize()
- {
- Swing_Thread.now {
- html_panel.resize(Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
- }
- }
-
- private def init_popup(snapshot: Document.Snapshot, x: Int, y: Int)
- {
- exit_popup()
-/* FIXME broken
- val offset = text_area.xyToOffset(x, y)
- val p = new Point(x, y); SwingUtilities.convertPointToScreen(p, text_area.getPainter)
-
- // FIXME snapshot.cumulate
- snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.popup) match {
- case Text.Info(_, Some(msg)) #:: _ =>
- val popup = PopupFactory.getSharedInstance().getPopup(text_area, html_panel, p.x, p.y + 60)
- html_panel.render_sync(List(msg))
- Thread.sleep(10) // FIXME !?
- popup.show
- html_popup = Some(popup)
- case _ =>
- }
-*/
- }
-
-
- /* subexpression highlighting */
-
- private def subexp_range(snapshot: Document.Snapshot, x: Int, y: Int)
- : Option[(Text.Range, Color)] =
- {
- val offset = text_area.xyToOffset(x, y)
- snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.subexp) match {
- case Text.Info(_, Some((range, color))) #:: _ => Some((snapshot.convert(range), color))
- case _ => None
- }
- }
-
- private var highlight_range: Option[(Text.Range, Color)] = None
-
-
- /* CONTROL-mouse management */
-
- private var control: Boolean = false
-
- private def exit_control()
- {
- exit_popup()
- highlight_range = None
- }
-
- private val focus_listener = new FocusAdapter {
- override def focusLost(e: FocusEvent) {
- highlight_range = None // FIXME exit_control !?
- }
- }
-
- private val window_listener = new WindowAdapter {
- override def windowIconified(e: WindowEvent) { exit_control() }
- override def windowDeactivated(e: WindowEvent) { exit_control() }
- }
-
- private val mouse_motion_listener = new MouseMotionAdapter {
- override def mouseMoved(e: MouseEvent) {
- control = if (OperatingSystem.isMacOS()) e.isMetaDown else e.isControlDown
- val x = e.getX()
- val y = e.getY()
-
- if (!model.buffer.isLoaded) exit_control()
- else
- Isabelle.swing_buffer_lock(model.buffer) {
- val snapshot = model.snapshot
-
- if (control) init_popup(snapshot, x, y)
-
- highlight_range map { case (range, _) => invalidate_line_range(range) }
- highlight_range = if (control) subexp_range(snapshot, x, y) else None
- highlight_range map { case (range, _) => invalidate_line_range(range) }
- }
- }
- }
-
-
- /* TextArea painters */
-
- private val background_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int)
- {
- Isabelle.swing_buffer_lock(model.buffer) {
- val snapshot = model.snapshot()
- val ascent = text_area.getPainter.getFontMetrics.getAscent
-
- for (i <- 0 until physical_lines.length) {
- if (physical_lines(i) != -1) {
- val line_range = proper_line_range(start(i), end(i))
-
- // background color: status
- val cmds = snapshot.node.command_range(snapshot.revert(line_range))
- for {
- (command, command_start) <- cmds if !command.is_ignored
- val range = line_range.restrict(snapshot.convert(command.range + command_start))
- r <- Isabelle.gfx_range(text_area, range)
- color <- Isabelle_Markup.status_color(snapshot, command)
- } {
- gfx.setColor(color)
- gfx.fillRect(r.x, y + i * line_height, r.length, line_height)
- }
-
- // background color (1): markup
- for {
- Text.Info(range, Some(color)) <-
- snapshot.select_markup(line_range)(Isabelle_Markup.background1).iterator
- r <- Isabelle.gfx_range(text_area, range)
- } {
- gfx.setColor(color)
- gfx.fillRect(r.x, y + i * line_height, r.length, line_height)
- }
-
- // background color (2): markup
- for {
- Text.Info(range, Some(color)) <-
- snapshot.select_markup(line_range)(Isabelle_Markup.background2).iterator
- r <- Isabelle.gfx_range(text_area, range)
- } {
- gfx.setColor(color)
- gfx.fillRect(r.x + 2, y + i * line_height + 2, r.length - 4, line_height - 4)
- }
-
- // sub-expression highlighting -- potentially from other snapshot
- highlight_range match {
- case Some((range, color)) if line_range.overlaps(range) =>
- Isabelle.gfx_range(text_area, line_range.restrict(range)) match {
- case None =>
- case Some(r) =>
- gfx.setColor(color)
- gfx.drawRect(r.x, y + i * line_height, r.length - 1, line_height - 1)
- }
- case _ =>
- }
-
- // squiggly underline
- for {
- Text.Info(range, Some(color)) <-
- snapshot.select_markup(line_range)(Isabelle_Markup.message).iterator
- r <- Isabelle.gfx_range(text_area, range)
- } {
- gfx.setColor(color)
- val x0 = (r.x / 2) * 2
- val y0 = r.y + ascent + 1
- for (x1 <- Range(x0, x0 + r.length, 2)) {
- val y1 = if (x1 % 4 < 2) y0 else y0 + 1
- gfx.drawLine(x1, y1, x1 + 1, y1)
- }
- }
- }
- }
- }
- }
-
- override def getToolTipText(x: Int, y: Int): String =
- {
- Isabelle.swing_buffer_lock(model.buffer) {
- val snapshot = model.snapshot()
- val offset = text_area.xyToOffset(x, y)
- if (control) {
- snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.tooltip) match
- {
- case Text.Info(_, Some(text)) #:: _ => Isabelle.tooltip(text)
- case _ => null
- }
- }
- else {
- // FIXME snapshot.cumulate
- snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Markup.popup) match
- {
- case Text.Info(_, Some(text)) #:: _ => Isabelle.tooltip(text)
- case _ => null
- }
- }
- }
- }
- }
-
- var use_text_painter = false
-
- private val text_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int)
- {
- if (use_text_painter) {
- Isabelle.swing_buffer_lock(model.buffer) {
- val painter = text_area.getPainter
- val fm = painter.getFontMetrics
-
- val all_chunks = line_chunks(Set[Int]() ++ physical_lines.iterator.filter(i => i != -1))
-
- val x0 = text_area.getHorizontalOffset
- var y0 = y + fm.getHeight - (fm.getLeading + 1) - fm.getDescent
- for (i <- 0 until physical_lines.length) {
- if (physical_lines(i) != -1) {
- all_chunks.get(start(i)) match {
- case Some(chunk) =>
- Chunk.paintChunkList(chunk, gfx, x0, y0, !Debug.DISABLE_GLYPH_VECTOR)
- case None =>
- }
- }
- y0 += line_height
- }
- }
- }
- }
- }
-
- private val gutter_painter = new TextAreaExtension
- {
- override def paintScreenLineRange(gfx: Graphics2D,
- first_line: Int, last_line: Int, physical_lines: Array[Int],
- start: Array[Int], end: Array[Int], y: Int, line_height: Int)
- {
- val gutter = text_area.getGutter
- val width = GutterOptionPane.getSelectionAreaWidth
- val border_width = jEdit.getIntegerProperty("view.gutter.borderWidth", 3)
- val FOLD_MARKER_SIZE = 12
-
- if (gutter.isSelectionAreaEnabled && !gutter.isExpanded && width >= 12 && line_height >= 12) {
- Isabelle.swing_buffer_lock(model.buffer) {
- val snapshot = model.snapshot()
- for (i <- 0 until physical_lines.length) {
- if (physical_lines(i) != -1) {
- val line_range = proper_line_range(start(i), end(i))
-
- // gutter icons
- val icons =
- (for (Text.Info(_, Some(icon)) <- // FIXME snapshot.cumulate
- snapshot.select_markup(line_range)(Isabelle_Markup.gutter_message).iterator)
- yield icon).toList.sortWith(_ >= _)
- icons match {
- case icon :: _ =>
- val icn = icon.icon
- val x0 = (FOLD_MARKER_SIZE + width - border_width - icn.getIconWidth) max 10
- val y0 = y + i * line_height + (((line_height - icn.getIconHeight) / 2) max 0)
- icn.paintIcon(gutter, gfx, x0, y0)
- case Nil =>
- }
- }
- }
- }
- }
- }
- }
-
-
- /* caret handling */
-
- def selected_command(): Option[Command] =
- {
- Swing_Thread.require()
- model.snapshot().node.proper_command_at(text_area.getCaretPosition)
- }
-
- private val caret_listener = new CaretListener {
- private val delay = Swing_Thread.delay_last(session.input_delay) {
- session.perspective.event(Session.Perspective)
- }
- override def caretUpdate(e: CaretEvent) { delay() }
- }
-
-
- /* overview of command status left of scrollbar */
-
- private val overview = new JPanel(new BorderLayout)
- {
- private val WIDTH = 10
- private val HEIGHT = 2
-
- setPreferredSize(new Dimension(WIDTH, 0))
-
- setRequestFocusEnabled(false)
-
- addMouseListener(new MouseAdapter {
- override def mousePressed(event: MouseEvent) {
- val line = y_to_line(event.getY)
- if (line >= 0 && line < text_area.getLineCount)
- text_area.setCaretPosition(text_area.getLineStartOffset(line))
- }
- })
-
- override def addNotify() {
- super.addNotify()
- ToolTipManager.sharedInstance.registerComponent(this)
- }
-
- override def removeNotify() {
- ToolTipManager.sharedInstance.unregisterComponent(this)
- super.removeNotify
- }
-
- override def paintComponent(gfx: Graphics)
- {
- super.paintComponent(gfx)
- Swing_Thread.assert()
-
- val buffer = model.buffer
- Isabelle.buffer_lock(buffer) {
- val snapshot = model.snapshot()
-
- def line_range(command: Command, start: Text.Offset): Option[(Int, Int)] =
- {
- try {
- val line1 = buffer.getLineOfOffset(snapshot.convert(start))
- val line2 = buffer.getLineOfOffset(snapshot.convert(start + command.length)) + 1
- Some((line1, line2))
- }
- catch { case e: ArrayIndexOutOfBoundsException => None }
- }
- for {
- (command, start) <- snapshot.node.command_starts
- if !command.is_ignored
- (line1, line2) <- line_range(command, start)
- val y = line_to_y(line1)
- val height = HEIGHT * (line2 - line1)
- color <- Isabelle_Markup.overview_color(snapshot, command)
- } {
- gfx.setColor(color)
- gfx.fillRect(0, y, getWidth - 1, height)
- }
- }
- }
-
- private def line_to_y(line: Int): Int =
- (line * getHeight) / (text_area.getBuffer.getLineCount max text_area.getVisibleLines)
-
- private def y_to_line(y: Int): Int =
- (y * (text_area.getBuffer.getLineCount max text_area.getVisibleLines)) / getHeight
- }
-
-
- /* main actor */
-
- private val main_actor = actor {
- loop {
- react {
- case Session.Commands_Changed(changed) =>
- val buffer = model.buffer
- Isabelle.swing_buffer_lock(buffer) {
- val snapshot = model.snapshot()
-
- if (changed.exists(snapshot.node.commands.contains))
- overview.repaint()
-
- val visible_range = screen_lines_range()
- val visible_cmds = snapshot.node.command_range(snapshot.revert(visible_range)).map(_._1)
- if (visible_cmds.exists(changed)) {
- for {
- line <- 0 until text_area.getVisibleLines
- val start = text_area.getScreenLineStartOffset(line) if start >= 0
- val end = text_area.getScreenLineEndOffset(line) if end >= 0
- val range = proper_line_range(start, end)
- val line_cmds = snapshot.node.command_range(snapshot.revert(range)).map(_._1)
- if line_cmds.exists(changed)
- } text_area.invalidateScreenLineRange(line, line)
-
- // FIXME danger of deadlock!?
- // FIXME potentially slow!?
- model.buffer.propertiesChanged()
- }
- }
-
- case Session.Global_Settings => html_panel_resize()
-
- case bad => System.err.println("command_change_actor: ignoring bad message " + bad)
- }
- }
- }
-
-
- /* activation */
-
- private def activate()
- {
- val painter = text_area.getPainter
- painter.addExtension(TextAreaPainter.LINE_BACKGROUND_LAYER + 1, background_painter)
- painter.addExtension(TextAreaPainter.TEXT_LAYER + 1, text_painter)
- text_area.getGutter.addExtension(gutter_painter)
- text_area.addFocusListener(focus_listener)
- text_area.getView.addWindowListener(window_listener)
- painter.addMouseMotionListener(mouse_motion_listener)
- text_area.addCaretListener(caret_listener)
- text_area.addLeftOfScrollBar(overview)
- session.commands_changed += main_actor
- session.global_settings += main_actor
- }
-
- private def deactivate()
- {
- val painter = text_area.getPainter
- session.commands_changed -= main_actor
- session.global_settings -= main_actor
- text_area.removeFocusListener(focus_listener)
- text_area.getView.removeWindowListener(window_listener)
- painter.removeMouseMotionListener(mouse_motion_listener)
- text_area.removeCaretListener(caret_listener)
- text_area.removeLeftOfScrollBar(overview)
- text_area.getGutter.removeExtension(gutter_painter)
- painter.removeExtension(text_painter)
- painter.removeExtension(background_painter)
- exit_popup()
- }
-}
--- a/src/Tools/jEdit/src/jedit/html_panel.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,206 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/html_panel.scala
- Author: Makarius
-
-HTML panel based on Lobo/Cobra.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import java.io.StringReader
-import java.awt.{Font, BorderLayout, Dimension, GraphicsEnvironment, Toolkit, FontMetrics}
-import java.awt.event.MouseEvent
-
-import java.util.logging.{Logger, Level}
-
-import org.w3c.dom.html2.HTMLElement
-
-import org.lobobrowser.html.parser.{DocumentBuilderImpl, InputSourceImpl}
-import org.lobobrowser.html.gui.HtmlPanel
-import org.lobobrowser.html.domimpl.{HTMLDocumentImpl, HTMLStyleElementImpl, NodeImpl}
-import org.lobobrowser.html.test.{SimpleHtmlRendererContext, SimpleUserAgentContext}
-
-import scala.actors.Actor._
-
-
-object HTML_Panel
-{
- sealed abstract class Event { val element: HTMLElement; val mouse: MouseEvent }
- case class Context_Menu(val element: HTMLElement, mouse: MouseEvent) extends Event
- case class Mouse_Click(val element: HTMLElement, mouse: MouseEvent) extends Event
- case class Double_Click(val element: HTMLElement, mouse: MouseEvent) extends Event
- case class Mouse_Over(val element: HTMLElement, mouse: MouseEvent) extends Event
- case class Mouse_Out(val element: HTMLElement, mouse: MouseEvent) extends Event
-}
-
-
-class HTML_Panel(
- system: Isabelle_System,
- initial_font_family: String,
- initial_font_size: Int)
- extends HtmlPanel
-{
- /** Lobo setup **/
-
- /* global logging */
-
- Logger.getLogger("org.lobobrowser").setLevel(Level.WARNING)
-
-
- /* pixel size -- cf. org.lobobrowser.html.style.HtmlValues.getFontSize */
-
- val screen_resolution =
- if (GraphicsEnvironment.isHeadless()) 72
- else Toolkit.getDefaultToolkit().getScreenResolution()
-
- def lobo_px(raw_px: Int): Int = raw_px * 96 / screen_resolution
- def raw_px(lobo_px: Int): Int = (lobo_px * screen_resolution + 95) / 96
-
-
- /* contexts and event handling */
-
- protected val handler: PartialFunction[HTML_Panel.Event, Unit] = Library.undefined
-
- private val ucontext = new SimpleUserAgentContext
- private val rcontext = new SimpleHtmlRendererContext(this, ucontext)
- {
- private def handle(event: HTML_Panel.Event): Boolean =
- if (handler.isDefinedAt(event)) { handler(event); false }
- else true
-
- override def onContextMenu(elem: HTMLElement, event: MouseEvent): Boolean =
- handle(HTML_Panel.Context_Menu(elem, event))
- override def onMouseClick(elem: HTMLElement, event: MouseEvent): Boolean =
- handle(HTML_Panel.Mouse_Click(elem, event))
- override def onDoubleClick(elem: HTMLElement, event: MouseEvent): Boolean =
- handle(HTML_Panel.Double_Click(elem, event))
- override def onMouseOver(elem: HTMLElement, event: MouseEvent)
- { handle(HTML_Panel.Mouse_Over(elem, event)) }
- override def onMouseOut(elem: HTMLElement, event: MouseEvent)
- { handle(HTML_Panel.Mouse_Out(elem, event)) }
- }
-
- private val builder = new DocumentBuilderImpl(ucontext, rcontext)
-
-
- /* document template with style sheets */
-
- private val template_head =
- """<?xml version="1.0" encoding="utf-8"?>
-<!DOCTYPE html PUBLIC "-//W3C//DTD XHTML 1.0 Transitional//EN"
- "http://www.w3.org/TR/xhtml1/DTD/xhtml1-transitional.dtd">
-<html xmlns="http://www.w3.org/1999/xhtml">
-<head>
-<style media="all" type="text/css">
-""" +
- system.try_read(system.getenv_strict("JEDIT_STYLE_SHEETS").split(":"))
-
- private val template_tail =
-"""
-</style>
-</head>
-<body/>
-</html>
-"""
-
- private def template(font_family: String, font_size: Int): String =
- template_head +
- "body { font-family: " + font_family + "; font-size: " + raw_px(font_size) + "px; }" +
- template_tail
-
-
- /** main actor **/
-
- /* internal messages */
-
- private case class Resize(font_family: String, font_size: Int)
- private case class Render_Document(text: String)
- private case class Render(body: XML.Body)
- private case class Render_Sync(body: XML.Body)
- private case object Refresh
-
- private val main_actor = actor {
-
- /* internal state */
-
- var current_font_metrics: FontMetrics = null
- var current_font_family = ""
- var current_font_size: Int = 0
- var current_margin: Int = 0
- var current_body: XML.Body = Nil
-
- def resize(font_family: String, font_size: Int)
- {
- val font = new Font(font_family, Font.PLAIN, lobo_px(raw_px(font_size)))
- val (font_metrics, margin) =
- Swing_Thread.now {
- val metrics = getFontMetrics(font)
- (metrics, (getWidth() / (metrics.charWidth(Symbol.spc) max 1) - 4) max 20)
- }
- if (current_font_metrics == null ||
- current_font_family != font_family ||
- current_font_size != font_size ||
- current_margin != margin)
- {
- current_font_metrics = font_metrics
- current_font_family = font_family
- current_font_size = font_size
- current_margin = margin
- refresh()
- }
- }
-
- def refresh() { render(current_body) }
-
- def render_document(text: String)
- {
- val doc = builder.parse(new InputSourceImpl(new StringReader(text), "http://localhost"))
- Swing_Thread.later { setDocument(doc, rcontext) }
- }
-
- def render(body: XML.Body)
- {
- current_body = body
- val html_body =
- current_body.flatMap(div =>
- Pretty.formatted(List(div), current_margin, Pretty.font_metric(current_font_metrics))
- .map(t =>
- XML.Elem(Markup(HTML.PRE, List((Markup.CLASS, Markup.MESSAGE))),
- HTML.spans(t, true))))
- val doc =
- builder.parse(
- new InputSourceImpl(
- new StringReader(template(current_font_family, current_font_size)), "http://localhost"))
- doc.removeChild(doc.getLastChild())
- doc.appendChild(XML.document_node(doc, XML.elem(HTML.BODY, html_body)))
- Swing_Thread.later { setDocument(doc, rcontext) }
- }
-
-
- /* main loop */
-
- resize(initial_font_family, initial_font_size)
-
- loop {
- react {
- case Resize(font_family, font_size) => resize(font_family, font_size)
- case Refresh => refresh()
- case Render_Document(text) => render_document(text)
- case Render(body) => render(body)
- case Render_Sync(body) => render(body); reply(())
- case bad => System.err.println("main_actor: ignoring bad message " + bad)
- }
- }
- }
-
-
- /* external methods */
-
- def resize(font_family: String, font_size: Int) { main_actor ! Resize(font_family, font_size) }
- def refresh() { main_actor ! Refresh }
- def render_document(text: String) { main_actor ! Render_Document(text) }
- def render(body: XML.Body) { main_actor ! Render(body) }
- def render_sync(body: XML.Body) { main_actor !? Render_Sync(body) }
-}
--- a/src/Tools/jEdit/src/jedit/isabelle_encoding.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,65 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/isabelle_encoding.scala
- Author: Makarius
-
-Isabelle encoding -- based on UTF-8.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import org.gjt.sp.jedit.io.Encoding
-import org.gjt.sp.jedit.buffer.JEditBuffer
-
-import java.nio.charset.{Charset, CodingErrorAction}
-import java.io.{InputStream, OutputStream, Reader, Writer, InputStreamReader, OutputStreamWriter,
- CharArrayReader, ByteArrayOutputStream}
-
-import scala.io.{Codec, Source, BufferedSource}
-
-
-object Isabelle_Encoding
-{
- val NAME = "UTF-8-Isabelle"
-
- def is_active(buffer: JEditBuffer): Boolean =
- buffer.getStringProperty(JEditBuffer.ENCODING).asInstanceOf[String] == NAME
-}
-
-class Isabelle_Encoding extends Encoding
-{
- private val charset = Charset.forName(Standard_System.charset)
- val BUFSIZE = 32768
-
- private def text_reader(in: InputStream, codec: Codec): Reader =
- {
- val source = new BufferedSource(in)(codec)
- new CharArrayReader(Isabelle.system.symbols.decode(source.mkString).toArray)
- }
-
- override def getTextReader(in: InputStream): Reader =
- text_reader(in, Standard_System.codec())
-
- override def getPermissiveTextReader(in: InputStream): Reader =
- {
- val codec = Standard_System.codec()
- codec.onMalformedInput(CodingErrorAction.REPLACE)
- codec.onUnmappableCharacter(CodingErrorAction.REPLACE)
- text_reader(in, codec)
- }
-
- override def getTextWriter(out: OutputStream): Writer =
- {
- val buffer = new ByteArrayOutputStream(BUFSIZE) {
- override def flush()
- {
- val text = Isabelle.system.symbols.encode(toString(Standard_System.charset))
- out.write(text.getBytes(Standard_System.charset))
- out.flush()
- }
- override def close() { out.close() }
- }
- new OutputStreamWriter(buffer, charset.newEncoder())
- }
-}
--- a/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,88 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/isabelle_hyperlinks.scala
- Author: Fabian Immler, TU Munich
-
-Hyperlink setup for Isabelle proof documents.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import java.io.File
-
-import gatchan.jedit.hyperlinks.{Hyperlink, HyperlinkSource, AbstractHyperlink}
-
-import org.gjt.sp.jedit.{View, jEdit, Buffer, TextUtilities}
-
-
-private class Internal_Hyperlink(start: Int, end: Int, line: Int, def_offset: Int)
- extends AbstractHyperlink(start, end, line, "")
-{
- override def click(view: View) {
- view.getTextArea.moveCaretPosition(def_offset)
- }
-}
-
-class External_Hyperlink(start: Int, end: Int, line: Int, def_file: String, def_line: Int)
- extends AbstractHyperlink(start, end, line, "")
-{
- override def click(view: View) = {
- Isabelle.system.source_file(def_file) match {
- case None =>
- Library.error_dialog(view, "File not found", "Could not find source file " + def_file)
- case Some(file) =>
- jEdit.openFiles(view, file.getParent, Array(file.getName, "+line:" + def_line))
- }
- }
-}
-
-class Isabelle_Hyperlinks extends HyperlinkSource
-{
- def getHyperlink(buffer: Buffer, buffer_offset: Int): Hyperlink =
- {
- Swing_Thread.assert()
- Isabelle.buffer_lock(buffer) {
- Document_Model(buffer) match {
- case Some(model) =>
- val snapshot = model.snapshot()
- val markup =
- snapshot.select_markup(Text.Range(buffer_offset, buffer_offset + 1)) {
- // FIXME Isar_Document.Hyperlink extractor
- case Text.Info(info_range,
- XML.Elem(Markup(Markup.ENTITY, props), _))
- if (props.find(
- { case (Markup.KIND, Markup.ML_OPEN) => true
- case (Markup.KIND, Markup.ML_STRUCT) => true
- case _ => false }).isEmpty) =>
- val Text.Range(begin, end) = info_range
- val line = buffer.getLineOfOffset(begin)
- (Position.File.unapply(props), Position.Line.unapply(props)) match {
- case (Some(def_file), Some(def_line)) =>
- new External_Hyperlink(begin, end, line, def_file, def_line)
- case _ =>
- (props, props) match {
- case (Position.Id(def_id), Position.Offset(def_offset)) =>
- snapshot.lookup_command(def_id) match {
- case Some(def_cmd) =>
- snapshot.node.command_start(def_cmd) match {
- case Some(def_cmd_start) =>
- new Internal_Hyperlink(begin, end, line,
- snapshot.convert(def_cmd_start + def_cmd.decode(def_offset)))
- case None => null
- }
- case None => null
- }
- case _ => null
- }
- }
- }
- markup match {
- case Text.Info(_, Some(link)) #:: _ => link
- case _ => null
- }
- case None => null
- }
- }
- }
-}
--- a/src/Tools/jEdit/src/jedit/isabelle_markup.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,216 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/isabelle_markup.scala
- Author: Makarius
-
-Isabelle specific physical rendering and markup selection.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import java.awt.Color
-
-import org.gjt.sp.jedit.syntax.Token
-
-
-object Isabelle_Markup
-{
- /* physical rendering */
-
- val outdated_color = new Color(240, 240, 240)
- val unfinished_color = new Color(255, 228, 225)
-
- val light_color = new Color(240, 240, 240)
- val regular_color = new Color(192, 192, 192)
- val warning_color = new Color(255, 140, 0)
- val error_color = new Color(178, 34, 34)
- val bad_color = new Color(255, 106, 106, 100)
- val hilite_color = new Color(255, 204, 102, 100)
-
- class Icon(val priority: Int, val icon: javax.swing.Icon)
- {
- def >= (that: Icon): Boolean = this.priority >= that.priority
- }
- val warning_icon = new Icon(1, Isabelle.load_icon("16x16/status/dialog-warning.png"))
- val error_icon = new Icon(2, Isabelle.load_icon("16x16/status/dialog-error.png"))
-
-
- /* command status */
-
- def status_color(snapshot: Document.Snapshot, command: Command): Option[Color] =
- {
- val state = snapshot.state(command)
- if (snapshot.is_outdated) Some(outdated_color)
- else
- Isar_Document.command_status(state.status) match {
- case Isar_Document.Forked(i) if i > 0 => Some(unfinished_color)
- case Isar_Document.Unprocessed => Some(unfinished_color)
- case _ => None
- }
- }
-
- def overview_color(snapshot: Document.Snapshot, command: Command): Option[Color] =
- {
- val state = snapshot.state(command)
- if (snapshot.is_outdated) None
- else
- Isar_Document.command_status(state.status) match {
- case Isar_Document.Forked(i) => if (i > 0) Some(unfinished_color) else None
- case Isar_Document.Unprocessed => Some(unfinished_color)
- case Isar_Document.Failed => Some(error_color)
- case Isar_Document.Finished =>
- if (state.results.exists(r => Isar_Document.is_error(r._2))) Some(error_color)
- else if (state.results.exists(r => Isar_Document.is_warning(r._2))) Some(warning_color)
- else None
- }
- }
-
-
- /* markup selectors */
-
- val message: Markup_Tree.Select[Color] =
- {
- case Text.Info(_, XML.Elem(Markup(Markup.WRITELN, _), _)) => regular_color
- case Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), _)) => warning_color
- case Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _)) => error_color
- }
-
- val popup: Markup_Tree.Select[String] =
- {
- case Text.Info(_, msg @ XML.Elem(Markup(markup, _), _))
- if markup == Markup.WRITELN || markup == Markup.WARNING || markup == Markup.ERROR =>
- Pretty.string_of(List(msg), margin = Isabelle.Int_Property("tooltip-margin"))
- }
-
- val gutter_message: Markup_Tree.Select[Icon] =
- {
- case Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), _)) => warning_icon
- case Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _)) => error_icon
- }
-
- val background1: Markup_Tree.Select[Color] =
- {
- case Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _)) => bad_color
- case Text.Info(_, XML.Elem(Markup(Markup.HILITE, _), _)) => hilite_color
- }
-
- val background2: Markup_Tree.Select[Color] =
- {
- case Text.Info(_, XML.Elem(Markup(Markup.TOKEN_RANGE, _), _)) => light_color
- }
-
- val tooltip: Markup_Tree.Select[String] =
- {
- case Text.Info(_, XML.Elem(Markup.Entity(kind, name), _)) => kind + " \"" + name + "\""
- case Text.Info(_, XML.Elem(Markup(Markup.ML_TYPING, _), body)) =>
- Pretty.string_of(List(Pretty.block(XML.Text("ML:") :: Pretty.Break(1) :: body)),
- margin = Isabelle.Int_Property("tooltip-margin"))
- case Text.Info(_, XML.Elem(Markup(Markup.SORT, _), _)) => "sort"
- case Text.Info(_, XML.Elem(Markup(Markup.TYP, _), _)) => "type"
- case Text.Info(_, XML.Elem(Markup(Markup.TERM, _), _)) => "term"
- case Text.Info(_, XML.Elem(Markup(Markup.PROP, _), _)) => "proposition"
- case Text.Info(_, XML.Elem(Markup(Markup.TOKEN_RANGE, _), _)) => "inner syntax token"
- case Text.Info(_, XML.Elem(Markup(Markup.FREE, _), _)) => "free variable (globally fixed)"
- case Text.Info(_, XML.Elem(Markup(Markup.SKOLEM, _), _)) => "skolem variable (locally fixed)"
- case Text.Info(_, XML.Elem(Markup(Markup.BOUND, _), _)) => "bound variable (internally fixed)"
- case Text.Info(_, XML.Elem(Markup(Markup.VAR, _), _)) => "schematic variable"
- case Text.Info(_, XML.Elem(Markup(Markup.TFREE, _), _)) => "free type variable"
- case Text.Info(_, XML.Elem(Markup(Markup.TVAR, _), _)) => "schematic type variable"
- }
-
- private val subexp_include =
- Set(Markup.SORT, Markup.TYP, Markup.TERM, Markup.PROP, Markup.ML_TYPING, Markup.TOKEN_RANGE,
- Markup.ENTITY, Markup.FREE, Markup.SKOLEM, Markup.BOUND, Markup.VAR,
- Markup.TFREE, Markup.TVAR)
-
- val subexp: Markup_Tree.Select[(Text.Range, Color)] =
- {
- case Text.Info(range, XML.Elem(Markup(name, _), _)) if subexp_include(name) =>
- (range, Color.black)
- }
-
-
- /* token markup -- text styles */
-
- private val command_style: Map[String, Byte] =
- {
- import Token._
- Map[String, Byte](
- Keyword.THY_END -> KEYWORD2,
- Keyword.THY_SCRIPT -> LABEL,
- Keyword.PRF_SCRIPT -> LABEL,
- Keyword.PRF_ASM -> KEYWORD3,
- Keyword.PRF_ASM_GOAL -> KEYWORD3
- ).withDefaultValue(KEYWORD1)
- }
-
- private val token_style: Map[String, Byte] =
- {
- import Token._
- Map[String, Byte](
- // logical entities
- Markup.TCLASS -> NULL,
- Markup.TYCON -> NULL,
- Markup.FIXED -> NULL,
- Markup.CONST -> LITERAL2,
- Markup.DYNAMIC_FACT -> LABEL,
- // inner syntax
- Markup.TFREE -> NULL,
- Markup.FREE -> MARKUP,
- Markup.TVAR -> NULL,
- Markup.SKOLEM -> COMMENT2,
- Markup.BOUND -> LABEL,
- Markup.VAR -> NULL,
- Markup.NUM -> DIGIT,
- Markup.FLOAT -> DIGIT,
- Markup.XNUM -> DIGIT,
- Markup.XSTR -> LITERAL4,
- Markup.LITERAL -> OPERATOR,
- Markup.INNER_COMMENT -> COMMENT1,
- Markup.SORT -> NULL,
- Markup.TYP -> NULL,
- Markup.TERM -> NULL,
- Markup.PROP -> NULL,
- // ML syntax
- Markup.ML_KEYWORD -> KEYWORD1,
- Markup.ML_DELIMITER -> OPERATOR,
- Markup.ML_IDENT -> NULL,
- Markup.ML_TVAR -> NULL,
- Markup.ML_NUMERAL -> DIGIT,
- Markup.ML_CHAR -> LITERAL1,
- Markup.ML_STRING -> LITERAL1,
- Markup.ML_COMMENT -> COMMENT1,
- Markup.ML_MALFORMED -> INVALID,
- // embedded source text
- Markup.ML_SOURCE -> COMMENT3,
- Markup.DOC_SOURCE -> COMMENT3,
- Markup.ANTIQ -> COMMENT4,
- Markup.ML_ANTIQ -> COMMENT4,
- Markup.DOC_ANTIQ -> COMMENT4,
- // outer syntax
- Markup.KEYWORD -> KEYWORD2,
- Markup.OPERATOR -> OPERATOR,
- Markup.COMMAND -> KEYWORD1,
- Markup.IDENT -> NULL,
- Markup.VERBATIM -> COMMENT3,
- Markup.COMMENT -> COMMENT1,
- Markup.CONTROL -> COMMENT3,
- Markup.MALFORMED -> INVALID,
- Markup.STRING -> LITERAL3,
- Markup.ALTSTRING -> LITERAL1
- ).withDefaultValue(NULL)
- }
-
- def tokens(syntax: Outer_Syntax): Markup_Tree.Select[Byte] =
- {
- case Text.Info(_, XML.Elem(Markup(Markup.COMMAND, List((Markup.NAME, name))), _))
- if syntax.keyword_kind(name).isDefined => command_style(syntax.keyword_kind(name).get)
-
- case Text.Info(_, XML.Elem(Markup(Markup.ENTITY, Markup.Kind(kind)), _))
- if token_style(kind) != Token.NULL => token_style(kind)
-
- case Text.Info(_, XML.Elem(Markup(name, _), _))
- if token_style(name) != Token.NULL => token_style(name)
- }
-}
--- a/src/Tools/jEdit/src/jedit/isabelle_options.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/isabelle_options.scala
- Author: Johannes Hölzl, TU Munich
-
-Editor pane for plugin options.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import javax.swing.JSpinner
-
-import scala.swing.CheckBox
-
-import org.gjt.sp.jedit.AbstractOptionPane
-
-
-class Isabelle_Options extends AbstractOptionPane("isabelle")
-{
- private val logic_selector = Isabelle.logic_selector(Isabelle.Property("logic"))
- private val auto_start = new CheckBox()
- private val relative_font_size = new JSpinner()
- private val tooltip_font_size = new JSpinner()
- private val tooltip_margin = new JSpinner()
- private val tooltip_dismiss_delay = new JSpinner()
-
- override def _init()
- {
- addComponent(Isabelle.Property("logic.title"), logic_selector.peer)
-
- addComponent(Isabelle.Property("auto-start.title"), auto_start.peer)
- auto_start.selected = Isabelle.Boolean_Property("auto-start")
-
- relative_font_size.setValue(Isabelle.Int_Property("relative-font-size", 100))
- addComponent(Isabelle.Property("relative-font-size.title"), relative_font_size)
-
- tooltip_font_size.setValue(Isabelle.Int_Property("tooltip-font-size", 10))
- addComponent(Isabelle.Property("tooltip-font-size.title"), tooltip_font_size)
-
- tooltip_margin.setValue(Isabelle.Int_Property("tooltip-margin", 40))
- addComponent(Isabelle.Property("tooltip-margin.title"), tooltip_margin)
-
- tooltip_dismiss_delay.setValue(
- Isabelle.Time_Property("tooltip-dismiss-delay", Time.seconds(8.0)).ms.toInt)
- addComponent(Isabelle.Property("tooltip-dismiss-delay.title"), tooltip_dismiss_delay)
- }
-
- override def _save()
- {
- Isabelle.Property("logic") = logic_selector.selection.item.name
-
- Isabelle.Boolean_Property("auto-start") = auto_start.selected
-
- Isabelle.Int_Property("relative-font-size") =
- relative_font_size.getValue().asInstanceOf[Int]
-
- Isabelle.Int_Property("tooltip-font-size") =
- tooltip_font_size.getValue().asInstanceOf[Int]
-
- Isabelle.Int_Property("tooltip-margin") =
- tooltip_margin.getValue().asInstanceOf[Int]
-
- Isabelle.Time_Property("tooltip-dismiss-delay") =
- Time.seconds(tooltip_dismiss_delay.getValue().asInstanceOf[Int])
- }
-}
--- a/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,176 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/isabelle_sidekick.scala
- Author: Fabian Immler, TU Munich
- Author: Makarius
-
-SideKick parsers for Isabelle proof documents.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.collection.Set
-import scala.collection.immutable.TreeSet
-
-import javax.swing.tree.DefaultMutableTreeNode
-import javax.swing.text.Position
-import javax.swing.Icon
-
-import org.gjt.sp.jedit.{Buffer, EditPane, TextUtilities, View}
-import errorlist.DefaultErrorSource
-import sidekick.{SideKickParser, SideKickParsedData, SideKickCompletion, IAsset}
-
-
-object Isabelle_Sidekick
-{
- def int_to_pos(offset: Text.Offset): Position =
- new Position { def getOffset = offset; override def toString = offset.toString }
-
- class Asset(name: String, start: Text.Offset, end: Text.Offset) extends IAsset
- {
- protected var _name = name
- protected var _start = int_to_pos(start)
- protected var _end = int_to_pos(end)
- override def getIcon: Icon = null
- override def getShortString: String = _name
- override def getLongString: String = _name
- override def getName: String = _name
- override def setName(name: String) = _name = name
- override def getStart: Position = _start
- override def setStart(start: Position) = _start = start
- override def getEnd: Position = _end
- override def setEnd(end: Position) = _end = end
- override def toString = _name
- }
-}
-
-
-abstract class Isabelle_Sidekick(name: String) extends SideKickParser(name)
-{
- /* parsing */
-
- @volatile protected var stopped = false
- override def stop() = { stopped = true }
-
- def parser(data: SideKickParsedData, model: Document_Model): Unit
-
- def parse(buffer: Buffer, error_source: DefaultErrorSource): SideKickParsedData =
- {
- stopped = false
-
- // FIXME lock buffer (!??)
- val data = new SideKickParsedData(buffer.getName)
- val root = data.root
- data.getAsset(root).setEnd(Isabelle_Sidekick.int_to_pos(buffer.getLength))
-
- Swing_Thread.now { Document_Model(buffer) } match {
- case Some(model) =>
- parser(data, model)
- if (stopped) root.add(new DefaultMutableTreeNode("<parser stopped>"))
- case None => root.add(new DefaultMutableTreeNode("<buffer inactive>"))
- }
- data
- }
-
-
- /* completion */
-
- override def supportsCompletion = true
- override def canCompleteAnywhere = true
-
- override def complete(pane: EditPane, caret: Text.Offset): SideKickCompletion =
- {
- val buffer = pane.getBuffer
- Isabelle.swing_buffer_lock(buffer) {
- Document_Model(buffer) match {
- case None => null
- case Some(model) =>
- val line = buffer.getLineOfOffset(caret)
- val start = buffer.getLineStartOffset(line)
- val text = buffer.getSegment(start, caret - start)
-
- val completion = model.session.current_syntax().completion
- completion.complete(text) match {
- case None => null
- case Some((word, cs)) =>
- val ds =
- (if (Isabelle_Encoding.is_active(buffer))
- cs.map(Isabelle.system.symbols.decode(_)).sortWith(_ < _)
- else cs).filter(_ != word)
- if (ds.isEmpty) null
- else new SideKickCompletion(
- pane.getView, word, ds.toArray.asInstanceOf[Array[Object]]) { }
- }
- }
- }
- }
-}
-
-
-class Isabelle_Sidekick_Default extends Isabelle_Sidekick("isabelle")
-{
- import Thy_Syntax.Structure
-
- def parser(data: SideKickParsedData, model: Document_Model)
- {
- val syntax = model.session.current_syntax()
-
- def make_tree(offset: Text.Offset, entry: Structure.Entry): List[DefaultMutableTreeNode] =
- entry match {
- case Structure.Block(name, body) =>
- val node =
- new DefaultMutableTreeNode(
- new Isabelle_Sidekick.Asset(Library.first_line(name), offset, offset + entry.length))
- (offset /: body)((i, e) =>
- {
- make_tree(i, e) foreach (nd => node.add(nd))
- i + e.length
- })
- List(node)
- case Structure.Atom(command)
- if command.is_command && syntax.heading_level(command).isEmpty =>
- val node =
- new DefaultMutableTreeNode(
- new Isabelle_Sidekick.Asset(command.name, offset, offset + entry.length))
- List(node)
- case _ => Nil
- }
-
- val text = Isabelle.buffer_text(model.buffer)
- val structure = Structure.parse(syntax, "theory " + model.thy_name, text)
-
- make_tree(0, structure) foreach (node => data.root.add(node))
- }
-}
-
-
-class Isabelle_Sidekick_Raw extends Isabelle_Sidekick("isabelle-raw")
-{
- def parser(data: SideKickParsedData, model: Document_Model)
- {
- val root = data.root
- val snapshot = Swing_Thread.now { model.snapshot() } // FIXME cover all nodes (!??)
- for ((command, command_start) <- snapshot.node.command_range() if !stopped) {
- snapshot.state(command).root_markup.swing_tree(root)((info: Text.Info[Any]) =>
- {
- val range = info.range + command_start
- val content = command.source(info.range).replace('\n', ' ')
- val info_text =
- info.info match {
- case elem @ XML.Elem(_, _) =>
- Pretty.formatted(List(elem), margin = 40).mkString("\n")
- case x => x.toString
- }
-
- new DefaultMutableTreeNode(
- new Isabelle_Sidekick.Asset(command.toString, range.start, range.stop) {
- override def getShortString: String = content
- override def getLongString: String = info_text
- override def toString = "\"" + content + "\" " + range.toString
- })
- })
- }
- }
-}
-
--- a/src/Tools/jEdit/src/jedit/output_dockable.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,163 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/output_dockable.scala
- Author: Makarius
-
-Dockable window with result message output.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.actors.Actor._
-
-import scala.swing.{FlowPanel, Button, CheckBox}
-import scala.swing.event.ButtonClicked
-
-import java.awt.BorderLayout
-import java.awt.event.{ComponentEvent, ComponentAdapter}
-
-import org.gjt.sp.jedit.View
-
-
-class Output_Dockable(view: View, position: String) extends Dockable(view, position)
-{
- Swing_Thread.require()
-
- private val html_panel =
- new HTML_Panel(Isabelle.system, Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
- {
- override val handler: PartialFunction[HTML_Panel.Event, Unit] = {
- case HTML_Panel.Mouse_Click(elem, event) if elem.getClassName() == "sendback" =>
- val text = elem.getFirstChild().getNodeValue()
-
- Document_View(view.getTextArea) match {
- case Some(doc_view) =>
- val cmd = current_command.get
- val start_ofs = doc_view.model.snapshot().node.command_start(cmd).get
- val buffer = view.getBuffer
- buffer.beginCompoundEdit()
- buffer.remove(start_ofs, cmd.length)
- buffer.insert(start_ofs, text)
- buffer.endCompoundEdit()
- case None =>
- }
- }
- }
-
- set_content(html_panel)
-
-
- /* component state -- owned by Swing thread */
-
- private var zoom_factor = 100
- private var show_tracing = false
- private var follow_caret = true
- private var current_command: Option[Command] = None
-
-
- private def handle_resize()
- {
- Swing_Thread.now {
- html_panel.resize(Isabelle.font_family(),
- scala.math.round(Isabelle.font_size() * zoom_factor / 100))
- }
- }
-
- private def handle_perspective(): Boolean =
- Swing_Thread.now {
- Document_View(view.getTextArea) match {
- case Some(doc_view) =>
- val cmd = doc_view.selected_command()
- if (current_command == cmd) false
- else { current_command = cmd; true }
- case None => false
- }
- }
-
- private def handle_update(restriction: Option[Set[Command]] = None)
- {
- Swing_Thread.now {
- if (follow_caret) handle_perspective()
- Document_View(view.getTextArea) match {
- case Some(doc_view) =>
- current_command match {
- case Some(cmd) if !restriction.isDefined || restriction.get.contains(cmd) =>
- val snapshot = doc_view.model.snapshot()
- val filtered_results =
- snapshot.state(cmd).results.iterator.map(_._2) filter {
- case XML.Elem(Markup(Markup.TRACING, _), _) => show_tracing // FIXME not scalable
- case _ => true
- }
- html_panel.render(filtered_results.toList)
- case _ =>
- }
- case None =>
- }
- }
- }
-
-
- /* main actor */
-
- private val main_actor = actor {
- loop {
- react {
- case Session.Global_Settings => handle_resize()
- case Session.Commands_Changed(changed) => handle_update(Some(changed))
- case Session.Perspective => if (follow_caret && handle_perspective()) handle_update()
- case bad => System.err.println("Output_Dockable: ignoring bad message " + bad)
- }
- }
- }
-
- override def init()
- {
- Isabelle.session.global_settings += main_actor
- Isabelle.session.commands_changed += main_actor
- Isabelle.session.perspective += main_actor
- }
-
- override def exit()
- {
- Isabelle.session.global_settings -= main_actor
- Isabelle.session.commands_changed -= main_actor
- Isabelle.session.perspective -= main_actor
- }
-
-
- /* resize */
-
- addComponentListener(new ComponentAdapter {
- val delay = Swing_Thread.delay_last(Isabelle.session.update_delay) { handle_resize() }
- override def componentResized(e: ComponentEvent) { delay() }
- })
-
-
- /* controls */
-
- private val zoom = new Library.Zoom_Box(factor => { zoom_factor = factor; handle_resize() })
- zoom.tooltip = "Zoom factor for basic font size"
-
- private val tracing = new CheckBox("Tracing") {
- reactions += { case ButtonClicked(_) => show_tracing = this.selected; handle_update() }
- }
- tracing.selected = show_tracing
- tracing.tooltip = "Indicate output of tracing messages"
-
- private val auto_update = new CheckBox("Auto update") {
- reactions += { case ButtonClicked(_) => follow_caret = this.selected; handle_update() }
- }
- auto_update.selected = follow_caret
- auto_update.tooltip = "Indicate automatic update following cursor movement"
-
- private val update = new Button("Update") {
- reactions += { case ButtonClicked(_) => handle_perspective(); handle_update() }
- }
- update.tooltip = "Update display according to the command at cursor position"
-
- private val controls = new FlowPanel(FlowPanel.Alignment.Right)(zoom, tracing, auto_update, update)
- add(controls.peer, BorderLayout.NORTH)
-
- handle_update()
-}
--- a/src/Tools/jEdit/src/jedit/plugin.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,401 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/plugin.scala
- Author: Makarius
-
-Main Isabelle/jEdit plugin setup.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import java.io.{FileInputStream, IOException}
-import java.awt.Font
-
-import scala.collection.mutable
-import scala.swing.ComboBox
-
-import org.gjt.sp.jedit.{jEdit, GUIUtilities, EBMessage, EBPlugin,
- Buffer, EditPane, ServiceManager, View}
-import org.gjt.sp.jedit.buffer.JEditBuffer
-import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea}
-import org.gjt.sp.jedit.msg.{EditorStarted, BufferUpdate, EditPaneUpdate, PropertiesChanged}
-import org.gjt.sp.jedit.gui.DockableWindowManager
-
-import org.gjt.sp.util.Log
-
-import scala.actors.Actor
-import Actor._
-
-
-object Isabelle
-{
- /* plugin instance */
-
- var system: Isabelle_System = null
- var session: Session = null
-
-
- /* properties */
-
- val OPTION_PREFIX = "options.isabelle."
-
- object Property
- {
- def apply(name: String): String =
- jEdit.getProperty(OPTION_PREFIX + name)
- def apply(name: String, default: String): String =
- jEdit.getProperty(OPTION_PREFIX + name, default)
- def update(name: String, value: String) =
- jEdit.setProperty(OPTION_PREFIX + name, value)
- }
-
- object Boolean_Property
- {
- def apply(name: String): Boolean =
- jEdit.getBooleanProperty(OPTION_PREFIX + name)
- def apply(name: String, default: Boolean): Boolean =
- jEdit.getBooleanProperty(OPTION_PREFIX + name, default)
- def update(name: String, value: Boolean) =
- jEdit.setBooleanProperty(OPTION_PREFIX + name, value)
- }
-
- object Int_Property
- {
- def apply(name: String): Int =
- jEdit.getIntegerProperty(OPTION_PREFIX + name)
- def apply(name: String, default: Int): Int =
- jEdit.getIntegerProperty(OPTION_PREFIX + name, default)
- def update(name: String, value: Int) =
- jEdit.setIntegerProperty(OPTION_PREFIX + name, value)
- }
-
- object Double_Property
- {
- def apply(name: String): Double =
- jEdit.getDoubleProperty(OPTION_PREFIX + name, 0.0)
- def apply(name: String, default: Double): Double =
- jEdit.getDoubleProperty(OPTION_PREFIX + name, default)
- def update(name: String, value: Double) =
- jEdit.setDoubleProperty(OPTION_PREFIX + name, value)
- }
-
- object Time_Property
- {
- def apply(name: String): Time =
- Time.seconds(Double_Property(name))
- def apply(name: String, default: Time): Time =
- Time.seconds(Double_Property(name, default.seconds))
- def update(name: String, value: Time) =
- Double_Property.update(name, value.seconds)
- }
-
-
- /* font */
-
- def font_family(): String = jEdit.getProperty("view.font")
-
- def font_size(): Float =
- (jEdit.getIntegerProperty("view.fontsize", 16) *
- Int_Property("relative-font-size", 100)).toFloat / 100
-
-
- /* text area ranges */
-
- case class Gfx_Range(val x: Int, val y: Int, val length: Int)
-
- def gfx_range(text_area: TextArea, range: Text.Range): Option[Gfx_Range] =
- {
- val p = text_area.offsetToXY(range.start)
- val q = text_area.offsetToXY(range.stop)
- if (p != null && q != null && p.y == q.y) Some(new Gfx_Range(p.x, p.y, q.x - p.x))
- else None
- }
-
-
- /* tooltip markup */
-
- def tooltip(text: String): String =
- "<html><pre style=\"font-family: " + font_family() + "; font-size: " +
- Int_Property("tooltip-font-size", 10).toString + "px; \">" + // FIXME proper scaling (!?)
- HTML.encode(text) + "</pre></html>"
-
- def tooltip_dismiss_delay(): Time =
- Time_Property("tooltip-dismiss-delay", Time.seconds(8.0)) max Time.seconds(0.5)
-
- def setup_tooltips()
- {
- Swing_Thread.now {
- val manager = javax.swing.ToolTipManager.sharedInstance
- manager.setDismissDelay(tooltip_dismiss_delay().ms.toInt)
- }
- }
-
-
- /* icons */
-
- def load_icon(name: String): javax.swing.Icon =
- {
- val icon = GUIUtilities.loadIcon(name)
- if (icon.getIconWidth < 0 || icon.getIconHeight < 0)
- Log.log(Log.ERROR, icon, "Bad icon: " + name)
- icon
- }
-
-
- /* check JVM */
-
- def check_jvm()
- {
- if (!Platform.is_hotspot) {
- Library.warning_dialog(jEdit.getActiveView, "Bad Java Virtual Machine",
- "This is " + Platform.jvm_name,
- "Isabelle/jEdit requires Java Hotspot from Sun/Oracle/Apple!")
- }
- }
-
-
- /* main jEdit components */
-
- def jedit_buffers(): Iterator[Buffer] = jEdit.getBuffers().iterator
-
- def jedit_views(): Iterator[View] = jEdit.getViews().iterator
-
- def jedit_text_areas(view: View): Iterator[JEditTextArea] =
- view.getEditPanes().iterator.map(_.getTextArea)
-
- def jedit_text_areas(): Iterator[JEditTextArea] =
- jedit_views().flatMap(jedit_text_areas(_))
-
- def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] =
- jedit_text_areas().filter(_.getBuffer == buffer)
-
- def buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
- {
- try { buffer.readLock(); body }
- finally { buffer.readUnlock() }
- }
-
- def swing_buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
- Swing_Thread.now { buffer_lock(buffer) { body } }
-
- def buffer_text(buffer: JEditBuffer): String =
- buffer_lock(buffer) { buffer.getText(0, buffer.getLength) }
-
-
- /* dockable windows */
-
- 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)
- case _ => None
- }
-
- def docked_output(view: View): Option[Output_Dockable] =
- wm(view).getDockableWindow("isabelle-output") match {
- case dockable: Output_Dockable => Some(dockable)
- case _ => None
- }
-
- def docked_raw_output(view: View): Option[Raw_Output_Dockable] =
- wm(view).getDockableWindow("isabelle-raw-output") match {
- case dockable: Raw_Output_Dockable => Some(dockable)
- case _ => None
- }
-
- def docked_protocol(view: View): Option[Protocol_Dockable] =
- wm(view).getDockableWindow("isabelle-protocol") match {
- case dockable: Protocol_Dockable => Some(dockable)
- case _ => None
- }
-
-
- /* logic image */
-
- def default_logic(): String =
- {
- val logic = system.getenv("JEDIT_LOGIC")
- if (logic != "") logic
- else system.getenv_strict("ISABELLE_LOGIC")
- }
-
- class Logic_Entry(val name: String, val description: String)
- {
- override def toString = description
- }
-
- def logic_selector(logic: String): ComboBox[Logic_Entry] =
- {
- val entries =
- new Logic_Entry("", "default (" + default_logic() + ")") ::
- system.find_logics().map(name => new Logic_Entry(name, name))
- val component = new ComboBox(entries)
- entries.find(_.name == logic) match {
- case None =>
- case Some(entry) => component.selection.item = entry
- }
- component.tooltip = "Isabelle logic image"
- component
- }
-
- def start_session()
- {
- val timeout = Time_Property("startup-timeout", Time.seconds(10)) max Time.seconds(5)
- val modes = system.getenv("JEDIT_PRINT_MODE").split(",").toList.map("-m" + _)
- val logic = {
- val logic = Property("logic")
- if (logic != null && logic != "") logic
- else Isabelle.default_logic()
- }
- session.start(timeout, modes ::: List(logic))
- }
-}
-
-
-class Plugin extends EBPlugin
-{
- /* session management */
-
- private def init_model(buffer: Buffer)
- {
- Isabelle.swing_buffer_lock(buffer) {
- val opt_model =
- Document_Model(buffer) match {
- case Some(model) => model.refresh; Some(model)
- case None =>
- Thy_Header.split_thy_path(Isabelle.system.posix_path(buffer.getPath)) match {
- case Some((dir, thy_name)) =>
- Some(Document_Model.init(Isabelle.session, buffer, dir + "/" + thy_name))
- case None => None
- }
- }
- if (opt_model.isDefined) {
- for (text_area <- Isabelle.jedit_text_areas(buffer)) {
- if (Document_View(text_area).map(_.model) != opt_model)
- Document_View.init(opt_model.get, text_area)
- }
- }
- }
- }
-
- private def exit_model(buffer: Buffer)
- {
- Isabelle.swing_buffer_lock(buffer) {
- Isabelle.jedit_text_areas(buffer).foreach(Document_View.exit)
- Document_Model.exit(buffer)
- }
- }
-
- private case class Init_Model(buffer: Buffer)
- private case class Exit_Model(buffer: Buffer)
- private case class Init_View(buffer: Buffer, text_area: JEditTextArea)
- private case class Exit_View(buffer: Buffer, text_area: JEditTextArea)
-
- private val session_manager = actor {
- var ready = false
- loop {
- react {
- case phase: Session.Phase =>
- ready = phase == Session.Ready
- phase match {
- case Session.Failed =>
- Swing_Thread.now {
- val text = new scala.swing.TextArea(Isabelle.session.syslog())
- text.editable = false
- Library.error_dialog(jEdit.getActiveView, "Failed to start Isabelle process", text)
- }
-
- case Session.Ready => Isabelle.jedit_buffers.foreach(init_model)
- case Session.Shutdown => Isabelle.jedit_buffers.foreach(exit_model)
- case _ =>
- }
-
- case Init_Model(buffer) =>
- if (ready) init_model(buffer)
-
- case Exit_Model(buffer) =>
- exit_model(buffer)
-
- case Init_View(buffer, text_area) =>
- if (ready) {
- Isabelle.swing_buffer_lock(buffer) {
- Document_Model(buffer) match {
- case Some(model) => Document_View.init(model, text_area)
- case None =>
- }
- }
- }
-
- case Exit_View(buffer, text_area) =>
- Isabelle.swing_buffer_lock(buffer) {
- Document_View.exit(text_area)
- }
-
- case bad => System.err.println("session_manager: ignoring bad message " + bad)
- }
- }
- }
-
-
- /* main plugin plumbing */
-
- override def handleMessage(message: EBMessage)
- {
- message match {
- case msg: EditorStarted =>
- Isabelle.check_jvm()
- if (Isabelle.Boolean_Property("auto-start")) Isabelle.start_session()
-
- case msg: BufferUpdate
- if msg.getWhat == BufferUpdate.PROPERTIES_CHANGED =>
-
- val buffer = msg.getBuffer
- if (buffer != null) session_manager ! Init_Model(buffer)
-
- case msg: EditPaneUpdate
- if (msg.getWhat == EditPaneUpdate.BUFFER_CHANGING ||
- msg.getWhat == EditPaneUpdate.BUFFER_CHANGED ||
- msg.getWhat == EditPaneUpdate.CREATED ||
- msg.getWhat == EditPaneUpdate.DESTROYED) =>
-
- val edit_pane = msg.getEditPane
- val buffer = edit_pane.getBuffer
- val text_area = edit_pane.getTextArea
-
- if (buffer != null && text_area != null) {
- if (msg.getWhat == EditPaneUpdate.BUFFER_CHANGED ||
- msg.getWhat == EditPaneUpdate.CREATED)
- session_manager ! Init_View(buffer, text_area)
- else
- session_manager ! Exit_View(buffer, text_area)
- }
-
- case msg: PropertiesChanged =>
- Swing_Thread.now {
- Isabelle.setup_tooltips()
- for (text_area <- Isabelle.jedit_text_areas if Document_View(text_area).isDefined)
- Document_View(text_area).get.extend_styles()
- }
- Isabelle.session.global_settings.event(Session.Global_Settings)
-
- case _ =>
- }
- }
-
- override def start()
- {
- Isabelle.setup_tooltips()
- Isabelle.system = new Isabelle_System
- Isabelle.system.install_fonts()
- Isabelle.session = new Session(Isabelle.system)
- Isabelle.session.phase_changed += session_manager
- }
-
- override def stop()
- {
- Isabelle.session.stop()
- Isabelle.session.phase_changed -= session_manager
- }
-}
--- a/src/Tools/jEdit/src/jedit/protocol_dockable.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/protocol_dockable.scala
- Author: Makarius
-
-Dockable window for protocol messages.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.actors.Actor._
-import scala.swing.{TextArea, ScrollPane}
-
-import org.gjt.sp.jedit.View
-
-
-class Protocol_Dockable(view: View, position: String) extends Dockable(view, position)
-{
- private val text_area = new TextArea
- set_content(new ScrollPane(text_area))
-
-
- /* main actor */
-
- private val main_actor = actor {
- loop {
- react {
- case result: Isabelle_Process.Result =>
- Swing_Thread.now { text_area.append(result.message.toString + "\n") }
-
- case bad => System.err.println("Protocol_Dockable: ignoring bad message " + bad)
- }
- }
- }
-
- override def init() { Isabelle.session.raw_messages += main_actor }
- override def exit() { Isabelle.session.raw_messages -= main_actor }
-}
--- a/src/Tools/jEdit/src/jedit/raw_output_dockable.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,42 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/raw_output_dockable.scala
- Author: Makarius
-
-Dockable window for raw process output (stdout).
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import scala.actors.Actor._
-import scala.swing.{TextArea, ScrollPane}
-
-import org.gjt.sp.jedit.View
-
-
-class Raw_Output_Dockable(view: View, position: String)
- extends Dockable(view: View, position: String)
-{
- private val text_area = new TextArea
- text_area.editable = false
- set_content(new ScrollPane(text_area))
-
-
- /* main actor */
-
- private val main_actor = actor {
- loop {
- react {
- case result: Isabelle_Process.Result =>
- if (result.is_stdout)
- Swing_Thread.now { text_area.append(XML.content(result.message).mkString) }
-
- case bad => System.err.println("Raw_Output_Dockable: ignoring bad message " + bad)
- }
- }
- }
-
- override def init() { Isabelle.session.raw_messages += main_actor }
- override def exit() { Isabelle.session.raw_messages -= main_actor }
-}
--- a/src/Tools/jEdit/src/jedit/scala_console.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,155 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/scala_console.scala
- Author: Makarius
-
-Scala instance of Console plugin.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import console.{Console, ConsolePane, Shell, Output}
-
-import org.gjt.sp.jedit.{jEdit, JARClassLoader}
-import org.gjt.sp.jedit.MiscUtilities
-
-import java.io.{File, OutputStream, Writer, PrintWriter}
-
-import scala.tools.nsc.{Interpreter, GenericRunnerSettings, NewLinePrintWriter, ConsoleWriter}
-import scala.collection.mutable
-
-
-class Scala_Console extends Shell("Scala")
-{
- /* reconstructed jEdit/plugin classpath */
-
- private def reconstruct_classpath(): String =
- {
- def find_jars(start: String): List[String] =
- if (start != null)
- Standard_System.find_files(new File(start),
- entry => entry.isFile && entry.getName.endsWith(".jar")).map(_.getAbsolutePath)
- else Nil
- val path = find_jars(jEdit.getSettingsDirectory) ::: find_jars(jEdit.getJEditHome)
- path.mkString(File.pathSeparator)
- }
-
-
- /* global state -- owned by Swing thread */
-
- private var interpreters = Map[Console, Interpreter]()
-
- private var global_console: Console = null
- private var global_out: Output = null
- private var global_err: Output = null
-
- private val console_stream = new OutputStream
- {
- val buf = new StringBuilder
- override def flush()
- {
- val str = Standard_System.decode_permissive_utf8(buf.toString)
- buf.clear
- if (global_out == null) System.out.print(str)
- else Swing_Thread.now { global_out.writeAttrs(null, str) }
- }
- override def close() { flush () }
- def write(byte: Int) { buf.append(byte.toChar) }
- }
-
- private val console_writer = new Writer
- {
- def flush() {}
- def close() {}
-
- def write(cbuf: Array[Char], off: Int, len: Int)
- {
- if (len > 0) write(new String(cbuf.slice(off, off + len)))
- }
-
- override def write(str: String)
- {
- if (global_out == null) System.out.println(str)
- else global_out.print(null, str)
- }
- }
-
- private def with_console[A](console: Console, out: Output, err: Output)(e: => A): A =
- {
- global_console = console
- global_out = out
- global_err = if (err == null) out else err
- val res = Exn.capture { scala.Console.withOut(console_stream)(e) }
- console_stream.flush
- global_console = null
- global_out = null
- global_err = null
- Exn.release(res)
- }
-
- private def report_error(str: String)
- {
- if (global_console == null || global_err == null) System.err.println(str)
- else Swing_Thread.now { global_err.print(global_console.getErrorColor, str) }
- }
-
-
- /* jEdit console methods */
-
- override def openConsole(console: Console)
- {
- val settings = new GenericRunnerSettings(report_error)
- settings.classpath.value = reconstruct_classpath()
-
- val interp = new Interpreter(settings, new PrintWriter(console_writer, true))
- {
- override def parentClassLoader = new JARClassLoader
- }
- interp.setContextClassLoader
- interp.bind("view", "org.gjt.sp.jedit.View", console.getView)
- interp.bind("console", "console.Console", console)
- interp.interpret("import isabelle.jedit.Isabelle")
-
- interpreters += (console -> interp)
- }
-
- override def closeConsole(console: Console)
- {
- interpreters -= console
- }
-
- override def printInfoMessage(out: Output)
- {
- out.print(null,
- "This shell evaluates Isabelle/Scala expressions.\n\n" +
- "The following special toplevel bindings are provided:\n" +
- " view -- current jEdit/Swing view (e.g. view.getBuffer, view.getTextArea)\n" +
- " console -- jEdit Console plugin instance\n" +
- " Isabelle -- Isabelle plugin instance (e.g. Isabelle.system, Isabelle.session)\n")
- }
-
- override def printPrompt(console: Console, out: Output)
- {
- out.writeAttrs(ConsolePane.colorAttributes(console.getInfoColor), "scala>")
- out.writeAttrs(ConsolePane.colorAttributes(console.getPlainColor), " ")
- }
-
- override def execute(console: Console, input: String, out: Output, err: Output, command: String)
- {
- val interp = interpreters(console)
- with_console(console, out, err) { interp.interpret(command) }
- if (err != null) err.commandDone()
- out.commandDone()
- }
-
- override def stop(console: Console)
- {
- closeConsole(console)
- console.clear
- openConsole(console)
- val out = console.getOutput
- out.commandDone
- printPrompt(console, out)
- }
-}
--- a/src/Tools/jEdit/src/jedit/session_dockable.scala Wed Jun 08 17:01:07 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,104 +0,0 @@
-/* Title: Tools/jEdit/src/jedit/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, ScrollPane, TabbedPane, Component, Swing}
-import scala.swing.event.{ButtonClicked, SelectionChanged}
-
-import java.awt.BorderLayout
-import javax.swing.border.{BevelBorder, SoftBevelBorder}
-
-import org.gjt.sp.jedit.View
-
-
-class Session_Dockable(view: View, position: String) extends Dockable(view: View, position: String)
-{
- /* main tabs */
-
- private val readme = new HTML_Panel(Isabelle.system, "SansSerif", 14)
- readme.render_document(Isabelle.system.try_read(List("$JEDIT_HOME/README.html")))
-
- private val syslog = new TextArea(Isabelle.session.syslog())
- syslog.editable = false
-
- private val tabs = new TabbedPane {
- pages += new TabbedPane.Page("README", Component.wrap(readme))
- pages += new TabbedPane.Page("System log", new ScrollPane(syslog))
-
- selection.index =
- {
- val index = Isabelle.Int_Property("session-panel.selection", 0)
- if (index >= pages.length) 0 else index
- }
- listenTo(selection)
- reactions += {
- case SelectionChanged(_) =>
- Isabelle.Int_Property("session-panel.selection") = selection.index
- }
- }
-
- set_content(tabs)
-
-
- /* controls */
-
- val session_phase = new Label(Isabelle.session.phase.toString)
- session_phase.border = new SoftBevelBorder(BevelBorder.LOWERED)
- session_phase.tooltip = "Prover status"
-
- private val interrupt = new Button("Interrupt") {
- reactions += { case ButtonClicked(_) => Isabelle.session.interrupt }
- }
- interrupt.tooltip = "Broadcast interrupt to all prover tasks"
-
- private val logic = Isabelle.logic_selector(Isabelle.Property("logic"))
- logic.listenTo(logic.selection)
- logic.reactions += {
- case SelectionChanged(_) => Isabelle.Property("logic") = logic.selection.item.name
- }
-
- private val controls =
- new FlowPanel(FlowPanel.Alignment.Right)(session_phase, interrupt, logic)
- add(controls.peer, BorderLayout.NORTH)
-
-
- /* main actor */
-
- private val main_actor = actor {
- loop {
- react {
- case result: Isabelle_Process.Result =>
- if (result.is_syslog)
- Swing_Thread.now {
- val text = Isabelle.session.syslog()
- if (text != syslog.text) {
- syslog.text = text
- }
- }
-
- case phase: Session.Phase =>
- Swing_Thread.now { session_phase.text = " " + phase.toString + " " }
-
- case bad => System.err.println("Session_Dockable: ignoring bad message " + bad)
- }
- }
- }
-
- override def init() {
- Isabelle.session.raw_messages += main_actor
- Isabelle.session.phase_changed += main_actor
- }
-
- override def exit() {
- Isabelle.session.raw_messages -= main_actor
- Isabelle.session.phase_changed -= main_actor
- }
-}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/modes/isabelle-session.xml Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,41 @@
+<?xml version="1.0"?>
+<!DOCTYPE MODE SYSTEM "xmode.dtd">
+
+<!-- Isabelle session mode -->
+<MODE>
+ <PROPS>
+ <PROPERTY NAME="commentStart" VALUE="(*"/>
+ <PROPERTY NAME="commentEnd" VALUE="*)"/>
+ <PROPERTY NAME="noWordSep" VALUE="_'.?"/>
+ <PROPERTY NAME="unalignedOpenBrackets" VALUE="(" />
+ <PROPERTY NAME="unalignedCloseBrackets" VALUE=")" />
+ <PROPERTY NAME="tabSize" VALUE="2" />
+ <PROPERTY NAME="indentSize" VALUE="2" />
+ </PROPS>
+ <RULES IGNORE_CASE="FALSE" HIGHLIGHT_DIGITS="FALSE" ESCAPE="\">
+ <SPAN TYPE="COMMENT1">
+ <BEGIN>(*</BEGIN>
+ <END>*)</END>
+ </SPAN>
+ <SPAN TYPE="COMMENT3">
+ <BEGIN>{*</BEGIN>
+ <END>*}</END>
+ </SPAN>
+ <SPAN TYPE="LITERAL1">
+ <BEGIN>`</BEGIN>
+ <END>`</END>
+ </SPAN>
+ <SPAN TYPE="LITERAL3">
+ <BEGIN>"</BEGIN>
+ <END>"</END>
+ </SPAN>
+ <KEYWORDS>
+ <KEYWORD1>session</KEYWORD1>
+ <KEYWORD2>parent</KEYWORD2>
+ <KEYWORD2>imports</KEYWORD2>
+ <KEYWORD2>uses</KEYWORD2>
+ <KEYWORD2>options</KEYWORD2>
+ <KEYWORD2>dependencies</KEYWORD2>
+ </KEYWORDS>
+ </RULES>
+</MODE>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/modes/isabelle.xml Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,41 @@
+<?xml version="1.0"?>
+<!DOCTYPE MODE SYSTEM "xmode.dtd">
+
+<!-- Isabelle theory mode -->
+<MODE>
+ <PROPS>
+ <PROPERTY NAME="commentStart" VALUE="(*"/>
+ <PROPERTY NAME="commentEnd" VALUE="*)"/>
+ <PROPERTY NAME="noWordSep" VALUE="_'.?"/>
+ <PROPERTY NAME="unalignedOpenBrackets" VALUE="{[(" />
+ <PROPERTY NAME="unalignedCloseBrackets" VALUE=")]}" />
+ <PROPERTY NAME="tabSize" VALUE="2" />
+ <PROPERTY NAME="indentSize" VALUE="2" />
+ </PROPS>
+ <RULES IGNORE_CASE="FALSE" HIGHLIGHT_DIGITS="FALSE" ESCAPE="\">
+ <SPAN TYPE="COMMENT1">
+ <BEGIN>(*</BEGIN>
+ <END>*)</END>
+ </SPAN>
+ <SPAN TYPE="COMMENT3">
+ <BEGIN>{*</BEGIN>
+ <END>*}</END>
+ </SPAN>
+ <SPAN TYPE="LITERAL1">
+ <BEGIN>`</BEGIN>
+ <END>`</END>
+ </SPAN>
+ <SPAN TYPE="LITERAL3">
+ <BEGIN>"</BEGIN>
+ <END>"</END>
+ </SPAN>
+ <KEYWORDS>
+ <KEYWORD2>header</KEYWORD2>
+ <KEYWORD1>theory</KEYWORD1>
+ <KEYWORD2>imports</KEYWORD2>
+ <KEYWORD2>uses</KEYWORD2>
+ <KEYWORD2>begin</KEYWORD2>
+ <KEYWORD2>end</KEYWORD2>
+ </KEYWORDS>
+ </RULES>
+</MODE>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/output_dockable.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,163 @@
+/* Title: Tools/jEdit/src/output_dockable.scala
+ Author: Makarius
+
+Dockable window with result message output.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.actors.Actor._
+
+import scala.swing.{FlowPanel, Button, CheckBox}
+import scala.swing.event.ButtonClicked
+
+import java.awt.BorderLayout
+import java.awt.event.{ComponentEvent, ComponentAdapter}
+
+import org.gjt.sp.jedit.View
+
+
+class Output_Dockable(view: View, position: String) extends Dockable(view, position)
+{
+ Swing_Thread.require()
+
+ private val html_panel =
+ new HTML_Panel(Isabelle.system, Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
+ {
+ override val handler: PartialFunction[HTML_Panel.Event, Unit] = {
+ case HTML_Panel.Mouse_Click(elem, event) if elem.getClassName() == "sendback" =>
+ val text = elem.getFirstChild().getNodeValue()
+
+ Document_View(view.getTextArea) match {
+ case Some(doc_view) =>
+ val cmd = current_command.get
+ val start_ofs = doc_view.model.snapshot().node.command_start(cmd).get
+ val buffer = view.getBuffer
+ buffer.beginCompoundEdit()
+ buffer.remove(start_ofs, cmd.length)
+ buffer.insert(start_ofs, text)
+ buffer.endCompoundEdit()
+ case None =>
+ }
+ }
+ }
+
+ set_content(html_panel)
+
+
+ /* component state -- owned by Swing thread */
+
+ private var zoom_factor = 100
+ private var show_tracing = false
+ private var follow_caret = true
+ private var current_command: Option[Command] = None
+
+
+ private def handle_resize()
+ {
+ Swing_Thread.now {
+ html_panel.resize(Isabelle.font_family(),
+ scala.math.round(Isabelle.font_size() * zoom_factor / 100))
+ }
+ }
+
+ private def handle_perspective(): Boolean =
+ Swing_Thread.now {
+ Document_View(view.getTextArea) match {
+ case Some(doc_view) =>
+ val cmd = doc_view.selected_command()
+ if (current_command == cmd) false
+ else { current_command = cmd; true }
+ case None => false
+ }
+ }
+
+ private def handle_update(restriction: Option[Set[Command]] = None)
+ {
+ Swing_Thread.now {
+ if (follow_caret) handle_perspective()
+ Document_View(view.getTextArea) match {
+ case Some(doc_view) =>
+ current_command match {
+ case Some(cmd) if !restriction.isDefined || restriction.get.contains(cmd) =>
+ val snapshot = doc_view.model.snapshot()
+ val filtered_results =
+ snapshot.state(cmd).results.iterator.map(_._2) filter {
+ case XML.Elem(Markup(Markup.TRACING, _), _) => show_tracing // FIXME not scalable
+ case _ => true
+ }
+ html_panel.render(filtered_results.toList)
+ case _ =>
+ }
+ case None =>
+ }
+ }
+ }
+
+
+ /* main actor */
+
+ private val main_actor = actor {
+ loop {
+ react {
+ case Session.Global_Settings => handle_resize()
+ case Session.Commands_Changed(changed) => handle_update(Some(changed))
+ case Session.Perspective => if (follow_caret && handle_perspective()) handle_update()
+ case bad => System.err.println("Output_Dockable: ignoring bad message " + bad)
+ }
+ }
+ }
+
+ override def init()
+ {
+ Isabelle.session.global_settings += main_actor
+ Isabelle.session.commands_changed += main_actor
+ Isabelle.session.perspective += main_actor
+ }
+
+ override def exit()
+ {
+ Isabelle.session.global_settings -= main_actor
+ Isabelle.session.commands_changed -= main_actor
+ Isabelle.session.perspective -= main_actor
+ }
+
+
+ /* resize */
+
+ addComponentListener(new ComponentAdapter {
+ val delay = Swing_Thread.delay_last(Isabelle.session.update_delay) { handle_resize() }
+ override def componentResized(e: ComponentEvent) { delay() }
+ })
+
+
+ /* controls */
+
+ private val zoom = new Library.Zoom_Box(factor => { zoom_factor = factor; handle_resize() })
+ zoom.tooltip = "Zoom factor for basic font size"
+
+ private val tracing = new CheckBox("Tracing") {
+ reactions += { case ButtonClicked(_) => show_tracing = this.selected; handle_update() }
+ }
+ tracing.selected = show_tracing
+ tracing.tooltip = "Indicate output of tracing messages"
+
+ private val auto_update = new CheckBox("Auto update") {
+ reactions += { case ButtonClicked(_) => follow_caret = this.selected; handle_update() }
+ }
+ auto_update.selected = follow_caret
+ auto_update.tooltip = "Indicate automatic update following cursor movement"
+
+ private val update = new Button("Update") {
+ reactions += { case ButtonClicked(_) => handle_perspective(); handle_update() }
+ }
+ update.tooltip = "Update display according to the command at cursor position"
+
+ private val controls = new FlowPanel(FlowPanel.Alignment.Right)(zoom, tracing, auto_update, update)
+ add(controls.peer, BorderLayout.NORTH)
+
+ handle_update()
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/plugin.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,401 @@
+/* Title: Tools/jEdit/src/plugin.scala
+ Author: Makarius
+
+Main Isabelle/jEdit plugin setup.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.io.{FileInputStream, IOException}
+import java.awt.Font
+
+import scala.collection.mutable
+import scala.swing.ComboBox
+
+import org.gjt.sp.jedit.{jEdit, GUIUtilities, EBMessage, EBPlugin,
+ Buffer, EditPane, ServiceManager, View}
+import org.gjt.sp.jedit.buffer.JEditBuffer
+import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea}
+import org.gjt.sp.jedit.msg.{EditorStarted, BufferUpdate, EditPaneUpdate, PropertiesChanged}
+import org.gjt.sp.jedit.gui.DockableWindowManager
+
+import org.gjt.sp.util.Log
+
+import scala.actors.Actor
+import Actor._
+
+
+object Isabelle
+{
+ /* plugin instance */
+
+ var system: Isabelle_System = null
+ var session: Session = null
+
+
+ /* properties */
+
+ val OPTION_PREFIX = "options.isabelle."
+
+ object Property
+ {
+ def apply(name: String): String =
+ jEdit.getProperty(OPTION_PREFIX + name)
+ def apply(name: String, default: String): String =
+ jEdit.getProperty(OPTION_PREFIX + name, default)
+ def update(name: String, value: String) =
+ jEdit.setProperty(OPTION_PREFIX + name, value)
+ }
+
+ object Boolean_Property
+ {
+ def apply(name: String): Boolean =
+ jEdit.getBooleanProperty(OPTION_PREFIX + name)
+ def apply(name: String, default: Boolean): Boolean =
+ jEdit.getBooleanProperty(OPTION_PREFIX + name, default)
+ def update(name: String, value: Boolean) =
+ jEdit.setBooleanProperty(OPTION_PREFIX + name, value)
+ }
+
+ object Int_Property
+ {
+ def apply(name: String): Int =
+ jEdit.getIntegerProperty(OPTION_PREFIX + name)
+ def apply(name: String, default: Int): Int =
+ jEdit.getIntegerProperty(OPTION_PREFIX + name, default)
+ def update(name: String, value: Int) =
+ jEdit.setIntegerProperty(OPTION_PREFIX + name, value)
+ }
+
+ object Double_Property
+ {
+ def apply(name: String): Double =
+ jEdit.getDoubleProperty(OPTION_PREFIX + name, 0.0)
+ def apply(name: String, default: Double): Double =
+ jEdit.getDoubleProperty(OPTION_PREFIX + name, default)
+ def update(name: String, value: Double) =
+ jEdit.setDoubleProperty(OPTION_PREFIX + name, value)
+ }
+
+ object Time_Property
+ {
+ def apply(name: String): Time =
+ Time.seconds(Double_Property(name))
+ def apply(name: String, default: Time): Time =
+ Time.seconds(Double_Property(name, default.seconds))
+ def update(name: String, value: Time) =
+ Double_Property.update(name, value.seconds)
+ }
+
+
+ /* font */
+
+ def font_family(): String = jEdit.getProperty("view.font")
+
+ def font_size(): Float =
+ (jEdit.getIntegerProperty("view.fontsize", 16) *
+ Int_Property("relative-font-size", 100)).toFloat / 100
+
+
+ /* text area ranges */
+
+ case class Gfx_Range(val x: Int, val y: Int, val length: Int)
+
+ def gfx_range(text_area: TextArea, range: Text.Range): Option[Gfx_Range] =
+ {
+ val p = text_area.offsetToXY(range.start)
+ val q = text_area.offsetToXY(range.stop)
+ if (p != null && q != null && p.y == q.y) Some(new Gfx_Range(p.x, p.y, q.x - p.x))
+ else None
+ }
+
+
+ /* tooltip markup */
+
+ def tooltip(text: String): String =
+ "<html><pre style=\"font-family: " + font_family() + "; font-size: " +
+ Int_Property("tooltip-font-size", 10).toString + "px; \">" + // FIXME proper scaling (!?)
+ HTML.encode(text) + "</pre></html>"
+
+ def tooltip_dismiss_delay(): Time =
+ Time_Property("tooltip-dismiss-delay", Time.seconds(8.0)) max Time.seconds(0.5)
+
+ def setup_tooltips()
+ {
+ Swing_Thread.now {
+ val manager = javax.swing.ToolTipManager.sharedInstance
+ manager.setDismissDelay(tooltip_dismiss_delay().ms.toInt)
+ }
+ }
+
+
+ /* icons */
+
+ def load_icon(name: String): javax.swing.Icon =
+ {
+ val icon = GUIUtilities.loadIcon(name)
+ if (icon.getIconWidth < 0 || icon.getIconHeight < 0)
+ Log.log(Log.ERROR, icon, "Bad icon: " + name)
+ icon
+ }
+
+
+ /* check JVM */
+
+ def check_jvm()
+ {
+ if (!Platform.is_hotspot) {
+ Library.warning_dialog(jEdit.getActiveView, "Bad Java Virtual Machine",
+ "This is " + Platform.jvm_name,
+ "Isabelle/jEdit requires Java Hotspot from Sun/Oracle/Apple!")
+ }
+ }
+
+
+ /* main jEdit components */
+
+ def jedit_buffers(): Iterator[Buffer] = jEdit.getBuffers().iterator
+
+ def jedit_views(): Iterator[View] = jEdit.getViews().iterator
+
+ def jedit_text_areas(view: View): Iterator[JEditTextArea] =
+ view.getEditPanes().iterator.map(_.getTextArea)
+
+ def jedit_text_areas(): Iterator[JEditTextArea] =
+ jedit_views().flatMap(jedit_text_areas(_))
+
+ def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] =
+ jedit_text_areas().filter(_.getBuffer == buffer)
+
+ def buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
+ {
+ try { buffer.readLock(); body }
+ finally { buffer.readUnlock() }
+ }
+
+ def swing_buffer_lock[A](buffer: JEditBuffer)(body: => A): A =
+ Swing_Thread.now { buffer_lock(buffer) { body } }
+
+ def buffer_text(buffer: JEditBuffer): String =
+ buffer_lock(buffer) { buffer.getText(0, buffer.getLength) }
+
+
+ /* dockable windows */
+
+ 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)
+ case _ => None
+ }
+
+ def docked_output(view: View): Option[Output_Dockable] =
+ wm(view).getDockableWindow("isabelle-output") match {
+ case dockable: Output_Dockable => Some(dockable)
+ case _ => None
+ }
+
+ def docked_raw_output(view: View): Option[Raw_Output_Dockable] =
+ wm(view).getDockableWindow("isabelle-raw-output") match {
+ case dockable: Raw_Output_Dockable => Some(dockable)
+ case _ => None
+ }
+
+ def docked_protocol(view: View): Option[Protocol_Dockable] =
+ wm(view).getDockableWindow("isabelle-protocol") match {
+ case dockable: Protocol_Dockable => Some(dockable)
+ case _ => None
+ }
+
+
+ /* logic image */
+
+ def default_logic(): String =
+ {
+ val logic = system.getenv("JEDIT_LOGIC")
+ if (logic != "") logic
+ else system.getenv_strict("ISABELLE_LOGIC")
+ }
+
+ class Logic_Entry(val name: String, val description: String)
+ {
+ override def toString = description
+ }
+
+ def logic_selector(logic: String): ComboBox[Logic_Entry] =
+ {
+ val entries =
+ new Logic_Entry("", "default (" + default_logic() + ")") ::
+ system.find_logics().map(name => new Logic_Entry(name, name))
+ val component = new ComboBox(entries)
+ entries.find(_.name == logic) match {
+ case None =>
+ case Some(entry) => component.selection.item = entry
+ }
+ component.tooltip = "Isabelle logic image"
+ component
+ }
+
+ def start_session()
+ {
+ val timeout = Time_Property("startup-timeout", Time.seconds(10)) max Time.seconds(5)
+ val modes = system.getenv("JEDIT_PRINT_MODE").split(",").toList.map("-m" + _)
+ val logic = {
+ val logic = Property("logic")
+ if (logic != null && logic != "") logic
+ else Isabelle.default_logic()
+ }
+ session.start(timeout, modes ::: List(logic))
+ }
+}
+
+
+class Plugin extends EBPlugin
+{
+ /* session management */
+
+ private def init_model(buffer: Buffer)
+ {
+ Isabelle.swing_buffer_lock(buffer) {
+ val opt_model =
+ Document_Model(buffer) match {
+ case Some(model) => model.refresh; Some(model)
+ case None =>
+ Thy_Header.split_thy_path(Isabelle.system.posix_path(buffer.getPath)) match {
+ case Some((dir, thy_name)) =>
+ Some(Document_Model.init(Isabelle.session, buffer, dir + "/" + thy_name))
+ case None => None
+ }
+ }
+ if (opt_model.isDefined) {
+ for (text_area <- Isabelle.jedit_text_areas(buffer)) {
+ if (Document_View(text_area).map(_.model) != opt_model)
+ Document_View.init(opt_model.get, text_area)
+ }
+ }
+ }
+ }
+
+ private def exit_model(buffer: Buffer)
+ {
+ Isabelle.swing_buffer_lock(buffer) {
+ Isabelle.jedit_text_areas(buffer).foreach(Document_View.exit)
+ Document_Model.exit(buffer)
+ }
+ }
+
+ private case class Init_Model(buffer: Buffer)
+ private case class Exit_Model(buffer: Buffer)
+ private case class Init_View(buffer: Buffer, text_area: JEditTextArea)
+ private case class Exit_View(buffer: Buffer, text_area: JEditTextArea)
+
+ private val session_manager = actor {
+ var ready = false
+ loop {
+ react {
+ case phase: Session.Phase =>
+ ready = phase == Session.Ready
+ phase match {
+ case Session.Failed =>
+ Swing_Thread.now {
+ val text = new scala.swing.TextArea(Isabelle.session.syslog())
+ text.editable = false
+ Library.error_dialog(jEdit.getActiveView, "Failed to start Isabelle process", text)
+ }
+
+ case Session.Ready => Isabelle.jedit_buffers.foreach(init_model)
+ case Session.Shutdown => Isabelle.jedit_buffers.foreach(exit_model)
+ case _ =>
+ }
+
+ case Init_Model(buffer) =>
+ if (ready) init_model(buffer)
+
+ case Exit_Model(buffer) =>
+ exit_model(buffer)
+
+ case Init_View(buffer, text_area) =>
+ if (ready) {
+ Isabelle.swing_buffer_lock(buffer) {
+ Document_Model(buffer) match {
+ case Some(model) => Document_View.init(model, text_area)
+ case None =>
+ }
+ }
+ }
+
+ case Exit_View(buffer, text_area) =>
+ Isabelle.swing_buffer_lock(buffer) {
+ Document_View.exit(text_area)
+ }
+
+ case bad => System.err.println("session_manager: ignoring bad message " + bad)
+ }
+ }
+ }
+
+
+ /* main plugin plumbing */
+
+ override def handleMessage(message: EBMessage)
+ {
+ message match {
+ case msg: EditorStarted =>
+ Isabelle.check_jvm()
+ if (Isabelle.Boolean_Property("auto-start")) Isabelle.start_session()
+
+ case msg: BufferUpdate
+ if msg.getWhat == BufferUpdate.PROPERTIES_CHANGED =>
+
+ val buffer = msg.getBuffer
+ if (buffer != null) session_manager ! Init_Model(buffer)
+
+ case msg: EditPaneUpdate
+ if (msg.getWhat == EditPaneUpdate.BUFFER_CHANGING ||
+ msg.getWhat == EditPaneUpdate.BUFFER_CHANGED ||
+ msg.getWhat == EditPaneUpdate.CREATED ||
+ msg.getWhat == EditPaneUpdate.DESTROYED) =>
+
+ val edit_pane = msg.getEditPane
+ val buffer = edit_pane.getBuffer
+ val text_area = edit_pane.getTextArea
+
+ if (buffer != null && text_area != null) {
+ if (msg.getWhat == EditPaneUpdate.BUFFER_CHANGED ||
+ msg.getWhat == EditPaneUpdate.CREATED)
+ session_manager ! Init_View(buffer, text_area)
+ else
+ session_manager ! Exit_View(buffer, text_area)
+ }
+
+ case msg: PropertiesChanged =>
+ Swing_Thread.now {
+ Isabelle.setup_tooltips()
+ for (text_area <- Isabelle.jedit_text_areas if Document_View(text_area).isDefined)
+ Document_View(text_area).get.extend_styles()
+ }
+ Isabelle.session.global_settings.event(Session.Global_Settings)
+
+ case _ =>
+ }
+ }
+
+ override def start()
+ {
+ Isabelle.setup_tooltips()
+ Isabelle.system = new Isabelle_System
+ Isabelle.system.install_fonts()
+ Isabelle.session = new Session(Isabelle.system)
+ Isabelle.session.phase_changed += session_manager
+ }
+
+ override def stop()
+ {
+ Isabelle.session.stop()
+ Isabelle.session.phase_changed -= session_manager
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/protocol_dockable.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,39 @@
+/* Title: Tools/jEdit/src/protocol_dockable.scala
+ Author: Makarius
+
+Dockable window for protocol messages.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.actors.Actor._
+import scala.swing.{TextArea, ScrollPane}
+
+import org.gjt.sp.jedit.View
+
+
+class Protocol_Dockable(view: View, position: String) extends Dockable(view, position)
+{
+ private val text_area = new TextArea
+ set_content(new ScrollPane(text_area))
+
+
+ /* main actor */
+
+ private val main_actor = actor {
+ loop {
+ react {
+ case result: Isabelle_Process.Result =>
+ Swing_Thread.now { text_area.append(result.message.toString + "\n") }
+
+ case bad => System.err.println("Protocol_Dockable: ignoring bad message " + bad)
+ }
+ }
+ }
+
+ override def init() { Isabelle.session.raw_messages += main_actor }
+ override def exit() { Isabelle.session.raw_messages -= main_actor }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/raw_output_dockable.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,42 @@
+/* Title: Tools/jEdit/src/raw_output_dockable.scala
+ Author: Makarius
+
+Dockable window for raw process output (stdout).
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import scala.actors.Actor._
+import scala.swing.{TextArea, ScrollPane}
+
+import org.gjt.sp.jedit.View
+
+
+class Raw_Output_Dockable(view: View, position: String)
+ extends Dockable(view: View, position: String)
+{
+ private val text_area = new TextArea
+ text_area.editable = false
+ set_content(new ScrollPane(text_area))
+
+
+ /* main actor */
+
+ private val main_actor = actor {
+ loop {
+ react {
+ case result: Isabelle_Process.Result =>
+ if (result.is_stdout)
+ Swing_Thread.now { text_area.append(XML.content(result.message).mkString) }
+
+ case bad => System.err.println("Raw_Output_Dockable: ignoring bad message " + bad)
+ }
+ }
+ }
+
+ override def init() { Isabelle.session.raw_messages += main_actor }
+ override def exit() { Isabelle.session.raw_messages -= main_actor }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/scala_console.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,155 @@
+/* Title: Tools/jEdit/src/scala_console.scala
+ Author: Makarius
+
+Scala instance of Console plugin.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import console.{Console, ConsolePane, Shell, Output}
+
+import org.gjt.sp.jedit.{jEdit, JARClassLoader}
+import org.gjt.sp.jedit.MiscUtilities
+
+import java.io.{File, OutputStream, Writer, PrintWriter}
+
+import scala.tools.nsc.{Interpreter, GenericRunnerSettings, NewLinePrintWriter, ConsoleWriter}
+import scala.collection.mutable
+
+
+class Scala_Console extends Shell("Scala")
+{
+ /* reconstructed jEdit/plugin classpath */
+
+ private def reconstruct_classpath(): String =
+ {
+ def find_jars(start: String): List[String] =
+ if (start != null)
+ Standard_System.find_files(new File(start),
+ entry => entry.isFile && entry.getName.endsWith(".jar")).map(_.getAbsolutePath)
+ else Nil
+ val path = find_jars(jEdit.getSettingsDirectory) ::: find_jars(jEdit.getJEditHome)
+ path.mkString(File.pathSeparator)
+ }
+
+
+ /* global state -- owned by Swing thread */
+
+ private var interpreters = Map[Console, Interpreter]()
+
+ private var global_console: Console = null
+ private var global_out: Output = null
+ private var global_err: Output = null
+
+ private val console_stream = new OutputStream
+ {
+ val buf = new StringBuilder
+ override def flush()
+ {
+ val str = Standard_System.decode_permissive_utf8(buf.toString)
+ buf.clear
+ if (global_out == null) System.out.print(str)
+ else Swing_Thread.now { global_out.writeAttrs(null, str) }
+ }
+ override def close() { flush () }
+ def write(byte: Int) { buf.append(byte.toChar) }
+ }
+
+ private val console_writer = new Writer
+ {
+ def flush() {}
+ def close() {}
+
+ def write(cbuf: Array[Char], off: Int, len: Int)
+ {
+ if (len > 0) write(new String(cbuf.slice(off, off + len)))
+ }
+
+ override def write(str: String)
+ {
+ if (global_out == null) System.out.println(str)
+ else global_out.print(null, str)
+ }
+ }
+
+ private def with_console[A](console: Console, out: Output, err: Output)(e: => A): A =
+ {
+ global_console = console
+ global_out = out
+ global_err = if (err == null) out else err
+ val res = Exn.capture { scala.Console.withOut(console_stream)(e) }
+ console_stream.flush
+ global_console = null
+ global_out = null
+ global_err = null
+ Exn.release(res)
+ }
+
+ private def report_error(str: String)
+ {
+ if (global_console == null || global_err == null) System.err.println(str)
+ else Swing_Thread.now { global_err.print(global_console.getErrorColor, str) }
+ }
+
+
+ /* jEdit console methods */
+
+ override def openConsole(console: Console)
+ {
+ val settings = new GenericRunnerSettings(report_error)
+ settings.classpath.value = reconstruct_classpath()
+
+ val interp = new Interpreter(settings, new PrintWriter(console_writer, true))
+ {
+ override def parentClassLoader = new JARClassLoader
+ }
+ interp.setContextClassLoader
+ interp.bind("view", "org.gjt.sp.jedit.View", console.getView)
+ interp.bind("console", "console.Console", console)
+ interp.interpret("import isabelle.jedit.Isabelle")
+
+ interpreters += (console -> interp)
+ }
+
+ override def closeConsole(console: Console)
+ {
+ interpreters -= console
+ }
+
+ override def printInfoMessage(out: Output)
+ {
+ out.print(null,
+ "This shell evaluates Isabelle/Scala expressions.\n\n" +
+ "The following special toplevel bindings are provided:\n" +
+ " view -- current jEdit/Swing view (e.g. view.getBuffer, view.getTextArea)\n" +
+ " console -- jEdit Console plugin instance\n" +
+ " Isabelle -- Isabelle plugin instance (e.g. Isabelle.system, Isabelle.session)\n")
+ }
+
+ override def printPrompt(console: Console, out: Output)
+ {
+ out.writeAttrs(ConsolePane.colorAttributes(console.getInfoColor), "scala>")
+ out.writeAttrs(ConsolePane.colorAttributes(console.getPlainColor), " ")
+ }
+
+ override def execute(console: Console, input: String, out: Output, err: Output, command: String)
+ {
+ val interp = interpreters(console)
+ with_console(console, out, err) { interp.interpret(command) }
+ if (err != null) err.commandDone()
+ out.commandDone()
+ }
+
+ override def stop(console: Console)
+ {
+ closeConsole(console)
+ console.clear
+ openConsole(console)
+ val out = console.getOutput
+ out.commandDone
+ printPrompt(console, out)
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/services.xml Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,20 @@
+<?xml version="1.0"?>
+<!DOCTYPE SERVICES SYSTEM "services.dtd">
+
+<SERVICES>
+ <SERVICE NAME="UTF-8-Isabelle" CLASS="org.gjt.sp.jedit.io.Encoding">
+ new isabelle.jedit.Isabelle_Encoding();
+ </SERVICE>
+ <SERVICE NAME="isabelle" CLASS="sidekick.SideKickParser">
+ new isabelle.jedit.Isabelle_Sidekick_Default();
+ </SERVICE>
+ <SERVICE NAME="isabelle-raw" CLASS="sidekick.SideKickParser">
+ new isabelle.jedit.Isabelle_Sidekick_Raw();
+ </SERVICE>
+ <SERVICE NAME="isabelle" CLASS="gatchan.jedit.hyperlinks.HyperlinkSource">
+ new isabelle.jedit.Isabelle_Hyperlinks();
+ </SERVICE>
+ <SERVICE CLASS="console.Shell" NAME="Scala">
+ new isabelle.jedit.Scala_Console();
+ </SERVICE>
+</SERVICES>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/session_dockable.scala Wed Jun 08 22:13:49 2011 +0200
@@ -0,0 +1,104 @@
+/* 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, ScrollPane, TabbedPane, Component, Swing}
+import scala.swing.event.{ButtonClicked, SelectionChanged}
+
+import java.awt.BorderLayout
+import javax.swing.border.{BevelBorder, SoftBevelBorder}
+
+import org.gjt.sp.jedit.View
+
+
+class Session_Dockable(view: View, position: String) extends Dockable(view: View, position: String)
+{
+ /* main tabs */
+
+ private val readme = new HTML_Panel(Isabelle.system, "SansSerif", 14)
+ readme.render_document(Isabelle.system.try_read(List("$JEDIT_HOME/README.html")))
+
+ private val syslog = new TextArea(Isabelle.session.syslog())
+ syslog.editable = false
+
+ private val tabs = new TabbedPane {
+ pages += new TabbedPane.Page("README", Component.wrap(readme))
+ pages += new TabbedPane.Page("System log", new ScrollPane(syslog))
+
+ selection.index =
+ {
+ val index = Isabelle.Int_Property("session-panel.selection", 0)
+ if (index >= pages.length) 0 else index
+ }
+ listenTo(selection)
+ reactions += {
+ case SelectionChanged(_) =>
+ Isabelle.Int_Property("session-panel.selection") = selection.index
+ }
+ }
+
+ set_content(tabs)
+
+
+ /* controls */
+
+ val session_phase = new Label(Isabelle.session.phase.toString)
+ session_phase.border = new SoftBevelBorder(BevelBorder.LOWERED)
+ session_phase.tooltip = "Prover status"
+
+ private val interrupt = new Button("Interrupt") {
+ reactions += { case ButtonClicked(_) => Isabelle.session.interrupt }
+ }
+ interrupt.tooltip = "Broadcast interrupt to all prover tasks"
+
+ private val logic = Isabelle.logic_selector(Isabelle.Property("logic"))
+ logic.listenTo(logic.selection)
+ logic.reactions += {
+ case SelectionChanged(_) => Isabelle.Property("logic") = logic.selection.item.name
+ }
+
+ private val controls =
+ new FlowPanel(FlowPanel.Alignment.Right)(session_phase, interrupt, logic)
+ add(controls.peer, BorderLayout.NORTH)
+
+
+ /* main actor */
+
+ private val main_actor = actor {
+ loop {
+ react {
+ case result: Isabelle_Process.Result =>
+ if (result.is_syslog)
+ Swing_Thread.now {
+ val text = Isabelle.session.syslog()
+ if (text != syslog.text) {
+ syslog.text = text
+ }
+ }
+
+ case phase: Session.Phase =>
+ Swing_Thread.now { session_phase.text = " " + phase.toString + " " }
+
+ case bad => System.err.println("Session_Dockable: ignoring bad message " + bad)
+ }
+ }
+ }
+
+ override def init() {
+ Isabelle.session.raw_messages += main_actor
+ Isabelle.session.phase_changed += main_actor
+ }
+
+ override def exit() {
+ Isabelle.session.raw_messages -= main_actor
+ Isabelle.session.phase_changed -= main_actor
+ }
+}
--- a/src/Tools/subtyping.ML Wed Jun 08 17:01:07 2011 +0200
+++ b/src/Tools/subtyping.ML Wed Jun 08 22:13:49 2011 +0200
@@ -246,7 +246,7 @@
| gen cs _ (Free (_, T)) tye_idx = (T, tye_idx, cs)
| gen cs _ (Var (_, T)) tye_idx = (T, tye_idx, cs)
| gen cs bs (Bound i) tye_idx =
- (snd (nth bs i handle Subscript => err_loose i), tye_idx, cs)
+ (snd (nth bs i handle General.Subscript => err_loose i), tye_idx, cs)
| gen cs bs (Abs (x, T, t)) tye_idx =
let val (U, tye_idx', cs') = gen cs ((x, T) :: bs) t tye_idx
in (T --> U, tye_idx', cs') end
@@ -636,7 +636,7 @@
let val T' = T;
in (Var (xi, T'), T') end
| insert bs (Bound i) =
- let val T = nth bs i handle Subscript => err_loose i;
+ let val T = nth bs i handle General.Subscript => err_loose i;
in (Bound i, T) end
| insert bs (Abs (x, T, t)) =
let
@@ -671,7 +671,7 @@
| inf _ (t as (Free (_, T))) tye_idx = (t, T, tye_idx)
| inf _ (t as (Var (_, T))) tye_idx = (t, T, tye_idx)
| inf bs (t as (Bound i)) tye_idx =
- (t, snd (nth bs i handle Subscript => err_loose i), tye_idx)
+ (t, snd (nth bs i handle General.Subscript => err_loose i), tye_idx)
| inf bs (Abs (x, T, t)) tye_idx =
let val (t', U, tye_idx') = inf ((x, T) :: bs) t tye_idx
in (Abs (x, T, t'), T --> U, tye_idx') end