merged
authorhaftmann
Thu, 26 Mar 2009 13:02:12 +0100
changeset 30735 e925e4a7f237
parent 30714 88bc86d7dec3 (diff)
parent 30734 ab05be086c4a (current diff)
child 30736 f5d9cc53f4c8
merged
doc-src/Codegen/codegen_process.pdf
doc-src/Codegen/codegen_process.ps
--- a/Admin/Mercurial/cvsids	Thu Mar 26 13:01:09 2009 +0100
+++ b/Admin/Mercurial/cvsids	Thu Mar 26 13:02:12 2009 +0100
@@ -1,6 +1,6 @@
 Identifiers of some old CVS file versions
 =========================================
 
+src/Pure/General/file.ML    1.18    6cdd6a8da9b9
+src/Pure/thm.ML             1.189   4b339d3907a0  (referenced in 25f28f9c28a3 as "2005-01-24 (revision 1.44)")
 src/Pure/type.ML            1.65    0d984ee030a1
-src/Pure/General/file.ML    1.18    6cdd6a8da9b9
-
--- a/Admin/isatest/isatest-stats	Thu Mar 26 13:01:09 2009 +0100
+++ b/Admin/isatest/isatest-stats	Thu Mar 26 13:02:12 2009 +0100
@@ -1,13 +1,12 @@
 #!/usr/bin/env bash
 #
-# $Id$
 # Author: Makarius
 #
 # DESCRIPTION: Standard statistics.
 
 THIS=$(cd "$(dirname "$0")"; pwd -P)
 
-PLATFORMS="at-poly at-sml-dev at64-poly at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp"
+PLATFORMS="at-poly at64-poly at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
 
 ISABELLE_SESSIONS="\
   HOL-Plain \
--- a/NEWS	Thu Mar 26 13:01:09 2009 +0100
+++ b/NEWS	Thu Mar 26 13:02:12 2009 +0100
@@ -330,6 +330,11 @@
 * Simplifier: simproc for let expressions now unfolds if bound variable
 occurs at most once in let expression body.  INCOMPATIBILITY.
 
+* New attribute "arith" for facts that should always be used automaticaly
+by arithmetic. It is intended to be used locally in proofs, eg
+assumes [arith]: "x > 0"
+Global usage is discouraged because of possible performance impact.
+
 * New classes "top" and "bot" with corresponding operations "top" and "bot"
 in theory Orderings;  instantiation of class "complete_lattice" requires
 instantiation of classes "top" and "bot".  INCOMPATIBILITY.
--- a/doc-src/HOL/HOL.tex	Thu Mar 26 13:01:09 2009 +0100
+++ b/doc-src/HOL/HOL.tex	Thu Mar 26 13:02:12 2009 +0100
@@ -1427,7 +1427,7 @@
 provides a decision procedure for \emph{linear arithmetic}: formulae involving
 addition and subtraction. The simplifier invokes a weak version of this
 decision procedure automatically. If this is not sufficent, you can invoke the
-full procedure \ttindex{arith_tac} explicitly.  It copes with arbitrary
+full procedure \ttindex{linear_arith_tac} explicitly.  It copes with arbitrary
 formulae involving {\tt=}, {\tt<}, {\tt<=}, {\tt+}, {\tt-}, {\tt Suc}, {\tt
   min}, {\tt max} and numerical constants. Other subterms are treated as
 atomic, while subformulae not involving numerical types are ignored. Quantified
@@ -1438,10 +1438,10 @@
 If {\tt k} is a numeral, then {\tt div k}, {\tt mod k} and
 {\tt k dvd} are also supported. The former two are eliminated
 by case distinctions, again blowing up the running time.
-If the formula involves explicit quantifiers, \texttt{arith_tac} may take
+If the formula involves explicit quantifiers, \texttt{linear_arith_tac} may take
 super-exponential time and space.
 
-If \texttt{arith_tac} fails, try to find relevant arithmetic results in
+If \texttt{linear_arith_tac} fails, try to find relevant arithmetic results in
 the library.  The theories \texttt{Nat} and \texttt{NatArith} contain
 theorems about {\tt<}, {\tt<=}, \texttt{+}, \texttt{-} and \texttt{*}.
 Theory \texttt{Divides} contains theorems about \texttt{div} and
--- a/src/HOL/Decision_Procs/Ferrack.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -1995,6 +1995,8 @@
   "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0)))
     (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))"
 
