merged
authorpaulson
Thu, 10 May 2018 15:41:45 +0100
changeset 68137 afcdc4c0ef0d
parent 68135 763f5a8f3f7f (diff)
parent 68136 f022083489d0 (current diff)
child 68138 c738f40e88d4
merged
--- a/etc/options	Thu May 10 15:41:34 2018 +0100
+++ b/etc/options	Thu May 10 15:41:45 2018 +0100
@@ -71,6 +71,8 @@
 option threads_stack_limit : real = 0.25
   -- "maximum stack size for worker threads (in giga words, 0 = unlimited)"
 
+public option parallel_limit : int = 0
+  -- "approximative limit for parallel tasks (0 = unlimited)"
 public option parallel_print : bool = true
   -- "parallel and asynchronous printing of results"
 public option parallel_proofs : int = 1
--- a/src/HOL/Analysis/Determinants.thy	Thu May 10 15:41:34 2018 +0100
+++ b/src/HOL/Analysis/Determinants.thy	Thu May 10 15:41:45 2018 +0100
@@ -33,26 +33,14 @@
   apply (simp add: mult.commute)
   done
 
-text \<open>Definition of determinant.\<close>
+subsubsection \<open>Definition of determinant\<close>
 
 definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
   "det A =
     sum (\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set))
       {p. p permutes (UNIV :: 'n set)}"
 
-text \<open>A few general lemmas we need below.\<close>
-
-lemma prod_permute:
-  assumes p: "p permutes S"
-  shows "prod f S = prod (f \<circ> p) S"
-  using assms by (fact prod.permute)
-
-lemma product_permute_nat_interval:
-  fixes m n :: nat
-  shows "p permutes {m..n} \<Longrightarrow> prod f {m..n} = prod (f \<circ> p) {m..n}"
-  by (blast intro!: prod_permute)
-
-text \<open>Basic determinant properties.\<close>
+text \<open>Basic determinant properties\<close>
 
 lemma det_transpose [simp]: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)"
 proof -
@@ -77,15 +65,10 @@
       unfolding prod.reindex[OF pi] ..
     also have "\<dots> = prod (\<lambda>i. ?di A i (p i)) ?U"
     proof -
-      {
-        fix i
-        assume i: "i \<in> ?U"
-        from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
-        have "((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) i = ?di A i (p i)"
-          unfolding transpose_def by (simp add: fun_eq_iff)
-      }
-      then show "prod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U =
-        prod (\<lambda>i. ?di A i (p i)) ?U"
+      have "((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) i = ?di A i (p i)" if "i \<in> ?U" for i
+        using that permutes_inv_o[OF pU] permutes_in_image[OF pU]
+        unfolding transpose_def by (simp add: fun_eq_iff)
+      then show "prod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U = prod (\<lambda>i. ?di A i (p i)) ?U"
         by (auto intro: prod.cong)
     qed
     finally have "of_int (sign (inv p)) * (prod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U) =
@@ -94,11 +77,7 @@
   }
   then show ?thesis
     unfolding det_def
-    apply (subst sum_permutations_inverse)
-    apply (rule sum.cong)
-    apply (rule refl)
-    apply blast
-    done
+    by (subst sum_permutations_inverse) (blast intro: sum.cong elim: )
 qed
 
 lemma det_lowerdiagonal:
@@ -111,24 +90,20 @@
   let ?pp = "\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
   have fU: "finite ?U"
     by simp
-  from finite_permutations[OF fU] have fPU: "finite ?PU" .
   have id0: "{id} \<subseteq> ?PU"
     by (auto simp add: permutes_id)
-  {
+  have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
+  proof
     fix p
-    assume p: "p \<in> ?PU - {id}"
-    from p have pU: "p permutes ?U" and pid: "p \<noteq> id"
-      by blast+
-    from permutes_natset_le[OF pU] pid obtain i where i: "p i > i"
-      by (metis not_le)
-    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0"
+    assume "p \<in> ?PU - {id}"
+    then obtain i where i: "p i > i"
+      by clarify (meson leI permutes_natset_le)
+    from ld[OF i] have "\<exists>i \<in> ?U. A$i$p i = 0"
       by blast
-    from prod_zero[OF fU ex] have "?pp p = 0"
-      by simp
-  }
-  then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
-    by blast
-  from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
+    with prod_zero[OF fU] show "?pp p = 0"
+      by force
+  qed
+  from sum.mono_neutral_cong_left[OF finite_permutations[OF fU] id0 p0] show ?thesis
     unfolding det_def by (simp add: sign_id)
 qed
 
@@ -142,24 +117,20 @@
   let ?pp = "(\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set))"
   have fU: "finite ?U"
     by simp
-  from finite_permutations[OF fU] have fPU: "finite ?PU" .
   have id0: "{id} \<subseteq> ?PU"
     by (auto simp add: permutes_id)
-  {
+  have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"
+  proof
     fix p
     assume p: "p \<in> ?PU - {id}"
-    from p have pU: "p permutes ?U" and pid: "p \<noteq> id"
-      by blast+
-    from permutes_natset_ge[OF pU] pid obtain i where i: "p i < i"
-      by (metis not_le)
-    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0"
+    then obtain i where i: "p i < i"
+      by clarify (meson leI permutes_natset_ge)
+    from ld[OF i] have "\<exists>i \<in> ?U. A$i$p i = 0"
       by blast
-    from prod_zero[OF fU ex] have "?pp p = 0"
-      by simp
-  }
-  then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"
-    by blast
-  from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
+    with prod_zero[OF fU]  show "?pp p = 0"
+      by force
+  qed
+  from sum.mono_neutral_cong_left[OF finite_permutations[OF fU] id0 p0] show ?thesis
     unfolding det_def by (simp add: sign_id)
 qed
 
@@ -175,20 +146,17 @@
   from finite_permutations[OF fU] have fPU: "finite ?PU" .
   have id0: "{id} \<subseteq> ?PU"
     by (auto simp add: permutes_id)
-  {
+  have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
+  proof
     fix p
     assume p: "p \<in> ?PU - {id}"
-    then have "p \<noteq> id"
-      by simp
     then obtain i where i: "p i \<noteq> i"
-      unfolding fun_eq_iff by auto
-    from ld [OF i [symmetric]] have ex:"\<exists>i \<in> ?U. A$i$p i = 0"
-      by blast
-    from prod_zero [OF fU ex] have "?pp p = 0"
-      by simp
-  }
-  then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
-    by blast
+      by fastforce
+    with ld have "\<exists>i \<in> ?U. A$i$p i = 0"
+      by (metis UNIV_I)
+    with prod_zero [OF fU] show "?pp p = 0"
+      by force
+  qed
   from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
     unfolding det_def by (simp add: sign_id)
 qed
@@ -203,32 +171,38 @@
   fixes A :: "'a::comm_ring_1^'n^'n"
   assumes p: "p permutes (UNIV :: 'n::finite set)"
   shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
