merged
authorwenzelm
Tue, 06 Aug 2013 23:24:10 +0200
changeset 52882 45678f8e7a0f
parent 52872 fd14b0ead643 (current diff)
parent 52881 4eb44754f1bb (diff)
child 52883 0a7c97c76f46
child 52885 9e4bae21494d
merged
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Tue Aug 06 23:24:10 2013 +0200
@@ -10,62 +10,61 @@
 
 text{* Application of polynomial as a real function. *}
 
-primrec poly :: "'a list => 'a  => ('a::{comm_ring})"
+primrec poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a::comm_ring"
 where
   poly_Nil:  "poly [] x = 0"
-| poly_Cons: "poly (h#t) x = h + x * poly t x"
+| poly_Cons: "poly (h # t) x = h + x * poly t x"
 
 
 subsection{*Arithmetic Operations on Polynomials*}
 
 text{*addition*}
-primrec padd :: "['a list, 'a list] => ('a::comm_ring_1) list"  (infixl "+++" 65)
+primrec padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"  (infixl "+++" 65)
 where
   padd_Nil:  "[] +++ l2 = l2"
-| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t
-                            else (h + hd l2)#(t +++ tl l2))"
+| padd_Cons: "(h # t) +++ l2 = (if l2 = [] then h # t else (h + hd l2) # (t +++ tl l2))"
 
 text{*Multiplication by a constant*}
-primrec cmult :: "['a :: comm_ring_1, 'a list] => 'a list"  (infixl "%*" 70)
+primrec cmult :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70)
 where
   cmult_Nil:  "c %* [] = []"
-| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
+| cmult_Cons: "c %* (h # t) = (c * h) # (c %* t)"
 
 text{*Multiplication by a polynomial*}
-primrec pmult :: "['a list, 'a list] => ('a::comm_ring_1) list"  (infixl "***" 70)
+primrec pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"  (infixl "***" 70)
 where
   pmult_Nil:  "[] *** l2 = []"
-| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
-                              else (h %* l2) +++ ((0) # (t *** l2)))"
+| pmult_Cons: "(h # t) *** l2 =
+    (if t = [] then h %* l2 else (h %* l2) +++ (0 # (t *** l2)))"
 
 text{*Repeated multiplication by a polynomial*}
-primrec mulexp :: "[nat, 'a list, 'a  list] => ('a ::comm_ring_1) list"
+primrec mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"
 where
   mulexp_zero:  "mulexp 0 p q = q"
 | mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
 
 text{*Exponential*}
-primrec pexp :: "['a list, nat] => ('a::comm_ring_1) list"  (infixl "%^" 80)
+primrec pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a::comm_ring_1 list"  (infixl "%^" 80)
 where
   pexp_0:   "p %^ 0 = [1]"
 | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
 
 text{*Quotient related value of dividing a polynomial by x + a*}
 (* Useful for divisor properties in inductive proofs *)
-primrec pquot :: "['a list, 'a::field] => 'a list"
+primrec pquot :: "'a list \<Rightarrow> 'a::field \<Rightarrow> 'a list"
 where
-  pquot_Nil: "pquot [] a= []"
-| pquot_Cons:
-    "pquot (h#t) a = (if t = [] then [h] else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
+  pquot_Nil: "pquot [] a = []"
+| pquot_Cons: "pquot (h # t) a =
+    (if t = [] then [h] else (inverse a * (h - hd (pquot t a))) # pquot t a)"
 
 
 text{*normalization of polynomials (remove extra 0 coeff)*}
-primrec pnormalize :: "('a::comm_ring_1) list => 'a list"
+primrec pnormalize :: "'a::comm_ring_1 list \<Rightarrow> 'a list"
 where
   pnormalize_Nil:  "pnormalize [] = []"
-| pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = [])
-                                     then (if (h = 0) then [] else [h])
-                                     else (h#(pnormalize p)))"
+| pnormalize_Cons: "pnormalize (h # p) =
+    (if (pnormalize p = []) then (if h = 0 then [] else [h])
+     else (h # pnormalize p))"
 
 definition "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
 definition "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
@@ -74,20 +73,18 @@
 definition poly_minus :: "'a list => ('a :: comm_ring_1) list"  ("-- _" [80] 80)
   where "-- p = (- 1) %* p"
 
-definition divides :: "[('a::comm_ring_1) list, 'a list] => bool"  (infixl "divides" 70)
-  where "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
+definition divides :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
+  where "p1 divides p2 = (\<exists>q. poly p2 = poly (p1 *** q))"
 
-definition order :: "('a::comm_ring_1) => 'a list => nat" --{*order of a polynomial*}
-  where "order a p = (SOME n. ([-a, 1] %^ n) divides p & ~ (([-a, 1] %^ (Suc n)) divides p))"
+definition order :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> nat" --{*order of a polynomial*}
+  where "order a p = (SOME n. ([-a, 1] %^ n) divides p & ~ (([-a, 1] %^ Suc n) divides p))"
 
-definition degree :: "('a::comm_ring_1) list => nat" --{*degree of a polynomial*}
+definition degree :: "'a::comm_ring_1 list \<Rightarrow> nat" --{*degree of a polynomial*}
   where "degree p = length (pnormalize p) - 1"
 
-definition
-  rsquarefree :: "('a::comm_ring_1) list => bool" where
-     --{*squarefree polynomials --- NB with respect to real roots only.*}
-  "rsquarefree p = (poly p \<noteq> poly [] &
-                     (\<forall>a. (order a p = 0) | (order a p = 1)))"
+definition rsquarefree :: "'a::comm_ring_1 list \<Rightarrow> bool"
+  where --{*squarefree polynomials --- NB with respect to real roots only.*}
+  "rsquarefree p = (poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1))"
 
 lemma padd_Nil2 [simp]: "p +++ [] = p"
   by (induct p) auto
@@ -110,63 +107,58 @@
 text{*Handy general properties*}
 
 lemma padd_commut: "b +++ a = a +++ b"
-  apply (subgoal_tac "\<forall>a. b +++ a = a +++ b")
-  apply (induct_tac [2] "b", auto)
+  apply (induct b arbitrary: a)
+  apply auto
   apply (rule padd_Cons [THEN ssubst])
-  apply (case_tac "aa", auto)
+  apply (case_tac aa)
+  apply auto
   done
 
-lemma padd_assoc [rule_format]: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
-  apply (induct "a", simp, clarify)
-  apply (case_tac b, simp_all)
+lemma padd_assoc: "(a +++ b) +++ c = a +++ (b +++ c)"
+  apply (induct a arbitrary: b c)
+  apply simp
+  apply (case_tac b)
+  apply simp_all
   done
 
-lemma poly_cmult_distr [rule_format]: "\<forall>q. a %* ( p +++ q) = (a %* p +++ a %* q)"
-  apply (induct p)
+lemma poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
+  apply (induct p arbitrary: q)
   apply simp
-  apply clarify
   apply (case_tac q)
   apply (simp_all add: distrib_left)
   done
 
-lemma pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
-  apply (induct t)
-  apply simp
-  apply (auto simp add: padd_commut)
-  done
+lemma pmult_by_x [simp]: "[0, 1] *** t = ((0)#t)"
+  by (induct t) (auto simp add: padd_commut)
 
 
 text{*properties of evaluation of polynomials.*}
 
 lemma poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
-  apply (subgoal_tac "\<forall>p2. poly (p1 +++ p2) x = poly (p1) x + poly (p2) x")
-  apply (induct_tac [2] "p1", auto)
+  apply (induct p1 arbitrary: p2)
+  apply auto
   apply (case_tac "p2")
   apply (auto simp add: distrib_left)
   done
 
 lemma poly_cmult: "poly (c %* p) x = c * poly p x"
-  apply (induct "p")
-  apply (case_tac [2] "x=0")
+  apply (induct p)
+  apply simp
+  apply (cases "x = 0")
   apply (auto simp add: distrib_left mult_ac)
   done
 
 lemma poly_minus: "poly (-- p) x = - (poly p x)"
-  apply (simp add: poly_minus_def)
-  apply (auto simp add: poly_cmult)
-  done
+  by (simp add: poly_minus_def poly_cmult)
 
 lemma poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
-  apply (subgoal_tac "\<forall>p2. poly (p1 *** p2) x = poly p1 x * poly p2 x")
-  apply (simp (no_asm_simp))
-  apply (induct "p1")
-  apply (auto simp add: poly_cmult)
+  apply (induct p1 arbitrary: p2)
   apply (case_tac p1)
   apply (auto simp add: poly_cmult poly_add distrib_right distrib_left mult_ac)
   done
 
 lemma poly_exp: "poly (p %^ n) (x::'a::comm_ring_1) = (poly p x) ^ n"
-  by (induct "n") (auto simp add: poly_cmult poly_mult)
+  by (induct n) (auto simp add: poly_cmult poly_mult)
 
 text{*More Polynomial Evaluation Lemmas*}
 
@@ -177,45 +169,49 @@
   by (simp add: poly_mult mult_assoc)
 
 lemma poly_mult_Nil2 [simp]: "poly (p *** []) x = 0"
-  by (induct "p") auto
+  by (induct p) auto
 
 lemma poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
-  by (induct "n") (auto simp add: poly_mult mult_assoc)
+  by (induct n) (auto simp add: poly_mult mult_assoc)
 
 subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
  @{term "p(x)"} *}
 
-lemma lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-  apply (induct "t", safe)
+lemma poly_linear_rem: "\<exists>q r. h # t = [r] +++ [-a, 1] *** q"
+  apply (induct t arbitrary: h)
   apply (rule_tac x = "[]" in exI)
-  apply (rule_tac x = h in exI, simp)
-  apply (drule_tac x = aa in spec, safe)
+  apply (rule_tac x = h in exI)
+  apply simp
+  apply (drule_tac x = aa in meta_spec)
+  apply safe
   apply (rule_tac x = "r#q" in exI)
   apply (rule_tac x = "a*r + h" in exI)
-  apply (case_tac "q", auto)
+  apply (case_tac q)
+  apply auto
   done
 
-lemma poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-  using lemma_poly_linear_rem [where t = t and a = a] by auto
-
-
-lemma poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
+lemma poly_linear_divides: "poly p a = 0 \<longleftrightarrow> p = [] \<or> (\<exists>q. p = [-a, 1] *** q)"
   apply (auto simp add: poly_add poly_cmult distrib_left)
-  apply (case_tac "p", simp)
-  apply (cut_tac h = aa and t = list and a = a in poly_linear_rem, safe)
-  apply (case_tac "q", auto)
-  apply (drule_tac x = "[]" in spec, simp)
+  apply (case_tac p)
+  apply simp
+  apply (cut_tac h = aa and t = list and a = a in poly_linear_rem)
+  apply safe
+  apply (case_tac q)
+  apply auto
+  apply (drule_tac x = "[]" in spec)
+  apply simp
   apply (auto simp add: poly_add poly_cmult add_assoc)
-  apply (drule_tac x = "aa#lista" in spec, auto)
+  apply (drule_tac x = "aa#lista" in spec)
+  apply auto
   done
 
-lemma lemma_poly_length_mult [simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
-  by (induct p) auto
+lemma lemma_poly_length_mult [simp]: "length (k %* p +++  (h # (a %* p))) = Suc (length p)"
+  by (induct p arbitrary: h k a) auto
 
-lemma lemma_poly_length_mult2 [simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
-  by (induct p) auto
+lemma lemma_poly_length_mult2 [simp]: "length (k %* p +++  (h # p)) = Suc (length p)"
+  by (induct p arbitrary: h k) auto
 
-lemma poly_length_mult [simp]: "length([-a,1] *** q) = Suc (length q)"
+lemma poly_length_mult [simp]: "length([-a, 1] *** q) = Suc (length q)"
   by auto
 
 
@@ -224,26 +220,23 @@
 lemma poly_cmult_length [simp]: "length (a %* p) = length p"
   by (induct p) auto
 
-lemma poly_add_length [rule_format]:
-  "\<forall>p2. length (p1 +++ p2) = (if (length p1 < length p2) then length p2 else length p1)"
-  apply (induct p1)
-  apply simp_all
-  apply arith
-  done
+lemma poly_add_length:
+  "length (p1 +++ p2) = (if (length p1 < length p2) then length p2 else length p1)"
+  by (induct p1 arbitrary: p2) auto
 
-lemma poly_root_mult_length [simp]: "length([a,b] *** p) = Suc (length p)"
+lemma poly_root_mult_length [simp]: "length ([a, b] *** p) = Suc (length p)"
   by simp
 
-lemma poly_mult_not_eq_poly_Nil [simp]: "(poly (p *** q) x \<noteq> poly [] x) =
-    (poly p x \<noteq> poly [] x & poly q x \<noteq> poly [] (x::'a::idom))"
+lemma poly_mult_not_eq_poly_Nil [simp]:
+  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] (x::'a::idom)"
   by (auto simp add: poly_mult)
 
-lemma poly_mult_eq_zero_disj: "(poly (p *** q) (x::'a::idom) = 0) = (poly p x = 0 | poly q x = 0)"
+lemma poly_mult_eq_zero_disj: "poly (p *** q) (x::'a::idom) = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
   by (auto simp add: poly_mult)
 
 text{*Normalisation Properties*}
 
-lemma poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
+lemma poly_normalized_nil: "pnormalize p = [] \<longrightarrow> poly p x = 0"
   by (induct p) auto
 
 text{*A nontrivial polynomial of degree n has no more than n roots*}
@@ -251,16 +244,21 @@
 lemma poly_roots_index_lemma0 [rule_format]:
    "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
     --> (\<exists>i. \<forall>x. (poly p x = (0::'a::idom)) --> (\<exists>m. (m \<le> n & x = i m)))"
-  apply (induct "n", safe)
+  apply (induct n)
+  apply safe
   apply (rule ccontr)
-  apply (subgoal_tac "\<exists>a. poly p a = 0", safe)
-  apply (drule poly_linear_divides [THEN iffD1], safe)
+  apply (subgoal_tac "\<exists>a. poly p a = 0")
+  apply safe
+  apply (drule poly_linear_divides [THEN iffD1])
+  apply safe
   apply (drule_tac x = q in spec)
   apply (drule_tac x = x in spec)
   apply (simp del: poly_Nil pmult_Cons)
   apply (erule exE)
-  apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec, safe)
-  apply (drule poly_mult_eq_zero_disj [THEN iffD1], safe)
+  apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec)
+  apply safe
+  apply (drule poly_mult_eq_zero_disj [THEN iffD1])
+  apply safe
   apply (drule_tac x = "Suc (length q)" in spec)
   apply (auto simp add: field_simps)
   apply (drule_tac x = xa in spec)
@@ -278,7 +276,8 @@
 lemma poly_roots_finite_lemma:
   "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
     \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N & x = i n)"
-  apply (drule poly_roots_index_length0, safe)
+  apply (drule poly_roots_index_length0)
+  apply safe
   apply (rule_tac x = "Suc (length p)" in exI)
   apply (rule_tac x = i in exI)
   apply (simp add: less_Suc_eq_le)
@@ -286,22 +285,25 @@
 
 
 lemma real_finite_lemma:
-  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j & x = j!n)"
+  assumes "\<forall>x. P x \<longrightarrow> (\<exists>n. n < length j & x = j!n)"
   shows "finite {(x::'a::idom). P x}"
 proof -
   let ?M = "{x. P x}"
   let ?N = "set j"
-  have "?M \<subseteq> ?N" using P by auto
+  have "?M \<subseteq> ?N" using assms by auto
   then show ?thesis using finite_subset by auto
 qed
 
 lemma poly_roots_index_lemma [rule_format]:
   "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
     \<longrightarrow> (\<exists>i. \<forall>x. (poly p x = (0::'a::{idom})) \<longrightarrow> x \<in> set i)"
-  apply (induct "n", safe)
+  apply (induct n)
+  apply safe
   apply (rule ccontr)
-  apply (subgoal_tac "\<exists>a. poly p a = 0", safe)
-  apply (drule poly_linear_divides [THEN iffD1], safe)
+  apply (subgoal_tac "\<exists>a. poly p a = 0")
+  apply safe
+  apply (drule poly_linear_divides [THEN iffD1])
+  apply safe
   apply (drule_tac x = q in spec)
   apply (drule_tac x = x in spec)
   apply (auto simp del: poly_Nil pmult_Cons)
@@ -348,21 +350,21 @@
 lemma UNIV_ring_char_0_infinte: "\<not> finite (UNIV:: ('a::ring_char_0) set)"
 proof
   assume F: "finite (UNIV :: 'a set)"
-  have th0: "of_nat ` UNIV \<subseteq> (UNIV:: 'a set)" by simp
+  have th0: "of_nat ` UNIV \<subseteq> (UNIV :: 'a set)" by simp
   from finite_subset[OF th0 F] have th: "finite (of_nat ` UNIV :: 'a set)" .
-  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) (UNIV)"
+  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) UNIV"
     unfolding inj_on_def by auto
   from finite_imageD[OF th th'] UNIV_nat_infinite
   show False by blast
 qed
 
-lemma poly_roots_finite: "(poly p \<noteq> poly []) = finite {x. poly p x = (0::'a::{idom, ring_char_0})}"
+lemma poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = (0::'a::{idom,ring_char_0})}"
 proof
-  assume H: "poly p \<noteq> poly []"
-  show "finite {x. poly p x = (0::'a)}"
-    using H
+  assume "poly p \<noteq> poly []"
+  then show "finite {x. poly p x = (0::'a)}"
     apply -
-    apply (erule contrapos_np, rule ext)
+    apply (erule contrapos_np)
+    apply (rule ext)
     apply (rule ccontr)
     apply (clarify dest!: poly_roots_finite_lemma')
     using finite_subset
@@ -376,7 +378,8 @@
   qed
 next
   assume "finite {x. poly p x = (0\<Colon>'a)}"
-  then show "poly p \<noteq> poly []" using UNIV_ring_char_0_infinte by auto
+  then show "poly p \<noteq> poly []"
+    using UNIV_ring_char_0_infinte by auto
 qed
 
 text{*Entirety and Cancellation for polynomials*}
@@ -387,21 +390,21 @@
   by (auto simp add: poly_roots_finite poly_mult Collect_disj_eq)
 
 lemma poly_entire:
-  "(poly (p *** q) =
-    poly ([]::('a::{idom,ring_char_0}) list)) = ((poly p = poly []) | (poly q = poly []))"
+  "poly (p *** q) = poly ([]::('a::{idom,ring_char_0}) list) \<longleftrightarrow>
+    (poly p = poly [] \<or> poly q = poly [])"
   apply (auto dest: fun_cong simp add: poly_entire_lemma poly_mult)
   apply (blast intro: ccontr dest: poly_entire_lemma poly_mult [THEN subst])
   done
 
 lemma poly_entire_neg:
-  "(poly (p *** q) \<noteq> poly ([]::('a::{idom,ring_char_0}) list)) =
-    ((poly p \<noteq> poly []) & (poly q \<noteq> poly []))"
+  "poly (p *** q) \<noteq> poly ([]::('a::{idom,ring_char_0}) list) \<longleftrightarrow>
+    poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
   by (simp add: poly_entire)
 
 lemma fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
   by auto
 
-lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
+lemma poly_add_minus_zero_iff: "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
   by (auto simp add: field_simps poly_add poly_minus_def fun_eq poly_cmult)
 
 lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
@@ -409,26 +412,27 @@
 
 lemma poly_mult_left_cancel:
   "(poly (p *** q) = poly (p *** r)) =
-    (poly p = poly ([]::('a::{idom, ring_char_0}) list) | poly q = poly r)"
+    (poly p = poly ([]::('a::{idom,ring_char_0}) list) | poly q = poly r)"
   apply (rule_tac p1 = "p *** q" in poly_add_minus_zero_iff [THEN subst])
   apply (auto simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
   done
 
 lemma poly_exp_eq_zero [simp]:
-  "(poly (p %^ n) = poly ([]::('a::idom) list)) = (poly p = poly [] & n \<noteq> 0)"
+  "poly (p %^ n) = poly ([]::('a::idom) list) \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
   apply (simp only: fun_eq add: HOL.all_simps [symmetric])
   apply (rule arg_cong [where f = All])
   apply (rule ext)
-  apply (induct_tac "n")
+  apply (induct_tac n)
   apply (auto simp add: poly_mult)
   done
 
-lemma poly_prime_eq_zero [simp]: "poly [a,(1::'a::comm_ring_1)] \<noteq> poly []"
+lemma poly_prime_eq_zero [simp]: "poly [a, 1::'a::comm_ring_1] \<noteq> poly []"
   apply (simp add: fun_eq)
-  apply (rule_tac x = "1 - a" in exI, simp)
+  apply (rule_tac x = "1 - a" in exI)
+  apply simp
   done
 
-lemma poly_exp_prime_eq_zero [simp]: "(poly ([a, (1::'a::idom)] %^ n) \<noteq> poly [])"
+lemma poly_exp_prime_eq_zero [simp]: "poly ([a, (1::'a::idom)] %^ n) \<noteq> poly []"
   by auto
 
 text{*A more constructive notion of polynomials being trivial*}
@@ -437,8 +441,10 @@
   "poly (h # t) = poly [] \<Longrightarrow> h = (0::'a::{idom,ring_char_0}) & poly t = poly []"
   apply (simp add: fun_eq)
   apply (case_tac "h = 0")
-  apply (drule_tac [2] x = 0 in spec, auto)
-  apply (case_tac "poly t = poly []", simp)
+  apply (drule_tac [2] x = 0 in spec)
+  apply auto
+  apply (case_tac "poly t = poly []")
+  apply simp
 proof -
   fix x
   assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
@@ -450,7 +456,7 @@
   show "poly t x = (0\<Colon>'a)" by simp
 qed
 
-lemma poly_zero: "(poly p = poly []) = list_all (%c. c = (0::'a::{idom,ring_char_0})) p"
+lemma poly_zero: "poly p = poly [] \<longleftrightarrow> list_all (\<lambda>c. c = (0::'a::{idom,ring_char_0})) p"
   apply (induct p)
   apply simp
   apply (rule iffI)
@@ -461,7 +467,7 @@
 
 text{*Basics of divisibility.*}
 
-lemma poly_primes: "([a, (1::'a::idom)] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)"
+lemma poly_primes: "[a, (1::'a::idom)] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
   apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult distrib_right [symmetric])
   apply (drule_tac x = "-a" in spec)
   apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
@@ -477,7 +483,8 @@
   done
 
 lemma poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
-  apply (simp add: divides_def, safe)
+  apply (simp add: divides_def)
+  apply safe
   apply (rule_tac x = "qa *** qaa" in exI)
   apply (auto simp add: poly_mult fun_eq mult_assoc)
   done
@@ -495,25 +502,26 @@
   by (blast intro: poly_divides_exp poly_divides_trans)
 
 lemma poly_divides_add: "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
-  apply (simp add: divides_def, auto)
+  apply (simp add: divides_def)
+  apply auto
   apply (rule_tac x = "qa +++ qaa" in exI)
   apply (auto simp add: poly_add fun_eq poly_mult distrib_left)
   done
 
 lemma poly_divides_diff: "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
-  apply (simp add: divides_def, auto)
+  apply (auto simp add: divides_def)
   apply (rule_tac x = "qaa +++ -- qa" in exI)
   apply (auto simp add: poly_add fun_eq poly_mult poly_minus algebra_simps)
   done
 
-lemma poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q"
+lemma poly_divides_diff2: "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
   apply (erule poly_divides_diff)
   apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
   done
 
-lemma poly_divides_zero: "poly p = poly [] ==> q divides p"
+lemma poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
   apply (simp add: divides_def)
-  apply (rule exI[where x="[]"])
+  apply (rule exI [where x = "[]"])
   apply (auto simp add: fun_eq poly_mult)
   done
 
@@ -525,33 +533,40 @@
 
 text{*At last, we can consider the order of a root.*}
 
-
 lemma poly_order_exists_lemma [rule_format]:
   "\<forall>p. length p = d \<longrightarrow> poly p \<noteq> poly [] \<longrightarrow>
     (\<exists>n q. p = mulexp n [-a, (1::'a::{idom,ring_char_0})] q & poly q a \<noteq> 0)"
   apply (induct "d")
-  apply (simp add: fun_eq, safe)
+  apply (simp add: fun_eq)
+  apply safe
   apply (case_tac "poly p a = 0")
-  apply (drule_tac poly_linear_divides [THEN iffD1], safe)
+  apply (drule_tac poly_linear_divides [THEN iffD1])
+  apply safe
   apply (drule_tac x = q in spec)
-  apply (drule_tac poly_entire_neg [THEN iffD1], safe, force)
+  apply (drule_tac poly_entire_neg [THEN iffD1])
+  apply safe
+  apply force
   apply (rule_tac x = "Suc n" in exI)
   apply (rule_tac x = qa in exI)
   apply (simp del: pmult_Cons)
-  apply (rule_tac x = 0 in exI, force)
+  apply (rule_tac x = 0 in exI)
+  apply force
   done
 
 (* FIXME: Tidy up *)
 lemma poly_order_exists:
-     "[| length p = d; poly p \<noteq> poly [] |]
-      ==> \<exists>n. ([-a, 1] %^ n) divides p &
-                ~(([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)"
-  apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)
-  apply (rule_tac x = n in exI, safe)
+  "length p = d \<Longrightarrow> poly p \<noteq> poly [] \<Longrightarrow>
+    \<exists>n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)"
+  apply (drule poly_order_exists_lemma [where a=a])
+  apply assumption
+  apply clarify
+  apply (rule_tac x = n in exI)
+  apply safe
   apply (unfold divides_def)
   apply (rule_tac x = q in exI)
-  apply (induct_tac n, simp)
-  apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult distrib_left mult_ac)
+  apply (induct_tac n)
+  apply simp
+  apply (simp add: poly_add poly_cmult poly_mult distrib_left mult_ac)
   apply safe
   apply (subgoal_tac "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** qa)")
   apply simp
@@ -567,19 +582,18 @@
 lemma poly_one_divides [simp]: "[1] divides p"
   by (auto simp: divides_def)
 
-lemma poly_order: "poly p \<noteq> poly []
-      ==> EX! n. ([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p &
-                 ~(([-a, 1] %^ (Suc n)) divides p)"
+lemma poly_order: "poly p \<noteq> poly [] \<Longrightarrow>
+    \<exists>! n. ([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
   apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
   apply (cut_tac x = y and y = n in less_linear)
   apply (drule_tac m = n in poly_exp_divides)
   apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
-              simp del: pmult_Cons pexp_Suc)
+    simp del: pmult_Cons pexp_Suc)
   done
 
 text{*Order*}
 
-lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n"
+lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> EX! n. P n \<Longrightarrow> P n"
   by (blast intro: someI2)
 
 lemma order:
@@ -597,14 +611,14 @@
     ~(([-a, 1] %^ (Suc(order a p))) divides p)"
   by (simp add: order del: pexp_Suc)
 
-lemma order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p;
-  ~(([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)
-      |] ==> (n = order a p)"
+lemma order_unique: "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow>
+  \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p) \<Longrightarrow> n = order a p"
   using order [of a n p] by auto
 
-lemma order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p &
-         ~(([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p))
-      ==> (n = order a p)"
+lemma order_unique_lemma:
+  "(poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and>
+    \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)) \<Longrightarrow>
+    n = order a p"
   by (blast intro: order_unique)
 
 lemma order_poly: "poly p = poly q ==> order a p = order a q"
@@ -613,40 +627,44 @@
 lemma pexp_one [simp]: "p %^ (Suc 0) = p"
   by (induct p) simp_all
 
-lemma lemma_order_root [rule_format]:
-  "\<forall>p a. 0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p
-    --> poly p a = 0"
-  apply (induct n)
+lemma lemma_order_root:
+  "0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
+  apply (induct n arbitrary: p a)
   apply blast
   apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
   done
 
-lemma order_root: "(poly p a = (0::'a::{idom,ring_char_0})) = ((poly p = poly []) | order a p \<noteq> 0)"
-  apply (case_tac "poly p = poly []", auto)
-  apply (simp add: poly_linear_divides del: pmult_Cons, safe)
+lemma order_root: "poly p a = (0::'a::{idom,ring_char_0}) \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
+  apply (cases "poly p = poly []")
+  apply auto
+  apply (simp add: poly_linear_divides del: pmult_Cons)
+  apply safe
   apply (drule_tac [!] a = a in order2)
   apply (rule ccontr)
-  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
-  using neq0_conv
-  apply (blast intro: lemma_order_root)
+  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons)
+  apply blast
+  using neq0_conv apply (blast intro: lemma_order_root)
   done
 
-lemma order_divides: "(([-a, 1::'a::{idom,ring_char_0}] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)"
-  apply (case_tac "poly p = poly []", auto)
+lemma order_divides: "([-a, 1::'a::{idom,ring_char_0}] %^ n) divides p \<longleftrightarrow>
+    poly p = poly [] \<or> n \<le> order a p"
+  apply (cases "poly p = poly []")
+  apply auto
   apply (simp add: divides_def fun_eq poly_mult)
   apply (rule_tac x = "[]" in exI)
-  apply (auto dest!: order2 [where a=a]
-              intro: poly_exp_divides simp del: pexp_Suc)
+  apply (auto dest!: order2 [where a = a] intro: poly_exp_divides simp del: pexp_Suc)
   done
 
 lemma order_decomp:
   "poly p \<noteq> poly [] \<Longrightarrow>
-    \<exists>q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) &
-      ~([-a, 1::'a::{idom,ring_char_0}] divides q)"
+    \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and>
+      \<not> ([-a, 1::'a::{idom,ring_char_0}] divides q)"
   apply (unfold divides_def)
   apply (drule order2 [where a = a])
-  apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
-  apply (rule_tac x = q in exI, safe)
+  apply (simp add: divides_def del: pexp_Suc pmult_Cons)
+  apply safe
+  apply (rule_tac x = q in exI)
+  apply safe
   apply (drule_tac x = qa in spec)
   apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
   done
@@ -659,25 +677,29 @@
   apply (auto simp add: poly_entire simp del: pmult_Cons)
   apply (drule_tac a = a in order2)+
   apply safe
-  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
+  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons)
+  apply safe
   apply (rule_tac x = "qa *** qaa" in exI)
   apply (simp add: poly_mult mult_ac del: pmult_Cons)
   apply (drule_tac a = a in order_decomp)+
   apply safe
-  apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
+  apply (subgoal_tac "[-a, 1] divides (qa *** qaa) ")
   apply (simp add: poly_primes del: pmult_Cons)
   apply (auto simp add: divides_def simp del: pmult_Cons)
   apply (rule_tac x = qb in exI)
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
-  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
-  apply (drule poly_mult_left_cancel [THEN iffD1], force)
+  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) =
+    poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
+  apply (drule poly_mult_left_cancel [THEN iffD1])
+  apply force
+  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) =
+    poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
+  apply (drule poly_mult_left_cancel [THEN iffD1])
+  apply force
   apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
   done
 
-lemma order_root2: "poly p \<noteq> poly [] ==> (poly p a = 0) = (order (a::'a::{idom,ring_char_0}) p \<noteq> 0)"
-  by (rule order_root [THEN ssubst], auto)
-
+lemma order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order (a::'a::{idom,ring_char_0}) p \<noteq> 0"
+  by (rule order_root [THEN ssubst]) auto
 
 lemma pmult_one [simp]: "[1] *** p = p"
   by auto
@@ -686,17 +708,20 @@
   by (simp add: fun_eq)
 
 lemma rsquarefree_decomp:
-  "[| rsquarefree p; poly p a = (0::'a::{idom,ring_char_0}) |]
-    ==> \<exists>q. (poly p = poly ([-a, 1] *** q)) & poly q a \<noteq> 0"
-  apply (simp add: rsquarefree_def, safe)
+  "rsquarefree p \<Longrightarrow> poly p a = (0::'a::{idom,ring_char_0}) \<Longrightarrow>
+    \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
+  apply (simp add: rsquarefree_def)
+  apply safe
   apply (frule_tac a = a in order_decomp)
   apply (drule_tac x = a in spec)
   apply (drule_tac a = a in order_root2 [symmetric])
   apply (auto simp del: pmult_Cons)
-  apply (rule_tac x = q in exI, safe)
+  apply (rule_tac x = q in exI)
+  apply safe
   apply (simp add: poly_mult fun_eq)
   apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
-  apply (simp add: divides_def del: pmult_Cons, safe)
+  apply (simp add: divides_def del: pmult_Cons)
+  apply safe
   apply (drule_tac x = "[]" in spec)
   apply (auto simp add: fun_eq)
   done
@@ -710,45 +735,55 @@
 
 text{*The degree of a polynomial.*}
 
-lemma lemma_degree_zero: "list_all (%c. c = 0) p \<longleftrightarrow>  pnormalize p = []"
+lemma lemma_degree_zero: "list_all (\<lambda>c. c = 0) p \<longleftrightarrow> pnormalize p = []"
   by (induct p) auto
 
-lemma degree_zero: "(poly p = poly ([]:: (('a::{idom,ring_char_0}) list))) \<Longrightarrow> (degree p = 0)"
+lemma degree_zero: "poly p = poly ([] :: 'a::{idom,ring_char_0} list) \<Longrightarrow> degree p = 0"
+  apply (cases "pnormalize p = []")
   apply (simp add: degree_def)
-  apply (case_tac "pnormalize p = []")
-  apply (auto simp add: poly_zero lemma_degree_zero )
+  apply (auto simp add: poly_zero lemma_degree_zero)
   done
 
-lemma pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
+lemma pnormalize_sing: "pnormalize [x] = [x] \<longleftrightarrow> x \<noteq> 0"
+  by simp
 
-lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
+lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
+  by simp
 
-lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
+lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c # p)"
   unfolding pnormal_def by simp
 
-lemma pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
+lemma pnormal_tail: "p \<noteq> [] \<Longrightarrow> pnormal (c # p) \<Longrightarrow> pnormal p"
   unfolding pnormal_def
-  apply (cases "pnormalize p = []", auto)
-  apply (cases "c = 0", auto)
+  apply (cases "pnormalize p = []")
+  apply auto
+  apply (cases "c = 0")
+  apply auto
   done
 
-lemma pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
-  apply (induct p, auto simp add: pnormal_def)
-  apply (case_tac "pnormalize p = []", auto)
-  apply (case_tac "a=0", auto)
+lemma pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
+  apply (induct p)
+  apply (auto simp add: pnormal_def)
+  apply (case_tac "pnormalize p = []")
+  apply auto
+  apply (case_tac "a = 0")
+  apply auto
   done
 
 lemma  pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
   unfolding pnormal_def length_greater_0_conv by blast
 
-lemma pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p"
-  apply (induct p, auto)
-  apply (case_tac "p = []", auto)
+lemma pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
+  apply (induct p)
+  apply auto
+  apply (case_tac "p = []")
+  apply auto
   apply (simp add: pnormal_def)
-  apply (rule pnormal_cons, auto)
+  apply (rule pnormal_cons)
+  apply auto
   done
 
-lemma pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)"
+lemma pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
   using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
 
 text{*Tidier versions of finiteness of roots.*}
@@ -759,13 +794,15 @@
 
 text{*bound for polynomial.*}
 
-lemma poly_mono: "abs(x) \<le> k ==> abs(poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
-  apply (induct "p", auto)
+lemma poly_mono: "abs x \<le> k \<Longrightarrow> abs (poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
+  apply (induct p)
+  apply auto
   apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
   apply (rule abs_triangle_ineq)
   apply (auto intro!: mult_mono simp add: abs_mult)
   done
 
-lemma poly_Sing: "poly [c] x = c" by simp
+lemma poly_Sing: "poly [c] x = c"
+  by simp
 
 end
--- a/src/Pure/PIDE/markup.scala	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Pure/PIDE/markup.scala	Tue Aug 06 23:24:10 2013 +0200
@@ -283,9 +283,10 @@
   val WARNING_MESSAGE = "warning_message"
   val ERROR_MESSAGE = "error_message"
 
-  val message: String => String =
+  val messages =
     Map(WRITELN -> WRITELN_MESSAGE, TRACING -> TRACING_MESSAGE,
-        WARNING -> WARNING_MESSAGE, ERROR -> ERROR_MESSAGE).withDefault((s: String) => s)
+        WARNING -> WARNING_MESSAGE, ERROR -> ERROR_MESSAGE)
+  val message: String => String = messages.withDefault((s: String) => s)
 
   val Return_Code = new Properties.Int("return_code")
 
--- a/src/Pure/PIDE/query_operation.ML	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Pure/PIDE/query_operation.ML	Tue Aug 06 23:24:10 2013 +0200
@@ -7,24 +7,40 @@
 
 signature QUERY_OPERATION =
 sig
+  val register_parallel: string -> (Toplevel.state -> string list -> (unit -> string) list) -> unit
   val register: string -> (Toplevel.state -> string list -> string) -> unit
 end;
 
 structure Query_Operation: QUERY_OPERATION =
 struct
 
-fun register name f =
+fun register_parallel name f =
   Command.print_function name
     (fn {args = instance :: args, ...} =>
         SOME {delay = NONE, pri = 0, persistent = false,
           print_fn = fn _ => fn state =>
             let
-              val msg = XML.Elem ((Markup.writelnN, []), [XML.Text (f state args)])
-                handle exn =>
-                  if Exn.is_interrupt exn then reraise exn
-                  else XML.Elem ((Markup.errorN, []), [XML.Text (ML_Compiler.exn_message exn)]);
-            in Output.result [(Markup.instanceN, instance)] (YXML.string_of msg) end}
+              fun result s = Output.result [(Markup.instanceN, instance)] s;
+              fun status m = result (Markup.markup_only m);
+              fun error_msg exn =
+                if Exn.is_interrupt exn then reraise exn
+                else
+                  XML.Elem ((Markup.errorN, []), [XML.Text (ML_Compiler.exn_message exn)])
+                  |> YXML.string_of |> result;
+
+              val _ = status Markup.running;
+              val outputs = f state args handle exn => (error_msg exn; []);
+              val _ = outputs |> Par_List.map (fn out =>
+                (case Exn.capture out () of
+                  Exn.Res s =>
+                    result (YXML.string_of (XML.Elem ((Markup.writelnN, []), [XML.Text s])))
+                | Exn.Exn exn => error_msg exn));
+              val _ = status Markup.finished;
+            in () end}
       | _ => NONE);
 
+fun register name f =
+  register_parallel name (fn st => fn args => [fn () => f st args]);
+
 end;
 
--- a/src/Tools/jEdit/etc/options	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Tools/jEdit/etc/options	Tue Aug 06 23:24:10 2013 +0200
@@ -83,4 +83,6 @@
 option gutter_warning_icon : string = "idea-icons/runConfigurations/testFailed.png"
 option gutter_legacy_icon : string = "idea-icons/general/balloonWarning.png"
 option gutter_error_icon : string = "idea-icons/runConfigurations/testError.png"
+option process_passive_icon : string = "idea-icons/process/step_passive.png"
+option process_active_icons : string = "idea-icons/process/step_1.png:idea-icons/process/step_2.png:idea-icons/process/step_3.png:idea-icons/process/step_4.png:idea-icons/process/step_5.png:idea-icons/process/step_6.png:idea-icons/process/step_7.png:idea-icons/process/step_8.png:idea-icons/process/step_9.png:idea-icons/process/step_10.png:idea-icons/process/step_11.png:idea-icons/process/step_12.png"
 
--- a/src/Tools/jEdit/src/find_dockable.scala	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Tools/jEdit/src/find_dockable.scala	Tue Aug 06 23:24:10 2013 +0200
@@ -99,12 +99,12 @@
 
   private val query_wrapped = Component.wrap(query)
 
-  private val find = new Button("Find") {
+  private val apply_query = new Button("Apply") {
     tooltip = "Find theorems meeting specified criteria"
     reactions += { case ButtonClicked(_) => find_theorems.apply_query(List(query.getText)) }
   }
 
-  private val locate = new Button("Locate") {
+  private val locate_query = new Button("Locate") {
     tooltip = "Locate context of current query within source text"
     reactions += { case ButtonClicked(_) => find_theorems.locate_query() }
   }
@@ -114,6 +114,7 @@
   }
 
   private val controls =
-    new FlowPanel(FlowPanel.Alignment.Right)(query_wrapped, find, locate, zoom)
+    new FlowPanel(FlowPanel.Alignment.Right)(
+      query_wrapped, find_theorems.animation, apply_query, locate_query, zoom)
   add(controls.peer, BorderLayout.NORTH)
 }
--- a/src/Tools/jEdit/src/jedit_lib.scala	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Tools/jEdit/src/jedit_lib.scala	Tue Aug 06 23:24:10 2013 +0200
@@ -10,10 +10,11 @@
 import isabelle._
 
 import java.awt.{Component, Container, Window, GraphicsEnvironment, Point, Rectangle}
+import javax.swing.{Icon, ImageIcon}
 
 import scala.annotation.tailrec
 
-import org.gjt.sp.jedit.{jEdit, Buffer, View}
+import org.gjt.sp.jedit.{jEdit, Buffer, View, GUIUtilities}
 import org.gjt.sp.jedit.buffer.JEditBuffer
 import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, TextAreaPainter}
 
@@ -222,5 +223,28 @@
       val average = string_width("mix") / (3 * unit)
       def apply(s: String): Double = if (s == "\n") 1.0 else string_width(s) / unit
     }
+
+
+  /* icons */
+
+  def load_icon(name: String): Icon =
+  {
+    val name1 =
+      if (name.startsWith("idea-icons/")) {
+        val file =
+          Isabelle_System.platform_file_url(Path.explode("$JEDIT_HOME/dist/jars/idea-icons.jar"))
+        "jar:" + file + "!/" + name
+      }
+      else name
+    val icon = GUIUtilities.loadIcon(name1)
+    if (icon.getIconWidth < 0 || icon.getIconHeight < 0) error("Bad icon: " + name)
+    else icon
+  }
+
+  def load_image_icon(name: String): ImageIcon =
+    load_icon(name) match {
+      case icon: ImageIcon => icon
+      case _ => error("Bad image icon: " + name)
+    }
 }
 
--- a/src/Tools/jEdit/src/plugin.scala	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Tools/jEdit/src/plugin.scala	Tue Aug 06 23:24:10 2013 +0200
@@ -50,6 +50,20 @@
 
   /* document model and view */
 
+  def document_snapshot(name: Document.Node.Name): Document.Snapshot =
+  {
+    Swing_Thread.require()
+
+    JEdit_Lib.jedit_buffer(name.node) match {
+      case Some(buffer) =>
+        document_model(buffer) match {
+          case Some(model) => model.snapshot
+          case None => session.snapshot(name)
+        }
+      case None => session.snapshot(name)
+    }
+  }
+
   def document_model(buffer: Buffer): Option[Document_Model] = Document_Model(buffer)
   def document_view(text_area: JEditTextArea): Option[Document_View] = Document_View(text_area)
 
--- a/src/Tools/jEdit/src/query_operation.scala	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Tools/jEdit/src/query_operation.scala	Tue Aug 06 23:24:10 2013 +0200
@@ -11,8 +11,10 @@
 import isabelle._
 
 import scala.actors.Actor._
+import scala.swing.Label
 
 import org.gjt.sp.jedit.View
+import org.gjt.sp.jedit.gui.AnimatedIcon
 
 
 object Query_Operation
@@ -34,12 +36,27 @@
 
   /* implicit state -- owned by Swing thread */
 
+  private object Status extends Enumeration
+  {
+    val WAITING = Value("waiting")
+    val RUNNING = Value("running")
+    val FINISHED = Value("finished")
+  }
+
   private var current_location: Option[Command] = None
   private var current_query: List[String] = Nil
-  private var current_result = false
-  private var current_snapshot = Document.State.init.snapshot()
-  private var current_state = Command.empty.init_state
+  private var current_update_pending = false
   private var current_output: List[XML.Tree] = Nil
+  private var current_status = Status.FINISHED
+
+  private def reset_state()
+  {
+    current_location = None
+    current_query = Nil
+    current_update_pending = false
+    current_output = Nil
+    current_status = Status.FINISHED
+  }
 
   private def remove_overlay()
   {
@@ -52,43 +69,111 @@
     } model.remove_overlay(command, operation_name, instance :: current_query)
   }
 
-  private def handle_result()
+
+  /* animation */
+
+  val animation = new Label
+
+  private val passive_icon =
+    JEdit_Lib.load_image_icon(PIDE.options.string("process_passive_icon")).getImage
+  private val active_icons =
+    space_explode(':', PIDE.options.string("process_active_icons")).map(name =>
+      JEdit_Lib.load_image_icon(name).getImage)
+
+  private val animation_icon =
+    new AnimatedIcon(passive_icon, active_icons.toArray, 5, animation.peer)
+  animation.icon = animation_icon
+
+  private def animation_update()
+  {
+    animation_icon.stop
+    current_status match {
+      case Status.WAITING =>
+        animation.tooltip = "Waiting for evaluation of query context ..."
+        animation_icon.setRate(5)
+        animation_icon.start
+      case Status.RUNNING =>
+        animation.tooltip = "Running query operation ..."
+        animation_icon.setRate(15)
+        animation_icon.start
+      case Status.FINISHED =>
+        animation.tooltip = null
+    }
+  }
+
+
+  /* content update */
+
+  private def content_update()
   {
     Swing_Thread.require()
 
-    val (new_snapshot, new_state) =
-      Document_View(view.getTextArea) match {
-        case Some(doc_view) =>
-          val snapshot = doc_view.model.snapshot()
-          current_location match {
-            case Some(cmd) =>
-              (snapshot, snapshot.state.command_state(snapshot.version, cmd))
-            case None =>
-              (Document.State.init.snapshot(), Command.empty.init_state)
-          }
-        case None => (current_snapshot, current_state)
+
+    /* snapshot */
+
+    val (snapshot, state) =
+      current_location match {
+        case Some(cmd) =>
+          val snapshot = PIDE.document_snapshot(cmd.node_name)
+          val state = snapshot.state.command_state(snapshot.version, cmd)
+          (snapshot, state)
+        case None =>
+          (Document.State.init.snapshot(), Command.empty.init_state)
       }
 
+    val results =
+      (for {
+        (_, XML.Elem(Markup(Markup.RESULT, props), body)) <- state.results.entries
+        if props.contains((Markup.INSTANCE, instance))
+      } yield body).toList
+
+
+    /* output */
+
     val new_output =
-      (for {
-        (_, XML.Elem(Markup(Markup.RESULT, props), List(XML.Elem(markup, body)))) <-
-          new_state.results.entries
-        if props.contains((Markup.INSTANCE, instance))
-      } yield XML.Elem(Markup(Markup.message(markup.name), markup.properties), body)).toList
+      for {
+        List(XML.Elem(markup, body)) <- results
+        if Markup.messages.contains(markup.name)
+      }
+      yield XML.Elem(Markup(Markup.message(markup.name), markup.properties), body)
+
+
+    /* status */
 
-    if (new_output != current_output)
-      consume_result(new_snapshot, new_state.results, new_output)
+    def get_status(name: String, status: Status.Value): Option[Status.Value] =
+      results.collectFirst({ case List(XML.Elem(m, _)) if m.name == name => status })
+
+    val new_status =
+      get_status(Markup.FINISHED, Status.FINISHED) orElse
+      get_status(Markup.RUNNING, Status.RUNNING) getOrElse
+      Status.WAITING
+
+
+    /* state update */
 
-    if (!new_output.isEmpty) {
-      current_result = true
-      remove_overlay()
-      PIDE.flush_buffers()
+    if (current_output != new_output || current_status != new_status) {
+      if (snapshot.is_outdated)
+        current_update_pending = true
+      else {
+        current_update_pending = false
+        if (current_output != new_output) {
+          current_output = new_output
+          consume_result(snapshot, state.results, new_output)
+        }
+        if (current_status != new_status) {
+          current_status = new_status
+          animation_update()
+          if (new_status == Status.FINISHED) {
+            remove_overlay()
+            PIDE.flush_buffers()
+          }
+        }
+      }
     }
+  }
 
-    current_snapshot = new_snapshot
-    current_state = new_state
-    current_output = new_output
-  }
+
+  /* apply query */
 
   def apply_query(query: List[String])
   {
@@ -98,30 +183,34 @@
       case Some(doc_view) =>
         val snapshot = doc_view.model.snapshot()
         remove_overlay()
-        current_location = None
-        current_query = Nil
-        current_result = false
+        reset_state()
         snapshot.node.command_at(doc_view.text_area.getCaretPosition).map(_._1) match {
           case Some(command) =>
             current_location = Some(command)
             current_query = query
+            current_status = Status.WAITING
             doc_view.model.insert_overlay(command, operation_name, instance :: query)
           case None =>
         }
+        animation_update()
         PIDE.flush_buffers()
       case None =>
     }
   }
 
+
+  /* locate query */
+
   def locate_query()
   {
     Swing_Thread.require()
 
     current_location match {
       case Some(command) =>
-        val snapshot = PIDE.session.snapshot(command.node_name)
+        val snapshot = PIDE.document_snapshot(command.node_name)
         val commands = snapshot.node.commands
         if (commands.contains(command)) {
+          // FIXME revert offset (!?)
           val sources = commands.iterator.takeWhile(_ != command).map(_.source)
           val (line, column) = ((1, 1) /: sources)(Symbol.advance_line_column)
           Hyperlink(command.node_name.node, line, column).follow(view)
@@ -138,8 +227,10 @@
       react {
         case changed: Session.Commands_Changed =>
           current_location match {
-            case Some(command) if !current_result && changed.commands.contains(command) =>
-              Swing_Thread.later { handle_result() }
+            case Some(command)
+            if current_update_pending ||
+              (current_status != Status.FINISHED && changed.commands.contains(command)) =>
+              Swing_Thread.later { content_update() }
             case _ =>
           }
         case bad =>
--- a/src/Tools/jEdit/src/rendering.scala	Tue Aug 06 15:50:23 2013 +0200
+++ b/src/Tools/jEdit/src/rendering.scala	Tue Aug 06 23:24:10 2013 +0200
@@ -14,7 +14,7 @@
 import javax.swing.Icon
 
 import org.gjt.sp.jedit.syntax.{Token => JEditToken}
-import org.gjt.sp.jedit.{jEdit, GUIUtilities}
+import org.gjt.sp.jedit.jEdit
 
 import scala.collection.immutable.SortedMap
 
@@ -41,23 +41,6 @@
     Markup.ERROR -> error_pri, Markup.ERROR_MESSAGE -> error_pri)
 
 
-  /* icons */
-
-  def load_icon(name: String): Icon =
-  {
-    val name1 =
-      if (name.startsWith("idea-icons/")) {
-        val file =
-          Isabelle_System.platform_file_url(Path.explode("$JEDIT_HOME/dist/jars/idea-icons.jar"))
-        "jar:" + file + "!/" + name
-      }
-      else name
-    val icon = GUIUtilities.loadIcon(name1)
-    if (icon.getIconWidth < 0 || icon.getIconHeight < 0) error("Bad icon: " + name)
-    else icon
-  }
-
-
   /* jEdit font */
 
   def font_family(): String = jEdit.getProperty("view.font")
@@ -381,15 +364,15 @@
   val tooltip_margin: Int = options.int("jedit_tooltip_margin")
   val tooltip_bounds: Double = (options.real("jedit_tooltip_bounds") max 0.2) min 0.8
 
-  lazy val tooltip_close_icon = Rendering.load_icon(options.string("tooltip_close_icon"))
-  lazy val tooltip_detach_icon = Rendering.load_icon(options.string("tooltip_detach_icon"))
+  lazy val tooltip_close_icon = JEdit_Lib.load_icon(options.string("tooltip_close_icon"))
+  lazy val tooltip_detach_icon = JEdit_Lib.load_icon(options.string("tooltip_detach_icon"))
 
 
   private lazy val gutter_icons = Map(
-    Rendering.information_pri -> Rendering.load_icon(options.string("gutter_information_icon")),
-    Rendering.warning_pri -> Rendering.load_icon(options.string("gutter_warning_icon")),
-    Rendering.legacy_pri -> Rendering.load_icon(options.string("gutter_legacy_icon")),
-    Rendering.error_pri -> Rendering.load_icon(options.string("gutter_error_icon")))
+    Rendering.information_pri -> JEdit_Lib.load_icon(options.string("gutter_information_icon")),
+    Rendering.warning_pri -> JEdit_Lib.load_icon(options.string("gutter_warning_icon")),
+    Rendering.legacy_pri -> JEdit_Lib.load_icon(options.string("gutter_legacy_icon")),
+    Rendering.error_pri -> JEdit_Lib.load_icon(options.string("gutter_error_icon")))
 
   private val gutter_elements = Set(Markup.WRITELN, Markup.WARNING, Markup.ERROR)