+code_reserved SML oo
+
 ML {* @{code ferrack_test} () *}
 
 oracle linr_oracle = {*
--- a/src/HOL/HoareParallel/OG_Examples.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/HoareParallel/OG_Examples.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -443,7 +443,7 @@
 --{* 32 subgoals left *}
 apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
 
-apply(tactic {* TRYALL (simple_arith_tac @{context}) *})
+apply(tactic {* TRYALL (linear_arith_tac @{context}) *})
 --{* 9 subgoals left *}
 apply (force simp add:less_Suc_eq)
 apply(drule sym)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -0,0 +1,11 @@
+(*  Title:      HOL/Imperative_HOL/Imperative_HOL_ex.thy
+    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
+*)
+
+header {* Mmonadic imperative HOL with examples *}
+
+theory Imperative_HOL_ex
+imports Imperative_HOL "ex/Imperative_Quicksort"
+begin
+
+end
--- a/src/HOL/Imperative_HOL/ROOT.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Imperative_HOL/ROOT.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -1,2 +1,2 @@
 
-use_thy "Imperative_HOL";
+use_thy "Imperative_HOL_ex";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -0,0 +1,639 @@
+(* Author: Lukas Bulwahn, TU Muenchen *)
+
+theory Imperative_Quicksort
+imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
+begin
+
+text {* We prove QuickSort correct in the Relational Calculus. *}
+
+definition swap :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
+where
+  "swap arr i j = (
+     do
+       x \<leftarrow> nth arr i;
+       y \<leftarrow> nth arr j;
+       upd i y arr;
+       upd j x arr;
+       return ()
+     done)"
+
+lemma swap_permutes:
+  assumes "crel (swap a i j) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+  using assms
+  unfolding swap_def
+  by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
+
+function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
+where
+  "part1 a left right p = (
+     if (right \<le> left) then return right
+     else (do
+       v \<leftarrow> nth a left;
+       (if (v \<le> p) then (part1 a (left + 1) right p)
+                    else (do swap a left right;
+  part1 a left (right - 1) p done))
+     done))"
+by pat_completeness auto
+
+termination
+by (relation "measure (\<lambda>(_,l,r,_). r - l )") auto
+
+declare part1.simps[simp del]
+
+lemma part_permutes:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+  using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  thus ?case
+    unfolding part1.simps [of a l r p]
+    by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes)
+qed
+
+lemma part_returns_index_in_bounds:
+  assumes "crel (part1 a l r p) h h' rs"
+  assumes "l \<le> r"
+  shows "l \<le> rs \<and> rs \<le> r"
+using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr `l \<le> r` show ?thesis
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+  next
+    case False (* recursive case *)
+    note rec_condition = this
+    let ?v = "get_array a h ! l"
+    show ?thesis
+    proof (cases "?v \<le> p")
+      case True
+      with cr False
+      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from rec_condition have "l + 1 \<le> r" by arith
+      from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`]
+      show ?thesis by simp
+    next
+      case False
+      with rec_condition cr
+      obtain h1 where swp: "crel (swap a l r) h h1 ()"
+        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from rec_condition have "l \<le> r - 1" by arith
+      from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastsimp
+    qed
+  qed
+qed
+
+lemma part_length_remains:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "Heap.length a h = Heap.length a h'"
+using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr show ?thesis
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+  next
+    case False (* recursive case *)
+    with cr 1 show ?thesis
+      unfolding part1.simps [of a l r p] swap_def
+      by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp
+  qed
+qed
+
+lemma part_outer_remains:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
+  using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr show ?thesis
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+  next
+    case False (* recursive case *)
+    note rec_condition = this
+    let ?v = "get_array a h ! l"
+    show ?thesis
+    proof (cases "?v \<le> p")
+      case True
+      with cr False
+      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from 1(1)[OF rec_condition True rec1]
+      show ?thesis by fastsimp
+    next
+      case False
+      with rec_condition cr
+      obtain h1 where swp: "crel (swap a l r) h h1 ()"
+        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from swp rec_condition have
+        "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
+	unfolding swap_def
+	by (elim crelE crel_nth crel_upd crel_return) auto
+      with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
+    qed
+  qed
+qed
+
+
+lemma part_partitions:
+  assumes "crel (part1 a l r p) h h' rs"
+  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> p)
+  \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! i \<ge> p)"
+  using assms
+proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
+  case (1 a l r p h h' rs)
+  note cr = `crel (part1 a l r p) h h' rs`
+  
+  show ?case
+  proof (cases "r \<le> l")
+    case True (* Terminating case *)
+    with cr have "rs = r"
+      unfolding part1.simps[of a l r p]
+      by (elim crelE crel_if crel_return crel_nth) auto
+    with True
+    show ?thesis by auto
+  next
+    case False (* recursive case *)
+    note lr = this
+    let ?v = "get_array a h ! l"
+    show ?thesis
+    proof (cases "?v \<le> p")
+      case True
+      with lr cr
+      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
+	by fastsimp
+      have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith
+      with 1(1)[OF False True rec1] a_l show ?thesis
+	by auto
+    next
+      case False
+      with lr cr
+      obtain h1 where swp: "crel (swap a l r) h h1 ()"
+        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
+        unfolding part1.simps[of a l r p]
+        by (elim crelE crel_nth crel_if crel_return) auto
+      from swp False have "get_array a h1 ! r \<ge> p"
+	unfolding swap_def
+	by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
+      with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
+	by fastsimp
+      have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
+      with 1(2)[OF lr False rec2] a_r show ?thesis
+	by auto
+    qed
+  qed
+qed
+
+
+fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
+where
+  "partition a left right = (do
+     pivot \<leftarrow> nth a right;
+     middle \<leftarrow> part1 a left (right - 1) pivot;
+     v \<leftarrow> nth a middle;
+     m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle);
+     swap a m right;
+     return m
+   done)"
+
+declare partition.simps[simp del]
+
+lemma partition_permutes:
+  assumes "crel (partition a l r) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+proof -
+    from assms part_permutes swap_permutes show ?thesis
+      unfolding partition.simps
+      by (elim crelE crel_return crel_nth crel_if crel_upd) auto
+qed
+
+lemma partition_length_remains:
+  assumes "crel (partition a l r) h h' rs"
+  shows "Heap.length a h = Heap.length a h'"
+proof -
+  from assms part_length_remains show ?thesis
+    unfolding partition.simps swap_def
+    by (elim crelE crel_return crel_nth crel_if crel_upd) auto
+qed
+
+lemma partition_outer_remains:
+  assumes "crel (partition a l r) h h' rs"
+  assumes "l < r"
+  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
+proof -
+  from assms part_outer_remains part_returns_index_in_bounds show ?thesis
+    unfolding partition.simps swap_def
+    by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp
+qed
+
+lemma partition_returns_index_in_bounds:
+  assumes crel: "crel (partition a l r) h h' rs"
+  assumes "l < r"
+  shows "l \<le> rs \<and> rs \<le> r"
+proof -
+  from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle"
+    and rs_equals: "rs = (if get_array a h'' ! middle \<le> get_array a h ! r then middle + 1
+         else middle)"
+    unfolding partition.simps
+    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
+  from `l < r` have "l \<le> r - 1" by arith
+  from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto
+qed
+
+lemma partition_partitions:
+  assumes crel: "crel (partition a l r) h h' rs"
+  assumes "l < r"
+  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> get_array a h' ! rs) \<and>
+  (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! rs \<le> get_array a h' ! i)"
+proof -
+  let ?pivot = "get_array a h ! r" 
+  from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle"
+    and swap: "crel (swap a rs r) h1 h' ()"
+    and rs_equals: "rs = (if get_array a h1 ! middle \<le> ?pivot then middle + 1
+         else middle)"
+    unfolding partition.simps
+    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
+  from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
+    (Heap.upd a rs (get_array a h1 ! r) h1)"
+    unfolding swap_def
+    by (elim crelE crel_return crel_nth crel_upd) simp
+  from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
+    unfolding swap_def
+    by (elim crelE crel_return crel_nth crel_upd) simp
+  from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
+    unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
+  from `l < r` have "l \<le> r - 1" by simp 
+  note middle_in_bounds = part_returns_index_in_bounds[OF part this]
+  from part_outer_remains[OF part] `l < r`
+  have "get_array a h ! r = get_array a h1 ! r"
+    by fastsimp
+  with swap
+  have right_remains: "get_array a h ! r = get_array a h' ! rs"
+    unfolding swap_def
+    by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
+  from part_partitions [OF part]
+  show ?thesis
+  proof (cases "get_array a h1 ! middle \<le> ?pivot")
+    case True
+    with rs_equals have rs_equals: "rs = middle + 1" by simp
+    { 
+      fix i
+      assume i_is_left: "l \<le> i \<and> i < rs"
+      with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
+      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+      from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
+      with part_partitions[OF part] right_remains True
+      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
+      with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
+	unfolding Heap.upd_def Heap.length_def by simp
+    }
+    moreover
+    {
+      fix i
+      assume "rs < i \<and> i \<le> r"
+
+      hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
+      hence "get_array a h' ! rs \<le> get_array a h' ! i"
+      proof
+	assume i_is: "rs < i \<and> i \<le> r - 1"
+	with swap_length_remains in_bounds middle_in_bounds rs_equals
+	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+	from part_partitions[OF part] rs_equals right_remains i_is
+	have "get_array a h' ! rs \<le> get_array a h1 ! i"
+	  by fastsimp
+	with i_props h'_def show ?thesis by fastsimp
+      next
+	assume i_is: "rs < i \<and> i = r"
+	with rs_equals have "Suc middle \<noteq> r" by arith
+	with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
+	with part_partitions[OF part] right_remains 
+	have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
+	  by fastsimp
+	with i_is True rs_equals right_remains h'_def
+	show ?thesis using in_bounds
+	  unfolding Heap.upd_def Heap.length_def
+	  by auto
+      qed
+    }
+    ultimately show ?thesis by auto
+  next
+    case False
+    with rs_equals have rs_equals: "middle = rs" by simp
+    { 
+      fix i
+      assume i_is_left: "l \<le> i \<and> i < rs"
+      with swap_length_remains in_bounds middle_in_bounds rs_equals
+      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+      from part_partitions[OF part] rs_equals right_remains i_is_left
+      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
+      with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
+	unfolding Heap.upd_def by simp
+    }
+    moreover
+    {
+      fix i
+      assume "rs < i \<and> i \<le> r"
+      hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
+      hence "get_array a h' ! rs \<le> get_array a h' ! i"
+      proof
+	assume i_is: "rs < i \<and> i \<le> r - 1"
+	with swap_length_remains in_bounds middle_in_bounds rs_equals
+	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+	from part_partitions[OF part] rs_equals right_remains i_is
+	have "get_array a h' ! rs \<le> get_array a h1 ! i"
+	  by fastsimp
+	with i_props h'_def show ?thesis by fastsimp
+      next
+	assume i_is: "i = r"
+	from i_is False rs_equals right_remains h'_def
+	show ?thesis using in_bounds
+	  unfolding Heap.upd_def Heap.length_def
+	  by auto
+      qed
+    }
+    ultimately
+    show ?thesis by auto
+  qed
+qed
+
+
+function quicksort :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
+where
+  "quicksort arr left right =
+     (if (right > left)  then
+        do
+          pivotNewIndex \<leftarrow> partition arr left right;
+          pivotNewIndex \<leftarrow> assert (\<lambda>x. left \<le> x \<and> x \<le> right) pivotNewIndex;
+          quicksort arr left (pivotNewIndex - 1);
+          quicksort arr (pivotNewIndex + 1) right
+        done
+     else return ())"
+by pat_completeness auto
+
+(* For termination, we must show that the pivotNewIndex is between left and right *) 
+termination
+by (relation "measure (\<lambda>(a, l, r). (r - l))") auto
+
+declare quicksort.simps[simp del]
+
+
+lemma quicksort_permutes:
+  assumes "crel (quicksort a l r) h h' rs"
+  shows "multiset_of (get_array a h') 
+  = multiset_of (get_array a h)"
+  using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  with partition_permutes show ?case
+    unfolding quicksort.simps [of a l r]
+    by (elim crel_if crelE crel_assert crel_return) auto
+qed
+
+lemma length_remains:
+  assumes "crel (quicksort a l r) h h' rs"
+  shows "Heap.length a h = Heap.length a h'"
+using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  with partition_length_remains show ?case
+    unfolding quicksort.simps [of a l r]
+    by (elim crel_if crelE crel_assert crel_return) auto
+qed
+
+lemma quicksort_outer_remains:
+  assumes "crel (quicksort a l r) h h' rs"
+   shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
+  using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  note cr = `crel (quicksort a l r) h h' rs`
+  thus ?case
+  proof (cases "r > l")
+    case False
+    with cr have "h' = h"
+      unfolding quicksort.simps [of a l r]
+      by (elim crel_if crel_return) auto
+    thus ?thesis by simp
+  next
+  case True
+   { 
+      fix h1 h2 p ret1 ret2 i
+      assume part: "crel (partition a l r) h h1 p"
+      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1"
+      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2"
+      assume pivot: "l \<le> p \<and> p \<le> r"
+      assume i_outer: "i < l \<or> r < i"
+      from  partition_outer_remains [OF part True] i_outer
+      have "get_array a h !i = get_array a h1 ! i" by fastsimp
+      moreover
+      with 1(1) [OF True pivot qs1] pivot i_outer
+      have "get_array a h1 ! i = get_array a h2 ! i" by auto
+      moreover
+      with qs2 1(2) [of p h2 h' ret2] True pivot i_outer
+      have "get_array a h2 ! i = get_array a h' ! i" by auto
+      ultimately have "get_array a h ! i= get_array a h' ! i" by simp
+    }
+    with cr show ?thesis
+      unfolding quicksort.simps [of a l r]
+      by (elim crel_if crelE crel_assert crel_return) auto
+  qed
+qed
+
+lemma quicksort_is_skip:
+  assumes "crel (quicksort a l r) h h' rs"
+  shows "r \<le> l \<longrightarrow> h = h'"
+  using assms
+  unfolding quicksort.simps [of a l r]
+  by (elim crel_if crel_return) auto
+ 
+lemma quicksort_sorts:
+  assumes "crel (quicksort a l r) h h' rs"
+  assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" 
+  shows "sorted (subarray l (r + 1) a h')"
+  using assms
+proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
+  case (1 a l r h h' rs)
+  note cr = `crel (quicksort a l r) h h' rs`
+  thus ?case
+  proof (cases "r > l")
+    case False
+    hence "l \<ge> r + 1 \<or> l = r" by arith 
+    with length_remains[OF cr] 1(5) show ?thesis
+      by (auto simp add: subarray_Nil subarray_single)
+  next
+    case True
+    { 
+      fix h1 h2 p
+      assume part: "crel (partition a l r) h h1 p"
+      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()"
+      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()"
+      from partition_returns_index_in_bounds [OF part True]
+      have pivot: "l\<le> p \<and> p \<le> r" .
+     note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part]
+      from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1]
+      have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto)
+        (*-- First of all, by induction hypothesis both sublists are sorted. *)
+      from 1(1)[OF True pivot qs1] length_remains pivot 1(5) 
+      have IH1: "sorted (subarray l p a h2)"  by (cases p, auto simp add: subarray_Nil)
+      from quicksort_outer_remains [OF qs2] length_remains
+      have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
+	by (simp add: subarray_eq_samelength_iff)
+      with IH1 have IH1': "sorted (subarray l p a h')" by simp
+      from 1(2)[OF True pivot qs2] pivot 1(5) length_remains
+      have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
+        by (cases "Suc p \<le> r", auto simp add: subarray_Nil)
+           (* -- Secondly, both sublists remain partitioned. *)
+      from partition_partitions[OF part True]
+      have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> get_array a h1 ! p "
+        and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> get_array a h1 ! p \<le> j"
+        by (auto simp add: all_in_set_subarray_conv)
+      from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
+        length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
+      have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
+	unfolding Heap.length_def subarray_def by (cases p, auto)
+      with left_subarray_remains part_conds1 pivot_unchanged
+      have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
+        by (simp, subst set_of_multiset_of[symmetric], simp)
+          (* -- These steps are the analogous for the right sublist \<dots> *)
+      from quicksort_outer_remains [OF qs1] length_remains
+      have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
+	by (auto simp add: subarray_eq_samelength_iff)
+      from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
+        length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
+      have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
+        unfolding Heap.length_def subarray_def by auto
+      with right_subarray_remains part_conds2 pivot_unchanged
+      have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
+        by (simp, subst set_of_multiset_of[symmetric], simp)
+          (* -- Thirdly and finally, we show that the array is sorted
+          following from the facts above. *)
+      from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
+	by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
+      with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
+	unfolding subarray_def
+	apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
+	by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
+    }
+    with True cr show ?thesis
+      unfolding quicksort.simps [of a l r]
+      by (elim crel_if crel_return crelE crel_assert) auto
+  qed
+qed
+
+
+lemma quicksort_is_sort:
+  assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
+  shows "get_array a h' = sort (get_array a h)"
+proof (cases "get_array a h = []")
+  case True
+  with quicksort_is_skip[OF crel] show ?thesis
+  unfolding Heap.length_def by simp
+next
+  case False
+  from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
+    unfolding Heap.length_def subarray_def by auto
+  with length_remains[OF crel] have "sorted (get_array a h')"
+    unfolding Heap.length_def by simp
+  with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
+qed
+
+subsection {* No Errors in quicksort *}
+text {* We have proved that quicksort sorts (if no exceptions occur).
+We will now show that exceptions do not occur. *}
+
+lemma noError_part1: 
+  assumes "l < Heap.length a h" "r < Heap.length a h"
+  shows "noError (part1 a l r p) h"
+  using assms
+proof (induct a l r p arbitrary: h rule: part1.induct)
+  case (1 a l r p)
+  thus ?case
+    unfolding part1.simps [of a l r] swap_def
+    by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return)
+qed
+
+lemma noError_partition:
+  assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
+  shows "noError (partition a l r) h"
+using assms
+unfolding partition.simps swap_def
+apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return)
+apply (frule part_length_remains)
+apply (frule part_returns_index_in_bounds)
+apply auto
+apply (frule part_length_remains)
+apply (frule part_returns_index_in_bounds)
+apply auto
+apply (frule part_length_remains)
+apply auto
+done
+
+lemma noError_quicksort:
+  assumes "l < Heap.length a h" "r < Heap.length a h"
+  shows "noError (quicksort a l r) h"
+using assms
+proof (induct a l r arbitrary: h rule: quicksort.induct)
+  case (1 a l ri h)
+  thus ?case
+    unfolding quicksort.simps [of a l ri]
+    apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition)
+    apply (frule partition_returns_index_in_bounds)
+    apply auto
+    apply (frule partition_returns_index_in_bounds)
+    apply auto
+    apply (auto elim!: crel_assert dest!: partition_length_remains length_remains)
+    apply (subgoal_tac "Suc r \<le> ri \<or> r = ri") 
+    apply (erule disjE)
+    apply auto
+    unfolding quicksort.simps [of a "Suc ri" ri]
+    apply (auto intro!: noError_if noError_return)
+    done
+qed
+
+
+subsection {* Example *}
+
+definition "qsort a = do
+    k \<leftarrow> length a;
+    quicksort a 0 (k - 1);
+    return a
+  done"
+
+ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
+
+export_code qsort in SML_imp module_name QSort
+export_code qsort in OCaml module_name QSort file -
+export_code qsort in OCaml_imp module_name QSort file -
+export_code qsort in Haskell module_name QSort file -
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ex/Subarray.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -0,0 +1,66 @@
+theory Subarray
+imports Array Sublist
+begin
+
+definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
+where
+  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
+
+lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
+apply (simp add: subarray_def Heap.upd_def)
+apply (simp add: sublist'_update1)
+done
+
+lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
+apply (simp add: subarray_def Heap.upd_def)
+apply (subst sublist'_update2)
+apply fastsimp
+apply simp
+done
+
+lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
+unfolding subarray_def Heap.upd_def
+by (simp add: sublist'_update3)
+
+lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
+by (simp add: subarray_def sublist'_Nil')
+
+lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
+by (simp add: subarray_def Heap.length_def sublist'_single)
+
+lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
+by (simp add: subarray_def Heap.length_def length_sublist')
+
+lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
+by (simp add: length_subarray)
+
+lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_front)
+
+lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_back)
+
+lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
+unfolding subarray_def
+by (simp add: sublist'_append)
+
+lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
+unfolding Heap.length_def subarray_def
+by (simp add: sublist'_all)
+
+lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
+unfolding Heap.length_def subarray_def
+by (simp add: nth_sublist')
+
+lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
+unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
+
+lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
+
+lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Imperative_HOL/ex/Sublist.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -0,0 +1,505 @@
+(* $Id$ *)
+
+header {* Slices of lists *}
+
+theory Sublist
+imports Multiset
+begin
+
+
+lemma sublist_split: "i \<le> j \<and> j \<le> k \<Longrightarrow> sublist xs {i..<j} @ sublist xs {j..<k} = sublist xs {i..<k}" 
+apply (induct xs arbitrary: i j k)
+apply simp
+apply (simp only: sublist_Cons)
+apply simp
+apply safe
+apply simp
+apply (erule_tac x="0" in meta_allE)
+apply (erule_tac x="j - 1" in meta_allE)
+apply (erule_tac x="k - 1" in meta_allE)
+apply (subgoal_tac "0 \<le> j - 1 \<and> j - 1 \<le> k - 1")
+apply simp
+apply (subgoal_tac "{ja. Suc ja < j} = {0..<j - Suc 0}")
+apply (subgoal_tac "{ja. j \<le> Suc ja \<and> Suc ja < k} = {j - Suc 0..<k - Suc 0}")
+apply (subgoal_tac "{j. Suc j < k} = {0..<k - Suc 0}")
+apply simp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply (erule_tac x="i - 1" in meta_allE)
+apply (erule_tac x="j - 1" in meta_allE)
+apply (erule_tac x="k - 1" in meta_allE)
+apply (subgoal_tac " {ja. i \<le> Suc ja \<and> Suc ja < j} = {i - 1 ..<j - 1}")
+apply (subgoal_tac " {ja. j \<le> Suc ja \<and> Suc ja < k} = {j - 1..<k - 1}")
+apply (subgoal_tac "{j. i \<le> Suc j \<and> Suc j < k} = {i - 1..<k - 1}")
+apply (subgoal_tac " i - 1 \<le> j - 1 \<and> j - 1 \<le> k - 1")
+apply simp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+apply fastsimp
+done
+
+lemma sublist_update1: "i \<notin> inds \<Longrightarrow> sublist (xs[i := v]) inds = sublist xs inds"
+apply (induct xs arbitrary: i inds)
+apply simp
+apply (case_tac i)
+apply (simp add: sublist_Cons)
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_update2: "i \<in> inds \<Longrightarrow> sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \<in> inds. k < i}):= v]"
+proof (induct xs arbitrary: i inds)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs)
+  thus ?case
+  proof (cases i)
+    case 0 with Cons show ?thesis by (simp add: sublist_Cons)
+  next
+    case (Suc i')
+    with Cons show ?thesis
+      apply simp
+      apply (simp add: sublist_Cons)
+      apply auto
+      apply (auto simp add: nat.split)
+      apply (simp add: card_less_Suc[symmetric])
+      apply (simp add: card_less_Suc2)
+      done
+  qed
+qed
+
+lemma sublist_update: "sublist (xs[i := v]) inds = (if i \<in> inds then (sublist xs inds)[(card {k \<in> inds. k < i}) := v] else sublist xs inds)"
+by (simp add: sublist_update1 sublist_update2)
+
+lemma sublist_take: "sublist xs {j. j < m} = take m xs"
+apply (induct xs arbitrary: m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_take': "sublist xs {0..<m} = take m xs"
+apply (induct xs arbitrary: m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons sublist_take)
+done
+
+lemma sublist_all[simp]: "sublist xs {j. j < length xs} = xs"
+apply (induct xs)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_all'[simp]: "sublist xs {0..<length xs} = xs"
+apply (induct xs)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_single: "a < length xs \<Longrightarrow> sublist xs {a} = [xs ! a]"
+apply (induct xs arbitrary: a)
+apply simp
+apply(case_tac aa)
+apply simp
+apply (simp add: sublist_Cons)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+lemma sublist_is_Nil: "\<forall>i \<in> inds. i \<ge> length xs \<Longrightarrow> sublist xs inds = []" 
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply auto
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply auto
+done
+
+lemma sublist_Nil': "sublist xs inds = [] \<Longrightarrow> \<forall>i \<in> inds. i \<ge> length xs"
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply (auto split: if_splits)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (case_tac x, auto)
+done
+
+lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\<forall>i \<in> inds. i \<ge> length xs)"
+apply (induct xs arbitrary: inds)
+apply simp
+apply (simp add: sublist_Cons)
+apply auto
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (case_tac x, auto)
+done
+
+lemma sublist_eq_subseteq: " \<lbrakk> inds' \<subseteq> inds; sublist xs inds = sublist ys inds \<rbrakk> \<Longrightarrow> sublist xs inds' = sublist ys inds'"
+apply (induct xs arbitrary: ys inds inds')
+apply simp
+apply (drule sym, rule sym)
+apply (simp add: sublist_Nil, fastsimp)
+apply (case_tac ys)
+apply (simp add: sublist_Nil, fastsimp)
+apply (auto simp add: sublist_Cons)
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
+apply fastsimp
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
+apply fastsimp
+done
+
+lemma sublist_eq: "\<lbrakk> \<forall>i \<in> inds. ((i < length xs) \<and> (i < length ys)) \<or> ((i \<ge> length xs ) \<and> (i \<ge> length ys)); \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
+apply (induct xs arbitrary: ys inds)
+apply simp
+apply (rule sym, simp add: sublist_Nil)
+apply (case_tac ys)
+apply (simp add: sublist_Nil)
+apply (auto simp add: sublist_Cons)
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply fastsimp
+apply (erule_tac x="list" in meta_allE)
+apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
+apply fastsimp
+done
+
+lemma sublist_eq_samelength: "\<lbrakk> length xs = length ys; \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
+by (rule sublist_eq, auto)
+
+lemma sublist_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist xs inds = sublist ys inds) = (\<forall>i \<in> inds. xs ! i = ys ! i)"
+apply (induct xs arbitrary: ys inds)
+apply simp
+apply (rule sym, simp add: sublist_Nil)
+apply (case_tac ys)
+apply (simp add: sublist_Nil)
+apply (auto simp add: sublist_Cons)
+apply (case_tac i)
+apply auto
+apply (case_tac i)
+apply auto
+done
+
+section {* Another sublist function *}
+
+function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  "sublist' n m [] = []"
+| "sublist' n 0 xs = []"
+| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)"
+| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs"
+by pat_completeness auto
+termination by lexicographic_order
+
+subsection {* Proving equivalence to the other sublist command *}
+
+lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac n)
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons)
+done
+
+
+lemma "sublist' n m xs = sublist xs {n..<m}"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac n, case_tac m)
+apply simp
+apply simp
+apply (simp add: sublist_take')
+apply (case_tac m)
+apply simp
+apply (simp add: sublist_Cons sublist'_sublist)
+done
+
+
+subsection {* Showing equivalence to use of drop and take for definition *}
+
+lemma "sublist' n m xs = take (m - n) (drop n xs)"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+subsection {* General lemma about sublist *}
+
+lemma sublist'_Nil[simp]: "sublist' i j [] = []"
+by simp
+
+lemma sublist'_Cons[simp]: "sublist' i (Suc j) (x#xs) = (case i of 0 \<Rightarrow> (x # sublist' 0 j xs) | Suc i' \<Rightarrow>  sublist' i' j xs)"
+by (cases i) auto
+
+lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))"
+apply (cases j)
+apply auto
+apply (cases i)
+apply auto
+done
+
+lemma sublist_n_0: "sublist' n 0 xs = []"
+by (induct xs, auto)
+
+lemma sublist'_Nil': "n \<ge> m \<Longrightarrow> sublist' n m xs = []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_Nil2: "n \<ge> length xs \<Longrightarrow> sublist' n m xs = []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \<ge> m) \<or> (n \<ge> length xs))"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_notNil: "\<lbrakk> n < length xs; n < m \<rbrakk> \<Longrightarrow> sublist' n m xs \<noteq> []"
+apply (induct xs arbitrary: n m)
+apply simp
+apply (case_tac m)
+apply simp
+apply (case_tac n)
+apply simp
+apply simp
+done
+
+lemma sublist'_single: "n < length xs \<Longrightarrow> sublist' n (Suc n) xs = [xs ! n]"
+apply (induct xs arbitrary: n)
+apply simp
+apply simp
+apply (case_tac n)
+apply (simp add: sublist_n_0)
+apply simp
+done
+
+lemma sublist'_update1: "i \<ge> m \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
+apply (induct xs arbitrary: n m i)
+apply simp
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_update2: "i < n \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
+apply (induct xs arbitrary: n m i)
+apply simp
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_update3: "\<lbrakk>n \<le> i; i < m\<rbrakk> \<Longrightarrow> sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]"
+proof (induct xs arbitrary: n m i)
+  case Nil thus ?case by auto
+next
+  case (Cons x xs)
+  thus ?case
+    apply -
+    apply auto
+    apply (cases i)
+    apply auto
+    apply (cases i)
+    apply auto
+    done
+qed
+
+lemma "\<lbrakk> sublist' i j xs = sublist' i j ys; n \<ge> i; m \<le> j \<rbrakk> \<Longrightarrow> sublist' n m xs = sublist' n m ys"
+proof (induct xs arbitrary: i j ys n m)
+  case Nil
+  thus ?case
+    apply -
+    apply (rule sym, drule sym)
+    apply (simp add: sublist'_Nil)
+    apply (simp add: sublist'_Nil3)
+    apply arith
+    done
+next
+  case (Cons x xs i j ys n m)
+  note c = this
+  thus ?case
+  proof (cases m)
+    case 0 thus ?thesis by (simp add: sublist_n_0)
+  next
+    case (Suc m')
+    note a = this
+    thus ?thesis
+    proof (cases n)
+      case 0 note b = this
+      show ?thesis
+      proof (cases ys)
+	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
+      next
+	case (Cons y ys)
+	show ?thesis
+	proof (cases j)
+	  case 0 with a b Cons.prems show ?thesis by simp
+	next
+	  case (Suc j') with a b Cons.prems Cons show ?thesis 
+	    apply -
+	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
+	    done
+	qed
+      qed
+    next
+      case (Suc n')
+      show ?thesis
+      proof (cases ys)
+	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
+      next
+	case (Cons y ys) with Suc a Cons.prems show ?thesis
+	  apply -
+	  apply simp
+	  apply (cases j)
+	  apply simp
+	  apply (cases i)
+	  apply simp
+	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
+	  apply simp
+	  apply simp
+	  apply simp
+	  apply simp
+	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
+	  apply simp
+	  apply simp
+	  apply simp
+	  done
+      qed
+    qed
+  qed
+qed
+
+lemma length_sublist': "j \<le> length xs \<Longrightarrow> length (sublist' i j xs) = j - i"
+by (induct xs arbitrary: i j, auto)
+
+lemma sublist'_front: "\<lbrakk> i < j; i < length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = xs ! i # sublist' (Suc i) j xs"
+apply (induct xs arbitrary: a i j)
+apply simp
+apply (case_tac j)
+apply simp
+apply (case_tac i)
+apply simp
+apply simp
+done
+
+lemma sublist'_back: "\<lbrakk> i < j; j \<le> length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]"
+apply (induct xs arbitrary: a i j)
+apply simp
+apply simp
+apply (case_tac j)
+apply simp
+apply auto
+apply (case_tac nat)
+apply auto
+done
+
+(* suffices that j \<le> length xs and length ys *) 
+lemma sublist'_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist' i j xs  = sublist' i j ys) = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> xs ! i' = ys ! i')"
+proof (induct xs arbitrary: ys i j)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs)
+  thus ?case
+    apply -
+    apply (cases ys)
+    apply simp
+    apply simp
+    apply auto
+    apply (case_tac i', auto)
+    apply (erule_tac x="Suc i'" in allE, auto)
+    apply (erule_tac x="i' - 1" in allE, auto)
+    apply (case_tac i', auto)
+    apply (erule_tac x="Suc i'" in allE, auto)
+    done
+qed
+
+lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs"
+by (induct xs, auto)
+
+lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" 
+by (induct xs arbitrary: i j n m) (auto simp add: min_diff)
+
+lemma sublist'_append: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow>(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs"
+by (induct xs arbitrary: i j k) auto
+
+lemma nth_sublist': "\<lbrakk> k < j - i; j \<le> length xs \<rbrakk> \<Longrightarrow> (sublist' i j xs) ! k = xs ! (i + k)"
+apply (induct xs arbitrary: i j k)
+apply auto
+apply (case_tac k)
+apply auto
+apply (case_tac i)
+apply auto
+done
+
+lemma set_sublist': "set (sublist' i j xs) = {x. \<exists>k. i \<le> k \<and> k < j \<and> k < List.length xs \<and> x = xs ! k}"
+apply (simp add: sublist'_sublist)
+apply (simp add: set_sublist)
+apply auto
+done
+
+lemma all_in_set_sublist'_conv: "(\<forall>j. j \<in> set (sublist' l r xs) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
+unfolding set_sublist' by blast
+
+lemma ball_in_set_sublist'_conv: "(\<forall>j \<in> set (sublist' l r xs). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
+unfolding set_sublist' by blast
+
+
+lemma multiset_of_sublist:
+assumes l_r: "l \<le> r \<and> r \<le> List.length xs"
+assumes left: "\<forall> i. i < l \<longrightarrow> (xs::'a list) ! i = ys ! i"
+assumes right: "\<forall> i. i \<ge> r \<longrightarrow> (xs::'a list) ! i = ys ! i"
+assumes multiset: "multiset_of xs = multiset_of ys"
+  shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)"
+proof -
+  from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") 
+    by (simp add: sublist'_append)
+  from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length)
+  with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") 
+    by (simp add: sublist'_append)
+  from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp
+  moreover
+  from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys"
+    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
+  moreover
+  from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys"
+    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
+  moreover
+  ultimately show ?thesis by (simp add: multiset_of_append)
+qed
+
+
+end
--- a/src/HOL/IsaMakefile	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/IsaMakefile	Thu Mar 26 13:02:12 2009 +0100
@@ -649,7 +649,11 @@
 $(LOG)/HOL-Imperative_HOL.gz: $(OUT)/HOL Imperative_HOL/Heap.thy \
   Imperative_HOL/Heap_Monad.thy Imperative_HOL/Array.thy \
   Imperative_HOL/Relational.thy \
-  Imperative_HOL/Ref.thy Imperative_HOL/Imperative_HOL.thy
+  Imperative_HOL/Ref.thy Imperative_HOL/Imperative_HOL.thy \
+  Imperative_HOL/Imperative_HOL_ex.thy \
+  Imperative_HOL/ex/Imperative_Quicksort.thy \
+  Imperative_HOL/ex/Subarray.thy \
+  Imperative_HOL/ex/Sublist.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Imperative_HOL
 
 
@@ -836,7 +840,7 @@
   ex/Formal_Power_Series_Examples.thy ex/Fundefs.thy			\
   ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
   ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
-  ex/Hilbert_Classical.thy ex/ImperativeQuicksort.thy			\
+  ex/Hilbert_Classical.thy			\
   ex/Induction_Scheme.thy ex/InductiveInvariant.thy			\
   ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy		\
   ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy ex/MergeSort.thy		\
@@ -845,8 +849,8 @@
   ex/Quickcheck_Examples.thy ex/Quickcheck_Generators.thy ex/ROOT.ML	\
   ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy			\
   ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy		\
-  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy ex/Subarray.thy		\
-  ex/Sublist.thy ex/Sudoku.thy ex/Tarski.thy ex/Term_Of_Syntax.thy	\
+  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy \
+  ex/Sudoku.thy ex/Tarski.thy ex/Term_Of_Syntax.thy	\
   ex/Termination.thy ex/Unification.thy ex/document/root.bib		\
   ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
   ex/Predicate_Compile.thy ex/predicate_compile.ML
--- a/src/HOL/NSA/hypreal_arith.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/NSA/hypreal_arith.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -30,10 +30,10 @@
     Simplifier.simproc (the_context ())
       "fast_hypreal_arith" 
       ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"]
-    (K LinArith.lin_arith_simproc);
+    (K Lin_Arith.lin_arith_simproc);
 
 val hypreal_arith_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = mult_mono_thms,
     inj_thms = real_inj_thms @ inj_thms,
--- a/src/HOL/Nat.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Nat.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -63,9 +63,8 @@
 end
 
 lemma Suc_not_Zero: "Suc m \<noteq> 0"
-  apply (simp add: Zero_nat_def Suc_def Abs_Nat_inject [unfolded mem_def]
+  by (simp add: Zero_nat_def Suc_def Abs_Nat_inject [unfolded mem_def]
     Rep_Nat [unfolded mem_def] Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep [unfolded mem_def])
-  done
 
 lemma Zero_not_Suc: "0 \<noteq> Suc m"
   by (rule not_sym, rule Suc_not_Zero not_sym)
@@ -82,7 +81,7 @@
   done
 
 lemma nat_induct [case_names 0 Suc, induct type: nat]:
-  -- {* for backward compatibility -- naming of variables differs *}
+  -- {* for backward compatibility -- names of variables differ *}
   fixes n
   assumes "P 0"
     and "\<And>n. P n \<Longrightarrow> P (Suc n)"
@@ -1345,19 +1344,13 @@
   shows "u = s"
   using 2 1 by (rule trans)
 
+setup Arith_Data.setup
+
 use "Tools/nat_arith.ML"
 declaration {* K Nat_Arith.setup *}
 
-ML{*
-structure ArithFacts =
-  NamedThmsFun(val name = "arith"
-               val description = "arith facts - only ground formulas");
-*}
-
-setup ArithFacts.setup
-
 use "Tools/lin_arith.ML"
-declaration {* K LinArith.setup *}
+declaration {* K Lin_Arith.setup *}
 
 lemmas [arith_split] = nat_diff_split split_min split_max
 
--- a/src/HOL/NatBin.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/NatBin.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -651,7 +651,7 @@
 val numeral_ss = @{simpset} addsimps @{thms numerals};
 
 val nat_bin_arith_setup =
- LinArith.map_data
+ Lin_Arith.map_data
    (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
      {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
       inj_thms = inj_thms,
--- a/src/HOL/OrderedGroup.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/OrderedGroup.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -466,7 +466,7 @@
   then show ?thesis by simp
 qed
 
-lemma add_neg_nonpos: 
+lemma add_neg_nonpos:
   assumes "a < 0" and "b \<le> 0" shows "a + b < 0"
 proof -
   have "a + b < 0 + 0"
@@ -494,6 +494,10 @@
   then show ?thesis by simp
 qed
 
+lemmas add_sign_intros =
+  add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg
+  add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos
+
 lemma add_nonneg_eq_0_iff:
   assumes x: "0 \<le> x" and y: "0 \<le> y"
   shows "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
--- a/src/HOL/Presburger.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Presburger.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -439,12 +439,7 @@
 
 use "Tools/Qelim/presburger.ML"
 
-declaration {* fn _ =>
-  arith_tactic_add
-    (mk_arith_tactic "presburger" (fn ctxt => fn i => fn st =>
-       (warning "Trying Presburger arithmetic ...";   
-    Presburger.cooper_tac true [] [] ctxt i st)))
-*}
+setup {* Arith_Data.add_tactic "Presburger arithmetic" (K (Presburger.cooper_tac true [] [])) *}
 
 method_setup presburger = {*
 let
--- a/src/HOL/Ring_and_Field.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Ring_and_Field.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -729,11 +729,15 @@
 subclass pordered_semiring ..
 
 lemma mult_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
-by (drule mult_left_mono [of zero b], auto)
+using mult_left_mono [of zero b a] by simp
 
 lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
-by (drule mult_left_mono [of b zero], auto)
-
+using mult_left_mono [of b zero a] by simp
+
+lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
+using mult_right_mono [of a zero b] by simp
+
+text {* Legacy - use @{text mult_nonpos_nonneg} *}
 lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0" 
 by (drule mult_right_mono [of b zero], auto)
 
@@ -786,31 +790,32 @@
   "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
 by (force simp add: mult_strict_right_mono not_less [symmetric])
 
-lemma mult_pos_pos:
-  "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
-by (drule mult_strict_left_mono [of zero b], auto)
-
-lemma mult_pos_neg:
-  "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
-by (drule mult_strict_left_mono [of b zero], auto)
-
-lemma mult_pos_neg2:
-  "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0" 
+lemma mult_pos_pos: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
+using mult_strict_left_mono [of zero b a] by simp
+
+lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
+using mult_strict_left_mono [of b zero a] by simp
+
+lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
+using mult_strict_right_mono [of a zero b] by simp
+
+text {* Legacy - use @{text mult_neg_pos} *}
+lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0" 
 by (drule mult_strict_right_mono [of b zero], auto)
 
 lemma zero_less_mult_pos:
   "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-apply (cases "b\<le>0") 
+apply (cases "b\<le>0")
  apply (auto simp add: le_less not_less)
-apply (drule_tac mult_pos_neg [of a b]) 
+apply (drule_tac mult_pos_neg [of a b])
  apply (auto dest: less_not_sym)
 done
 
 lemma zero_less_mult_pos2:
   "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-apply (cases "b\<le>0") 
+apply (cases "b\<le>0")
  apply (auto simp add: le_less not_less)
-apply (drule_tac mult_pos_neg2 [of a b]) 
+apply (drule_tac mult_pos_neg2 [of a b])
  apply (auto dest: less_not_sym)
 done
 
@@ -819,9 +824,9 @@
   assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
   shows "a * c < b * d"
   using assms apply (cases "c=0")
-  apply (simp add: mult_pos_pos) 
+  apply (simp add: mult_pos_pos)
   apply (erule mult_strict_right_mono [THEN less_trans])
-  apply (force simp add: le_less) 
+  apply (force simp add: le_less)
   apply (erule mult_strict_left_mono, assumption)
   done
 
@@ -960,9 +965,8 @@
   apply (simp_all add: minus_mult_right [symmetric]) 
   done
 
-lemma mult_nonpos_nonpos:
-  "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
-by (drule mult_right_mono_neg [of a zero b]) auto
+lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
+using mult_right_mono_neg [of a zero b] by simp
 
 lemma split_mult_pos_le:
   "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
@@ -1006,21 +1010,14 @@
 
 subclass ordered_ring ..
 
-lemma mult_strict_left_mono_neg:
-  "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
-  apply (drule mult_strict_left_mono [of _ _ "uminus c"])
-  apply (simp_all add: minus_mult_left [symmetric]) 
-  done
-
-lemma mult_strict_right_mono_neg:
-  "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
-  apply (drule mult_strict_right_mono [of _ _ "uminus c"])
-  apply (simp_all add: minus_mult_right [symmetric]) 
-  done
-
-lemma mult_neg_neg:
-  "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
-by (drule mult_strict_right_mono_neg, auto)
+lemma mult_strict_left_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
+using mult_strict_left_mono [of b a "- c"] by simp
+
+lemma mult_strict_right_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
+using mult_strict_right_mono [of b a "- c"] by simp
+
+lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
+using mult_strict_right_mono_neg [of a zero b] by simp
 
 subclass ring_no_zero_divisors
 proof
@@ -1144,11 +1141,6 @@
   "c < 0 \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> b \<le> a"
 by (auto simp: mult_le_cancel_left)
 
-end
-
-context ordered_ring_strict
-begin
-
 lemma mult_less_cancel_left_pos:
   "0 < c \<Longrightarrow> c * a < c * b \<longleftrightarrow> a < b"
 by (auto simp: mult_less_cancel_left)
@@ -1162,6 +1154,11 @@
 text{*Legacy - use @{text algebra_simps} *}
 lemmas ring_simps[noatp] = algebra_simps
 
+lemmas mult_sign_intros =
+  mult_nonneg_nonneg mult_nonneg_nonpos
+  mult_nonpos_nonneg mult_nonpos_nonpos
+  mult_pos_pos mult_pos_neg
+  mult_neg_pos mult_neg_neg
 
 class pordered_comm_ring = comm_ring + pordered_comm_semiring
 begin
--- a/src/HOL/Tools/Qelim/cooper.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -172,7 +172,7 @@
 
     (* Canonical linear form for terms, formulae etc.. *)
 fun provelin ctxt t = Goal.prove ctxt [] [] t 
-  (fn _ => EVERY [simp_tac lin_ss 1, TRY (simple_arith_tac ctxt 1)]);
+  (fn _ => EVERY [simp_tac lin_ss 1, TRY (linear_arith_tac ctxt 1)]);
 fun linear_cmul 0 tm = zero 
   | linear_cmul n tm = case tm of  
       Const (@{const_name HOL.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
--- a/src/HOL/Tools/Qelim/presburger.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/Qelim/presburger.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -163,8 +163,10 @@
 
 fun cooper_tac elim add_ths del_ths ctxt =
 let val ss = Simplifier.context ctxt (fst (CooperData.get ctxt)) delsimps del_ths addsimps add_ths
+    val aprems = Arith_Data.get_arith_facts ctxt
 in
-  ObjectLogic.full_atomize_tac
+  Method.insert_tac aprems
+  THEN_ALL_NEW ObjectLogic.full_atomize_tac
   THEN_ALL_NEW CONVERSION Thm.eta_long_conversion
   THEN_ALL_NEW simp_tac ss
   THEN_ALL_NEW (TRY o generalize_tac (int_nat_terms ctxt))
--- a/src/HOL/Tools/TFL/post.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/TFL/post.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -55,7 +55,7 @@
   Prim.postprocess strict
    {wf_tac     = REPEAT (ares_tac wfs 1),
     terminator = asm_simp_tac ss 1
-                 THEN TRY (silent_arith_tac (Simplifier.the_context ss) 1 ORELSE
+                 THEN TRY (Arith_Data.arith_tac (Simplifier.the_context ss) 1 ORELSE
                            fast_tac (cs addSDs [@{thm not0_implies_Suc}] addss ss) 1),
     simplifier = Rules.simpl_conv ss []};
 
--- a/src/HOL/Tools/arith_data.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/arith_data.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -6,6 +6,11 @@
 
 signature ARITH_DATA =
 sig
+  val arith_tac: Proof.context -> int -> tactic
+  val verbose_arith_tac: Proof.context -> int -> tactic
+  val add_tactic: string -> (bool -> Proof.context -> int -> tactic) -> theory -> theory
+  val get_arith_facts: Proof.context -> thm list
+
   val prove_conv_nohyps: tactic list -> Proof.context -> term * term -> thm option
   val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option
   val prove_conv2: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm
@@ -14,11 +19,54 @@
   val trans_tac: thm option -> tactic
   val prep_simproc: string * string list * (theory -> simpset -> term -> thm option)
     -> simproc
+
+  val setup: theory -> theory
 end;
 
 structure Arith_Data: ARITH_DATA =
 struct
 
+(* slots for pluging in arithmetic facts and tactics *)
+
+structure Arith_Facts = NamedThmsFun(
+  val name = "arith"
+  val description = "arith facts - only ground formulas"
+);
+
+val get_arith_facts = Arith_Facts.get;
+
+structure Arith_Tactics = TheoryDataFun
+(
+  type T = (serial * (string * (bool -> Proof.context -> int -> tactic))) list;
+  val empty = [];
+  val copy = I;
+  val extend = I;
+  fun merge _ = AList.merge (op =) (K true);
+);
+
+fun add_tactic name tac = Arith_Tactics.map (cons (serial (), (name, tac)));
+
+fun gen_arith_tac verbose ctxt =
+  let
+    val tactics = (Arith_Tactics.get o ProofContext.theory_of) ctxt
+    fun invoke (_, (name, tac)) k st = (if verbose
+      then warning ("Trying " ^ name ^ "...") else ();
+      tac verbose ctxt k st);
+  in FIRST' (map invoke (rev tactics)) end;
+
+val arith_tac = gen_arith_tac false;
+val verbose_arith_tac = gen_arith_tac true;
+
+val arith_method = Args.bang_facts >> (fn prems => fn ctxt =>
+  METHOD (fn facts => HEADGOAL (Method.insert_tac (prems @ get_arith_facts ctxt @ facts)
+    THEN' verbose_arith_tac ctxt)));
+
+val setup = Arith_Facts.setup
+  #> Method.setup @{binding arith} arith_method "various arithmetic decision procedures";
+
+
+(* various auxiliary and legacy *)
+
 fun prove_conv_nohyps tacs ctxt (t, u) =
   if t aconv u then NONE
   else let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))
--- a/src/HOL/Tools/function_package/scnp_reconstruct.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/function_package/scnp_reconstruct.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -197,7 +197,7 @@
               else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1}
           in
             rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
-            THEN (if tag_flag then arith_tac ctxt 1 else all_tac)
+            THEN (if tag_flag then Arith_Data.verbose_arith_tac ctxt 1 else all_tac)
           end
 
         fun steps_tac MAX strict lq lp =
--- a/src/HOL/Tools/int_arith.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/int_arith.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -530,7 +530,7 @@
   :: Int_Numeral_Simprocs.cancel_numerals;
 
 val setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms,
     inj_thms = nat_inj_thms @ inj_thms,
@@ -547,7 +547,7 @@
   "fast_int_arith" 
      ["(m::'a::{ordered_idom,number_ring}) < n",
       "(m::'a::{ordered_idom,number_ring}) <= n",
-      "(m::'a::{ordered_idom,number_ring}) = n"] (K LinArith.lin_arith_simproc);
+      "(m::'a::{ordered_idom,number_ring}) = n"] (K Lin_Arith.lin_arith_simproc);
 
 end;
 
--- a/src/HOL/Tools/int_factor_simprocs.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/int_factor_simprocs.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -232,7 +232,7 @@
       val less = Const(@{const_name HOL.less}, [T,T] ---> HOLogic.boolT);
       val pos = less $ zero $ t and neg = less $ t $ zero
       fun prove p =
-        Option.map Eq_True_elim (LinArith.lin_arith_simproc ss p)
+        Option.map Eq_True_elim (Lin_Arith.lin_arith_simproc ss p)
         handle THM _ => NONE
     in case prove pos of
          SOME th => SOME(th RS pos_th)
--- a/src/HOL/Tools/lin_arith.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/lin_arith.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -6,13 +6,9 @@
 
 signature BASIC_LIN_ARITH =
 sig
-  type arith_tactic
-  val mk_arith_tactic: string -> (Proof.context -> int -> tactic) -> arith_tactic
-  val eq_arith_tactic: arith_tactic * arith_tactic -> bool
   val arith_split_add: attribute
   val arith_discrete: string -> Context.generic -> Context.generic
   val arith_inj_const: string * typ -> Context.generic -> Context.generic
-  val arith_tactic_add: arith_tactic -> Context.generic -> Context.generic
   val fast_arith_split_limit: int Config.T
   val fast_arith_neq_limit: int Config.T
   val lin_arith_pre_tac: Proof.context -> int -> tactic
@@ -21,9 +17,7 @@
   val trace_arith: bool ref
   val lin_arith_simproc: simpset -> term -> thm option
   val fast_nat_arith_simproc: simproc
-  val simple_arith_tac: Proof.context -> int -> tactic
-  val arith_tac: Proof.context -> int -> tactic
-  val silent_arith_tac: Proof.context -> int -> tactic
+  val linear_arith_tac: Proof.context -> int -> tactic
 end;
 
 signature LIN_ARITH =
@@ -39,7 +33,7 @@
   val setup: Context.generic -> Context.generic
 end;
 
-structure LinArith: LIN_ARITH =
+structure Lin_Arith: LIN_ARITH =
 struct
 
 (* Parameters data for general linear arithmetic functor *)
@@ -72,7 +66,7 @@
   let val _ $ t = Thm.prop_of thm
   in t = Const("False",HOLogic.boolT) end;
 
-fun is_nat(t) = fastype_of1 t = HOLogic.natT;
+fun is_nat t = (fastype_of1 t = HOLogic.natT);
 
 fun mk_nat_thm sg t =
   let val ct = cterm_of sg t  and cn = cterm_of sg (Var(("n",0),HOLogic.natT))
@@ -83,49 +77,35 @@
 
 (* arith context data *)
 
-datatype arith_tactic =
-  ArithTactic of {name: string, tactic: Proof.context -> int -> tactic, id: stamp};
-
-fun mk_arith_tactic name tactic = ArithTactic {name = name, tactic = tactic, id = stamp ()};
-
-fun eq_arith_tactic (ArithTactic {id = id1, ...}, ArithTactic {id = id2, ...}) = (id1 = id2);
-
 structure ArithContextData = GenericDataFun
 (
   type T = {splits: thm list,
             inj_consts: (string * typ) list,
-            discrete: string list,
-            tactics: arith_tactic list};
-  val empty = {splits = [], inj_consts = [], discrete = [], tactics = []};
+            discrete: string list};
+  val empty = {splits = [], inj_consts = [], discrete = []};
   val extend = I;
   fun merge _
-   ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, tactics= tactics1},
-    {splits= splits2, inj_consts= inj_consts2, discrete= discrete2, tactics= tactics2}) : T =
+   ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1},
+    {splits= splits2, inj_consts= inj_consts2, discrete= discrete2}) : T =
    {splits = Library.merge Thm.eq_thm_prop (splits1, splits2),
     inj_consts = Library.merge (op =) (inj_consts1, inj_consts2),
-    discrete = Library.merge (op =) (discrete1, discrete2),
-    tactics = Library.merge eq_arith_tactic (tactics1, tactics2)};
+    discrete = Library.merge (op =) (discrete1, discrete2)};
 );
 
 val get_arith_data = ArithContextData.get o Context.Proof;
 
 val arith_split_add = Thm.declaration_attribute (fn thm =>
-  ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
+  ArithContextData.map (fn {splits, inj_consts, discrete} =>
     {splits = update Thm.eq_thm_prop thm splits,
-     inj_consts = inj_consts, discrete = discrete, tactics = tactics}));
-
-fun arith_discrete d = ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
-  {splits = splits, inj_consts = inj_consts,
-   discrete = update (op =) d discrete, tactics = tactics});
+     inj_consts = inj_consts, discrete = discrete}));
 
-fun arith_inj_const c = ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
-  {splits = splits, inj_consts = update (op =) c inj_consts,
-   discrete = discrete, tactics= tactics});
+fun arith_discrete d = ArithContextData.map (fn {splits, inj_consts, discrete} =>
+  {splits = splits, inj_consts = inj_consts,
+   discrete = update (op =) d discrete});
 