-  apply (simp add: det_def sum_distrib_left mult.assoc[symmetric])
-  apply (subst sum_permutations_compose_right[OF p])
-proof (rule sum.cong)
+proof -
   let ?U = "UNIV :: 'n set"
   let ?PU = "{p. p permutes ?U}"
-  fix q
-  assume qPU: "q \<in> ?PU"
-  have fU: "finite ?U"
-    by simp
-  from qPU have q: "q permutes ?U"
-    by blast
-  from p q have pp: "permutation p" and qp: "permutation q"
-    by (metis fU permutation_permutes)+
-  from permutes_inv[OF p] have ip: "inv p permutes ?U" .
-  have "prod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = prod ((\<lambda>i. A$p i$(q \<circ> p) i) \<circ> inv p) ?U"
-    by (simp only: prod_permute[OF ip, symmetric])
-  also have "\<dots> = prod (\<lambda>i. A $ (p \<circ> inv p) i $ (q \<circ> (p \<circ> inv p)) i) ?U"
-    by (simp only: o_def)
-  also have "\<dots> = prod (\<lambda>i. A$i$q i) ?U"
-    by (simp only: o_def permutes_inverses[OF p])
-  finally have thp: "prod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = prod (\<lambda>i. A$i$q i) ?U"
-    by blast
-  show "of_int (sign (q \<circ> p)) * prod (\<lambda>i. A$ p i$ (q \<circ> p) i) ?U =
-    of_int (sign p) * of_int (sign q) * prod (\<lambda>i. A$i$q i) ?U"
-    by (simp only: thp sign_compose[OF qp pp] mult.commute of_int_mult)
-qed rule
+  have *: "(\<Sum>q\<in>?PU. of_int (sign (q \<circ> p)) * (\<Prod>i\<in>?U. A $ p i $ (q \<circ> p) i)) =
+           (\<Sum>n\<in>?PU. of_int (sign p) * of_int (sign n) * (\<Prod>i\<in>?U. A $ i $ n i))"
+  proof (rule sum.cong)
+    fix q
+    assume qPU: "q \<in> ?PU"
+    have fU: "finite ?U"
+      by simp
+    from qPU have q: "q permutes ?U"
+      by blast
+    have "prod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = prod ((\<lambda>i. A$p i$(q \<circ> p) i) \<circ> inv p) ?U"
+      by (simp only: prod.permute[OF permutes_inv[OF p], symmetric])
+    also have "\<dots> = prod (\<lambda>i. A $ (p \<circ> inv p) i $ (q \<circ> (p \<circ> inv p)) i) ?U"
+      by (simp only: o_def)
+    also have "\<dots> = prod (\<lambda>i. A$i$q i) ?U"
+      by (simp only: o_def permutes_inverses[OF p])
+    finally have thp: "prod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = prod (\<lambda>i. A$i$q i) ?U"
+      by blast
+    from p q have pp: "permutation p" and qp: "permutation q"
+      by (metis fU permutation_permutes)+
+    show "of_int (sign (q \<circ> p)) * prod (\<lambda>i. A$ p i$ (q \<circ> p) i) ?U =
+          of_int (sign p) * of_int (sign q) * prod (\<lambda>i. A$i$q i) ?U"
+      by (simp only: thp sign_compose[OF qp pp] mult.commute of_int_mult)
+  qed auto
+  show ?thesis
+    apply (simp add: det_def sum_distrib_left mult.assoc[symmetric])
+    apply (subst sum_permutations_compose_right[OF p])
+    apply (rule *)
+    done
+qed 
 
 lemma det_permute_columns:
   fixes A :: "'a::comm_ring_1^'n^'n"
@@ -279,19 +253,19 @@
     also have "... = - sign x" using sign_tjk by simp
     also have "... \<noteq> sign x" unfolding sign_def by simp
     finally have "sign (?t_jk \<circ> x) \<noteq> sign x" and "(?t_jk \<circ> x) \<in> ?S2"
-      by (auto, metis (lifting, full_types) mem_Collect_eq x)
+      using x by force+
   }
-  hence disjoint: "?S1 \<inter> ?S2 = {}" by (auto, metis sign_def)
+  hence disjoint: "?S1 \<inter> ?S2 = {}"
+    by (force simp: sign_def)
   have PU_decomposition: "?PU = ?S1 \<union> ?S2"
   proof (auto)
     fix x
     assume x: "x permutes ?U" and "\<forall>p. p permutes ?U \<longrightarrow> x = Fun.swap j k id \<circ> p \<longrightarrow> \<not> evenperm p"
-    from this obtain p where p: "p permutes UNIV" and x_eq: "x = Fun.swap j k id \<circ> p"
+    then obtain p where p: "p permutes UNIV" and x_eq: "x = Fun.swap j k id \<circ> p"
       and odd_p: "\<not> evenperm p"
-      by (metis (no_types) comp_assoc id_comp inv_swap_id permutes_compose
-          permutes_inv_o(1) tjk_permutes)
+      by (metis (mono_tags) id_o o_assoc permutes_compose swap_id_idempotent tjk_permutes)
     thus "evenperm x"
-      by (metis evenperm_comp evenperm_swap finite_class.finite_UNIV
+      by (meson evenperm_comp evenperm_swap finite_class.finite_UNIV
           jk permutation_permutes permutation_swap_id)
   next
     fix p assume p: "p permutes ?U"
@@ -328,13 +302,9 @@
 
 lemma det_identical_rows:
   fixes A :: "'a::comm_ring_1^'n^'n"
-  assumes ij: "i \<noteq> j"
-  and r: "row i A = row j A"
+  assumes ij: "i \<noteq> j" and r: "row i A = row j A"
   shows "det A = 0"
-  apply (subst det_transpose[symmetric])
-  apply (rule det_identical_columns[OF ij])
-  apply (metis column_transpose r)
-  done
+  by (metis column_transpose det_identical_columns det_transpose ij r)
 
 lemma det_zero_row:
   fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m"
@@ -366,38 +336,27 @@
     by blast
   have kU: "?U = insert k ?Uk"
     by blast
-  {
-    fix j
-    assume j: "j \<in> ?Uk"
-    from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j"
-      by simp_all
-  }
-  then have th1: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?g i $ p i) ?Uk"
-    and th2: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?h i $ p i) ?Uk"
-    apply -
-    apply (rule prod.cong, simp_all)+
-    done
-  have th3: "finite ?Uk" "k \<notin> ?Uk"
+  have eq: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?g i $ p i) ?Uk"
+           "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?h i $ p i) ?Uk"
+    by auto
+  have Uk: "finite ?Uk" "k \<notin> ?Uk"
     by auto
   have "prod (\<lambda>i. ?f i $ p i) ?U = prod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
     unfolding kU[symmetric] ..
   also have "\<dots> = ?f k $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk"
-    apply (rule prod.insert)
-    apply simp
-    apply blast
-    done
+    by (rule prod.insert) auto
   also have "\<dots> = (a k $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * prod (\<lambda>i. ?f i $ p i) ?Uk)"
     by (simp add: field_simps)
   also have "\<dots> = (a k $ p k * prod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * prod (\<lambda>i. ?h i $ p i) ?Uk)"