-fun arith_tactic_add tac = ArithContextData.map (fn {splits, inj_consts, discrete, tactics} =>
-  {splits = splits, inj_consts = inj_consts, discrete = discrete,
-   tactics = update eq_arith_tactic tac tactics});
-
+fun arith_inj_const c = ArithContextData.map (fn {splits, inj_consts, discrete} =>
+  {splits = splits, inj_consts = update (op =) c inj_consts,
+   discrete = discrete});
 
 val (fast_arith_split_limit, setup1) = Attrib.config_int "fast_arith_split_limit" 9;
 val (fast_arith_neq_limit, setup2) = Attrib.config_int "fast_arith_neq_limit" 9;
@@ -794,7 +774,7 @@
    Most of the work is done by the cancel tactics. *)
 
 val init_arith_data =
- Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, ...} =>
+ map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, ...} =>
    {add_mono_thms = add_mono_thms @
     @{thms add_mono_thms_ordered_semiring} @ @{thms add_mono_thms_ordered_field},
     mult_mono_thms = mult_mono_thms,
@@ -815,7 +795,7 @@
   arith_discrete "nat";
 
 fun add_arith_facts ss =
-  add_prems (ArithFacts.get (MetaSimplifier.the_context ss)) ss;
+  add_prems (Arith_Data.get_arith_facts (MetaSimplifier.the_context ss)) ss;
 
 val lin_arith_simproc = add_arith_facts #> Fast_Arith.lin_arith_simproc;
 
@@ -895,27 +875,16 @@
     (REPEAT_DETERM o split_tac (#splits (get_arith_data ctxt)))
     (fast_ex_arith_tac ctxt ex);
 
-fun more_arith_tacs ctxt =
-  let val tactics = #tactics (get_arith_data ctxt)
-  in FIRST' (map (fn ArithTactic {tactic, ...} => tactic ctxt) tactics) end;
-
 in
 
-fun simple_arith_tac ctxt = FIRST' [fast_arith_tac ctxt,
-  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt true];
-
-fun arith_tac ctxt = FIRST' [fast_arith_tac ctxt,
-  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt true,
-  more_arith_tacs ctxt];
+fun gen_linear_arith_tac ex ctxt = FIRST' [fast_arith_tac ctxt,
+  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt ex];
 
-fun silent_arith_tac ctxt = FIRST' [fast_arith_tac ctxt,
-  ObjectLogic.full_atomize_tac THEN' (REPEAT_DETERM o rtac impI) THEN' raw_arith_tac ctxt false,
-  more_arith_tacs ctxt];
+val linear_arith_tac = gen_linear_arith_tac true;
 
-val arith_method = Args.bang_facts >>
-  (fn prems => fn ctxt => METHOD (fn facts =>
-      HEADGOAL (Method.insert_tac (prems @ ArithFacts.get ctxt @ facts)
-      THEN' arith_tac ctxt)));
+val linarith_method = Args.bang_facts >> (fn prems => fn ctxt =>
+  METHOD (fn facts => HEADGOAL (Method.insert_tac (prems @ Arith_Data.get_arith_facts ctxt @ facts)
+    THEN' linear_arith_tac ctxt)));
 
 end;
 
@@ -929,11 +898,12 @@
       (add_arith_facts #> Fast_Arith.cut_lin_arith_tac))) #>
   Context.mapping
    (setup_options #>
-    Method.setup @{binding arith} arith_method "decide linear arithmetic" #>
+    Arith_Data.add_tactic "linear arithmetic" gen_linear_arith_tac #>
+    Method.setup @{binding linarith} linarith_method "linear arithmetic" #>
     Attrib.setup @{binding arith_split} (Scan.succeed arith_split_add)
       "declaration of split rules for arithmetic procedure") I;
 
 end;
 
-structure BasicLinArith: BASIC_LIN_ARITH = LinArith;
-open BasicLinArith;
+structure Basic_Lin_Arith: BASIC_LIN_ARITH = Lin_Arith;
+open Basic_Lin_Arith;
--- a/src/HOL/Tools/nat_simprocs.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/nat_simprocs.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -565,7 +565,7 @@
 in
 
 val nat_simprocs_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
     inj_thms = inj_thms, lessD = lessD, neqE = neqE,
     simpset = simpset addsimps add_rules
--- a/src/HOL/Tools/rat_arith.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/rat_arith.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -35,7 +35,7 @@
 in
 
 val rat_arith_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = mult_mono_thms,
     inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms,
--- a/src/HOL/Tools/real_arith.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Tools/real_arith.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -29,7 +29,7 @@
 in
 
 val real_arith_setup =
-  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
    {add_mono_thms = add_mono_thms,
     mult_mono_thms = mult_mono_thms,
     inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms,
--- a/src/HOL/Word/WordArith.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/Word/WordArith.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -512,7 +512,7 @@
 
 fun uint_arith_tacs ctxt = 
   let
-    fun arith_tac' n t = arith_tac ctxt n t handle COOPER => Seq.empty;
+    fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty;
     val cs = local_claset_of ctxt;
     val ss = local_simpset_of ctxt;
   in 
@@ -1075,7 +1075,7 @@
 
 fun unat_arith_tacs ctxt =   
   let
-    fun arith_tac' n t = arith_tac ctxt n t handle COOPER => Seq.empty;
+    fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty;
     val cs = local_claset_of ctxt;
     val ss = local_simpset_of ctxt;
   in 
--- a/src/HOL/ex/Arith_Examples.thy	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/ex/Arith_Examples.thy	Thu Mar 26 13:02:12 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:  HOL/ex/Arith_Examples.thy
-    ID:     $Id$
     Author: Tjark Weber
 *)
 
@@ -14,13 +13,13 @@
 
   @{ML fast_arith_tac} is a very basic version of the tactic.  It performs no
   meta-to-object-logic conversion, and only some splitting of operators.
-  @{ML simple_arith_tac} performs meta-to-object-logic conversion, full
+  @{ML linear_arith_tac} performs meta-to-object-logic conversion, full
   splitting of operators, and NNF normalization of the goal.  The @{text arith}
   method combines them both, and tries other methods (e.g.~@{text presburger})
   as well.  This is the one that you should use in your proofs!
 
   An @{text arith}-based simproc is available as well (see @{ML
-  LinArith.lin_arith_simproc}), which---for performance
+  Lin_Arith.lin_arith_simproc}), which---for performance
   reasons---however does even less splitting than @{ML fast_arith_tac}
   at the moment (namely inequalities only).  (On the other hand, it
   does take apart conjunctions, which @{ML fast_arith_tac} currently
@@ -83,7 +82,7 @@
   by (tactic {* fast_arith_tac @{context} 1 *})
 
 lemma "!!x. ((x::nat) <= y) = (x - y = 0)"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) < y; d < 1 |] ==> x - y = d"
   by (tactic {* fast_arith_tac @{context} 1 *})
@@ -140,34 +139,34 @@
 subsection {* Meta-Logic *}
 
 lemma "x < Suc y == x <= y"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "((x::nat) == z ==> x ~= y) ==> x ~= y | z ~= y"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 
 subsection {* Various Other Examples *}
 
 lemma "(x < Suc y) = (x <= y)"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) < y; y < z |] ==> x < z"
   by (tactic {* fast_arith_tac @{context} 1 *})
 
 lemma "(x::nat) < y & y < z ==> x < z"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 text {* This example involves no arithmetic at all, but is solved by
   preprocessing (i.e. NNF normalization) alone. *}
 
 lemma "(P::bool) = Q ==> Q = P"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> min (x::nat) y = 0"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> max (x::nat) y = x + y"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) ~= y; a + 2 = b; a < y; y < b; a < x; x < b |] ==> False"
   by (tactic {* fast_arith_tac @{context} 1 *})
@@ -185,7 +184,7 @@
   by (tactic {* fast_arith_tac @{context} 1 *})
 
 lemma "[| (x::nat) < y; P (x - y) |] ==> P 0"
-  by (tactic {* simple_arith_tac @{context} 1 *})
+  by (tactic {* linear_arith_tac @{context} 1 *})
 
 lemma "(x - y) - (x::nat) = (x - x) - y"
   by (tactic {* fast_arith_tac @{context} 1 *})
@@ -207,7 +206,7 @@
 (*        preprocessing negates the goal and tries to compute its negation *)
 (*        normal form, which creates lots of separate cases for this       *)
 (*        disjunction of conjunctions                                      *)
-(* by (tactic {* simple_arith_tac 1 *}) *)
+(* by (tactic {* linear_arith_tac 1 *}) *)
 oops
 
 lemma "2 * (x::nat) ~= 1"