-    by (metis th1 th2)
+    by (metis eq)
   also have "\<dots> = prod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + prod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
-    unfolding  prod.insert[OF th3] by simp
+    unfolding  prod.insert[OF Uk] by simp
   finally have "prod (\<lambda>i. ?f i $ p i) ?U = prod (\<lambda>i. ?g i $ p i) ?U + prod (\<lambda>i. ?h i $ p i) ?U"
     unfolding kU[symmetric] .
   then show "of_int (sign p) * prod (\<lambda>i. ?f i $ p i) ?U =
     of_int (sign p) * prod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * prod (\<lambda>i. ?h i $ p i) ?U"
     by (simp add: field_simps)
-qed rule
+qed auto
 
 lemma det_row_mul:
   fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
@@ -416,38 +375,25 @@
     by blast
   have kU: "?U = insert k ?Uk"
     by blast
-  {
-    fix j
-    assume j: "j \<in> ?Uk"
-    from j have "?f j $ p j = ?g j $ p j"
-      by simp
-  }
-  then have th1: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?g i $ p i) ?Uk"
-    apply -
-    apply (rule prod.cong)
-    apply simp_all
-    done
-  have th3: "finite ?Uk" "k \<notin> ?Uk"
+  have eq: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?g i $ p i) ?Uk"
+    by (auto simp: )
+  have Uk: "finite ?Uk" "k \<notin> ?Uk"
     by auto
   have "prod (\<lambda>i. ?f i $ p i) ?U = prod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
     unfolding kU[symmetric] ..
   also have "\<dots> = ?f k $ p k  * prod (\<lambda>i. ?f i $ p i) ?Uk"
-    apply (rule prod.insert)
-    apply simp
-    apply blast
-    done
+    by (rule prod.insert) auto
   also have "\<dots> = (c*s a k) $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk"
     by (simp add: field_simps)
   also have "\<dots> = c* (a k $ p k * prod (\<lambda>i. ?g i $ p i) ?Uk)"
-    unfolding th1 by (simp add: ac_simps)
+    unfolding eq by (simp add: ac_simps)
   also have "\<dots> = c* (prod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
-    unfolding prod.insert[OF th3] by simp
+    unfolding prod.insert[OF Uk] by simp
   finally have "prod (\<lambda>i. ?f i $ p i) ?U = c* (prod (\<lambda>i. ?g i $ p i) ?U)"
     unfolding kU[symmetric] .
-  then show "of_int (sign p) * prod (\<lambda>i. ?f i $ p i) ?U =
-    c * (of_int (sign p) * prod (\<lambda>i. ?g i $ p i) ?U)"
+  then show "of_int (sign p) * prod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * prod (\<lambda>i. ?g i $ p i) ?U)"
     by (simp add: field_simps)
-qed rule
+qed auto
 
 lemma det_row_0:
   fixes b :: "'n::finite \<Rightarrow> _ ^ 'n"
@@ -478,16 +424,10 @@
   using x
 proof (induction rule: vec.span_induct_alt)
   case base
-  {
-    fix k
-    have "(if k = i then row i A + 0 else row k A) = row k A"
-      by simp
-  }
+  have "(if k = i then row i A + 0 else row k A) = row k A" for k
+    by simp
   then show ?case
-    apply -
-    apply (rule cong[of det, OF refl])
-    apply (vector row_def)
-    done
+    by (simp add: row_def)
 next
   case (step c z y)
   then obtain j where j: "z = row j A" "i \<noteq> j"
@@ -531,29 +471,24 @@
   let ?U = "UNIV :: 'n set"
   from d obtain i where i: "row i A \<in> vec.span (rows A - {row i A})"
     unfolding vec.dependent_def rows_def by blast
-  {
-    fix j k
-    assume jk: "j \<noteq> k" and c: "row j A = row k A"
-    from det_identical_rows[OF jk c] have ?thesis .
-  }
-  moreover
-  {
-    assume H: "\<And> i j. i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
-    have th0: "- row i A \<in> vec.span {row j A|j. j \<noteq> i}"
-      apply (rule vec.span_neg)
-      apply (rule set_rev_mp)
-      apply (rule i)
-      apply (rule vec.span_mono)
-      using H i
-      apply (auto simp add: rows_def)
-      done
-    from det_row_span[OF th0]
+  show ?thesis
+  proof (cases "\<forall>i j. i \<noteq> j \<longrightarrow> row i A \<noteq> row j A")
+    case True
+    with i have "vec.span (rows A - {row i A}) \<subseteq> vec.span {row j A |j. j \<noteq> i}"
+      by (auto simp add: rows_def intro!: vec.span_mono)
+    then have "- row i A \<in> vec.span {row j A|j. j \<noteq> i}"
+      by (meson i subsetCE vec.span_neg)
+    from det_row_span[OF this]
     have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)"
       unfolding right_minus vector_smult_lzero ..
-    with det_row_mul[of i "0::'a" "\<lambda>i. 1"]
-    have "det A = 0" by simp
-  }
-  ultimately show ?thesis by blast
+    with det_row_mul[of i 0 "\<lambda>i. 1"]
+    show ?thesis by simp
+  next
+    case False
+    then obtain j k where jk: "j \<noteq> k" "row j A = row k A"
+      by auto
+    from det_identical_rows[OF jk] show ?thesis .
+  qed
 qed
 
 lemma det_dependent_columns:
@@ -561,45 +496,35 @@
   shows "det A = 0"
   by (metis d det_dependent_rows rows_transpose det_transpose)
 
-text \<open>Multilinearity and the multiplication formula.\<close>
+text \<open>Multilinearity and the multiplication formula\<close>
 
 lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)"
-  by (rule iffD1[OF vec_lambda_unique]) vector
+  by auto
 
 lemma det_linear_row_sum:
   assumes fS: "finite S"
   shows "det ((\<chi> i. if i = k then sum (a i) S else c i)::'a::comm_ring_1^'n^'n) =
     sum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
-proof (induct rule: finite_induct[OF fS])
-  case 1
-  then show ?case
-    apply simp
-    unfolding sum.empty det_row_0[of k]
-    apply rule
-    done
-next
-  case (2 x F)
-  then show ?case
-    by (simp add: det_row_add cong del: if_weak_cong)
-qed
+  using fS  by (induct rule: finite_induct; simp add: det_row_0 det_row_add cong: if_cong)
 
 lemma finite_bounded_functions:
   assumes fS: "finite S"
   shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
 proof (induct k)
   case 0
-  have th: "{f. \<forall>i. f i = i} = {id}"
+  have *: "{f. \<forall>i. f i = i} = {id}"
     by auto
   show ?case
-    by (auto simp add: th)
+    by (auto simp add: *)
 next
   case (Suc k)
   let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i"
   let ?S = "?f ` (S \<times> {f. (\<forall>i\<in>{1..k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)})"
   have "?S = {f. (\<forall>i\<in>{1.. Suc k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. Suc k} \<longrightarrow> f i = i)}"
     apply (auto simp add: image_iff)
-    apply (rule_tac x="x (Suc k)" in bexI)
-    apply (rule_tac x = "\<lambda>i. if i = Suc k then i else x i" in exI)
+    apply (rename_tac f)
+    apply (rule_tac x="f (Suc k)" in bexI)
+    apply (rule_tac x = "\<lambda>i. if i = Suc k then i else f i" in exI)
     apply auto
     done
   with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f]
@@ -633,17 +558,14 @@
   have thif2: "\<And>a b c d e. (if a then b else if c then d else e) =
      (if c then (if a then b else d) else (if a then b else e))"
     by simp
-  from \<open>z \<notin> T\<close> have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False"
+  from \<open>z \<notin> T\<close> have nz: "\<And>i. i \<in> T \<Longrightarrow> i \<noteq> z"
     by auto
   have "det (\<chi> i. if i \<in> insert z T then sum (a i) S else c i) =
     det (\<chi> i. if i = z then sum (a i) S else if i \<in> T then sum (a i) S else c i)"
     unfolding insert_iff thif ..
   also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then sum (a i) S else if i = z then a i j else c i))"
     unfolding det_linear_row_sum[OF fS]
-    apply (subst thif2)
-    using nz
-    apply (simp cong del: if_weak_cong cong add: if_cong)
-    done
+    by (subst thif2) (simp add: nz cong: if_cong)
   finally have tha:
     "det (\<chi> i. if i \<in> insert z T then sum (a i) S else c i) =
      (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i)
@@ -696,18 +618,10 @@
   shows "det (A ** B) = det A * det B"
 proof -
   let ?U = "UNIV :: 'n set"
-  let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
+  let ?F = "{f. (\<forall>i \<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
   let ?PU = "{p. p permutes ?U}"
-  have fU: "finite ?U"
+  have "p \<in> ?F" if "p permutes ?U" for p
     by simp
-  have fF: "finite ?F"
-    by (rule finite)
-  {
-    fix p
-    assume p: "p permutes ?U"
-    have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p]
-      using p[unfolded permutes_def] by simp
-  }
   then have PUF: "?PU \<subseteq> ?F" by blast
   {
     fix f
@@ -723,34 +637,20 @@
       assume fni: "\<not> inj_on f ?U"
       then obtain i j where ij: "f i = f j" "i \<noteq> j"
         unfolding inj_on_def by blast
-      from ij
-      have rth: "row i ?B = row j ?B"
+      then have "row i ?B = row j ?B"
         by (vector row_def)
-      from det_identical_rows[OF ij(2) rth]
+      with det_identical_rows[OF ij(2)]
       have "det (\<chi> i. A$i$f i *s B$f i) = 0"
-        unfolding det_rows_mul by simp
+        unfolding det_rows_mul by force
     }
     moreover
     {
       assume fi: "inj_on f ?U"
       from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j"
         unfolding inj_on_def by metis
-      note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]]
-      {
-        fix y
-        from fs f have "\<exists>x. f x = y"
-          by blast
-        then obtain x where x: "f x = y"
-          by blast
-        {
-          fix z
-          assume z: "f z = y"
-          from fith x z have "z = x"
-            by metis
-        }
-        with x have "\<exists>!x. f x = y"
-          by blast
-      }
+      note fs = fi[unfolded surjective_iff_injective_gen[OF finite finite refl fUU, symmetric]]
+      have "\<exists>!x. f x = y" for y
+        using fith fs by blast
       with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0"
         by blast
     }
@@ -784,9 +684,9 @@
         by (simp_all add: sign_idempotent)
       have ths: "?s q = ?s p * ?s (q \<circ> inv p)"
         using pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-        by (simp add:  th00 ac_simps sign_idempotent sign_compose)
+        by (simp add: th00 ac_simps sign_idempotent sign_compose)
       have th001: "prod (\<lambda>i. B$i$ q (inv p i)) ?U = prod ((\<lambda>i. B$i$ q (inv p i)) \<circ> p) ?U"
-        by (rule prod_permute[OF p])
+        by (rule prod.permute[OF p])
       have thp: "prod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U =
         prod (\<lambda>i. A$i$p i) ?U * prod (\<lambda>i. B$i$ q (inv p i)) ?U"
         unfolding th001 prod.distrib[symmetric] o_def permutes_inverses[OF p]
@@ -804,64 +704,51 @@
     unfolding det_def sum_product
     by (rule sum.cong [OF refl])
   have "det (A**B) = sum (\<lambda>f.  det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
-    unfolding matrix_mul_sum_alt det_linear_rows_sum[OF fU]
+    unfolding matrix_mul_sum_alt det_linear_rows_sum[OF finite]
     by simp
   also have "\<dots> = sum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
-    using sum.mono_neutral_cong_left[OF fF PUF zth, symmetric]
+    using sum.mono_neutral_cong_left[OF finite PUF zth, symmetric]
     unfolding det_rows_mul by auto
   finally show ?thesis unfolding th2 .
 qed
 
 
-subsection \<open>Relation to invertibility.\<close>
+subsection \<open>Relation to invertibility\<close>
 
 lemma invertible_det_nz:
   fixes A::"'a::{field}^'n^'n"
   shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
-proof -
-  {
-    assume "invertible A"
-    then obtain B :: "'a^'n^'n" where B: "A ** B = mat 1"
-      unfolding invertible_right_inverse by blast
-    then have "det (A ** B) = det (mat 1 :: 'a^'n^'n)"
-      by simp
-    then have "det A \<noteq> 0"
-      by (simp add: det_mul) algebra
-  }
-  moreover
-  {
-    assume H: "\<not> invertible A"
-    let ?U = "UNIV :: 'n set"
-    have fU: "finite ?U"
-      by simp
-    from H obtain c i where c: "sum (\<lambda>i. c i *s row i A) ?U = 0"
-      and iU: "i \<in> ?U"
-      and ci: "c i \<noteq> 0"
-      unfolding invertible_right_inverse
-      unfolding matrix_right_invertible_independent_rows
-      by blast
-    have thr0: "- row i A = sum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
-      unfolding sum_cmul
-      using c ci   
-      by (auto simp add: sum.remove[OF fU iU] eq_vector_fraction_iff add_eq_0_iff)
-    have thr: "- row i A \<in> vec.span {row j A| j. j \<noteq> i}"
-      unfolding thr0
-      apply (rule vec.span_sum)
-      apply simp
-      apply (rule vec.span_scale)
-      apply (rule vec.span_base)
-      apply auto
-      done
-    let ?B = "(\<chi> k. if k = i then 0 else row k A) :: 'a^'n^'n"
-    have thrb: "row i ?B = 0" using iU by (vector row_def)
-    have "det A = 0"
-      unfolding det_row_span[OF thr, symmetric] right_minus
-      unfolding det_zero_row(2)[OF thrb] ..
-  }
-  ultimately show ?thesis
+proof (cases "invertible A")
+  case True
+  then obtain B :: "'a^'n^'n" where B: "A ** B = mat 1"
+    unfolding invertible_right_inverse by blast
+  then have "det (A ** B) = det (mat 1 :: 'a^'n^'n)"
+    by simp
+  then show ?thesis
+    by (metis True det_I det_mul mult_zero_left one_neq_zero)
+next
+  case False
+  let ?U = "UNIV :: 'n set"
+  have fU: "finite ?U"
+    by simp
+  from False obtain c i where c: "sum (\<lambda>i. c i *s row i A) ?U = 0" and iU: "i \<in> ?U" and ci: "c i \<noteq> 0"
+    unfolding invertible_right_inverse matrix_right_invertible_independent_rows
     by blast