--- a/src/HOL/ex/ImperativeQuicksort.thy	Thu Mar 26 13:01:09 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,637 +0,0 @@
-theory ImperativeQuicksort
-imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat
-begin
-
-text {* We prove QuickSort correct in the Relational Calculus. *}
-
-definition swap :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
-where
-  "swap arr i j = (
-     do
-       x \<leftarrow> nth arr i;
-       y \<leftarrow> nth arr j;
-       upd i y arr;
-       upd j x arr;
-       return ()
-     done)"
-
-lemma swap_permutes:
-  assumes "crel (swap a i j) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-  using assms
-  unfolding swap_def
-  by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
-
-function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
-where
-  "part1 a left right p = (
-     if (right \<le> left) then return right
-     else (do
-       v \<leftarrow> nth a left;
-       (if (v \<le> p) then (part1 a (left + 1) right p)
-                    else (do swap a left right;
-  part1 a left (right - 1) p done))
-     done))"
-by pat_completeness auto
-
-termination
-by (relation "measure (\<lambda>(_,l,r,_). r - l )") auto
-
-declare part1.simps[simp del]
-
-lemma part_permutes:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-  using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  thus ?case
-    unfolding part1.simps [of a l r p]
-    by (elim crelE crel_if crel_return crel_nth) (auto simp add: swap_permutes)
-qed
-
-lemma part_returns_index_in_bounds:
-  assumes "crel (part1 a l r p) h h' rs"
-  assumes "l \<le> r"
-  shows "l \<le> rs \<and> rs \<le> r"
-using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr `l \<le> r` show ?thesis
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-  next
-    case False (* recursive case *)
-    note rec_condition = this
-    let ?v = "get_array a h ! l"
-    show ?thesis
-    proof (cases "?v \<le> p")
-      case True
-      with cr False
-      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from rec_condition have "l + 1 \<le> r" by arith
-      from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`]
-      show ?thesis by simp
-    next
-      case False
-      with rec_condition cr
-      obtain h1 where swp: "crel (swap a l r) h h1 ()"
-        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from rec_condition have "l \<le> r - 1" by arith
-      from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastsimp
-    qed
-  qed
-qed
-
-lemma part_length_remains:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
-using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr show ?thesis
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-  next
-    case False (* recursive case *)
-    with cr 1 show ?thesis
-      unfolding part1.simps [of a l r p] swap_def
-      by (auto elim!: crelE crel_if crel_nth crel_return crel_upd) fastsimp
-  qed
-qed
-
-lemma part_outer_remains:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
-  using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr show ?thesis
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-  next
-    case False (* recursive case *)
-    note rec_condition = this
-    let ?v = "get_array a h ! l"
-    show ?thesis
-    proof (cases "?v \<le> p")
-      case True
-      with cr False
-      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from 1(1)[OF rec_condition True rec1]
-      show ?thesis by fastsimp
-    next
-      case False
-      with rec_condition cr
-      obtain h1 where swp: "crel (swap a l r) h h1 ()"
-        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from swp rec_condition have
-        "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array a h ! i = get_array a h1 ! i"
-	unfolding swap_def
-	by (elim crelE crel_nth crel_upd crel_return) auto
-      with 1(2) [OF rec_condition False rec2] show ?thesis by fastsimp
-    qed
-  qed
-qed
-
-
-lemma part_partitions:
-  assumes "crel (part1 a l r p) h h' rs"
-  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> p)
-  \<and> (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! i \<ge> p)"
-  using assms
-proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
-  case (1 a l r p h h' rs)
-  note cr = `crel (part1 a l r p) h h' rs`
-  
-  show ?case
-  proof (cases "r \<le> l")
-    case True (* Terminating case *)
-    with cr have "rs = r"
-      unfolding part1.simps[of a l r p]
-      by (elim crelE crel_if crel_return crel_nth) auto
-    with True
-    show ?thesis by auto
-  next
-    case False (* recursive case *)
-    note lr = this
-    let ?v = "get_array a h ! l"
-    show ?thesis
-    proof (cases "?v \<le> p")
-      case True
-      with lr cr
-      have rec1: "crel (part1 a (l + 1) r p) h h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from True part_outer_remains[OF rec1] have a_l: "get_array a h' ! l \<le> p"
-	by fastsimp
-      have "\<forall>i. (l \<le> i = (l = i \<or> Suc l \<le> i))" by arith
-      with 1(1)[OF False True rec1] a_l show ?thesis
-	by auto
-    next
-      case False
-      with lr cr
-      obtain h1 where swp: "crel (swap a l r) h h1 ()"
-        and rec2: "crel (part1 a l (r - 1) p) h1 h' rs"
-        unfolding part1.simps[of a l r p]
-        by (elim crelE crel_nth crel_if crel_return) auto
-      from swp False have "get_array a h1 ! r \<ge> p"
-	unfolding swap_def
-	by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
-      with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
-	by fastsimp
-      have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
-      with 1(2)[OF lr False rec2] a_r show ?thesis
-	by auto
-    qed
-  qed
-qed
-
-
-fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
-where
-  "partition a left right = (do
-     pivot \<leftarrow> nth a right;
-     middle \<leftarrow> part1 a left (right - 1) pivot;
-     v \<leftarrow> nth a middle;
-     m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle);
-     swap a m right;
-     return m
-   done)"
-
-declare partition.simps[simp del]
-
-lemma partition_permutes:
-  assumes "crel (partition a l r) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-proof -
-    from assms part_permutes swap_permutes show ?thesis
-      unfolding partition.simps
-      by (elim crelE crel_return crel_nth crel_if crel_upd) auto
-qed
-
-lemma partition_length_remains:
-  assumes "crel (partition a l r) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
-proof -
-  from assms part_length_remains show ?thesis
-    unfolding partition.simps swap_def
-    by (elim crelE crel_return crel_nth crel_if crel_upd) auto
-qed
-
-lemma partition_outer_remains:
-  assumes "crel (partition a l r) h h' rs"
-  assumes "l < r"
-  shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
-proof -
-  from assms part_outer_remains part_returns_index_in_bounds show ?thesis
-    unfolding partition.simps swap_def
-    by (elim crelE crel_return crel_nth crel_if crel_upd) fastsimp
-qed
-
-lemma partition_returns_index_in_bounds:
-  assumes crel: "crel (partition a l r) h h' rs"
-  assumes "l < r"
-  shows "l \<le> rs \<and> rs \<le> r"
-proof -
-  from crel obtain middle h'' p where part: "crel (part1 a l (r - 1) p) h h'' middle"
-    and rs_equals: "rs = (if get_array a h'' ! middle \<le> get_array a h ! r then middle + 1
-         else middle)"
-    unfolding partition.simps
-    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
-  from `l < r` have "l \<le> r - 1" by arith
-  from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto
-qed
-
-lemma partition_partitions:
-  assumes crel: "crel (partition a l r) h h' rs"
-  assumes "l < r"
-  shows "(\<forall>i. l \<le> i \<and> i < rs \<longrightarrow> get_array (a::nat array) h' ! i \<le> get_array a h' ! rs) \<and>
-  (\<forall>i. rs < i \<and> i \<le> r \<longrightarrow> get_array a h' ! rs \<le> get_array a h' ! i)"
-proof -
-  let ?pivot = "get_array a h ! r" 
-  from crel obtain middle h1 where part: "crel (part1 a l (r - 1) ?pivot) h h1 middle"
-    and swap: "crel (swap a rs r) h1 h' ()"
-    and rs_equals: "rs = (if get_array a h1 ! middle \<le> ?pivot then middle + 1
-         else middle)"
-    unfolding partition.simps
-    by (elim crelE crel_return crel_nth crel_if crel_upd) simp
-  from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
-    (Heap.upd a rs (get_array a h1 ! r) h1)"
-    unfolding swap_def
-    by (elim crelE crel_return crel_nth crel_upd) simp
-  from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
-    unfolding swap_def
-    by (elim crelE crel_return crel_nth crel_upd) simp
-  from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
-    unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
-  from `l < r` have "l \<le> r - 1" by simp 
-  note middle_in_bounds = part_returns_index_in_bounds[OF part this]
-  from part_outer_remains[OF part] `l < r`
-  have "get_array a h ! r = get_array a h1 ! r"
-    by fastsimp
-  with swap
-  have right_remains: "get_array a h ! r = get_array a h' ! rs"
-    unfolding swap_def
-    by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
-  from part_partitions [OF part]
-  show ?thesis
-  proof (cases "get_array a h1 ! middle \<le> ?pivot")
-    case True
-    with rs_equals have rs_equals: "rs = middle + 1" by simp
-    { 
-      fix i
-      assume i_is_left: "l \<le> i \<and> i < rs"
-      with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
-      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-      from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
-      with part_partitions[OF part] right_remains True
-      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
-      with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
-	unfolding Heap.upd_def Heap.length_def by simp
-    }
-    moreover
-    {
-      fix i
-      assume "rs < i \<and> i \<le> r"
-
-      hence "(rs < i \<and> i \<le> r - 1) \<or> (rs < i \<and> i = r)" by arith
-      hence "get_array a h' ! rs \<le> get_array a h' ! i"
-      proof
-	assume i_is: "rs < i \<and> i \<le> r - 1"
-	with swap_length_remains in_bounds middle_in_bounds rs_equals
-	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-	from part_partitions[OF part] rs_equals right_remains i_is
-	have "get_array a h' ! rs \<le> get_array a h1 ! i"
-	  by fastsimp
-	with i_props h'_def show ?thesis by fastsimp
-      next
-	assume i_is: "rs < i \<and> i = r"
-	with rs_equals have "Suc middle \<noteq> r" by arith
-	with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
-	with part_partitions[OF part] right_remains 
-	have "get_array a h' ! rs \<le> get_array a h1 ! (Suc middle)"
-	  by fastsimp
-	with i_is True rs_equals right_remains h'_def
-	show ?thesis using in_bounds
-	  unfolding Heap.upd_def Heap.length_def
-	  by auto
-      qed
-    }
-    ultimately show ?thesis by auto
-  next
-    case False
-    with rs_equals have rs_equals: "middle = rs" by simp
-    { 
-      fix i
-      assume i_is_left: "l \<le> i \<and> i < rs"
-      with swap_length_remains in_bounds middle_in_bounds rs_equals
-      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-      from part_partitions[OF part] rs_equals right_remains i_is_left
-      have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
-      with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
-	unfolding Heap.upd_def by simp
-    }
-    moreover
-    {
-      fix i
-      assume "rs < i \<and> i \<le> r"
-      hence "(rs < i \<and> i \<le> r - 1) \<or> i = r" by arith
-      hence "get_array a h' ! rs \<le> get_array a h' ! i"
-      proof
-	assume i_is: "rs < i \<and> i \<le> r - 1"
-	with swap_length_remains in_bounds middle_in_bounds rs_equals
-	have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
-	from part_partitions[OF part] rs_equals right_remains i_is
-	have "get_array a h' ! rs \<le> get_array a h1 ! i"
-	  by fastsimp
-	with i_props h'_def show ?thesis by fastsimp
-      next
-	assume i_is: "i = r"
-	from i_is False rs_equals right_remains h'_def
-	show ?thesis using in_bounds
-	  unfolding Heap.upd_def Heap.length_def
-	  by auto
-      qed
-    }
-    ultimately
-    show ?thesis by auto
-  qed
-qed
-
-
-function quicksort :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
-where
-  "quicksort arr left right =
-     (if (right > left)  then
-        do
-          pivotNewIndex \<leftarrow> partition arr left right;
-          pivotNewIndex \<leftarrow> assert (\<lambda>x. left \<le> x \<and> x \<le> right) pivotNewIndex;
-          quicksort arr left (pivotNewIndex - 1);
-          quicksort arr (pivotNewIndex + 1) right
-        done
-     else return ())"
-by pat_completeness auto
-
-(* For termination, we must show that the pivotNewIndex is between left and right *) 
-termination
-by (relation "measure (\<lambda>(a, l, r). (r - l))") auto
-
-declare quicksort.simps[simp del]
-
-
-lemma quicksort_permutes:
-  assumes "crel (quicksort a l r) h h' rs"
-  shows "multiset_of (get_array a h') 
-  = multiset_of (get_array a h)"
-  using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  with partition_permutes show ?case
-    unfolding quicksort.simps [of a l r]
-    by (elim crel_if crelE crel_assert crel_return) auto
-qed
-
-lemma length_remains:
-  assumes "crel (quicksort a l r) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
-using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  with partition_length_remains show ?case
-    unfolding quicksort.simps [of a l r]
-    by (elim crel_if crelE crel_assert crel_return) auto
-qed
-
-lemma quicksort_outer_remains:
-  assumes "crel (quicksort a l r) h h' rs"
-   shows "\<forall>i. i < l \<or> r < i \<longrightarrow> get_array (a::nat array) h ! i = get_array a h' ! i"
-  using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  note cr = `crel (quicksort a l r) h h' rs`
-  thus ?case
-  proof (cases "r > l")
-    case False
-    with cr have "h' = h"
-      unfolding quicksort.simps [of a l r]
-      by (elim crel_if crel_return) auto
-    thus ?thesis by simp
-  next
-  case True
-   { 
-      fix h1 h2 p ret1 ret2 i
-      assume part: "crel (partition a l r) h h1 p"
-      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ret1"
-      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ret2"
-      assume pivot: "l \<le> p \<and> p \<le> r"
-      assume i_outer: "i < l \<or> r < i"
-      from  partition_outer_remains [OF part True] i_outer
-      have "get_array a h !i = get_array a h1 ! i" by fastsimp
-      moreover
-      with 1(1) [OF True pivot qs1] pivot i_outer
-      have "get_array a h1 ! i = get_array a h2 ! i" by auto
-      moreover
-      with qs2 1(2) [of p h2 h' ret2] True pivot i_outer
-      have "get_array a h2 ! i = get_array a h' ! i" by auto
-      ultimately have "get_array a h ! i= get_array a h' ! i" by simp
-    }
-    with cr show ?thesis
-      unfolding quicksort.simps [of a l r]
-      by (elim crel_if crelE crel_assert crel_return) auto
-  qed
-qed
-
-lemma quicksort_is_skip:
-  assumes "crel (quicksort a l r) h h' rs"
-  shows "r \<le> l \<longrightarrow> h = h'"
-  using assms
-  unfolding quicksort.simps [of a l r]
-  by (elim crel_if crel_return) auto
- 
-lemma quicksort_sorts:
-  assumes "crel (quicksort a l r) h h' rs"
-  assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" 
-  shows "sorted (subarray l (r + 1) a h')"
-  using assms
-proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
-  case (1 a l r h h' rs)
-  note cr = `crel (quicksort a l r) h h' rs`
-  thus ?case
-  proof (cases "r > l")
-    case False
-    hence "l \<ge> r + 1 \<or> l = r" by arith 
-    with length_remains[OF cr] 1(5) show ?thesis
-      by (auto simp add: subarray_Nil subarray_single)
-  next
-    case True
-    { 
-      fix h1 h2 p
-      assume part: "crel (partition a l r) h h1 p"
-      assume qs1: "crel (quicksort a l (p - 1)) h1 h2 ()"
-      assume qs2: "crel (quicksort a (p + 1) r) h2 h' ()"
-      from partition_returns_index_in_bounds [OF part True]
-      have pivot: "l\<le> p \<and> p \<le> r" .
-     note length_remains = length_remains[OF qs2] length_remains[OF qs1] partition_length_remains[OF part]
-      from quicksort_outer_remains [OF qs2] quicksort_outer_remains [OF qs1] pivot quicksort_is_skip[OF qs1]
-      have pivot_unchanged: "get_array a h1 ! p = get_array a h' ! p" by (cases p, auto)
-        (*-- First of all, by induction hypothesis both sublists are sorted. *)
-      from 1(1)[OF True pivot qs1] length_remains pivot 1(5) 
-      have IH1: "sorted (subarray l p a h2)"  by (cases p, auto simp add: subarray_Nil)
-      from quicksort_outer_remains [OF qs2] length_remains
-      have left_subarray_remains: "subarray l p a h2 = subarray l p a h'"
-	by (simp add: subarray_eq_samelength_iff)
-      with IH1 have IH1': "sorted (subarray l p a h')" by simp
-      from 1(2)[OF True pivot qs2] pivot 1(5) length_remains
-      have IH2: "sorted (subarray (p + 1) (r + 1) a h')"
-        by (cases "Suc p \<le> r", auto simp add: subarray_Nil)
-           (* -- Secondly, both sublists remain partitioned. *)
-      from partition_partitions[OF part True]
-      have part_conds1: "\<forall>j. j \<in> set (subarray l p a h1) \<longrightarrow> j \<le> get_array a h1 ! p "
-        and part_conds2: "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h1) \<longrightarrow> get_array a h1 ! p \<le> j"
-        by (auto simp add: all_in_set_subarray_conv)
-      from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
-        length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
-      have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
-	unfolding Heap.length_def subarray_def by (cases p, auto)
-      with left_subarray_remains part_conds1 pivot_unchanged
-      have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
-        by (simp, subst set_of_multiset_of[symmetric], simp)
-          (* -- These steps are the analogous for the right sublist \<dots> *)
-      from quicksort_outer_remains [OF qs1] length_remains
-      have right_subarray_remains: "subarray (p + 1) (r + 1) a h1 = subarray (p + 1) (r + 1) a h2"
-	by (auto simp add: subarray_eq_samelength_iff)
-      from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
-        length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
-      have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
-        unfolding Heap.length_def subarray_def by auto
-      with right_subarray_remains part_conds2 pivot_unchanged
-      have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
-        by (simp, subst set_of_multiset_of[symmetric], simp)
-          (* -- Thirdly and finally, we show that the array is sorted
-          following from the facts above. *)
-      from True pivot 1(5) length_remains have "subarray l (r + 1) a h' = subarray l p a h' @ [get_array a h' ! p] @ subarray (p + 1) (r + 1) a h'"
-	by (simp add: subarray_nth_array_Cons, cases "l < p") (auto simp add: subarray_append subarray_Nil)
-      with IH1' IH2 part_conds1' part_conds2' pivot have ?thesis
-	unfolding subarray_def
-	apply (auto simp add: sorted_append sorted_Cons all_in_set_sublist'_conv)
-	by (auto simp add: set_sublist' dest: le_trans [of _ "get_array a h' ! p"])
-    }
-    with True cr show ?thesis
-      unfolding quicksort.simps [of a l r]
-      by (elim crel_if crel_return crelE crel_assert) auto
-  qed
-qed
-
-
-lemma quicksort_is_sort:
-  assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
-  shows "get_array a h' = sort (get_array a h)"
-proof (cases "get_array a h = []")
-  case True
-  with quicksort_is_skip[OF crel] show ?thesis
-  unfolding Heap.length_def by simp
-next
-  case False
-  from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
-    unfolding Heap.length_def subarray_def by auto
-  with length_remains[OF crel] have "sorted (get_array a h')"
-    unfolding Heap.length_def by simp
-  with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
-qed
-
-subsection {* No Errors in quicksort *}
-text {* We have proved that quicksort sorts (if no exceptions occur).
-We will now show that exceptions do not occur. *}
-
-lemma noError_part1: 
-  assumes "l < Heap.length a h" "r < Heap.length a h"
-  shows "noError (part1 a l r p) h"
-  using assms
-proof (induct a l r p arbitrary: h rule: part1.induct)
-  case (1 a l r p)
-  thus ?case
-    unfolding part1.simps [of a l r] swap_def
-    by (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd elim!: crelE crel_upd crel_nth crel_return)
-qed
-
-lemma noError_partition:
-  assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
-  shows "noError (partition a l r) h"
-using assms
-unfolding partition.simps swap_def
-apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_part1 elim!: crelE crel_upd crel_nth crel_return)
-apply (frule part_length_remains)
-apply (frule part_returns_index_in_bounds)
-apply auto
-apply (frule part_length_remains)
-apply (frule part_returns_index_in_bounds)
-apply auto
-apply (frule part_length_remains)
-apply auto
-done
-
-lemma noError_quicksort:
-  assumes "l < Heap.length a h" "r < Heap.length a h"
-  shows "noError (quicksort a l r) h"
-using assms
-proof (induct a l r arbitrary: h rule: quicksort.induct)
-  case (1 a l ri h)
-  thus ?case
-    unfolding quicksort.simps [of a l ri]
-    apply (auto intro!: noError_if noErrorI noError_return noError_nth noError_upd noError_assert noError_partition)
-    apply (frule partition_returns_index_in_bounds)
-    apply auto
-    apply (frule partition_returns_index_in_bounds)
-    apply auto
-    apply (auto elim!: crel_assert dest!: partition_length_remains length_remains)
-    apply (subgoal_tac "Suc r \<le> ri \<or> r = ri") 
-    apply (erule disjE)
-    apply auto
-    unfolding quicksort.simps [of a "Suc ri" ri]
-    apply (auto intro!: noError_if noError_return)
-    done
-qed
-
-
-subsection {* Example *}
-
-definition "qsort a = do
-    k \<leftarrow> length a;
-    quicksort a 0 (k - 1);
-    return a
-  done"
-
-ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
-
-export_code qsort in SML_imp module_name QSort
-export_code qsort in OCaml module_name QSort file -
-export_code qsort in OCaml_imp module_name QSort file -
-export_code qsort in Haskell module_name QSort file -
-
-end
\ No newline at end of file
--- a/src/HOL/ex/ROOT.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/HOL/ex/ROOT.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -21,7 +21,6 @@
 
 use_thys [
   "Numeral",
-  "ImperativeQuicksort",
   "Higher_Order_Logic",
   "Abstract_NAT",
   "Guess",
--- a/src/HOL/ex/Subarray.thy	Thu Mar 26 13:01:09 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,66 +0,0 @@
-theory Subarray
-imports Array Sublist
-begin
-
-definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
-where
-  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
-
-lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
-apply (simp add: sublist'_update1)
-done
-
-lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
-apply (subst sublist'_update2)
-apply fastsimp
-apply simp
-done
-
-lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
-unfolding subarray_def Heap.upd_def
-by (simp add: sublist'_update3)
-
-lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
-by (simp add: subarray_def sublist'_Nil')
-
-lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
-by (simp add: subarray_def Heap.length_def sublist'_single)
-
-lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
-by (simp add: subarray_def Heap.length_def length_sublist')
-
-lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
-by (simp add: length_subarray)
-
-lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_front)
-
-lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_back)
-
-lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
-unfolding subarray_def
-by (simp add: sublist'_append)
-
-lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
-unfolding Heap.length_def subarray_def
-by (simp add: sublist'_all)
-
-lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
-unfolding Heap.length_def subarray_def
-by (simp add: nth_sublist')
-
-lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
-unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
-
-lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
-
-lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
-
-end
\ No newline at end of file
--- a/src/HOL/ex/Sublist.thy	Thu Mar 26 13:01:09 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,505 +0,0 @@
-(* $Id$ *)
-
-header {* Slices of lists *}
-
-theory Sublist
-imports Multiset
-begin
-
-
-lemma sublist_split: "i \<le> j \<and> j \<le> k \<Longrightarrow> sublist xs {i..<j} @ sublist xs {j..<k} = sublist xs {i..<k}" 
-apply (induct xs arbitrary: i j k)
-apply simp
-apply (simp only: sublist_Cons)
-apply simp
-apply safe
-apply simp
-apply (erule_tac x="0" in meta_allE)
-apply (erule_tac x="j - 1" in meta_allE)
-apply (erule_tac x="k - 1" in meta_allE)
-apply (subgoal_tac "0 \<le> j - 1 \<and> j - 1 \<le> k - 1")
-apply simp
-apply (subgoal_tac "{ja. Suc ja < j} = {0..<j - Suc 0}")
-apply (subgoal_tac "{ja. j \<le> Suc ja \<and> Suc ja < k} = {j - Suc 0..<k - Suc 0}")
-apply (subgoal_tac "{j. Suc j < k} = {0..<k - Suc 0}")
-apply simp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply (erule_tac x="i - 1" in meta_allE)
-apply (erule_tac x="j - 1" in meta_allE)
-apply (erule_tac x="k - 1" in meta_allE)
-apply (subgoal_tac " {ja. i \<le> Suc ja \<and> Suc ja < j} = {i - 1 ..<j - 1}")
-apply (subgoal_tac " {ja. j \<le> Suc ja \<and> Suc ja < k} = {j - 1..<k - 1}")
-apply (subgoal_tac "{j. i \<le> Suc j \<and> Suc j < k} = {i - 1..<k - 1}")
-apply (subgoal_tac " i - 1 \<le> j - 1 \<and> j - 1 \<le> k - 1")
-apply simp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-apply fastsimp
-done
-
-lemma sublist_update1: "i \<notin> inds \<Longrightarrow> sublist (xs[i := v]) inds = sublist xs inds"
-apply (induct xs arbitrary: i inds)
-apply simp
-apply (case_tac i)
-apply (simp add: sublist_Cons)
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_update2: "i \<in> inds \<Longrightarrow> sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \<in> inds. k < i}):= v]"
-proof (induct xs arbitrary: i inds)
-  case Nil thus ?case by simp
-next
-  case (Cons x xs)
-  thus ?case
-  proof (cases i)
-    case 0 with Cons show ?thesis by (simp add: sublist_Cons)
-  next
-    case (Suc i')
-    with Cons show ?thesis
-      apply simp
-      apply (simp add: sublist_Cons)
-      apply auto
-      apply (auto simp add: nat.split)
-      apply (simp add: card_less_Suc[symmetric])
-      apply (simp add: card_less_Suc2)
-      done
-  qed
-qed
-
-lemma sublist_update: "sublist (xs[i := v]) inds = (if i \<in> inds then (sublist xs inds)[(card {k \<in> inds. k < i}) := v] else sublist xs inds)"
-by (simp add: sublist_update1 sublist_update2)
-
-lemma sublist_take: "sublist xs {j. j < m} = take m xs"
-apply (induct xs arbitrary: m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_take': "sublist xs {0..<m} = take m xs"
-apply (induct xs arbitrary: m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons sublist_take)
-done
-
-lemma sublist_all[simp]: "sublist xs {j. j < length xs} = xs"
-apply (induct xs)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_all'[simp]: "sublist xs {0..<length xs} = xs"
-apply (induct xs)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_single: "a < length xs \<Longrightarrow> sublist xs {a} = [xs ! a]"
-apply (induct xs arbitrary: a)
-apply simp
-apply(case_tac aa)
-apply simp
-apply (simp add: sublist_Cons)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-lemma sublist_is_Nil: "\<forall>i \<in> inds. i \<ge> length xs \<Longrightarrow> sublist xs inds = []" 
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply auto
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply auto
-done
-
-lemma sublist_Nil': "sublist xs inds = [] \<Longrightarrow> \<forall>i \<in> inds. i \<ge> length xs"
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply (auto split: if_splits)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (case_tac x, auto)
-done
-
-lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\<forall>i \<in> inds. i \<ge> length xs)"
-apply (induct xs arbitrary: inds)
-apply simp
-apply (simp add: sublist_Cons)
-apply auto
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (case_tac x, auto)
-done
-
-lemma sublist_eq_subseteq: " \<lbrakk> inds' \<subseteq> inds; sublist xs inds = sublist ys inds \<rbrakk> \<Longrightarrow> sublist xs inds' = sublist ys inds'"
-apply (induct xs arbitrary: ys inds inds')
-apply simp
-apply (drule sym, rule sym)
-apply (simp add: sublist_Nil, fastsimp)
-apply (case_tac ys)
-apply (simp add: sublist_Nil, fastsimp)
-apply (auto simp add: sublist_Cons)
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
-apply fastsimp
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds'}" in meta_allE)
-apply fastsimp
-done
-
-lemma sublist_eq: "\<lbrakk> \<forall>i \<in> inds. ((i < length xs) \<and> (i < length ys)) \<or> ((i \<ge> length xs ) \<and> (i \<ge> length ys)); \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
-apply (induct xs arbitrary: ys inds)
-apply simp
-apply (rule sym, simp add: sublist_Nil)
-apply (case_tac ys)
-apply (simp add: sublist_Nil)
-apply (auto simp add: sublist_Cons)
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply fastsimp
-apply (erule_tac x="list" in meta_allE)
-apply (erule_tac x="{j. Suc j \<in> inds}" in meta_allE)
-apply fastsimp
-done
-
-lemma sublist_eq_samelength: "\<lbrakk> length xs = length ys; \<forall>i \<in> inds. xs ! i = ys ! i \<rbrakk> \<Longrightarrow> sublist xs inds = sublist ys inds"
-by (rule sublist_eq, auto)
-
-lemma sublist_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist xs inds = sublist ys inds) = (\<forall>i \<in> inds. xs ! i = ys ! i)"
-apply (induct xs arbitrary: ys inds)
-apply simp
-apply (rule sym, simp add: sublist_Nil)
-apply (case_tac ys)
-apply (simp add: sublist_Nil)
-apply (auto simp add: sublist_Cons)
-apply (case_tac i)
-apply auto
-apply (case_tac i)
-apply auto
-done
-
-section {* Another sublist function *}
-
-function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "sublist' n m [] = []"
-| "sublist' n 0 xs = []"
-| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)"
-| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs"
-by pat_completeness auto
-termination by lexicographic_order
-
-subsection {* Proving equivalence to the other sublist command *}
-
-lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac n)
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons)
-done
-
-
-lemma "sublist' n m xs = sublist xs {n..<m}"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac n, case_tac m)
-apply simp
-apply simp
-apply (simp add: sublist_take')
-apply (case_tac m)
-apply simp
-apply (simp add: sublist_Cons sublist'_sublist)
-done
-
-
-subsection {* Showing equivalence to use of drop and take for definition *}
-
-lemma "sublist' n m xs = take (m - n) (drop n xs)"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-subsection {* General lemma about sublist *}
-
-lemma sublist'_Nil[simp]: "sublist' i j [] = []"
-by simp
-
-lemma sublist'_Cons[simp]: "sublist' i (Suc j) (x#xs) = (case i of 0 \<Rightarrow> (x # sublist' 0 j xs) | Suc i' \<Rightarrow>  sublist' i' j xs)"
-by (cases i) auto
-
-lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))"
-apply (cases j)
-apply auto
-apply (cases i)
-apply auto
-done
-
-lemma sublist_n_0: "sublist' n 0 xs = []"
-by (induct xs, auto)
-
-lemma sublist'_Nil': "n \<ge> m \<Longrightarrow> sublist' n m xs = []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_Nil2: "n \<ge> length xs \<Longrightarrow> sublist' n m xs = []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \<ge> m) \<or> (n \<ge> length xs))"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_notNil: "\<lbrakk> n < length xs; n < m \<rbrakk> \<Longrightarrow> sublist' n m xs \<noteq> []"
-apply (induct xs arbitrary: n m)
-apply simp
-apply (case_tac m)
-apply simp
-apply (case_tac n)
-apply simp
-apply simp
-done
-
-lemma sublist'_single: "n < length xs \<Longrightarrow> sublist' n (Suc n) xs = [xs ! n]"
-apply (induct xs arbitrary: n)
-apply simp
-apply simp
-apply (case_tac n)
-apply (simp add: sublist_n_0)
-apply simp
-done
-
-lemma sublist'_update1: "i \<ge> m \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
-apply (induct xs arbitrary: n m i)
-apply simp
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_update2: "i < n \<Longrightarrow> sublist' n m (xs[i:=v]) = sublist' n m xs"
-apply (induct xs arbitrary: n m i)
-apply simp
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_update3: "\<lbrakk>n \<le> i; i < m\<rbrakk> \<Longrightarrow> sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]"
-proof (induct xs arbitrary: n m i)
-  case Nil thus ?case by auto
-next
-  case (Cons x xs)
-  thus ?case
-    apply -
-    apply auto
-    apply (cases i)
-    apply auto
-    apply (cases i)
-    apply auto
-    done
-qed
-
-lemma "\<lbrakk> sublist' i j xs = sublist' i j ys; n \<ge> i; m \<le> j \<rbrakk> \<Longrightarrow> sublist' n m xs = sublist' n m ys"
-proof (induct xs arbitrary: i j ys n m)
-  case Nil
-  thus ?case
-    apply -
-    apply (rule sym, drule sym)
-    apply (simp add: sublist'_Nil)
-    apply (simp add: sublist'_Nil3)
-    apply arith
-    done
-next
-  case (Cons x xs i j ys n m)
-  note c = this
-  thus ?case
-  proof (cases m)
-    case 0 thus ?thesis by (simp add: sublist_n_0)
-  next
-    case (Suc m')
-    note a = this
-    thus ?thesis
-    proof (cases n)
-      case 0 note b = this
-      show ?thesis
-      proof (cases ys)
-	case Nil  with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3)
-      next
-	case (Cons y ys)
-	show ?thesis
-	proof (cases j)
-	  case 0 with a b Cons.prems show ?thesis by simp
-	next
-	  case (Suc j') with a b Cons.prems Cons show ?thesis 
-	    apply -
-	    apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto)
-	    done
-	qed
-      qed
-    next
-      case (Suc n')
-      show ?thesis
-      proof (cases ys)
-	case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3)
-      next
-	case (Cons y ys) with Suc a Cons.prems show ?thesis
-	  apply -
-	  apply simp
-	  apply (cases j)
-	  apply simp
-	  apply (cases i)
-	  apply simp
-	  apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply simp
-	  apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"])
-	  apply simp
-	  apply simp
-	  apply simp
-	  done
-      qed
-    qed
-  qed
-qed
-
-lemma length_sublist': "j \<le> length xs \<Longrightarrow> length (sublist' i j xs) = j - i"
-by (induct xs arbitrary: i j, auto)
-
-lemma sublist'_front: "\<lbrakk> i < j; i < length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = xs ! i # sublist' (Suc i) j xs"
-apply (induct xs arbitrary: a i j)
-apply simp
-apply (case_tac j)
-apply simp
-apply (case_tac i)
-apply simp
-apply simp
-done
-
-lemma sublist'_back: "\<lbrakk> i < j; j \<le> length xs \<rbrakk> \<Longrightarrow> sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]"
-apply (induct xs arbitrary: a i j)
-apply simp
-apply simp
-apply (case_tac j)
-apply simp
-apply auto
-apply (case_tac nat)
-apply auto
-done
-
-(* suffices that j \<le> length xs and length ys *) 
-lemma sublist'_eq_samelength_iff: "length xs = length ys \<Longrightarrow> (sublist' i j xs  = sublist' i j ys) = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> xs ! i' = ys ! i')"
-proof (induct xs arbitrary: ys i j)
-  case Nil thus ?case by simp
-next
-  case (Cons x xs)
-  thus ?case
-    apply -
-    apply (cases ys)
-    apply simp
-    apply simp
-    apply auto
-    apply (case_tac i', auto)
-    apply (erule_tac x="Suc i'" in allE, auto)
-    apply (erule_tac x="i' - 1" in allE, auto)
-    apply (case_tac i', auto)
-    apply (erule_tac x="Suc i'" in allE, auto)
-    done
-qed
-
-lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs"
-by (induct xs, auto)
-
-lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" 
-by (induct xs arbitrary: i j n m) (auto simp add: min_diff)
-
-lemma sublist'_append: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow>(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs"
-by (induct xs arbitrary: i j k) auto
-
-lemma nth_sublist': "\<lbrakk> k < j - i; j \<le> length xs \<rbrakk> \<Longrightarrow> (sublist' i j xs) ! k = xs ! (i + k)"
-apply (induct xs arbitrary: i j k)
-apply auto
-apply (case_tac k)
-apply auto
-apply (case_tac i)
-apply auto
-done
-
-lemma set_sublist': "set (sublist' i j xs) = {x. \<exists>k. i \<le> k \<and> k < j \<and> k < List.length xs \<and> x = xs ! k}"
-apply (simp add: sublist'_sublist)
-apply (simp add: set_sublist)
-apply auto
-done
-
-lemma all_in_set_sublist'_conv: "(\<forall>j. j \<in> set (sublist' l r xs) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
-unfolding set_sublist' by blast
-
-lemma ball_in_set_sublist'_conv: "(\<forall>j \<in> set (sublist' l r xs). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < List.length xs \<longrightarrow> P (xs ! k))"
-unfolding set_sublist' by blast
-
-
-lemma multiset_of_sublist:
-assumes l_r: "l \<le> r \<and> r \<le> List.length xs"
-assumes left: "\<forall> i. i < l \<longrightarrow> (xs::'a list) ! i = ys ! i"
-assumes right: "\<forall> i. i \<ge> r \<longrightarrow> (xs::'a list) ! i = ys ! i"
-assumes multiset: "multiset_of xs = multiset_of ys"
-  shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)"
-proof -
-  from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") 
-    by (simp add: sublist'_append)
-  from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length)
-  with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") 
-    by (simp add: sublist'_append)
-  from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp
-  moreover
-  from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys"
-    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
-  moreover
-  from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys"
-    by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI)
-  moreover
-  ultimately show ?thesis by (simp add: multiset_of_append)
-qed
-
-
-end
--- a/src/Provers/Arith/fast_lin_arith.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -466,7 +466,7 @@
                      NONE => ( the (try_add ([thm2] RL inj_thms) thm1)
                                handle Option =>
                                (trace_thm "" thm1; trace_thm "" thm2;
-                                sys_error "Lin.arith. failed to add thms")
+                                sys_error "Linear arithmetic: failed to add thms")
                              )
                    | SOME thm => thm)
         | SOME thm => thm;
@@ -588,8 +588,8 @@
           handle NoEx => NONE
       in
         case ex of
-          SOME s => (warning "arith failed - see trace for a counterexample"; tracing s)
-        | NONE => warning "arith failed"
+          SOME s => (warning "Linear arithmetic failed - see trace for a counterexample."; tracing s)
+        | NONE => warning "Linear arithmetic failed"
       end;
 
 (* ------------------------------------------------------------------------- *)
--- a/src/Pure/Concurrent/future.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/Concurrent/future.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -212,7 +212,8 @@
     val _ = if continue then () else scheduler := NONE;
 
     val _ = notify_all ();
-    val _ = interruptible (fn () => wait_timeout (Time.fromSeconds 1)) ()
+    val _ = interruptible (fn () =>
+        wait_timeout (Time.fromMilliseconds (if null (! canceled) then 1000 else 50))) ()
       handle Exn.Interrupt => List.app do_cancel (Task_Queue.cancel_all (! queue));
   in continue end;
 
--- a/src/Pure/General/antiquote.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/antiquote.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -8,7 +8,7 @@
 sig
   datatype 'a antiquote =
     Text of 'a |
-    Antiq of Symbol_Pos.T list * Position.T |
+    Antiq of Symbol_Pos.T list * Position.range |
     Open of Position.T |
     Close of Position.T
   val is_text: 'a antiquote -> bool
@@ -26,7 +26,7 @@
 
 datatype 'a antiquote =
   Text of 'a |
-  Antiq of Symbol_Pos.T list * Position.T |
+  Antiq of Symbol_Pos.T list * Position.range |
   Open of Position.T |
   Close of Position.T;
 
@@ -39,7 +39,7 @@
 val report_antiq = Position.report Markup.antiq;
 
 fun report report_text (Text x) = report_text x
-  | report _ (Antiq (_, pos)) = report_antiq pos
+  | report _ (Antiq (_, (pos, _))) = report_antiq pos
   | report _ (Open pos) = report_antiq pos
   | report _ (Close pos) = report_antiq pos;
 
@@ -79,7 +79,7 @@
   Symbol_Pos.scan_pos -- ($$$ "@" |-- $$$ "{" |--
     Symbol_Pos.!!! "missing closing brace of antiquotation"
       (Scan.repeat scan_ant -- ($$$ "}" |-- Symbol_Pos.scan_pos)))
-  >> (fn (pos1, (body, pos2)) => (flat body, Position.encode_range (pos1, pos2)));
+  >> (fn (pos1, (body, pos2)) => (flat body, Position.range pos1 pos2));
 
 val scan_open = Symbol_Pos.scan_pos --| $$$ "\\<lbrace>";
 val scan_close = Symbol_Pos.scan_pos --| $$$ "\\<rbrace>";
--- a/src/Pure/General/markup.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/markup.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -70,6 +70,9 @@
   val ML_stringN: string val ML_string: T
   val ML_commentN: string val ML_comment: T
   val ML_malformedN: string val ML_malformed: T
+  val ML_defN: string val ML_def: T
+  val ML_refN: string val ML_ref: T
+  val ML_typingN: string val ML_typing: T
   val ML_sourceN: string val ML_source: T
   val doc_sourceN: string val doc_source: T
   val antiqN: string val antiq: T
@@ -232,6 +235,10 @@
 val (ML_commentN, ML_comment) = markup_elem "ML_comment";
 val (ML_malformedN, ML_malformed) = markup_elem "ML_malformed";
 
+val (ML_defN, ML_def) = markup_elem "ML_def";
+val (ML_refN, ML_ref) = markup_elem "ML_ref";
+val (ML_typingN, ML_typing) = markup_elem "ML_typing";
+
 
 (* embedded source text *)
 
--- a/src/Pure/General/markup.scala	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/markup.scala	Thu Mar 26 13:02:12 2009 +0100
@@ -91,6 +91,10 @@
   val ML_COMMENT = "ML_comment"
   val ML_MALFORMED = "ML_malformed"
 
+  val ML_DEF = "ML_def"
+  val ML_REF = "ML_ref"
+  val ML_TYPING = "ML_typing"
+
 
   /* outer syntax */
 
--- a/src/Pure/General/output.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/output.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -47,7 +47,6 @@
   val debugging: bool ref
   val no_warnings: ('a -> 'b) -> 'a -> 'b
   val debug: (unit -> string) -> unit
-  val ml_output: (string -> unit) * (string -> 'a)
   val accumulated_time: unit -> unit
 end;
 
@@ -101,7 +100,7 @@
 val error_fn = ref (std_output o suffix "\n" o prefix_lines "*** ");
 val debug_fn = ref (std_output o suffix "\n" o prefix_lines "::: ");
 val prompt_fn = ref std_output;
-val status_fn = ref (fn s => ! writeln_fn s);
+val status_fn = ref (fn _: string => ());
 
 fun writeln s = ! writeln_fn (output s);
 fun priority s = ! priority_fn (output s);
@@ -120,8 +119,6 @@
 val debugging = ref false;
 fun debug s = if ! debugging then ! debug_fn (output (s ())) else ()
 
-val ml_output = (writeln, error);
-
 
 
 (** timing **)
--- a/src/Pure/General/position.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/position.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -24,6 +24,7 @@
   val of_properties: Properties.T -> T
   val properties_of: T -> Properties.T
   val default_properties: T -> Properties.T -> Properties.T
+  val report_text: Markup.T -> T -> string -> unit
   val report: Markup.T -> T -> unit
   val str_of: T -> string
   type range = T * T
@@ -121,9 +122,11 @@
   if exists (member (op =) Markup.position_properties o #1) props then props
   else properties_of default @ props;
 
-fun report markup (pos as Pos (count, _)) =
+fun report_text markup (pos as Pos (count, _)) txt =
   if invalid_count count then ()
-  else Output.status (Markup.markup (Markup.properties (properties_of pos) markup) "");
+  else Output.status (Markup.markup (Markup.properties (properties_of pos) markup) txt);
+
+fun report markup pos = report_text markup pos "";
 
 
 (* str_of *)
--- a/src/Pure/General/pretty.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/pretty.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -104,9 +104,9 @@
 (** printing items: compound phrases, strings, and breaks **)
 
 datatype T =
-  Block of Markup.T * T list * int * int |  (*markup, body, indentation, length*)
-  String of output * int |                  (*text, length*)
-  Break of bool * int;                      (*mandatory flag, width if not taken*)
+  Block of (output * output) * T list * int * int |  (*markup output, body, indentation, length*)
+  String of output * int |                           (*text, length*)
+  Break of bool * int;                               (*mandatory flag, width if not taken*)
 
 fun length (Block (_, _, _, len)) = len
   | length (String (_, len)) = len
@@ -124,12 +124,14 @@
 fun breaks prts = Library.separate (brk 1) prts;
 fun fbreaks prts = Library.separate fbrk prts;
 
-fun markup_block m (indent, es) =
+fun block_markup m (indent, es) =
   let
     fun sum [] k = k
       | sum (e :: es) k = sum es (length e + k);
   in Block (m, es, indent, sum es 0) end;
 
+fun markup_block m arg = block_markup (Markup.output m) arg;
+
 val blk = markup_block Markup.none;
 fun block prts = blk (2, prts);
 val strs = block o breaks o map str;
@@ -197,7 +199,7 @@
 local
   fun pruning dp (Block (m, bes, indent, wd)) =
         if dp > 0
-        then markup_block m (indent, map (pruning (dp - 1)) bes)
+        then block_markup m (indent, map (pruning (dp - 1)) bes)
         else str "..."
     | pruning dp e = e
 in
@@ -263,7 +265,7 @@
 fun format ([], _, _) text = text
   | format (e :: es, block as (blockind, blockin), after) (text as {ind, pos, nl, ...}) =
       (case e of
-        Block (markup, bes, indent, wd) =>
+        Block ((bg, en), bes, indent, wd) =>
           let
             val {emergencypos, ...} = ! margin_info;
             val pos' = pos + indent;
@@ -271,7 +273,6 @@
             val block' =
               if pos' < emergencypos then (ind |> add_indent indent, pos')
               else (add_indent pos'' Buffer.empty, pos'');
-            val (bg, en) = Markup.output markup;
             val btext: text = text
               |> control bg
               |> format (bes, block', breakdist (es, after))
@@ -303,9 +304,9 @@
 (*symbolic markup -- no formatting*)
 fun symbolic prt =
   let
-    fun out (Block (m, [], _, _)) = Buffer.markup m I
-      | out (Block (m, prts, indent, _)) =
-          Buffer.markup m (Buffer.markup (Markup.block indent) (fold out prts))
+    fun out (Block ((bg, en), [], _, _)) = Buffer.add bg #> Buffer.add en
+      | out (Block ((bg, en), prts, indent, _)) =
+          Buffer.add bg #> Buffer.markup (Markup.block indent) (fold out prts) #> Buffer.add en
       | out (String (s, _)) = Buffer.add s
       | out (Break (false, wd)) = Buffer.markup (Markup.break wd) (Buffer.add (output_spaces wd))
       | out (Break (true, _)) = Buffer.markup Markup.fbreak (Buffer.add (output_spaces 1));
@@ -314,7 +315,7 @@
 (*unformatted output*)
 fun unformatted prt =
   let
-    fun fmt (Block (m, prts, _, _)) = Buffer.markup m (fold fmt prts)
+    fun fmt (Block ((bg, en), prts, _, _)) = Buffer.add bg #> fold fmt prts #> Buffer.add en
       | fmt (String (s, _)) = Buffer.add s
       | fmt (Break (false, wd)) = Buffer.add (output_spaces wd)
       | fmt (Break (true, _)) = Buffer.add (output_spaces 1);
@@ -323,11 +324,11 @@
 
 (* ML toplevel pretty printing *)
 
-fun to_ML (Block (m, prts, ind, _)) = ML_Pretty.Block (Markup.output m, map to_ML prts, ind)
+fun to_ML (Block (m, prts, ind, _)) = ML_Pretty.Block (m, map to_ML prts, ind)
   | to_ML (String s) = ML_Pretty.String s
   | to_ML (Break b) = ML_Pretty.Break b;
 
-fun from_ML (ML_Pretty.Block (_, prts, ind)) = blk (ind, map from_ML prts)
+fun from_ML (ML_Pretty.Block (m, prts, ind)) = block_markup m (ind, map from_ML prts)
   | from_ML (ML_Pretty.String s) = String s
   | from_ML (ML_Pretty.Break b) = Break b;
 
--- a/src/Pure/General/secure.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/General/secure.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -9,11 +9,8 @@
   val set_secure: unit -> unit
   val is_secure: unit -> bool
   val deny_secure: string -> unit
-  val use_text: ML_NameSpace.nameSpace -> int * string ->
-    (string -> unit) * (string -> 'a) -> bool -> string -> unit
-  val use_file: ML_NameSpace.nameSpace ->
-    (string -> unit) * (string -> 'a) -> bool -> string -> unit
-  val use: string -> unit
+  val use_text: use_context -> int * string -> bool -> string -> unit
+  val use_file: use_context -> bool -> string -> unit
   val toplevel_pp: string list -> string -> unit
   val commit: unit -> unit
   val system_out: string -> string * int
@@ -40,23 +37,17 @@
 
 fun secure_mltext () = deny_secure "Cannot evaluate ML source in secure mode";
 
-fun raw_use_text ns = use_text ML_Parse.fix_ints (Position.str_of oo Position.line_file) ns;
-fun raw_use_file ns = use_file ML_Parse.fix_ints (Position.str_of oo Position.line_file) ns;
-fun raw_toplevel_pp x =
-  toplevel_pp ML_Parse.fix_ints (Position.str_of oo Position.line_file) Output.ml_output x;
+val raw_use_text = use_text;
+val raw_use_file = use_file;
+val raw_toplevel_pp = toplevel_pp;
 
-fun use_text ns pos pr verbose txt =
-  (secure_mltext (); raw_use_text ns pos pr verbose txt);
+fun use_text context pos verbose txt = (secure_mltext (); raw_use_text context pos verbose txt);
+fun use_file context verbose name = (secure_mltext (); raw_use_file context verbose name);
 
-fun use_file ns pr verbose name =
-  (secure_mltext (); raw_use_file ns pr verbose name);
-
-fun use name = use_file ML_NameSpace.global Output.ml_output true name;
-
-fun toplevel_pp path pp = (secure_mltext (); raw_toplevel_pp path pp);
+fun toplevel_pp path pp = (secure_mltext (); raw_toplevel_pp ML_Parse.global_context path pp);
 
 (*commit is dynamically bound!*)
-fun commit () = raw_use_text ML_NameSpace.global (0, "") Output.ml_output false "commit();";
+fun commit () = raw_use_text ML_Parse.global_context (0, "") false "commit();";
 
 
 (* shell commands *)
@@ -77,7 +68,8 @@
 (*override previous toplevel bindings!*)
 val use_text = Secure.use_text;
 val use_file = Secure.use_file;
-fun use s = Secure.use s handle ERROR msg => (writeln msg; raise Fail "ML error");
+fun use s = Secure.use_file ML_Parse.global_context true s
+  handle ERROR msg => (writeln msg; raise Fail "ML error");
 val toplevel_pp = Secure.toplevel_pp;
 val system_out = Secure.system_out;
 val system = Secure.system;
--- a/src/Pure/IsaMakefile	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/IsaMakefile	Thu Mar 26 13:02:12 2009 +0100
@@ -70,7 +70,8 @@
   Isar/toplevel.ML Isar/value_parse.ML ML/ml_antiquote.ML		\
   ML/ml_context.ML ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML		\
   ML/ml_test.ML ML/ml_thms.ML ML-Systems/install_pp_polyml.ML		\
-  ML-Systems/install_pp_polyml-experimental.ML Proof/extraction.ML	\
+  ML-Systems/install_pp_polyml-experimental.ML				\
+  ML-Systems/use_context.ML Proof/extraction.ML				\
   Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML			\
   Proof/proofchecker.ML Proof/reconstruct.ML ProofGeneral/ROOT.ML	\
   ProofGeneral/pgip.ML ProofGeneral/pgip_input.ML			\
--- a/src/Pure/Isar/code_unit.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/Isar/code_unit.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -218,7 +218,7 @@
     |> burrow_thms (canonical_tvars thy purify_tvar)
     |> map (canonical_vars thy purify_var)
     |> map (canonical_absvars purify_var)
-    |> map Drule.zero_var_indexes
+    |> Drule.zero_var_indexes_list
   end;
 
 
--- a/src/Pure/Isar/isar_syn.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -295,28 +295,28 @@
 
 (* use ML text *)
 
-fun inherit_env (context as Context.Proof lthy) =
+fun propagate_env (context as Context.Proof lthy) =
       Context.Proof (LocalTheory.map_contexts (ML_Context.inherit_env context) lthy)
-  | inherit_env context = context;
+  | propagate_env context = context;
 
-fun inherit_env_prf prf = Proof.map_contexts
+fun propagate_env_prf prf = Proof.map_contexts
   (Context.proof_map (ML_Context.inherit_env (Context.Proof (Proof.context_of prf)))) prf;
 
 val _ =
   OuterSyntax.command "use" "ML text from file" (K.tag_ml K.thy_decl)
-    (P.path >> (fn path => Toplevel.generic_theory (ThyInfo.exec_file false path #> inherit_env)));
+    (P.path >> (fn path => Toplevel.generic_theory (ThyInfo.exec_file false path #> propagate_env)));
 
 val _ =
   OuterSyntax.command "ML" "ML text within theory or local theory" (K.tag_ml K.thy_decl)
     (P.ML_source >> (fn (txt, pos) =>
       Toplevel.generic_theory
-        (ML_Context.exec (fn () => ML_Context.eval true pos txt) #> inherit_env)));
+        (ML_Context.exec (fn () => ML_Context.eval true pos txt) #> propagate_env)));
 
 val _ =
   OuterSyntax.command "ML_prf" "ML text within proof" (K.tag_proof K.prf_decl)
     (P.ML_source >> (fn (txt, pos) =>
       Toplevel.proof (Proof.map_context (Context.proof_map
-        (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> inherit_env_prf)));
+        (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> propagate_env_prf)));
 
 val _ =
   OuterSyntax.command "ML_val" "diagnostic ML text" (K.tag_ml K.diag)
--- a/src/Pure/Isar/outer_keyword.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/Isar/outer_keyword.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -150,23 +150,25 @@
   Pretty.mark (Markup.command_decl name (kind_of kind))
     (Pretty.str ("Outer syntax command: " ^ quote name ^ " (" ^ kind_of kind ^ ")"));
 
+fun status_writeln s = (Output.status s; writeln s);
+
 fun report () =
   let val (keywords, commands) = CRITICAL (fn () =>
     (dest_keywords (), sort_wrt #1 (Symtab.dest (get_commands ()))))
   in map report_keyword keywords @ map report_command commands end
-  |> Pretty.chunks |> Pretty.string_of |> Output.status;
+  |> Pretty.chunks |> Pretty.string_of |> status_writeln;
 
 
 (* augment tables *)
 
 fun keyword name =
  (change_lexicons (apfst (Scan.extend_lexicon (Symbol.explode name)));
-  Output.status (Pretty.string_of (report_keyword name)));
+  status_writeln (Pretty.string_of (report_keyword name)));
 
 fun command name kind =
  (change_lexicons (apsnd (Scan.extend_lexicon (Symbol.explode name)));
   change_commands (Symtab.update (name, kind));
-  Output.status (Pretty.string_of (report_command (name, kind))));
+  status_writeln (Pretty.string_of (report_command (name, kind))));
 
 
 (* command categories *)
--- a/src/Pure/ML-Systems/install_pp_polyml-experimental.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/install_pp_polyml-experimental.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -4,24 +4,15 @@
 Poly/ML 5.3.
 *)
 
-local
-
-fun pretty_future depth (pretty: 'a * int -> pretty) (x: 'a future) =
+addPrettyPrinter (fn depth => fn pretty => fn x =>
   (case Future.peek x of
     NONE => PrettyString "<future>"
   | SOME (Exn.Exn _) => PrettyString "<failed>"
-  | SOME (Exn.Result y) => pretty (y, depth));
+  | SOME (Exn.Result y) => pretty (y, depth)));
 
-fun pretty_lazy depth (pretty: 'a * int -> pretty) (x: 'a lazy) =
+addPrettyPrinter (fn depth => fn pretty => fn x =>
   (case Lazy.peek x of
     NONE => PrettyString "<lazy>"
   | SOME (Exn.Exn _) => PrettyString "<failed>"
-  | SOME (Exn.Result y) => pretty (y, depth));
-
-in
+  | SOME (Exn.Result y) => pretty (y, depth)));
 
-val _ = addPrettyPrinter pretty_future;
-val _ = addPrettyPrinter pretty_lazy;
-
-end;
-
--- a/src/Pure/ML-Systems/ml_name_space.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/ml_name_space.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -4,7 +4,7 @@
 ML name space -- dummy version of Poly/ML 5.2 facility.
 *)
 
-structure ML_NameSpace =
+structure ML_Name_Space =
 struct
 
 type valueVal = unit;
@@ -14,7 +14,7 @@
 type signatureVal = unit;
 type functorVal = unit;
 
-type nameSpace =
+type T =
  {lookupVal:    string -> valueVal option,
   lookupType:   string -> typeVal option,
   lookupFix:    string -> fixityVal option,
@@ -34,7 +34,7 @@
   allSig:       unit -> (string * signatureVal) list,
   allFunct:     unit -> (string * functorVal) list};
 
-val global: nameSpace =
+val global: T =
  {lookupVal = fn _ => NONE,
   lookupType = fn _ => NONE,
   lookupFix = fn _ => NONE,
--- a/src/Pure/ML-Systems/mosml.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/mosml.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -46,6 +46,7 @@
 use "ML-Systems/time_limit.ML";
 use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/ml_pretty.ML";
+use "ML-Systems/use_context.ML";
 
 
 (*low-level pointer equality*)
@@ -120,7 +121,7 @@
 end;
 
 (*dummy implementation*)
-fun toplevel_pp _ _ _ _ _ = ();
+fun toplevel_pp _ _ _ = ();
 
 (*dummy implementation*)
 fun ml_prompts p1 p2 = ();
@@ -185,18 +186,18 @@
 (* ML command execution *)
 
 (*Can one redirect 'use' directly to an instream?*)
-fun use_text (tune: string -> string) _ _ _ _ _ txt =
+fun use_text ({tune_source, ...}: use_context) _ _ txt =
   let
     val tmp_name = FileSys.tmpName ();
     val tmp_file = TextIO.openOut tmp_name;
   in
-    TextIO.output (tmp_file, tune txt);
+    TextIO.output (tmp_file, tune_source txt);
     TextIO.closeOut tmp_file;
     use tmp_name;
     FileSys.remove tmp_name
   end;
 
-fun use_file _ _ _ _ _ name = use name;
+fun use_file _ _ name = use name;
 
 
 
--- a/src/Pure/ML-Systems/polyml-4.1.3.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.1.3.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -6,6 +6,7 @@
 use "ML-Systems/polyml_old_basis.ML";
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler4.ML";
 use "ML-Systems/polyml_pp.ML";
--- a/src/Pure/ML-Systems/polyml-4.1.4.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.1.4.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -6,6 +6,7 @@
 use "ML-Systems/polyml_old_basis.ML";
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler4.ML";
 use "ML-Systems/polyml_pp.ML";
--- a/src/Pure/ML-Systems/polyml-4.2.0.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-4.2.0.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -5,6 +5,7 @@
 
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler4.ML";
 use "ML-Systems/polyml_pp.ML";
--- a/src/Pure/ML-Systems/polyml-5.0.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-5.0.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -5,6 +5,7 @@
 
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler5.ML";
 use "ML-Systems/polyml_pp.ML";
--- a/src/Pure/ML-Systems/polyml-5.1.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-5.1.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -4,6 +4,7 @@
 *)
 
 use "ML-Systems/thread_dummy.ML";
+use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/polyml_old_compiler5.ML";
 use "ML-Systems/polyml_pp.ML";
--- a/src/Pure/ML-Systems/polyml-experimental.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml-experimental.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -4,6 +4,14 @@
 *)
 
 open Thread;
+
+structure ML_Name_Space =
+struct
+  open PolyML.NameSpace;
+  type T = PolyML.NameSpace.nameSpace;
+  val global = PolyML.globalNameSpace;
+end;
+
 use "ML-Systems/polyml_common.ML";
 use "ML-Systems/multithreading_polyml.ML";
 
@@ -14,12 +22,6 @@
 
 (* runtime compilation *)
 
-structure ML_NameSpace =
-struct
-  open PolyML.NameSpace;
-  val global = PolyML.globalNameSpace;
-end;
-
 local
 
 fun drop_newline s =
@@ -28,11 +30,11 @@
 
 in
 
-fun use_text (tune: string -> string) str_of_pos
-    name_space (start_line, name) (print, err) verbose txt =
+fun use_text ({tune_source, name_space, str_of_pos, print, error, ...}: use_context)
+    (start_line, name) verbose txt =
   let
     val current_line = ref start_line;
-    val in_buffer = ref (String.explode (tune txt));
+    val in_buffer = ref (String.explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = drop_newline (implode (rev (! out_buffer)));
 
@@ -42,10 +44,11 @@
       | c :: cs =>
           (in_buffer := cs; if c = #"\n" then current_line := ! current_line + 1 else (); SOME c));
     fun put s = out_buffer := s :: ! out_buffer;
-    fun put_message {message, hard, location = {startLine = line, ...}, context} =
+    fun put_message {message = msg1, hard, location = {startLine = line, ...}, context} =
      (put (if hard then "Error: " else "Warning: ");
-      PolyML.prettyPrint (put, 76) message;
-      put (str_of_pos line name ^ "\n"));
+      PolyML.prettyPrint (put, 76) msg1;
+      (case context of NONE => () | SOME msg2 => PolyML.prettyPrint (put, 76) msg2);
+      put ("At" ^ str_of_pos line name ^ "\n"));
 
     val parameters =
      [PolyML.Compiler.CPOutStream put,
@@ -58,30 +61,50 @@
         PolyML.compiler (get, parameters) ())
       handle exn =>
        (put ("Exception- " ^ General.exnMessage exn ^ " raised");
-        err (output ()); raise exn);
+        error (output ()); raise exn);
   in if verbose then print (output ()) else () end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
 
 end;
 
 
 (* toplevel pretty printing *)
 
-fun pretty_ml (PrettyBlock (ind, _, _, prts)) = ML_Pretty.Block (("", ""), map pretty_ml prts, ind)
-  | pretty_ml (PrettyString s) = ML_Pretty.String (s, size s)
-  | pretty_ml (PrettyBreak (wd, _)) = ML_Pretty.Break (if wd < 99999 then (false, wd) else (true, 2));
+val pretty_ml =
+  let
+    fun convert len (PrettyBlock (ind, _, context, prts)) =
+          let
+            fun property name default =
+              (case List.find (fn ContextProperty (a, _) => name = a | _ => false) context of
+                SOME (ContextProperty (_, b)) => b
+              | NONE => default);
+            val bg = property "begin" "";
+            val en = property "end" "";
+            val len' = property "length" len;
+          in ML_Pretty.Block ((bg, en), map (convert len') prts, ind) end
+      | convert len (PrettyString s) =
+          ML_Pretty.String (s, case Int.fromString len of SOME i => i | NONE => size s)
+      | convert _ (PrettyBreak (wd, _)) =
+          ML_Pretty.Break (if wd < 99999 then (false, wd) else (true, 2));
+  in convert "" end;
 
-fun ml_pretty (ML_Pretty.Block (_, prts, ind)) = PrettyBlock (ind, false, [], map ml_pretty prts)
-  | ml_pretty (ML_Pretty.String (s, _)) = PrettyString s
+fun ml_pretty (ML_Pretty.Block ((bg, en), prts, ind)) =
+      let val context =
+        (if bg = "" then [] else [ContextProperty ("begin", bg)]) @
+        (if en = "" then [] else [ContextProperty ("end", en)])
+      in PrettyBlock (ind, false, context, map ml_pretty prts) end
+  | ml_pretty (ML_Pretty.String (s, len)) =
+      if len = size s then PrettyString s
+      else PrettyBlock (0, false, [ContextProperty ("length", Int.toString len)], [PrettyString s])
   | ml_pretty (ML_Pretty.Break (false, wd)) = PrettyBreak (wd, 0)
   | ml_pretty (ML_Pretty.Break (true, _)) = PrettyBreak (99999, 0);
 
-fun toplevel_pp tune str_of_pos output (_: string list) pp =
-  use_text tune str_of_pos ML_NameSpace.global (1, "pp") output false
+fun toplevel_pp context (_: string list) pp =
+  use_text context (1, "pp") false
     ("addPrettyPrinter (fn _ => fn _ => ml_pretty o Pretty.to_ML o (" ^ pp ^ "))");
 
--- a/src/Pure/ML-Systems/polyml.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -4,6 +4,14 @@
 *)
 
 open Thread;
+
+structure ML_Name_Space =
+struct
+  open PolyML.NameSpace;
+  type T = PolyML.NameSpace.nameSpace;
+  val global = PolyML.globalNameSpace;
+end;
+
 use "ML-Systems/polyml_common.ML";
 
 if ml_system = "polyml-5.2"
@@ -17,12 +25,6 @@
 
 (* runtime compilation *)
 
-structure ML_NameSpace =
-struct
-  open PolyML.NameSpace;
-  val global = PolyML.globalNameSpace;
-end;
-
 local
 
 fun drop_newline s =
@@ -31,11 +33,11 @@
 
 in
 
-fun use_text (tune: string -> string) str_of_pos
-    name_space (start_line, name) (print, err) verbose txt =
+fun use_text ({tune_source, name_space, str_of_pos, print, error, ...}: use_context)
+    (start_line, name) verbose txt =
   let
     val current_line = ref start_line;
-    val in_buffer = ref (String.explode (tune txt));
+    val in_buffer = ref (String.explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = drop_newline (implode (rev (! out_buffer)));
 
@@ -58,14 +60,14 @@
         PolyML.compiler (get, parameters) ())
       handle exn =>
        (put ("Exception- " ^ General.exnMessage exn ^ " raised");
-        err (output ()); raise exn);
+        error (output ()); raise exn);
   in if verbose then print (output ()) else () end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
 
 end;
 
--- a/src/Pure/ML-Systems/polyml_common.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_common.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -9,8 +9,8 @@
 use "ML-Systems/multithreading.ML";
 use "ML-Systems/time_limit.ML";
 use "ML-Systems/system_shell.ML";
-use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/ml_pretty.ML";
+use "ML-Systems/use_context.ML";
 
 
 (** ML system and platform related **)
--- a/src/Pure/ML-Systems/polyml_old_compiler4.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_compiler4.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -3,9 +3,9 @@
 Runtime compilation -- for old PolyML.compiler (version 4.x).
 *)
 
-fun use_text (tune: string -> string) _ _ (line: int, name) (print, err) verbose txt =
+fun use_text ({tune_source, print, error, ...}: use_context) (line: int, name) verbose txt =
   let
-    val in_buffer = ref (explode (tune txt));
+    val in_buffer = ref (explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = implode (rev (case ! out_buffer of "\n" :: cs => cs | cs => cs));
 
@@ -21,12 +21,12 @@
       | _ => (PolyML.compiler (get, put) (); exec ()));
   in
     exec () handle exn =>
-      (err ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
+      (error ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
     if verbose then print (output ()) else ()
   end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
--- a/src/Pure/ML-Systems/polyml_old_compiler5.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_old_compiler5.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -3,9 +3,9 @@
 Runtime compilation -- for old PolyML.compilerEx (version 5.0, 5.1).
 *)
 
-fun use_text (tune: string -> string) _ _ (line, name) (print, err) verbose txt =
+fun use_text ({tune_source, print, error, ...}: use_context) (line, name) verbose txt =
   let
-    val in_buffer = ref (explode (tune txt));
+    val in_buffer = ref (explode (tune_source txt));
     val out_buffer = ref ([]: string list);
     fun output () = implode (rev (case ! out_buffer of "\n" :: cs => cs | cs => cs));
 
@@ -21,12 +21,12 @@
         [] => ()
       | _ => (PolyML.compilerEx (get, put, fn () => ! current_line, name) (); exec ()));
   in
-    exec () handle exn => (err (output ()); raise exn);
+    exec () handle exn => (error (output ()); raise exn);
     if verbose then print (output ()) else ()
   end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
--- a/src/Pure/ML-Systems/polyml_pp.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/polyml_pp.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -14,7 +14,7 @@
       | pprint (ML_Pretty.Break (true, _)) = brk (99999, 0);
   in pprint end;
 
-fun toplevel_pp tune str_of_pos output (_: string list) pp =
-  use_text tune str_of_pos ML_NameSpace.global (1, "pp") output false
+fun toplevel_pp context (_: string list) pp =
+  use_text context (1, "pp") false
     ("install_pp (fn args => fn _ => fn _ => ml_pprint args o Pretty.to_ML o (" ^ pp ^ "))");
 
--- a/src/Pure/ML-Systems/smlnj.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML-Systems/smlnj.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -14,6 +14,7 @@
 use "ML-Systems/system_shell.ML";
 use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/ml_pretty.ML";
+use "ML-Systems/use_context.ML";
 
 
 (*low-level pointer equality*)
@@ -100,7 +101,7 @@
 
 (* ML command execution *)
 
-fun use_text (tune: string -> string) _ _ (line: int, name) (print, err) verbose txt =
+fun use_text ({tune_source, print, error, ...}: use_context) (line, name) verbose txt =
   let
     val ref out_orig = Control.Print.out;
 
@@ -111,22 +112,20 @@
       in String.substring (str, 0, Int.max (0, size str - 1)) end;
   in
     Control.Print.out := out;
-    Backend.Interact.useStream (TextIO.openString (tune txt)) handle exn =>
+    Backend.Interact.useStream (TextIO.openString (tune_source txt)) handle exn =>
       (Control.Print.out := out_orig;
-        err ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
+        error ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);
     Control.Print.out := out_orig;
     if verbose then print (output ()) else ()
   end;
 
-fun use_file tune str_of_pos name_space output verbose name =
+fun use_file context verbose name =
   let
     val instream = TextIO.openIn name;
     val txt = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);
-  in use_text tune str_of_pos name_space (1, name) output verbose txt end;
+  in use_text context (1, name) verbose txt end;
 
-fun forget_structure name =
-  use_text (fn x => x) (fn _ => "") () (1, "ML") (TextIO.print, fn s => raise Fail s) false
-    ("structure " ^ name ^ " = struct end");
+fun forget_structure _ = ();
 
 
 (* toplevel pretty printing *)
@@ -143,8 +142,8 @@
       | pprint (ML_Pretty.Break (true, _)) = PrettyPrint.newline pps;
   in pprint end;
 
-fun toplevel_pp tune str_of_pos output path pp =
-  use_text tune str_of_pos ML_NameSpace.global (1, "pp") output false
+fun toplevel_pp context path pp =
+  use_text context (1, "pp") false
     ("CompilerPPTable.install_pp [" ^ String.concatWith "," (map (fn s => "\"" ^ s ^ "\"")  path) ^
       "] (fn pps => ml_pprint pps o Pretty.to_ML o (" ^ pp ^ "))");
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/use_context.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -0,0 +1,13 @@
+(*  Title:      Pure/ML-Systems/use_context.ML
+    Author:     Makarius
+
+Common context for "use" operations (compiler invocation).
+*)
+
+type use_context =
+ {tune_source: string -> string,
+  name_space: ML_Name_Space.T,
+  str_of_pos: int -> string -> string,
+  print: string -> unit,
+  error: string -> unit};
+
--- a/src/Pure/ML/ml_context.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML/ml_context.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -20,7 +20,8 @@
   val the_local_context: unit -> Proof.context
   val exec: (unit -> unit) -> Context.generic -> Context.generic
   val inherit_env: Context.generic -> Context.generic -> Context.generic
-  val name_space: ML_NameSpace.nameSpace
+  val name_space: ML_Name_Space.T
+  val local_context: use_context
   val stored_thms: thm list ref
   val ml_store_thm: string * thm -> unit
   val ml_store_thms: string * thm list -> unit
@@ -31,13 +32,10 @@
   val trace: bool ref
   val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T ->
     Context.generic option -> (ML_Lex.token list * ML_Lex.token list) * Context.generic option
-  val eval_wrapper: (string -> unit) * (string -> 'a) -> bool ->
-    Position.T -> Symbol_Pos.text -> unit
   val eval: bool -> Position.T -> Symbol_Pos.text -> unit
   val eval_file: Path.T -> unit
   val eval_in: Proof.context option -> bool -> Position.T -> Symbol_Pos.text -> unit
-  val evaluate: Proof.context -> (string -> unit) * (string -> 'b) -> bool ->
-    string * (unit -> 'a) option ref -> string -> 'a
+  val evaluate: Proof.context -> bool -> string * (unit -> 'a) option ref -> string -> 'a
   val expression: Position.T -> string -> string -> string -> Context.generic -> Context.generic
 end
 
@@ -61,12 +59,12 @@
 structure ML_Env = GenericDataFun
 (
   type T =
-    ML_NameSpace.valueVal Symtab.table *
-    ML_NameSpace.typeVal Symtab.table *
-    ML_NameSpace.fixityVal Symtab.table *
-    ML_NameSpace.structureVal Symtab.table *
-    ML_NameSpace.signatureVal Symtab.table *
-    ML_NameSpace.functorVal Symtab.table;
+    ML_Name_Space.valueVal Symtab.table *
+    ML_Name_Space.typeVal Symtab.table *
+    ML_Name_Space.fixityVal Symtab.table *
+    ML_Name_Space.structureVal Symtab.table *
+    ML_Name_Space.signatureVal Symtab.table *
+    ML_Name_Space.functorVal Symtab.table;
   val empty = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty);
   val extend = I;
   fun merge _
@@ -82,23 +80,23 @@
 
 val inherit_env = ML_Env.put o ML_Env.get;
 
-val name_space: ML_NameSpace.nameSpace =
+val name_space: ML_Name_Space.T =
   let
     fun lookup sel1 sel2 name =
       Context.thread_data ()
       |> (fn NONE => NONE | SOME context => Symtab.lookup (sel1 (ML_Env.get context)) name)
-      |> (fn NONE => sel2 ML_NameSpace.global name | some => some);
+      |> (fn NONE => sel2 ML_Name_Space.global name | some => some);
 
     fun all sel1 sel2 () =
       Context.thread_data ()
       |> (fn NONE => [] | SOME context => Symtab.dest (sel1 (ML_Env.get context)))
-      |> append (sel2 ML_NameSpace.global ())
+      |> append (sel2 ML_Name_Space.global ())
       |> sort_distinct (string_ord o pairself #1);
 
     fun enter ap1 sel2 entry =
       if is_some (Context.thread_data ()) then
         Context.>> (ML_Env.map (ap1 (Symtab.update entry)))
-      else sel2 ML_NameSpace.global entry;
+      else sel2 ML_Name_Space.global entry;
   in
    {lookupVal    = lookup #1 #lookupVal,
     lookupType   = lookup #2 #lookupType,
@@ -120,6 +118,13 @@
     allFunct     = all #6 #allFunct}
   end;
 
+val local_context: use_context =
+ {tune_source = ML_Parse.fix_ints,
+  name_space = name_space,
+  str_of_pos = Position.str_of oo Position.line_file,
+  print = writeln,
+  error = error};
+
 
 (* theorem bindings *)
 
@@ -134,7 +139,7 @@
       else if not (ML_Syntax.is_identifier name) then
         error ("Cannot bind theorem(s) " ^ quote name ^ " as ML value")
       else setmp stored_thms ths' (fn () =>
-        use_text name_space (0, "") Output.ml_output true
+        use_text local_context (0, "") true
           ("val " ^ name ^ " = " ^ sel ^ "(! ML_Context.stored_thms);")) ();
   in () end;
 
@@ -214,12 +219,12 @@
           fun no_decl _ = ([], []);
 
           fun expand (Antiquote.Text tok) state = (K ([], [tok]), state)
-            | expand (Antiquote.Antiq x) (scope, background) =
+            | expand (Antiquote.Antiq (ss, range)) (scope, background) =
                 let
                   val context = Stack.top scope;
-                  val (f, context') = antiquotation (T.read_antiq lex antiq x) context;
+                  val (f, context') = antiquotation (T.read_antiq lex antiq (ss, #1 range)) context;
                   val (decl, background') = f {background = background, struct_name = struct_name};
-                  val decl' = pairself ML_Lex.tokenize o decl;
+                  val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range));
                 in (decl', (Stack.map_top (K context') scope, background')) end
             | expand (Antiquote.Open _) (scope, background) =
                 (no_decl, (Stack.push scope, background))
@@ -233,10 +238,10 @@
 
 val trace = ref false;
 
-fun eval_wrapper pr verbose pos txt =
+fun eval verbose pos txt =
   let
-    fun eval_raw p = use_text name_space
-      (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p)) pr;
+    fun eval_raw p = use_text local_context
+      (the_default 1 (Position.line_of p), the_default "ML" (Position.file_of p));
 
     (*prepare source text*)
     val _ = Position.report Markup.ML_source pos;
@@ -260,20 +265,18 @@
 end;
 
 
-(* ML evaluation *)
+(* derived versions *)
 
-val eval = eval_wrapper Output.ml_output;
 fun eval_file path = eval true (Path.position path) (File.read path);
 
 fun eval_in ctxt verbose pos txt =
   Context.setmp_thread_data (Option.map Context.Proof ctxt) (fn () => eval verbose pos txt) ();
 
-fun evaluate ctxt pr verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
+fun evaluate ctxt verbose (ref_name, r) txt = NAMED_CRITICAL "ML" (fn () =>
   let
     val _ = r := NONE;
     val _ = Context.setmp_thread_data (SOME (Context.Proof ctxt)) (fn () =>
-      eval_wrapper pr verbose Position.none
-        ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
+      eval verbose Position.none ("val _ = (" ^ ref_name ^ " := SOME (fn () => " ^ txt ^ "))")) ();
   in (case ! r of NONE => error ("Bad evaluation for " ^ ref_name) | SOME e => e) end) ();
 
 fun expression pos bind body txt =
--- a/src/Pure/ML/ml_lex.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML/ml_lex.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -13,6 +13,7 @@
   val stopper: token Scan.stopper
   val is_regular: token -> bool
   val is_improper: token -> bool
+  val set_range: Position.range -> token -> token
   val pos_of: token -> Position.T
   val kind_of: token -> token_kind
   val content_of: token -> string
@@ -42,6 +43,8 @@
 
 (* position *)
 
+fun set_range range (Token (_, x)) = Token (range, x);
+
 fun pos_of (Token ((pos, _), _)) = pos;
 fun end_pos_of (Token ((_, pos), _)) = pos;
 
--- a/src/Pure/ML/ml_parse.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML/ml_parse.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -7,38 +7,38 @@
 signature ML_PARSE =
 sig
   val fix_ints: string -> string
+  val global_context: use_context
 end;
 
 structure ML_Parse: ML_PARSE =
 struct
 
-structure T = ML_Lex;
-
-
 (** error handling **)
 
 fun !!! scan =
   let
     fun get_pos [] = " (past end-of-file!)"
-      | get_pos (tok :: _) = Position.str_of (T.pos_of tok);
+      | get_pos (tok :: _) = Position.str_of (ML_Lex.pos_of tok);
 
     fun err (toks, NONE) = "SML syntax error" ^ get_pos toks
       | err (toks, SOME msg) = "SML syntax error" ^ get_pos toks ^ ": " ^ msg;
   in Scan.!! err scan end;
 
 fun bad_input x =
-  (Scan.some (fn tok => (case T.kind_of tok of T.Error msg => SOME msg | _ => NONE)) :|--
+  (Scan.some (fn tok => (case ML_Lex.kind_of tok of ML_Lex.Error msg => SOME msg | _ => NONE)) :|--
     (fn msg => Scan.fail_with (K msg))) x;
 
 
 (** basic parsers **)
 
 fun $$$ x =
-  Scan.one (fn tok => T.kind_of tok = T.Keyword andalso T.content_of tok = x) >> T.content_of;
-val int = Scan.one (fn tok => T.kind_of tok = T.Int) >> T.content_of;
+  Scan.one (fn tok => ML_Lex.kind_of tok = ML_Lex.Keyword andalso ML_Lex.content_of tok = x)
+    >> ML_Lex.content_of;
 
-val regular = Scan.one T.is_regular >> T.content_of;
-val improper = Scan.one T.is_improper >> T.content_of;
+val int = Scan.one (fn tok => ML_Lex.kind_of tok = ML_Lex.Int) >> ML_Lex.content_of;
+
+val regular = Scan.one ML_Lex.is_regular >> ML_Lex.content_of;
+val improper = Scan.one ML_Lex.is_improper >> ML_Lex.content_of;
 
 val blanks = Scan.repeat improper >> implode;
 
@@ -55,11 +55,21 @@
 
 fun do_fix_ints s =
   Source.of_string s
-  |> T.source
-  |> Source.source T.stopper (Scan.bulk (!!! fix_int)) NONE
+  |> ML_Lex.source
+  |> Source.source ML_Lex.stopper (Scan.bulk (!!! fix_int)) NONE
   |> Source.exhaust
   |> implode;
 
 val fix_ints = if ml_system_fix_ints then do_fix_ints else I;
 
+
+(* global use_context *)
+
+val global_context: use_context =
+ {tune_source = fix_ints,
+  name_space = ML_Name_Space.global,
+  str_of_pos = Position.str_of oo Position.line_file,
+  print = writeln,
+  error = error};
+
 end;
--- a/src/Pure/ML/ml_test.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ML/ml_test.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -6,112 +6,197 @@
 
 signature ML_TEST =
 sig
-  val get_result: Proof.context -> PolyML.parseTree option
-  val eval: bool -> bool -> Position.T -> ML_Lex.token list -> unit
+  val eval: bool -> Position.T -> ML_Lex.token list -> unit
 end
 
 structure ML_Test: ML_TEST =
 struct
 
-(* eval ML source tokens *)
+(* extra ML environment *)
 
-structure Result = GenericDataFun
+structure Extra_Env = GenericDataFun
 (
-  type T = PolyML.parseTree option;
-  val empty = NONE;
-  fun extend _ = NONE;
-  fun merge _ _ = NONE;
+  type T = Position.T Inttab.table;  (*position of registered tokens*)
+  val empty = Inttab.empty;
+  val extend = I;
+  fun merge _ = Inttab.merge (K true);
 );
 
-val get_result = Result.get o Context.Proof;
+fun inherit_env context =
+  ML_Context.inherit_env context #>
+  Extra_Env.put (Extra_Env.get context);
 
-fun eval do_run verbose pos toks =
+fun register_tokens toks context =
   let
-    val (print, err) = Output.ml_output;
+    val regs = map (fn tok => (serial (), tok)) toks;
+    val context' = context
+      |> Extra_Env.map (fold (fn (i, tok) => Inttab.update (i, ML_Lex.pos_of tok)) regs);
+  in (regs, context') end;
 
-    val input = toks |> map (fn tok =>
-      (serial (), (String.explode (ML_Lex.text_of tok), ML_Lex.pos_of tok)));
-    val index_pos = Inttab.lookup (Inttab.make (map (apsnd snd) input));
+fun pos_of_location context
+    ({file, startLine = line, startPosition = i, endPosition = j, ...}: location) =
+  (case pairself (Inttab.lookup (Extra_Env.get context)) (i, j) of
+    (SOME pos1, SOME pos2) => Position.encode_range (pos1, pos2)
+  | (SOME pos, NONE) => pos
+  | _ => Position.line_file line file);
+
+
+(* parse trees *)
 
-    fun pos_of ({file, startLine = line, startPosition = i, endPosition = j, ...}: location) =
-      (case (index_pos i, index_pos j) of
-        (SOME p, SOME q) => Position.encode_range (p, q)
-      | (SOME p, NONE) => p
-      | _ => Position.line_file line file);
+fun report_parse_tree context depth =
+  let
+    val pos_of = pos_of_location context;
+    fun report loc (PTtype types) =
+          PolyML.NameSpace.displayTypeExpression (types, depth)
+          |> pretty_ml |> Pretty.from_ML |> Pretty.string_of
+          |> Position.report_text Markup.ML_typing (pos_of loc)
+      | report loc (PTdeclaredAt decl) =
+          Markup.markup (Markup.properties (Position.properties_of (pos_of decl)) Markup.ML_def) ""
+          |> Position.report_text Markup.ML_ref (pos_of loc)
+      | report _ (PTnextSibling tree) = report_tree (tree ())
+      | report _ (PTfirstChild tree) = report_tree (tree ())
+      | report _ _ = ()
+    and report_tree (loc, props) = List.app (report loc) props;
+  in report_tree end;
+
+
+(* eval ML source tokens *)
 
-    val in_buffer = ref (map (apsnd fst) input);
+fun use_text ({name_space = space, print, error, ...}: use_context) verbose pos toks =
+  let
+    (* input *)
+
+    val input = Context.>>> (register_tokens toks);
+    val input_buffer = ref (map (apsnd (String.explode o ML_Lex.text_of)) input);
+
     val current_line = ref (the_default 1 (Position.line_of pos));
+
+    fun get_index () =
+      the_default 0 (get_first (fn (_, []) => NONE | (i, _) => SOME i) (! input_buffer));
+
     fun get () =
-      (case ! in_buffer of
+      (case ! input_buffer of
         [] => NONE
-      | (_, []) :: rest => (in_buffer := rest; get ())
+      | (_, []) :: rest => (input_buffer := rest; get ())
       | (i, c :: cs) :: rest =>
-          (in_buffer := (i, cs) :: rest;
+          (input_buffer := (i, cs) :: rest;
            if c = #"\n" then current_line := ! current_line + 1 else ();
            SOME c));
-    fun get_index () = (case ! in_buffer of [] => 0 | (i, _) :: _ => i);
+
+
+    (* output *)
 
-    val out_buffer = ref ([]: string list);
-    fun put s = out_buffer := s :: ! out_buffer;
-    fun output () = implode (rev (! out_buffer));
+    val output_buffer = ref Buffer.empty;
+    fun output () = Buffer.content (! output_buffer);
+    fun put s = change output_buffer (Buffer.add s);
 
     fun put_message {message, hard, location, context = _} =
      (put (if hard then "Error: " else "Warning: ");
       put (Pretty.string_of (Pretty.from_ML (pretty_ml message)));
-      put (Position.str_of (pos_of location) ^ "\n"));
+      put (Position.str_of (pos_of_location (the (Context.thread_data ())) location) ^ "\n"));
+
+
+    (* results *)
+
+    val depth = get_print_depth ();
+    val with_struct = ! PolyML.Compiler.printTypesWithStructureName;
+
+    fun apply_result {fixes, types, signatures, structures, functors, values} =
+      let
+        fun add_prefix prefix (PrettyBlock (ind, consistent, context, prts)) =
+              let
+                fun make_prefix context =
+                  (case get_first (fn ContextParentStructure p => SOME p | _ => NONE) context of
+                    SOME (name, sub_context) => make_prefix sub_context ^ name ^ "."
+                  | NONE => prefix);
+                val this_prefix = make_prefix context;
+              in PrettyBlock (ind, consistent, context, map (add_prefix this_prefix) prts) end
+          | add_prefix prefix (prt as PrettyString s) =
+              if prefix = "" then prt else PrettyString (prefix ^ s)
+          | add_prefix _ (prt as PrettyBreak _) = prt;
+
+        fun display disp x =
+          if depth > 0 then
+            (disp x
+              |> with_struct ? add_prefix ""
+              |> pretty_ml |> Pretty.from_ML |> Pretty.string_of |> put; put "\n")
+          else ();
 
-    fun result_fun (parse_tree, code) () =
-      (Context.>> (Result.put parse_tree); (if is_none code then warning "Static Errors" else ()));
+        fun apply_fix (a, b) =
+          (display PolyML.NameSpace.displayFix (a, b); #enterFix space (a, b));
+        fun apply_type (a, b) =
+          (display PolyML.NameSpace.displayType (b, depth); #enterType space (a, b));
+        fun apply_sig (a, b) =
+          (display PolyML.NameSpace.displaySig (b, depth, space); #enterSig space (a, b));
+        fun apply_struct (a, b) =
+          (display PolyML.NameSpace.displayStruct (b, depth, space); #enterStruct space (a, b));
+        fun apply_funct (a, b) =
+          (display PolyML.NameSpace.displayFunct (b, depth, space); #enterFunct space (a, b));
+        fun apply_val (a, b) =
+          (display PolyML.NameSpace.displayVal (b, depth, space); #enterVal space (a, b));
+      in
+        List.app apply_fix fixes;
+        List.app apply_type types;
+        List.app apply_sig signatures;
+        List.app apply_struct structures;
+        List.app apply_funct functors;
+        List.app apply_val values
+      end;
+
+    fun result_fun (phase1, phase2) () =
+     (case phase1 of NONE => ()
+      | SOME parse_tree => report_parse_tree (the (Context.thread_data ())) depth parse_tree;
+      case phase2 of NONE => error "Static Errors"
+      | SOME code => apply_result (Toplevel.program code));
+
+
+    (* compiler invocation *)
 
     val parameters =
      [PolyML.Compiler.CPOutStream put,
-      PolyML.Compiler.CPNameSpace ML_Context.name_space,
+      PolyML.Compiler.CPNameSpace space,
       PolyML.Compiler.CPErrorMessageProc put_message,
       PolyML.Compiler.CPLineNo (fn () => ! current_line),
+      PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)),
       PolyML.Compiler.CPLineOffset get_index,
-      PolyML.Compiler.CPFileName (the_default "ML" (Position.file_of pos)),
-      PolyML.Compiler.CPPrintInAlphabeticalOrder false] @
-     (if do_run then [] else [PolyML.Compiler.CPCompilerResultFun result_fun]);
+      PolyML.Compiler.CPCompilerResultFun result_fun];
     val _ =
-      (while not (List.null (! in_buffer)) do
+      (while not (List.null (! input_buffer)) do
         PolyML.compiler (get, parameters) ())
       handle exn =>
        (put ("Exception- " ^ General.exnMessage exn ^ " raised");
-        err (output ()); raise exn);
+        error (output ()); raise exn);
   in if verbose then print (output ()) else () end;
 
+val eval = use_text ML_Context.local_context;
+
 
 (* ML test command *)
 
-fun ML_test do_run (txt, pos) =
+fun ML_test (txt, pos) =
   let
     val _ = Position.report Markup.ML_source pos;
     val ants = ML_Lex.read_antiq (Symbol_Pos.explode (txt, pos), pos);
     val ((env, body), env_ctxt) = ML_Context.eval_antiquotes (ants, pos) (Context.thread_data ());
 
     val _ = Context.setmp_thread_data env_ctxt
-        (fn () => (eval true false Position.none env; Context.thread_data ())) ()
-      |> (fn NONE => () | SOME context' => Context.>> (ML_Context.inherit_env context'));
-    val _ = eval do_run true pos body;
-    val _ = eval true false Position.none (ML_Lex.tokenize "structure Isabelle = struct end");
+        (fn () => (eval false Position.none env; Context.thread_data ())) ()
+      |> (fn NONE => () | SOME context' => Context.>> (inherit_env context'));
+    val _ = eval true pos body;
+    val _ = eval false Position.none (ML_Lex.tokenize "structure Isabelle = struct end");
   in () end;
 
 
 local structure P = OuterParse and K = OuterKeyword in
 
-fun inherit_env (context as Context.Proof lthy) =
-      Context.Proof (LocalTheory.map_contexts (ML_Context.inherit_env context) lthy)
-  | inherit_env context = context;
+fun propagate_env (context as Context.Proof lthy) =
+      Context.Proof (LocalTheory.map_contexts (inherit_env context) lthy)
+  | propagate_env context = context;
 
 val _ =
   OuterSyntax.command "ML_test" "advanced ML compiler test" (K.tag_ml K.thy_decl)
     (P.ML_source >> (fn src =>
-      Toplevel.generic_theory (ML_Context.exec (fn () => ML_test true src) #> inherit_env)));
-
-val _ =
-  OuterSyntax.command "ML_parse" "advanced ML compiler test (parse only)" (K.tag_ml K.thy_decl)
-    (P.ML_source >> (fn src =>
-      Toplevel.generic_theory (ML_Context.exec (fn () => ML_test false src) #> inherit_env)));
+      Toplevel.generic_theory (ML_Context.exec (fn () => ML_test src) #> propagate_env)));
 
 end;
 
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -76,7 +76,7 @@
 
 fun setup_messages () =
  (Output.writeln_fn := message "" "" "";
-  Output.status_fn := (fn s => ! Output.priority_fn s);
+  Output.status_fn := (fn _ => ());
   Output.priority_fn := message (special "I") (special "J") "";
   Output.tracing_fn := message (special "I" ^ special "V") (special "J") "";
   Output.debug_fn := message (special "K") (special "L") "+++ ";
--- a/src/Pure/Thy/thy_output.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/Thy/thy_output.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -148,8 +148,8 @@
 fun eval_antiquote lex state (txt, pos) =
   let
     fun expand (Antiquote.Text ss) = Symbol_Pos.content ss
-      | expand (Antiquote.Antiq x) =
-          let val (opts, src) = T.read_antiq lex antiq x in
+      | expand (Antiquote.Antiq (ss, (pos, _))) =
+          let val (opts, src) = T.read_antiq lex antiq (ss, pos) in
             options opts (fn () => command src state) ();  (*preview errors!*)
             PrintMode.with_modes (! modes @ Latex.modes)
               (Output.no_warnings (options opts (fn () => command src state))) ()
--- a/src/Pure/Tools/find_theorems.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/Tools/find_theorems.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -336,7 +336,9 @@
 
 fun find_theorems ctxt opt_goal rem_dups raw_criteria =
   let
-    val add_prems = Seq.hd o (TRY (Method.insert_tac (Assumption.all_prems_of ctxt) 1));
+    val assms = ProofContext.get_fact ctxt (Facts.named "local.assms")
+                handle ERROR _ => [];
+    val add_prems = Seq.hd o (TRY (Method.insert_tac assms 1));
     val opt_goal' = Option.map add_prems opt_goal;
 
     val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
--- a/src/Pure/display.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/display.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -57,6 +57,18 @@
 fun pretty_flexpair pp (t, u) = Pretty.block
   [Pretty.term pp t, Pretty.str " =?=", Pretty.brk 1, Pretty.term pp u];
 
+fun display_status th =
+  let
+    val {oracle = oracle0, unfinished, failed} = Thm.status_of th;
+    val oracle = oracle0 andalso (not (! quick_and_dirty) orelse ! show_hyps);
+  in
+    if failed then "!!"
+    else if oracle andalso unfinished then "!?"
+    else if oracle then "!"
+    else if unfinished then "?"
+    else ""
+  end;
+
 fun pretty_thm_aux pp quote show_hyps' asms raw_th =
   let
     val th = Thm.strip_shyps raw_th;
@@ -68,20 +80,17 @@
     val prt_term = q o Pretty.term pp;
 
     val hyps' = if ! show_hyps then hyps else subtract (op aconv) asms hyps;
-(* FIXME
-    val ora' = Thm.oracle_of th andalso (! show_hyps orelse not (! quick_and_dirty)); *)
-    val ora' = false;
+    val status = display_status th;
 
     val hlen = length xshyps + length hyps' + length tpairs;
     val hsymbs =
-      if hlen = 0 andalso not ora' then []
+      if hlen = 0 andalso status = "" then []
       else if ! show_hyps orelse show_hyps' then
         [Pretty.brk 2, Pretty.list "[" "]"
           (map (q o pretty_flexpair pp) tpairs @ map prt_term hyps' @
            map (Pretty.sort pp) xshyps @
-            (if ora' then [Pretty.str "!"] else []))]
-      else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^
-        (if ora' then "!" else "") ^ "]")];
+            (if status = "" then [] else [Pretty.str status]))]
+      else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ status ^ "]")];
     val tsymbs =
       if null tags orelse not (! show_tags) then []
       else [Pretty.brk 1, pretty_tags tags];
--- a/src/Pure/proofterm.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/proofterm.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -40,8 +40,9 @@
     {oracles: oracle OrdList.T, thms: pthm OrdList.T, proof: proof}
   val join_proof: proof_body future -> proof
   val proof_of: proof_body -> proof
+  val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
   val fold_body_thms: (string * term * proof_body -> 'a -> 'a) -> proof_body list -> 'a -> 'a
-  val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
+  val status_of: proof_body list -> {failed: bool, oracle: bool, unfinished: bool}
 
   val oracle_ord: oracle * oracle -> order
   val thm_ord: pthm * pthm -> order
@@ -159,17 +160,6 @@
 
 (***** proof atoms *****)
 
-fun fold_body_thms f =
-  let
-    fun app (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) =>
-      if Inttab.defined seen i then (x, seen)
-      else
-        let
-          val body' = Future.join body;
-          val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
-        in (f (name, prop, body') x', seen') end);
-  in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
-
 fun fold_proof_atoms all f =
   let
     fun app (Abst (_, _, prf)) = app prf
@@ -185,6 +175,39 @@
       | app prf = (fn (x, seen) => (f prf x, seen));
   in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end;
 
+fun fold_body_thms f =
+  let
+    fun app (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) =>
+      if Inttab.defined seen i then (x, seen)
+      else
+        let
+          val body' = Future.join body;
+          val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
+        in (f (name, prop, body') x', seen') end);
+  in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
+
+fun status_of bodies =
+  let
+    fun status (PBody {oracles, thms, ...}) x =
+      let
+        val ((oracle, unfinished, failed), seen) =
+          (thms, x) |-> fold (fn (i, (_, _, body)) => fn (st, seen) =>
+            if Inttab.defined seen i then (st, seen)
+            else
+              let val seen' = Inttab.update (i, ()) seen in
+                (case Future.peek body of
+                  SOME (Exn.Result body') => status body' (st, seen')
+                | SOME (Exn.Exn _) =>
+                    let val (oracle, unfinished, _) = st
+                    in ((oracle, unfinished, true), seen') end
+                | NONE =>
+                    let val (oracle, _, failed) = st
+                    in ((oracle, true, failed), seen') end)
+              end);
+      in ((oracle orelse not (null oracles), unfinished, failed), seen) end;
+    val (oracle, unfinished, failed) = #1 (fold status bodies ((false, false, false), Inttab.empty));
+  in {oracle = oracle, unfinished = unfinished, failed = failed} end;
+
 
 (* proof body *)
 
--- a/src/Pure/thm.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Pure/thm.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -144,6 +144,7 @@
   val freezeT: thm -> thm
   val future: thm future -> cterm -> thm
   val pending_groups: thm -> Task_Queue.group list -> Task_Queue.group list
+  val status_of: thm -> {oracle: bool, unfinished: bool, failed: bool}
   val proof_body_of: thm -> proof_body
   val proof_of: thm -> proof
   val join_proof: thm -> unit
@@ -1635,16 +1636,29 @@
   end;
 
 
-(* pending task groups *)
+(* derivation status *)
+
+fun raw_proof_body_of (Thm (Deriv {body, ...}, _)) = body;
+val raw_proof_of = Proofterm.proof_of o raw_proof_body_of;
 
 fun pending_groups (Thm (Deriv {open_promises, ...}, _)) =
   fold (insert Task_Queue.eq_group o Future.group_of o #2) open_promises;
 
+fun status_of (Thm (Deriv {promises, body, ...}, _)) =
+  let
+    val ps = map (Future.peek o snd) promises;
+    val bodies = body ::
+      map_filter (fn SOME (Exn.Result th) => SOME (raw_proof_body_of th) | _ => NONE) ps;
+    val {oracle, unfinished, failed} = Pt.status_of bodies;
+  in
+   {oracle = oracle,
+    unfinished = unfinished orelse exists is_none ps,
+    failed = failed orelse exists (fn SOME (Exn.Exn _) => true | _ => false) ps}
+  end;
+
 
 (* fulfilled proofs *)
 
-fun raw_proof_of (Thm (Deriv {body, ...}, _)) = Proofterm.proof_of body;
-
 fun proof_body_of (Thm (Deriv {open_promises, promises, body, ...}, {thy_ref, ...})) =
   let
     val _ = Exn.release_all (map (Future.join_result o #2) (rev open_promises));
--- a/src/Tools/Compute_Oracle/am_compiler.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Tools/Compute_Oracle/am_compiler.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -185,7 +185,7 @@
 
     in
 	compiled_rewriter := NONE;	
-	use_text ML_Context.name_space (1, "") Output.ml_output false (!buffer);
+	use_text ML_Context.local_context (1, "") false (!buffer);
 	case !compiled_rewriter of 
 	    NONE => raise (Compile "cannot communicate with compiled function")
 	  | SOME r => (compiled_rewriter := NONE; r)
--- a/src/Tools/Compute_Oracle/am_sml.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Tools/Compute_Oracle/am_sml.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -492,7 +492,7 @@
 
 fun writeTextFile name s = File.write (Path.explode name) s
 
-fun use_source src = use_text ML_Context.name_space (1, "") Output.ml_output false src
+fun use_source src = use_text ML_Context.local_context (1, "") false src
     
 fun compile cache_patterns const_arity eqs = 
     let
--- a/src/Tools/code/code_ml.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Tools/code/code_ml.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -969,7 +969,7 @@
         val (value_code, [SOME value_name']) = ml_code_of thy naming program' [value_name];
         val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name'
           ^ space_implode " " (map (enclose "(" ")") args) ^ " end";
-      in ML_Context.evaluate ctxt Output.ml_output false reff sml_code end;
+      in ML_Context.evaluate ctxt false reff sml_code end;
   in eval'' thy (rpair eval') ct end;
 
 fun eval_term reff = eval Code_Thingol.eval_term I reff;
@@ -1037,7 +1037,7 @@
 fun isar_seri_sml module_name =
   Code_Target.parse_args (Scan.succeed ())
   #> (fn () => serialize_ml target_SML
-      (SOME (use_text ML_Context.name_space (1, "generated code") Output.ml_output false))
+      (SOME (use_text ML_Context.local_context (1, "generated code") false))
       pr_sml_module pr_sml_stmt module_name);
 
 fun isar_seri_ocaml module_name =
--- a/src/Tools/nbe.ML	Thu Mar 26 13:01:09 2009 +0100
+++ b/src/Tools/nbe.ML	Thu Mar 26 13:02:12 2009 +0100
@@ -277,14 +277,12 @@
       in
         s
         |> tracing (fn s => "\n--- code to be evaluated:\n" ^ s)
-        |> ML_Context.evaluate ctxt
-            (Output.tracing o enclose "\n---compiler echo:\n" "\n---\n",
-            Output.tracing o enclose "\n--- compiler echo (with error):\n" "\n---\n")
-            (!trace) univs_cookie
+        |> ML_Context.evaluate ctxt (!trace) univs_cookie
         |> (fn f => f deps_vals)
         |> (fn univs => cs ~~ univs)
       end;
 
+
 (* preparing function equations *)
 
 fun eqns_of_stmt (_, Code_Thingol.Fun (_, (_, []))) =