+  have thr0: "- row i A = sum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
+    unfolding sum_cmul  using c ci   
+    by (auto simp add: sum.remove[OF fU iU] eq_vector_fraction_iff add_eq_0_iff)
+  have thr: "- row i A \<in> vec.span {row j A| j. j \<noteq> i}"
+    unfolding thr0 by (auto intro: vec.span_base vec.span_scale vec.span_sum)
+  let ?B = "(\<chi> k. if k = i then 0 else row k A) :: 'a^'n^'n"
+  have thrb: "row i ?B = 0" using iU by (vector row_def)
+  have "det A = 0"
+    unfolding det_row_span[OF thr, symmetric] right_minus
+    unfolding det_zero_row(2)[OF thrb] ..
+  then show ?thesis
+    by (simp add: False)
 qed
 
+
 lemma det_nz_iff_inj_gen:
   fixes f :: "'a::field^'n \<Rightarrow> 'a::field^'n"
   assumes "Vector_Spaces.linear ( *s) ( *s) f"
@@ -975,7 +862,7 @@
       vec.linear_injective_left_inverse vec.linear_surjective_right_inverse)
 
 
-subsection \<open>Cramer's rule.\<close>
+subsection \<open>Cramer's rule\<close>
 
 lemma cramer_lemma_transpose:
   fixes A:: "real^'n^'n"
@@ -988,8 +875,6 @@
   let ?Uk = "?U - {k}"
   have U: "?U = insert k ?Uk"
     by blast
-  have fUk: "finite ?Uk"
-    by simp
   have kUk: "k \<notin> ?Uk"
     by simp
   have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
@@ -1000,15 +885,10 @@
   then have thd1: "det (\<chi> i. row i A) = det A"
     by simp
   have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A"
-    apply (rule det_row_span)
-    apply (rule vec.span_sum)
-    apply (rule vec.span_scale)
-    apply (rule vec.span_base)
-    apply auto
-    done
+    by (force intro: det_row_span vec.span_sum vec.span_scale vec.span_base)
   show "?lhs = x$k * det A"
     apply (subst U)
-    unfolding sum.insert[OF fUk kUk]
+    unfolding sum.insert[OF finite kUk]
     apply (subst th00)
     unfolding add.assoc
     apply (subst det_row_add)
@@ -1154,7 +1034,7 @@
   let ?m1 = "mat 1 :: real ^'n^'n"
   {
     assume ot: ?ot
-    from ot have lf: "Vector_Spaces.linear ( *s) ( *s) f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w"
+    from ot have lf: "Vector_Spaces.linear ( *s) ( *s) f" and fd: "\<And>v w. f v \<bullet> f w = v \<bullet> w"
       unfolding orthogonal_transformation_def orthogonal_matrix linear_def scalar_mult_eq_scaleR
       by blast+
     {
@@ -1163,7 +1043,7 @@
       have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
         "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
         by simp_all
-      from fd[rule_format, of "axis i 1" "axis j 1",
+      from fd[of "axis i 1" "axis j 1",
         simplified matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
       have "?A$i$j = ?m1 $ i $ j"
         by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def
@@ -1235,25 +1115,22 @@
 proof -
   {
     fix v w
-    {
-      fix x
-      note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right]
-    }
-    note th0 = this
-    have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)"
+    have "norm (f x) = c * norm x" for x
+      by (metis dist_0_norm f0 fd)
+    then have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)"
       unfolding dot_norm_neg dist_norm[symmetric]
-      unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
-  note fc = this
-  show ?thesis
+      by (simp add: fd power2_eq_square field_simps)
+  }
+  then show ?thesis
     unfolding linear_iff vector_eq[where 'a="'a"] scalar_mult_eq_scaleR
-    by (simp add: inner_add fc field_simps)
+    by (simp add: inner_add field_simps)
 qed
 
 lemma isometry_linear:
   "f (0::'a::real_inner) = (0::'a) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f"
   by (rule scaling_linear[where c=1]) simp_all
 
-text \<open>Hence another formulation of orthogonal transformation.\<close>
+text \<open>Hence another formulation of orthogonal transformation\<close>
 
 lemma orthogonal_transformation_isometry:
   "orthogonal_transformation f \<longleftrightarrow> f(0::'a::real_inner) = (0::'a) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
@@ -1498,7 +1375,7 @@
   shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
   by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
 
-text \<open>Explicit formulas for low dimensions.\<close>
+text \<open>Explicit formulas for low dimensions\<close>
 
 lemma prod_neutral_const: "prod f {(1::nat)..1} = f 1"
   by simp
@@ -1544,7 +1421,7 @@
     by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
 qed
 
-text\<open> Slightly stronger results giving rotation, but only in two or more dimensions.\<close>
+text\<open> Slightly stronger results giving rotation, but only in two or more dimensions\<close>
 
 lemma rotation_matrix_exists_basis:
   fixes a :: "real^'n"
--- a/src/HOL/ROOT	Thu May 10 15:41:34 2018 +0100
+++ b/src/HOL/ROOT	Thu May 10 15:41:45 2018 +0100
@@ -15,7 +15,7 @@
   description {*
     HOL-Main with explicit proof terms.
   *}
-  options [quick_and_dirty = false, record_proofs = 2, parallel_proofs = 0]
+  options [quick_and_dirty = false, record_proofs = 2, parallel_proofs = 0, parallel_limit = 500]
   sessions
     "HOL-Library"
   theories
--- a/src/Pure/Concurrent/future.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Concurrent/future.ML	Thu May 10 15:41:45 2018 +0100
@@ -34,6 +34,10 @@
   val forked_results: {name: string, deps: Task_Queue.task list} ->
     (unit -> 'a) list -> 'a Exn.result list
   val task_context: string -> group -> ('a -> 'b) -> 'a -> 'b
+  val enabled: unit -> bool
+  val relevant: 'a list -> bool
+  val proofs_enabled: int -> bool
+  val proofs_enabled_timing: Time.time -> bool
   val value_result: 'a Exn.result -> 'a future
   val value: 'a -> 'a future
   val cond_forks: params -> (unit -> 'a) list -> 'a future list
@@ -171,10 +175,10 @@
 fun report_status () = (*requires SYNCHRONIZED*)
   if ! ML_statistics then
     let
-      val {ready, pending, running, passive, urgent} = Task_Queue.status (! queue);
-      val total = length (! workers);
-      val active = count_workers Working;
-      val waiting = count_workers Waiting;
+      val {ready, pending, running, passive, urgent, total} = Task_Queue.status (! queue);
+      val workers_total = length (! workers);
+      val workers_active = count_workers Working;
+      val workers_waiting = count_workers Waiting;
       val stats =
        [("now", Value.print_real (Time.toReal (Time.now ()))),
         ("tasks_ready", Value.print_int ready),
@@ -182,9 +186,10 @@
         ("tasks_running", Value.print_int running),
         ("tasks_passive", Value.print_int passive),
         ("tasks_urgent", Value.print_int urgent),
-        ("workers_total", Value.print_int total),
-        ("workers_active", Value.print_int active),
-        ("workers_waiting", Value.print_int waiting)] @
+        ("tasks_total", Value.print_int total),
+        ("workers_total", Value.print_int workers_total),
+        ("workers_active", Value.print_int workers_active),
+        ("workers_waiting", Value.print_int workers_waiting)] @
         ML_Statistics.get ();
     in Output.try_protocol_message (Markup.ML_statistics :: stats) [] end
   else ();
@@ -576,6 +581,22 @@
     end);
 
 
+(* scheduling parameters *)
+
+fun enabled () =
+  Multithreading.max_threads () > 1 andalso
+    let val lim = Options.default_int "parallel_limit"
+    in lim <= 0 orelse SYNCHRONIZED "enabled" (fn () => Task_Queue.total_jobs (! queue) < lim) end;
+
+val relevant = (fn [] => false | [_] => false | _ => enabled ());
+
+fun proofs_enabled n =
+  ! Multithreading.parallel_proofs >= n andalso is_some (worker_task ()) andalso enabled ();
+
+fun proofs_enabled_timing t =
+  proofs_enabled 1 andalso Time.toReal t >= Options.default_real "parallel_subproofs_threshold";
+
+
 (* fast-path operations -- bypass task queue if possible *)
 
 fun value_result (res: 'a Exn.result) =
@@ -589,7 +610,7 @@
 fun value x = value_result (Exn.Res x);
 
 fun cond_forks args es =
-  if Multithreading.enabled () then forks args es
+  if enabled () then forks args es
   else map (fn e => value_result (Exn.interruptible_capture e ())) es;
 
 fun map_future f x =
--- a/src/Pure/Concurrent/lazy.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Concurrent/lazy.ML	Thu May 10 15:41:45 2018 +0100
@@ -129,7 +129,7 @@
     val pending =
       xs |> map_filter (fn x => if is_pending x then SOME (fn () => force_result x) else NONE);
     val _ =
-      if Multithreading.relevant pending then
+      if Future.relevant pending then
         ignore (Future.forked_results {name = "Lazy.consolidate", deps = []} pending)
       else List.app (fn e => ignore (e ())) pending;
   in xs end;
--- a/src/Pure/Concurrent/multithreading.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Concurrent/multithreading.ML	Thu May 10 15:41:45 2018 +0100
@@ -8,10 +8,7 @@
 sig
   val max_threads: unit -> int
   val max_threads_update: int -> unit
-  val enabled: unit -> bool
-  val relevant: 'a list -> bool
   val parallel_proofs: int ref
-  val parallel_proofs_enabled: int -> bool
   val sync_wait: Time.time option -> ConditionVar.conditionVar -> Mutex.mutex -> bool Exn.result
   val trace: int ref
   val tracing: int -> (unit -> string) -> unit
@@ -42,9 +39,6 @@
 
 fun max_threads () = ! max_threads_state;
 fun max_threads_update m = max_threads_state := max_threads_result m;
-fun enabled () = max_threads () > 1;
-
-val relevant = (fn [] => false | [_] => false | _ => enabled ());
 
 end;
 
@@ -53,9 +47,6 @@
 
 val parallel_proofs = ref 1;
 
-fun parallel_proofs_enabled n =
-  enabled () andalso ! parallel_proofs >= n;
-
 
 (* synchronous wait *)
 
--- a/src/Pure/Concurrent/par_list.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Concurrent/par_list.ML	Thu May 10 15:41:45 2018 +0100
@@ -29,7 +29,7 @@
 struct
 
 fun forked_results name f xs =
-  if Multithreading.relevant xs
+  if Future.relevant xs
   then Future.forked_results {name = name, deps = []} (map (fn x => fn () => f x) xs)
   else map (Exn.capture f) xs;
 
--- a/src/Pure/Concurrent/task_queue.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Concurrent/task_queue.ML	Thu May 10 15:41:45 2018 +0100
@@ -32,7 +32,9 @@
   val group_tasks: queue -> group list -> task list
   val known_task: queue -> task -> bool
   val all_passive: queue -> bool
-  val status: queue -> {ready: int, pending: int, running: int, passive: int, urgent: int}
+  val total_jobs: queue -> int
+  val status: queue ->
+    {ready: int, pending: int, running: int, passive: int, urgent: int, total: int}
   val cancel: queue -> group -> Thread.thread list
   val cancel_all: queue -> group list * Thread.thread list
   val finish: task -> queue -> bool * queue
@@ -217,10 +219,12 @@
 
 (* queue *)
 
-datatype queue = Queue of {groups: groups, jobs: jobs, urgent: int};
+datatype queue = Queue of {groups: groups, jobs: jobs, urgent: int, total: int};
 
-fun make_queue groups jobs urgent = Queue {groups = groups, jobs = jobs, urgent = urgent};
-val empty = make_queue Inttab.empty Task_Graph.empty 0;
+fun make_queue groups jobs urgent total =
+  Queue {groups = groups, jobs = jobs, urgent = urgent, total = total};
+
+val empty = make_queue Inttab.empty Task_Graph.empty 0 0;
 
 fun group_tasks (Queue {groups, ...}) gs =
   fold (fn g => fn tasks => Tasks.merge (op =) (tasks, get_tasks groups (group_id g)))
@@ -249,10 +253,12 @@
 
 fun all_passive (Queue {jobs, ...}) = is_none (Task_Graph.get_first active_job jobs);
 
+fun total_jobs (Queue {total, ...}) = total;
+
 
 (* queue status *)
 
-fun status (Queue {jobs, urgent, ...}) =
+fun status (Queue {jobs, urgent, total, ...}) =
   let
     val (x, y, z, w) =
       Task_Graph.fold (fn (_, (job, (deps, _))) => fn (x, y, z, w) =>
@@ -261,7 +267,7 @@
           | Running _ => (x, y, z + 1, w)
           | Passive _ => (x, y, z, w + 1)))
         jobs (0, 0, 0, 0);
-  in {ready = x, pending = y, running = z, passive = w, urgent = urgent} end;
+  in {ready = x, pending = y, running = z, passive = w, urgent = urgent, total = total} end;
 
 
 
@@ -295,35 +301,38 @@
 
 (* finish *)
 
-fun finish task (Queue {groups, jobs, urgent}) =
+fun finish task (Queue {groups, jobs, urgent, total}) =
   let
     val group = group_of_task task;
     val groups' = fold_groups (fn g => del_task (group_id g, task)) group groups;
     val jobs' = Task_Graph.del_node task jobs;
+    val total' = total - 1;
     val maximal = Task_Graph.is_maximal jobs task;
-  in (maximal, make_queue groups' jobs' urgent) end;
+  in (maximal, make_queue groups' jobs' urgent total') end;
 
 
 (* enroll *)
 
-fun enroll thread name group (Queue {groups, jobs, urgent}) =
+fun enroll thread name group (Queue {groups, jobs, urgent, total}) =
   let
     val task = new_task group name NONE;
     val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups;
     val jobs' = jobs |> Task_Graph.new_node (task, Running thread);
-  in (task, make_queue groups' jobs' urgent) end;
+    val total' = total + 1;
+  in (task, make_queue groups' jobs' urgent total') end;
 
 
 (* enqueue *)
 
-fun enqueue_passive group name abort (Queue {groups, jobs, urgent}) =
+fun enqueue_passive group name abort (Queue {groups, jobs, urgent, total}) =
   let
     val task = new_task group name NONE;
     val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups;
     val jobs' = jobs |> Task_Graph.new_node (task, Passive abort);
-  in (task, make_queue groups' jobs' urgent) end;
+    val total' = total + 1;
+  in (task, make_queue groups' jobs' urgent total') end;
 
-fun enqueue name group deps pri job (Queue {groups, jobs, urgent}) =
+fun enqueue name group deps pri job (Queue {groups, jobs, urgent, total}) =
   let
     val task = new_task group name (SOME pri);
     val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups;
@@ -331,39 +340,40 @@
       |> Task_Graph.new_node (task, Job [job])
       |> fold (add_job task) deps;
     val urgent' = if pri >= urgent_pri then urgent + 1 else urgent;
-  in (task, make_queue groups' jobs' urgent') end;
+    val total' = total + 1;
+  in (task, make_queue groups' jobs' urgent' total') end;
 
-fun extend task job (Queue {groups, jobs, urgent}) =
+fun extend task job (Queue {groups, jobs, urgent, total}) =
   (case try (get_job jobs) task of
-    SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs) urgent)
+    SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs) urgent total)
   | _ => NONE);
 
 
 (* dequeue *)
 
-fun dequeue_passive thread task (queue as Queue {groups, jobs, urgent}) =
+fun dequeue_passive thread task (queue as Queue {groups, jobs, urgent, total}) =
   (case try (get_job jobs) task of
     SOME (Passive _) =>
       let val jobs' = set_job task (Running thread) jobs
-      in (SOME true, make_queue groups jobs' urgent) end
+      in (SOME true, make_queue groups jobs' urgent total) end
   | SOME _ => (SOME false, queue)
   | NONE => (NONE, queue));
 
-fun dequeue thread urgent_only (queue as Queue {groups, jobs, urgent}) =
+fun dequeue thread urgent_only (queue as Queue {groups, jobs, urgent, total}) =
   if not urgent_only orelse urgent > 0 then
     (case Task_Graph.get_first (ready_job_urgent urgent_only) jobs of
       SOME (result as (task, _)) =>
         let
           val jobs' = set_job task (Running thread) jobs;
           val urgent' = if pri_of_task task >= urgent_pri then urgent - 1 else urgent;
-        in (SOME result, make_queue groups jobs' urgent') end
+        in (SOME result, make_queue groups jobs' urgent' total) end
     | NONE => (NONE, queue))
   else (NONE, queue);
 
 
 (* dequeue wrt. dynamic dependencies *)
 
-fun dequeue_deps thread deps (queue as Queue {groups, jobs, urgent}) =
+fun dequeue_deps thread deps (queue as Queue {groups, jobs, urgent, total}) =
   let
     fun ready [] rest = (NONE, rev rest)
       | ready (task :: tasks) rest =
@@ -388,7 +398,7 @@
       let
         val jobs' = set_job task (Running thread) jobs;
         val urgent' = if pri_of_task task >= urgent_pri then urgent - 1 else urgent;
-      in ((SOME res, deps'), make_queue groups jobs' urgent') end;
+      in ((SOME res, deps'), make_queue groups jobs' urgent' total) end;
   in
     (case ready deps [] of
       (SOME res, deps') => result res deps'
--- a/src/Pure/General/exn.scala	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/General/exn.scala	Thu May 10 15:41:45 2018 +0100
@@ -31,13 +31,11 @@
 
   def error(message: String): Nothing = throw ERROR(message)
 
-  def cat_message(msg1: String, msg2: String): String =
-    if (msg1 == "") msg2
-    else if (msg2 == "") msg1
-    else msg1 + "\n" + msg2
+  def cat_message(msgs: String*): String =
+    cat_lines(msgs.iterator.filterNot(_ == ""))
 
-  def cat_error(msg1: String, msg2: String): Nothing =
-    error(cat_message(msg1, msg2))
+  def cat_error(msgs: String*): Nothing =
+    error(cat_message(msgs:_*))
 
 
   /* exceptions as values */
--- a/src/Pure/Isar/proof.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Isar/proof.ML	Thu May 10 15:41:45 2018 +0100
@@ -1301,7 +1301,7 @@
 local
 
 fun future_terminal_proof proof1 proof2 done state =
-  if Goal.future_enabled 3 andalso not (is_relevant state) then
+  if Future.proofs_enabled 3 andalso not (is_relevant state) then
     state |> future_proof (fn state' =>
       let val pos = Position.thread_data () in
         Execution.fork {name = "Proof.future_terminal_proof", pos = pos, pri = ~1}
--- a/src/Pure/Isar/toplevel.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Isar/toplevel.ML	Thu May 10 15:41:45 2018 +0100
@@ -668,19 +668,19 @@
   let val trs = tl (Thy_Syntax.flat_element elem)
   in fold (fn tr => fn t => get_timing tr + t) trs Time.zeroTime end;
 
-fun proof_future_enabled estimate st =
+fun future_proofs_enabled estimate st =
   (case try proof_of st of
     NONE => false
   | SOME state =>
       not (Proof.is_relevant state) andalso
        (if can (Proof.assert_bottom true) state
-        then Goal.future_enabled 1
-        else Goal.future_enabled 2 orelse Goal.future_enabled_timing estimate));
+        then Future.proofs_enabled 1
+        else Future.proofs_enabled 2 orelse Future.proofs_enabled_timing estimate));
 
 fun atom_result keywords tr st =
   let
     val st' =
-      if Goal.future_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then
+      if Future.proofs_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then
         (Execution.fork
           {name = "Toplevel.diag", pos = pos_of tr, pri = ~1}
           (fn () => command tr st); st)
@@ -696,7 +696,7 @@
         val (body_elems, end_tr) = element_rest;
         val estimate = timing_estimate elem;
       in
-        if not (proof_future_enabled estimate st')
+        if not (future_proofs_enabled estimate st')
         then
           let
             val proof_trs = maps Thy_Syntax.flat_element body_elems @ [end_tr];
--- a/src/Pure/ML/ml_compiler.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/ML/ml_compiler.ML	Thu May 10 15:41:45 2018 +0100
@@ -120,7 +120,7 @@
       |> map (fn (pos, markup, text) => Position.reported_text pos (markup ()) (text ()))
       |> Output.report;
     val _ =
-      if not (null persistent_reports) andalso redirect andalso Multithreading.enabled ()
+      if not (null persistent_reports) andalso redirect andalso Future.enabled ()
       then
         Execution.print
           {name = "ML_Compiler.report", pos = Position.thread_data (), pri = Task_Queue.urgent_pri}
--- a/src/Pure/ML/ml_statistics.scala	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/ML/ml_statistics.scala	Thu May 10 15:41:45 2018 +0100
@@ -39,7 +39,8 @@
 
   val tasks_fields: Fields =
     ("Future tasks",
-      List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive", "tasks_urgent"))
+      List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive",
+        "tasks_urgent", "tasks_total"))
 
   val workers_fields: Fields =
     ("Worker threads", List("workers_total", "workers_active", "workers_waiting"))
--- a/src/Pure/PIDE/command.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/PIDE/command.ML	Thu May 10 15:41:45 2018 +0100
@@ -190,7 +190,7 @@
   end) ();
 
 fun run keywords int tr st =
-  if Goal.future_enabled 1 andalso Keyword.is_diag keywords (Toplevel.name_of tr) then
+  if Future.proofs_enabled 1 andalso Keyword.is_diag keywords (Toplevel.name_of tr) then
     (Execution.fork {name = "Toplevel.diag", pos = Toplevel.pos_of tr, pri = ~1}
       (fn () => Toplevel.command_exception int tr st); ([], SOME st))
   else Toplevel.command_errors int tr st;
@@ -429,7 +429,7 @@
 
 fun run_print execution_id (Print {name, delay, pri, exec_id, print_process, ...}) =
   if ignore_process print_process then ()
-  else if pri <= 0 orelse (Multithreading.enabled () andalso Options.default_bool "parallel_print")
+  else if pri <= 0 orelse (Future.enabled () andalso Options.default_bool "parallel_print")
   then
     let
       val group = Future.worker_subgroup ();
--- a/src/Pure/PIDE/execution.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/PIDE/execution.ML	Thu May 10 15:41:45 2018 +0100
@@ -162,7 +162,7 @@
 fun fork_prints exec_id =
   (case Inttab.lookup (#2 (get_state ())) exec_id of
     SOME (_, prints) =>
-      if Multithreading.relevant prints then
+      if Future.relevant prints then
         let val pos = Position.thread_data () in
           List.app (fn {name, pri, body} =>
             ignore (fork {name = name, pos = pos, pri = pri} body)) (rev prints)
--- a/src/Pure/ROOT.scala	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/ROOT.scala	Thu May 10 15:41:45 2018 +0100
@@ -8,7 +8,7 @@
 {
   val ERROR = Exn.ERROR
   val error = Exn.error _
-  val cat_error = Exn.cat_error _
+  def cat_error(msgs: String*): Nothing = Exn.cat_error(msgs:_*)
 
   def using[A <: { def close() }, B](x: A)(f: A => B): B = Library.using(x)(f)
   val space_explode = Library.space_explode _
--- a/src/Pure/Thy/present.scala	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Thy/present.scala	Thu May 10 15:41:45 2018 +0100
@@ -272,6 +272,7 @@
 
     if (!result.ok) {
       cat_error(
+        Library.trim_line(result.err),
         cat_lines(Latex.latex_errors(dir, root_name) ::: Bibtex.bibtex_errors(dir, root_name)),
         "Failed to build document in " + File.path(dir.absolute_file))
     }
--- a/src/Pure/Thy/thy_info.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/Thy/thy_info.ML	Thu May 10 15:41:45 2018 +0100
@@ -388,7 +388,7 @@
       {document = document, symbols = symbols, bibtex_entries = bibtex_entries,
         last_timing = last_timing};
     val (_, tasks) = require_thys context [] qualifier master_dir imports String_Graph.empty;
-  in if Multithreading.enabled () then schedule_futures tasks else schedule_seq tasks end;
+  in if Multithreading.max_threads () > 1 then schedule_futures tasks else schedule_seq tasks end;
 
 fun use_thy name =
   use_theories
--- a/src/Pure/goal.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/goal.ML	Thu May 10 15:41:45 2018 +0100
@@ -23,8 +23,6 @@
   val finish: Proof.context -> thm -> thm
   val norm_result: Proof.context -> thm -> thm
   val skip_proofs_enabled: unit -> bool
-  val future_enabled: int -> bool
-  val future_enabled_timing: Time.time -> bool
   val future_result: Proof.context -> thm future -> term -> thm
   val prove_internal: Proof.context -> cterm list -> cterm -> (thm list -> tactic) -> thm
   val is_schematic: term -> bool
@@ -111,14 +109,6 @@
     else skip
   end;
 
-fun future_enabled n =
-  Multithreading.parallel_proofs_enabled n andalso
-  is_some (Future.worker_task ());
-
-fun future_enabled_timing t =
-  future_enabled 1 andalso
-    Time.toReal t >= Options.default_real "parallel_subproofs_threshold";
-
 
 (* future_result *)
 
@@ -176,7 +166,7 @@
 
     val schematic = exists is_schematic props;
     val immediate = is_none fork_pri;
-    val future = future_enabled 1;
+    val future = Future.proofs_enabled 1;
     val skip = not immediate andalso not schematic andalso future andalso skip_proofs_enabled ();
 
     val pos = Position.thread_data ();
@@ -258,7 +248,7 @@
 fun prove_sorry ctxt xs asms prop tac =
   if Config.get ctxt quick_and_dirty then
     prove ctxt xs asms prop (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt))
-  else (if future_enabled 1 then prove_future_pri ctxt ~2 else prove ctxt) xs asms prop tac;
+  else (if Future.proofs_enabled 1 then prove_future_pri ctxt ~2 else prove ctxt) xs asms prop tac;
 
 fun prove_sorry_global thy xs asms prop tac =
   Drule.export_without_context
--- a/src/Pure/par_tactical.ML	Thu May 10 15:41:34 2018 +0100
+++ b/src/Pure/par_tactical.ML	Thu May 10 15:41:45 2018 +0100
@@ -46,7 +46,7 @@
       Drule.strip_imp_prems (Thm.cprop_of st)
       |> map (Thm.adjust_maxidx_cterm ~1);
   in
-    if Multithreading.relevant goals andalso forall (fn g => Thm.maxidx_of_cterm g = ~1) goals then
+    if Future.relevant goals andalso forall (fn g => Thm.maxidx_of_cterm g = ~1) goals then
       let
         fun try_goal g =
           (case SINGLE tac (Goal.init g) of