Merged
authorManuel Eberl <eberlm@in.tum.de>
Fri, 30 Sep 2016 10:00:49 +0200
changeset 63966 957ba35d1338
parent 63965 d510b816ea41 (current diff)
parent 63964 9f0308e80366 (diff)
child 63967 2aa42596edc3
Merged
--- a/CONTRIBUTORS	Thu Sep 29 16:49:42 2016 +0200
+++ b/CONTRIBUTORS	Fri Sep 30 10:00:49 2016 +0200
@@ -30,6 +30,12 @@
   Algebraic foundation for primes; generalization from nat
   to general factorial rings
 
+* September 2016: Sascha Boehme
+  Proof method 'argo' based on SMT technology for a combination of
+  quantifier-free propositional logic, equality and linear real
+  arithmetic
+
+
 Contributions to Isabelle2016
 -----------------------------
 
--- a/NEWS	Thu Sep 29 16:49:42 2016 +0200
+++ b/NEWS	Fri Sep 30 10:00:49 2016 +0200
@@ -240,6 +240,13 @@
 
 *** HOL ***
 
+* New proof method "argo" using the built-in Argo solver based on SMT technology.
+The method can be used to prove goals of quantifier-free propositional logic,
+goals based on a combination of quantifier-free propositional logic with equality,
+and goals based on a combination of quantifier-free propositional logic with
+linear real arithmetic including min/max/abs. See HOL/ex/Argo_Examples.thy for
+examples.
+
 * Type class "div" with operation "mod" renamed to type class "modulo" with
 operation "modulo", analogously to type class "divide".  This eliminates the
 need to qualify any of those names in the presence of infix "mod" syntax.
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -129,24 +129,28 @@
     have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg")
     proof (cases x c rule: le_cases)
       case le show ?diff_fg
-        apply (rule differentiable_transform_within [where d = "dist x c" and f = f])
-        using x le st
-        apply (simp_all add: dist_real_def)
-        apply (rule differentiable_at_withinI)
-        apply (rule differentiable_within_open [where s = "{a<..<c} - s", THEN iffD1], simp_all)
-        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
-        apply (force elim!: differentiable_subset)+
-        done
+      proof (rule differentiable_transform_within [where d = "dist x c" and f = f])
+        have "f differentiable at x within ({a<..<c} - s)"
+          apply (rule differentiable_at_withinI)
+          using x le st
+          by (metis (no_types, lifting) DiffD1 DiffD2 DiffI UnCI atLeastAtMost_diff_ends atLeastAtMost_iff at_within_interior insert_iff interior_atLeastAtMost le st(3) x)
+        moreover have "open ({a<..<c} - s)"
+          by (blast intro: open_greaterThanLessThan \<open>finite s\<close> finite_imp_closed)
+        ultimately show "f differentiable at x within {a..b}"
+          using x le by (auto simp add: at_within_open_NO_MATCH differentiable_at_withinI) 
+      qed (use x le st dist_real_def in auto)
     next
       case ge show ?diff_fg
-        apply (rule differentiable_transform_within [where d = "dist x c" and f = g])
-        using x ge st
-        apply (simp_all add: dist_real_def)
-        apply (rule differentiable_at_withinI)
-        apply (rule differentiable_within_open [where s = "{c<..<b} - t", THEN iffD1], simp_all)
-        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
-        apply (force elim!: differentiable_subset)+
-        done
+      proof (rule differentiable_transform_within [where d = "dist x c" and f = g])
+        have "g differentiable at x within ({c<..<b} - t)"
+          apply (rule differentiable_at_withinI)
+          using x ge st
+          by (metis DiffD1 DiffD2 DiffI UnCI atLeastAtMost_diff_ends atLeastAtMost_iff at_within_interior insert_iff interior_atLeastAtMost)
+        moreover have "open ({c<..<b} - t)"
+          by (blast intro: open_greaterThanLessThan \<open>finite t\<close> finite_imp_closed)
+        ultimately show "g differentiable at x within {a..b}"
+          using x ge by (auto simp add: at_within_open_NO_MATCH differentiable_at_withinI) 
+      qed (use x ge st dist_real_def in auto)
     qed
   }
   then have "\<exists>s. finite s \<and>
@@ -3801,7 +3805,7 @@
   moreover have "{a<..<b} - k \<subseteq> {a..b} - k"
     by force
   ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x"
-    by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open)
+    by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open)
   { fix w
     assume "w \<noteq> z"
     have "continuous_on (ball w (cmod (w - z))) (\<lambda>w. 1 / (w - z))"
@@ -7527,4 +7531,32 @@
 by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path)
 
 
+lemma simply_connected_imp_winding_number_zero:
+  assumes "simply_connected s" "path g"
+           "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
+    shows "winding_number g z = 0"
+proof -
+  have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z"
+    apply (rule winding_number_homotopic_paths)
+    apply (rule homotopic_loops_imp_homotopic_paths_null [where a = "pathstart g"])
+    apply (rule homotopic_loops_subset [of s])
+    using assms
+    apply (auto simp: homotopic_paths_imp_homotopic_loops path_defs simply_connected_eq_contractible_path)
+    done
+  also have "... = 0"
+    using assms by (force intro: winding_number_trivial)
+  finally show ?thesis .
+qed
+
+lemma Cauchy_theorem_simply_connected:
+  assumes "open s" "simply_connected s" "f holomorphic_on s" "valid_path g"
+           "path_image g \<subseteq> s" "pathfinish g = pathstart g"
+    shows "(f has_contour_integral 0) g"
+using assms
+apply (simp add: simply_connected_eq_contractible_path)
+apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"]
+                         homotopic_paths_imp_homotopic_loops)
+using valid_path_imp_path by blast
+
+
 end
--- a/src/HOL/Analysis/Complete_Measure.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Complete_Measure.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -85,6 +85,9 @@
   "A \<in> sets M \<Longrightarrow> A \<in> sets (completion M)"
   unfolding sets_completion by force
 
+lemma measurable_completion: "f \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> f \<in> completion M \<rightarrow>\<^sub>M N"
+  by (auto simp: measurable_def)
+
 lemma null_sets_completion:
   assumes "N' \<in> null_sets M" "N \<subseteq> N'" shows "N \<in> sets (completion M)"
   using assms by (intro sets_completionI[of N "{}" N N']) auto
@@ -305,9 +308,6 @@
 lemma null_sets_completion_iff: "N \<in> sets M \<Longrightarrow> N \<in> null_sets (completion M) \<longleftrightarrow> N \<in> null_sets M"
   by (auto simp: null_sets_def)
 
-lemma AE_completion_iff: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in completion M. P x)"
-  by (simp add: AE_iff_null null_sets_completion_iff)
-
 lemma sets_completion_AE: "(AE x in M. \<not> P x) \<Longrightarrow> Measurable.pred (completion M) P"
   unfolding pred_def sets_completion eventually_ae_filter
   by auto
@@ -339,6 +339,12 @@
   "B \<subseteq> A \<Longrightarrow> A \<in> null_sets (completion M) \<Longrightarrow> B \<in> null_sets (completion M)"
   unfolding null_sets_completion_iff2 by auto
 
+interpretation completion: complete_measure "completion M" for M
+proof
+  show "B \<subseteq> A \<Longrightarrow> A \<in> null_sets (completion M) \<Longrightarrow> B \<in> sets (completion M)" for B A
+    using null_sets_completion_subset[of B A M] by (simp add: null_sets_def)
+qed
+
 lemma null_sets_restrict_space:
   "\<Omega> \<in> sets M \<Longrightarrow> A \<in> null_sets (restrict_space M \<Omega>) \<longleftrightarrow> A \<subseteq> \<Omega> \<and> A \<in> null_sets M"
   by (auto simp: null_sets_def emeasure_restrict_space sets_restrict_space)
@@ -416,6 +422,16 @@
     by (intro bexI[of _ s]) (auto simp: simple_integral_completion simple_function_completion)
 qed
 
+lemma integrable_completion:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> integrable (completion M) f \<longleftrightarrow> integrable M f"
+  by (rule integrable_subalgebra[symmetric]) auto
+
+lemma integral_completion:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" shows "integral\<^sup>L (completion M) f = integral\<^sup>L M f"
+  by (rule integral_subalgebra[symmetric]) (auto intro: f)
+
 locale semifinite_measure =
   fixes M :: "'a measure"
   assumes semifinite:
@@ -685,6 +701,291 @@
   finally show ?thesis .
 qed
 
+lemma (in complete_measure) complete_sets_sandwich_fmeasurable:
+  assumes [measurable]: "A \<in> fmeasurable M" "C \<in> fmeasurable M" and subset: "A \<subseteq> B" "B \<subseteq> C"
+    and measure: "measure M A = measure M C"
+  shows "B \<in> fmeasurable M"
+proof (rule fmeasurableI2)
+  show "B \<subseteq> C" "C \<in> fmeasurable M" by fact+
+  show "B \<in> sets M"
+  proof (rule complete_sets_sandwich)
+    show "A \<in> sets M" "C \<in> sets M" "A \<subseteq> B" "B \<subseteq> C"
+      using assms by auto
+    show "emeasure M A < \<infinity>"
+      using \<open>A \<in> fmeasurable M\<close> by (auto simp: fmeasurable_def)
+    show "emeasure M A = emeasure M C"
+      using assms by (simp add: emeasure_eq_measure2)
+  qed
+qed
+
+lemma AE_completion_iff: "(AE x in completion M. P x) \<longleftrightarrow> (AE x in M. P x)"
+proof
+  assume "AE x in completion M. P x"
+  then obtain N where "N \<in> null_sets (completion M)" and P: "{x\<in>space M. \<not> P x} \<subseteq> N"
+    unfolding eventually_ae_filter by auto
+  then obtain N' where N': "N' \<in> null_sets M" and "N \<subseteq> N'"
+    unfolding null_sets_completion_iff2 by auto
+  from P \<open>N \<subseteq> N'\<close> have "{x\<in>space M. \<not> P x} \<subseteq> N'"
+    by auto
+  with N' show "AE x in M. P x"
+    unfolding eventually_ae_filter by auto
+qed (rule AE_completion)
+
+lemma null_part_null_sets: "S \<in> completion M \<Longrightarrow> null_part M S \<in> null_sets (completion M)"
+  by (auto dest!: null_part intro: null_sets_completionI null_sets_completion_subset)
+
+lemma AE_notin_null_part: "S \<in> completion M \<Longrightarrow> (AE x in M. x \<notin> null_part M S)"
+  by (auto dest!: null_part_null_sets AE_not_in simp: AE_completion_iff)
+
+lemma completion_upper:
+  assumes A: "A \<in> sets (completion M)"
+  shows "\<exists>A'\<in>sets M. A \<subseteq> A' \<and> emeasure (completion M) A = emeasure M A'"
+proof -
+  from AE_notin_null_part[OF A] obtain N where N: "N \<in> null_sets M" "null_part M A \<subseteq> N"
+    unfolding eventually_ae_filter using null_part_null_sets[OF A, THEN null_setsD2, THEN sets.sets_into_space] by auto
+  show ?thesis
+  proof (intro bexI conjI)
+    show "A \<subseteq> main_part M A \<union> N"
+      using \<open>null_part M A \<subseteq> N\<close> by (subst main_part_null_part_Un[symmetric, OF A]) auto
+    show "emeasure (completion M) A = emeasure M (main_part M A \<union> N)"
+      using A \<open>N \<in> null_sets M\<close> by (simp add: emeasure_Un_null_set)
+  qed (use A N in auto)
+qed
+
+lemma AE_in_main_part:
+  assumes A: "A \<in> completion M" shows "AE x in M. x \<in> main_part M A \<longleftrightarrow> x \<in> A"
+  using AE_notin_null_part[OF A]
+  by (subst (2) main_part_null_part_Un[symmetric, OF A]) auto
+
+lemma completion_density_eq:
+  assumes ae: "AE x in M. 0 < f x" and [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel"
+  shows "completion (density M f) = density (completion M) f"
+proof (intro measure_eqI)
+  have "N' \<in> sets M \<and> (AE x\<in>N' in M. f x = 0) \<longleftrightarrow> N' \<in> null_sets M" for N'
+  proof safe
+    assume N': "N' \<in> sets M" and ae_N': "AE x\<in>N' in M. f x = 0"
+    from ae_N' ae have "AE x in M. x \<notin> N'"
+      by eventually_elim auto
+    then show "N' \<in> null_sets M"
+      using N' by (simp add: AE_iff_null_sets)
+  next
+    assume N': "N' \<in> null_sets M" then show "N' \<in> sets M" "AE x\<in>N' in M. f x = 0"
+      using ae AE_not_in[OF N'] by (auto simp: less_le)
+  qed
+  then show sets_eq: "sets (completion (density M f)) = sets (density (completion M) f)"
+    by (simp add: sets_completion null_sets_density_iff)
+
+  fix A assume A: \<open>A \<in> completion (density M f)\<close>
+  moreover
+  have "A \<in> completion M"
+    using A unfolding sets_eq by simp
+  moreover
+  have "main_part (density M f) A \<in> M"
+    using A main_part_sets[of A "density M f"] unfolding sets_density sets_eq by simp
+  moreover have "AE x in M. x \<in> main_part (density M f) A \<longleftrightarrow> x \<in> A"
+    using AE_in_main_part[OF \<open>A \<in> completion (density M f)\<close>] ae by (auto simp add: AE_density)
+  ultimately show "emeasure (completion (density M f)) A = emeasure (density (completion M) f) A"
+    by (auto simp add: emeasure_density measurable_completion nn_integral_completion intro!: nn_integral_cong_AE)
+qed
+
+lemma null_sets_subset: "A \<subseteq> B \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> A \<in> null_sets M"
+  using emeasure_mono[of A B M] by (simp add: null_sets_def)
+
+lemma (in complete_measure) complete2: "A \<subseteq> B \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> A \<in> null_sets M"
+  using complete[of A B] null_sets_subset[of A B M] by simp
+
+lemma (in complete_measure) AE_iff_null_sets: "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
+  unfolding eventually_ae_filter by (auto intro: complete2)
+
+lemma (in complete_measure) null_sets_iff_AE: "A \<in> null_sets M \<longleftrightarrow> ((AE x in M. x \<notin> A) \<and> A \<subseteq> space M)"
+  unfolding AE_iff_null_sets by (auto cong: rev_conj_cong dest: sets.sets_into_space simp: subset_eq)
+
+lemma (in complete_measure) in_sets_AE:
+  assumes ae: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B" and A: "A \<in> sets M" and B: "B \<subseteq> space M"
+  shows "B \<in> sets M"
+proof -
+  have "(AE x in M. x \<notin> B - A \<and> x \<notin> A - B)"
+    using ae by eventually_elim auto
+  then have "B - A \<in> null_sets M" "A - B \<in> null_sets M"
+    using A B unfolding null_sets_iff_AE by (auto dest: sets.sets_into_space)
+  then have "A \<union> (B - A) - (A - B) \<in> sets M"
+    using A by blast
+  also have "A \<union> (B - A) - (A - B) = B"
+    by auto
+  finally show "B \<in> sets M" .
+qed
+
+lemma (in complete_measure) vimage_null_part_null_sets:
+  assumes f: "f \<in> M \<rightarrow>\<^sub>M N" and eq: "null_sets N \<subseteq> null_sets (distr M N f)"
+    and A: "A \<in> completion N"
+  shows "f -` null_part N A \<inter> space M \<in> null_sets M"
+proof -
+  obtain N' where "N' \<in> null_sets N" "null_part N A \<subseteq> N'"
+    using null_part[OF A] by auto
+  then have N': "N' \<in> null_sets (distr M N f)"
+    using eq by auto
+  show ?thesis
+  proof (rule complete2)
+    show "f -` null_part N A \<inter> space M \<subseteq> f -` N' \<inter> space M"
+      using \<open>null_part N A \<subseteq> N'\<close> by auto
+    show "f -` N' \<inter> space M \<in> null_sets M"
+      using f N' by (auto simp: null_sets_def emeasure_distr)
+  qed
+qed
+
+lemma (in complete_measure) vimage_null_part_sets:
+  "f \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> null_sets N \<subseteq> null_sets (distr M N f) \<Longrightarrow> A \<in> completion N \<Longrightarrow>
+  f -` null_part N A \<inter> space M \<in> sets M"
+  using vimage_null_part_null_sets[of f N A] by auto
+
+lemma (in complete_measure) measurable_completion2:
+  assumes f: "f \<in> M \<rightarrow>\<^sub>M N" and null_sets: "null_sets N \<subseteq> null_sets (distr M N f)"
+  shows "f \<in> M \<rightarrow>\<^sub>M completion N"
+proof (rule measurableI)
+  show "x \<in> space M \<Longrightarrow> f x \<in> space (completion N)" for x
+    using f[THEN measurable_space] by auto
+  fix A assume A: "A \<in> sets (completion N)"
+  have "f -` A \<inter> space M = (f -` main_part N A \<inter> space M) \<union> (f -` null_part N A \<inter> space M)"
+    using main_part_null_part_Un[OF A] by auto
+  then show "f -` A \<inter> space M \<in> sets M"
+    using f A null_sets by (auto intro: vimage_null_part_sets measurable_sets)
+qed
+
+lemma (in complete_measure) completion_distr_eq:
+  assumes X: "X \<in> M \<rightarrow>\<^sub>M N" and null_sets: "null_sets (distr M N X) = null_sets N"
+  shows "completion (distr M N X) = distr M (completion N) X"
+proof (rule measure_eqI)
+  show eq: "sets (completion (distr M N X)) = sets (distr M (completion N) X)"
+    by (simp add: sets_completion null_sets)
+
+  fix A assume A: "A \<in> completion (distr M N X)"
+  moreover have A': "A \<in> completion N"
+    using A by (simp add: eq)
+  moreover have "main_part (distr M N X) A \<in> sets N"
+    using main_part_sets[OF A] by simp
+  moreover have "X -` A \<inter> space M = (X -` main_part (distr M N X) A \<inter> space M) \<union> (X -` null_part (distr M N X) A \<inter> space M)"
+    using main_part_null_part_Un[OF A] by auto
+  moreover have "X -` null_part (distr M N X) A \<inter> space M \<in> null_sets M"
+    using X A by (intro vimage_null_part_null_sets) (auto cong: distr_cong)
+  ultimately show "emeasure (completion (distr M N X)) A = emeasure (distr M (completion N) X) A"
+    using X by (auto simp: emeasure_distr measurable_completion null_sets measurable_completion2
+                     intro!: emeasure_Un_null_set[symmetric])
+qed
+
+lemma distr_completion: "X \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> distr (completion M) N X = distr M N X"
+  by (intro measure_eqI) (auto simp: emeasure_distr measurable_completion)
+
+proposition (in complete_measure) fmeasurable_inner_outer:
+  "S \<in> fmeasurable M \<longleftrightarrow>
+    (\<forall>e>0. \<exists>T\<in>fmeasurable M. \<exists>U\<in>fmeasurable M. T \<subseteq> S \<and> S \<subseteq> U \<and> \<bar>measure M T - measure M U\<bar> < e)"
+  (is "_ \<longleftrightarrow> ?approx")
+proof safe
+  let ?\<mu> = "measure M" let ?D = "\<lambda>T U . \<bar>?\<mu> T - ?\<mu> U\<bar>"
+  assume ?approx
+  have "\<exists>A. \<forall>n. (fst (A n) \<in> fmeasurable M \<and> snd (A n) \<in> fmeasurable M \<and> fst (A n) \<subseteq> S \<and> S \<subseteq> snd (A n) \<and>
+    ?D (fst (A n)) (snd (A n)) < 1/Suc n) \<and> (fst (A n) \<subseteq> fst (A (Suc n)) \<and> snd (A (Suc n)) \<subseteq> snd (A n))"
+    (is "\<exists>A. \<forall>n. ?P n (A n) \<and> ?Q (A n) (A (Suc n))")
+  proof (intro dependent_nat_choice)
+    show "\<exists>A. ?P 0 A"
+      using \<open>?approx\<close>[THEN spec, of 1] by auto
+  next
+    fix A n assume "?P n A"
+    moreover
+    from \<open>?approx\<close>[THEN spec, of "1/Suc (Suc n)"]
+    obtain T U where *: "T \<in> fmeasurable M" "U \<in> fmeasurable M" "T \<subseteq> S" "S \<subseteq> U" "?D T U < 1 / Suc (Suc n)"
+      by auto
+    ultimately have "?\<mu> T \<le> ?\<mu> (T \<union> fst A)" "?\<mu> (U \<inter> snd A) \<le> ?\<mu> U"
+      "?\<mu> T \<le> ?\<mu> U" "?\<mu> (T \<union> fst A) \<le> ?\<mu> (U \<inter> snd A)"
+      by (auto intro!: measure_mono_fmeasurable)
+    then have "?D (T \<union> fst A) (U \<inter> snd A) \<le> ?D T U"
+      by auto
+    also have "?D T U < 1/Suc (Suc n)" by fact
+    finally show "\<exists>B. ?P (Suc n) B \<and> ?Q A B"
+      using \<open>?P n A\<close> *
+      by (intro exI[of _ "(T \<union> fst A, U \<inter> snd A)"] conjI) auto
+  qed
+  then obtain A
+    where lm: "\<And>n. fst (A n) \<in> fmeasurable M" "\<And>n. snd (A n) \<in> fmeasurable M"
+      and set_bound: "\<And>n. fst (A n) \<subseteq> S" "\<And>n. S \<subseteq> snd (A n)"
+      and mono: "\<And>n. fst (A n) \<subseteq> fst (A (Suc n))" "\<And>n. snd (A (Suc n)) \<subseteq> snd (A n)"
+      and bound: "\<And>n. ?D (fst (A n)) (snd (A n)) < 1/Suc n"
+    by metis
+
+  have INT_sA: "(\<Inter>n. snd (A n)) \<in> fmeasurable M"
+    using lm by (intro fmeasurable_INT[of _ 0]) auto
+  have UN_fA: "(\<Union>n. fst (A n)) \<in> fmeasurable M"
+    using lm order_trans[OF set_bound(1) set_bound(2)[of 0]] by (intro fmeasurable_UN[of _ _ "snd (A 0)"]) auto
+
+  have "(\<lambda>n. ?\<mu> (fst (A n)) - ?\<mu> (snd (A n))) \<longlonglongrightarrow> 0"
+    using bound
+    by (subst tendsto_rabs_zero_iff[symmetric])
+       (intro tendsto_sandwich[OF _ _ tendsto_const LIMSEQ_inverse_real_of_nat];
+        auto intro!: always_eventually less_imp_le simp: divide_inverse)
+  moreover
+  have "(\<lambda>n. ?\<mu> (fst (A n)) - ?\<mu> (snd (A n))) \<longlonglongrightarrow> ?\<mu> (\<Union>n. fst (A n)) - ?\<mu> (\<Inter>n. snd (A n))"
+  proof (intro tendsto_diff Lim_measure_incseq Lim_measure_decseq)
+    show "range (\<lambda>i. fst (A i)) \<subseteq> sets M" "range (\<lambda>i. snd (A i)) \<subseteq> sets M"
+      "incseq (\<lambda>i. fst (A i))" "decseq (\<lambda>n. snd (A n))"
+      using mono lm by (auto simp: incseq_Suc_iff decseq_Suc_iff intro!: measure_mono_fmeasurable)
+    show "emeasure M (\<Union>x. fst (A x)) \<noteq> \<infinity>" "emeasure M (snd (A n)) \<noteq> \<infinity>" for n
+      using lm(2)[of n] UN_fA by (auto simp: fmeasurable_def)
+  qed
+  ultimately have eq: "0 = ?\<mu> (\<Union>n. fst (A n)) - ?\<mu> (\<Inter>n. snd (A n))"
+    by (rule LIMSEQ_unique)
+
+  show "S \<in> fmeasurable M"
+    using UN_fA INT_sA
+  proof (rule complete_sets_sandwich_fmeasurable)
+    show "(\<Union>n. fst (A n)) \<subseteq> S" "S \<subseteq> (\<Inter>n. snd (A n))"
+      using set_bound by auto
+    show "?\<mu> (\<Union>n. fst (A n)) = ?\<mu> (\<Inter>n. snd (A n))"
+      using eq by auto
+  qed
+qed (auto intro!: bexI[of _ S])
+
+lemma (in complete_measure) fmeasurable_measure_inner_outer:
+   "(S \<in> fmeasurable M \<and> m = measure M S) \<longleftrightarrow>
+      (\<forall>e>0. \<exists>T\<in>fmeasurable M. T \<subseteq> S \<and> m - e < measure M T) \<and>
+      (\<forall>e>0. \<exists>U\<in>fmeasurable M. S \<subseteq> U \<and> measure M U < m + e)"
+    (is "?lhs = ?rhs")
+proof
+  assume RHS: ?rhs
+  then have T: "\<And>e. 0 < e \<longrightarrow> (\<exists>T\<in>fmeasurable M. T \<subseteq> S \<and> m - e < measure M T)"
+        and U: "\<And>e. 0 < e \<longrightarrow> (\<exists>U\<in>fmeasurable M. S \<subseteq> U \<and> measure M U < m + e)"
+    by auto
+  have "S \<in> fmeasurable M"
+  proof (subst fmeasurable_inner_outer, safe)
+    fix e::real assume "0 < e"
+    with RHS obtain T U where T: "T \<in> fmeasurable M" "T \<subseteq> S" "m - e/2 < measure M T"
+                          and U: "U \<in> fmeasurable M" "S \<subseteq> U" "measure M U < m + e/2"
+      by (meson half_gt_zero)+
+    moreover have "measure M U - measure M T < (m + e/2) - (m - e/2)"
+      by (intro diff_strict_mono) fact+
+    moreover have "measure M T \<le> measure M U"
+      using T U by (intro measure_mono_fmeasurable) auto
+    ultimately show "\<exists>T\<in>fmeasurable M. \<exists>U\<in>fmeasurable M. T \<subseteq> S \<and> S \<subseteq> U \<and> \<bar>measure M T - measure M U\<bar> < e"
+      apply (rule_tac bexI[OF _ \<open>T \<in> fmeasurable M\<close>])
+      apply (rule_tac bexI[OF _ \<open>U \<in> fmeasurable M\<close>])
+      by auto
+  qed
+  moreover have "m = measure M S"
+    using \<open>S \<in> fmeasurable M\<close> U[of "measure M S - m"] T[of "m - measure M S"]
+    by (cases m "measure M S" rule: linorder_cases)
+       (auto simp: not_le[symmetric] measure_mono_fmeasurable)
+  ultimately show ?lhs
+    by simp
+qed (auto intro!: bexI[of _ S])
+
+lemma (in complete_measure) null_sets_outer:
+  "S \<in> null_sets M \<longleftrightarrow> (\<forall>e>0. \<exists>T\<in>fmeasurable M. S \<subseteq> T \<and> measure M T < e)"
+proof -
+  have "S \<in> null_sets M \<longleftrightarrow> (S \<in> fmeasurable M \<and> 0 = measure M S)"
+    by (auto simp: null_sets_def emeasure_eq_measure2 intro: fmeasurableI) (simp add: measure_def)
+  also have "\<dots> = (\<forall>e>0. \<exists>T\<in>fmeasurable M. S \<subseteq> T \<and> measure M T < e)"
+    unfolding fmeasurable_measure_inner_outer by auto
+  finally show ?thesis .
+qed
+
 lemma (in cld_measure) notin_sets_outer_measure_of_cover:
   assumes E: "E \<subseteq> space M" "E \<notin> sets M"
   shows "\<exists>B\<in>sets M. 0 < emeasure M B \<and> emeasure M B < \<infinity> \<and>
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -6417,6 +6417,26 @@
   unfolding segment_convex_hull
   by (auto intro!: hull_subset[unfolded subset_eq, rule_format])
 
+lemma eventually_closed_segment:
+  fixes x0::"'a::real_normed_vector"
+  assumes "open X0" "x0 \<in> X0"
+  shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0"
+proof -
+  from openE[OF assms]
+  obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" .
+  then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e"
+    by (auto simp: dist_commute eventually_at)
+  then show ?thesis
+  proof eventually_elim
+    case (elim x)
+    have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp
+    from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim]
+    have "closed_segment x0 x \<subseteq> ball x0 e" .
+    also note \<open>\<dots> \<subseteq> X0\<close>
+    finally show ?case .
+  qed
+qed
+
 lemma segment_furthest_le:
   fixes a b x y :: "'a::euclidean_space"
   assumes "x \<in> closed_segment a b"
@@ -8697,6 +8717,59 @@
 apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def)
 by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen)
 
+lemma rel_interior_convex_Int_affine:
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "affine T" "interior S \<inter> T \<noteq> {}"
+    shows "rel_interior(S \<inter> T) = interior S \<inter> T"
+proof -
+  obtain a where aS: "a \<in> interior S" and aT:"a \<in> T"
+    using assms by force
+  have "rel_interior S = interior S"
+    by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior)
+  then show ?thesis
+    by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull)
+qed
+
+lemma closure_convex_Int_affine:
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "affine T" "rel_interior S \<inter> T \<noteq> {}"
+  shows "closure(S \<inter> T) = closure S \<inter> T"
+proof
+  have "closure (S \<inter> T) \<subseteq> closure T"
+    by (simp add: closure_mono)
+  also have "... \<subseteq> T"
+    by (simp add: affine_closed assms)
+  finally show "closure(S \<inter> T) \<subseteq> closure S \<inter> T"
+    by (simp add: closure_mono)
+next
+  obtain a where "a \<in> rel_interior S" "a \<in> T"
+    using assms by auto
+  then have ssT: "subspace ((\<lambda>x. (-a)+x) ` T)" and "a \<in> S"
+    using affine_diffs_subspace rel_interior_subset assms by blast+
+  show "closure S \<inter> T \<subseteq> closure (S \<inter> T)"
+  proof
+    fix x  assume "x \<in> closure S \<inter> T"
+    show "x \<in> closure (S \<inter> T)"
+    proof (cases "x = a")
+      case True
+      then show ?thesis
+        using \<open>a \<in> S\<close> \<open>a \<in> T\<close> closure_subset by fastforce
+    next
+      case False
+      then have "x \<in> closure(open_segment a x)"
+        by auto
+      then show ?thesis
+        using \<open>x \<in> closure S \<inter> T\<close> assms convex_affine_closure_Int by blast
+    qed
+  qed
+qed
+
+lemma rel_frontier_convex_Int_affine:
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "affine T" "interior S \<inter> T \<noteq> {}"
+    shows "rel_frontier(S \<inter> T) = frontier S \<inter> T"
+by (simp add: assms convex_affine_rel_frontier_Int)
+
 lemma subset_rel_interior_convex:
   fixes S T :: "'n::euclidean_space set"
   assumes "convex S"
@@ -10447,7 +10520,7 @@
 
 lemma collinear_3_expand:
    "collinear{a,b,c} \<longleftrightarrow> a = c \<or> (\<exists>u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)"
-proof -                          
+proof -
   have "collinear{a,b,c} = collinear{a,c,b}"
     by (simp add: insert_commute)
   also have "... = collinear {0, a - c, b - c}"
--- a/src/HOL/Analysis/Derivative.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Derivative.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -180,13 +180,6 @@
   shows "f differentiable F \<longleftrightarrow> (\<exists>d. (f has_derivative (\<lambda>x. x *\<^sub>R d)) F)"
   by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR)
 
-lemma differentiable_within_open: (* TODO: delete *)
-  assumes "a \<in> s"
-    and "open s"
-  shows "f differentiable (at a within s) \<longleftrightarrow> f differentiable (at a)"
-  using assms
-  by (simp only: at_within_interior interior_open)
-
 lemma differentiable_on_eq_differentiable_at:
   "open s \<Longrightarrow> f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable at x)"
   unfolding differentiable_on_def
@@ -207,6 +200,15 @@
 lemma differentiable_on_id [simp, derivative_intros]: "id differentiable_on S"
   by (simp add: id_def)
 
+lemma differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. c) differentiable_on S"
+  by (simp add: differentiable_on_def)
+
+lemma differentiable_on_mult [simp, derivative_intros]:
+  fixes f :: "'M::real_normed_vector \<Rightarrow> 'a::real_normed_algebra"
+  shows "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) differentiable_on S"
+  apply (simp add: differentiable_on_def differentiable_def)
+  using differentiable_def differentiable_mult by blast
+
 lemma differentiable_on_compose:
    "\<lbrakk>g differentiable_on S; f differentiable_on (g ` S)\<rbrakk> \<Longrightarrow> (\<lambda>x. f (g x)) differentiable_on S"
 by (simp add: differentiable_in_compose differentiable_on_def)
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -22,12 +22,6 @@
     by simp
 qed
 
-abbreviation lebesgue :: "'a::euclidean_space measure"
-  where "lebesgue \<equiv> completion lborel"
-
-abbreviation lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
-  where "lebesgue_on \<Omega> \<equiv> restrict_space (completion lborel) \<Omega>"
-
 lemma has_integral_implies_lebesgue_measurable_cbox:
   fixes f :: "'a :: euclidean_space \<Rightarrow> real"
   assumes f: "(f has_integral I) (cbox x y)"
@@ -779,19 +773,6 @@
 
 end
 
-lemma measurable_completion: "f \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> f \<in> completion M \<rightarrow>\<^sub>M N"
-  by (auto simp: measurable_def)
-
-lemma integrable_completion:
-  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
-  shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> integrable (completion M) f \<longleftrightarrow> integrable M f"
-  by (rule integrable_subalgebra[symmetric]) auto
-
-lemma integral_completion:
-  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
-  assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" shows "integral\<^sup>L (completion M) f = integral\<^sup>L M f"
-  by (rule integral_subalgebra[symmetric]) (auto intro: f)
-
 context
 begin
 
@@ -917,6 +898,11 @@
   shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f"
   using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def)
 
+lemma lmeasurable_iff_has_integral:
+  "S \<in> lmeasurable \<longleftrightarrow> ((indicator S) has_integral measure lebesgue S) UNIV"
+  by (subst has_integral_iff_nn_integral_lebesgue)
+     (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI)
+
 abbreviation
   absolutely_integrable_on :: "('a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
   (infixr "absolutely'_integrable'_on" 46)
@@ -948,11 +934,205 @@
   qed
 qed
 
+lemma absolutely_integrable_on_iff_nonneg:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> real"
+  assumes "\<And>x. 0 \<le> f x" shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s"
+proof -
+  from assms have "(\<lambda>x. \<bar>f x\<bar>) = f"
+    by (intro ext) auto
+  then show ?thesis
+    unfolding absolutely_integrable_on_def by simp
+qed
+
 lemma absolutely_integrable_onI:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f integrable_on s \<Longrightarrow> (\<lambda>x. norm (f x)) integrable_on s \<Longrightarrow> f absolutely_integrable_on s"
   unfolding absolutely_integrable_on_def by auto
 
+lemma lmeasurable_iff_integrable_on: "S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. 1::real) integrable_on S"
+  by (subst absolutely_integrable_on_iff_nonneg[symmetric])
+     (simp_all add: lmeasurable_iff_integrable)
+
+lemma lmeasure_integral_UNIV: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral UNIV (indicator S)"
+  by (simp add: lmeasurable_iff_has_integral integral_unique)
+
+lemma lmeasure_integral: "S \<in> lmeasurable \<Longrightarrow> measure lebesgue S = integral S (\<lambda>x. 1::real)"
+  by (auto simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on)
+
+lemma
+  assumes \<D>: "\<D> division_of S"
+  shows lmeasurable_division: "S \<in> lmeasurable" (is ?l)
+    and content_divsion: "(\<Sum>k\<in>\<D>. measure lebesgue k) = measure lebesgue S" (is ?m)
+proof -
+  { fix d1 d2 assume *: "d1 \<in> \<D>" "d2 \<in> \<D>" "d1 \<noteq> d2"
+    then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d"
+      using division_ofD(4)[OF \<D>] by blast
+    with division_ofD(5)[OF \<D> *]
+    have "d1 \<in> sets lborel" "d2 \<in> sets lborel" "d1 \<inter> d2 \<subseteq> (cbox a b - box a b) \<union> (cbox c d - box c d)"
+      by auto
+    moreover have "(cbox a b - box a b) \<union> (cbox c d - box c d) \<in> null_sets lborel"
+      by (intro null_sets.Un null_sets_cbox_Diff_box)
+    ultimately have "d1 \<inter> d2 \<in> null_sets lborel"
+      by (blast intro: null_sets_subset) }
+  then show ?l ?m
+    unfolding division_ofD(6)[OF \<D>, symmetric]
+    using division_ofD(1,4)[OF \<D>]
+    by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def)
+qed
+
+text \<open>This should be an abbreviation for negligible.\<close>
+lemma negligible_iff_null_sets: "negligible S \<longleftrightarrow> S \<in> null_sets lebesgue"
+proof
+  assume "negligible S"
+  then have "(indicator S has_integral (0::real)) UNIV"
+    by (auto simp: negligible)
+  then show "S \<in> null_sets lebesgue"
+    by (subst (asm) has_integral_iff_nn_integral_lebesgue)
+        (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff)
+next
+  assume S: "S \<in> null_sets lebesgue"
+  show "negligible S"
+    unfolding negligible_def
+  proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2]
+                      has_integral_restrict_univ[where s="cbox _ _", THEN iffD1])
+    fix a b
+    show "(\<lambda>x. if x \<in> cbox a b then indicator S x else 0) \<in> lebesgue \<rightarrow>\<^sub>M borel"
+      using S by (auto intro!: measurable_If)
+    then show "(\<integral>\<^sup>+ x. ennreal (if x \<in> cbox a b then indicator S x else 0) \<partial>lebesgue) = ennreal 0"
+      using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2])
+  qed auto
+qed
+
+lemma starlike_negligible:
+  assumes "closed S"
+      and eq1: "\<And>c x. \<lbrakk>(a + c *\<^sub>R x) \<in> S; 0 \<le> c; a + x \<in> S\<rbrakk> \<Longrightarrow> c = 1"
+    shows "negligible S"
+proof -
+  have "negligible (op + (-a) ` S)"
+  proof (subst negligible_on_intervals, intro allI)
+    fix u v
+    show "negligible (op + (- a) ` S \<inter> cbox u v)"
+      unfolding negligible_iff_null_sets
+      apply (rule starlike_negligible_compact)
+       apply (simp add: assms closed_translation closed_Int_compact, clarify)
+      by (metis eq1 minus_add_cancel)
+  qed
+  then show ?thesis
+    by (rule negligible_translation_rev)
+qed
+
+lemma starlike_negligible_strong:
+  assumes "closed S"
+      and star: "\<And>c x. \<lbrakk>0 \<le> c; c < 1; a+x \<in> S\<rbrakk> \<Longrightarrow> a + c *\<^sub>R x \<notin> S"
+    shows "negligible S"
+proof -
+  show ?thesis
+  proof (rule starlike_negligible [OF \<open>closed S\<close>, of a])
+    fix c x
+    assume cx: "a + c *\<^sub>R x \<in> S" "0 \<le> c" "a + x \<in> S"
+    with star have "~ (c < 1)" by auto
+    moreover have "~ (c > 1)"
+      using star [of "1/c" "c *\<^sub>R x"] cx by force
+    ultimately show "c = 1" by arith
+  qed
+qed
+
+subsection\<open>Applications\<close>
+
+lemma negligible_hyperplane:
+  assumes "a \<noteq> 0 \<or> b \<noteq> 0" shows "negligible {x. a \<bullet> x = b}"
+proof -
+  obtain x where x: "a \<bullet> x \<noteq> b"
+    using assms
+    apply auto
+     apply (metis inner_eq_zero_iff inner_zero_right)
+    using inner_zero_right by fastforce
+  show ?thesis
+    apply (rule starlike_negligible [OF closed_hyperplane, of x])
+    using x apply (auto simp: algebra_simps)
+    done
+qed
+
+lemma negligible_lowdim:
+  fixes S :: "'N :: euclidean_space set"
+  assumes "dim S < DIM('N)"
+    shows "negligible S"
+proof -
+  obtain a where "a \<noteq> 0" and a: "span S \<subseteq> {x. a \<bullet> x = 0}"
+    using lowdim_subset_hyperplane [OF assms] by blast
+  have "negligible (span S)"
+    using \<open>a \<noteq> 0\<close> a negligible_hyperplane by (blast intro: negligible_subset)
+  then show ?thesis
+    using span_inc by (blast intro: negligible_subset)
+qed
+
+proposition negligible_convex_frontier:
+  fixes S :: "'N :: euclidean_space set"
+  assumes "convex S"
+    shows "negligible(frontier S)"
+proof -
+  have nf: "negligible(frontier S)" if "convex S" "0 \<in> S" for S :: "'N set"
+  proof -
+    obtain B where "B \<subseteq> S" and indB: "independent B"
+               and spanB: "S \<subseteq> span B" and cardB: "card B = dim S"
+      by (metis basis_exists)
+    consider "dim S < DIM('N)" | "dim S = DIM('N)"
+      using dim_subset_UNIV le_eq_less_or_eq by blast
+    then show ?thesis
+    proof cases
+      case 1
+      show ?thesis
+        by (rule negligible_subset [of "closure S"])
+           (simp_all add: Diff_subset frontier_def negligible_lowdim 1)
+    next
+      case 2
+      obtain a where a: "a \<in> interior S"
+        apply (rule interior_simplex_nonempty [OF indB])
+          apply (simp add: indB independent_finite)
+         apply (simp add: cardB 2)
+        apply (metis \<open>B \<subseteq> S\<close> \<open>0 \<in> S\<close> \<open>convex S\<close> insert_absorb insert_subset interior_mono subset_hull)
+        done
+      show ?thesis
+      proof (rule starlike_negligible_strong [where a=a])
+        fix c::real and x
+        have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)"
+          by (simp add: algebra_simps)
+        assume "0 \<le> c" "c < 1" "a + x \<in> frontier S"
+        then show "a + c *\<^sub>R x \<notin> frontier S"
+          apply (clarsimp simp: frontier_def)
+          apply (subst eq)
+          apply (rule mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a, of _ "1-c"], auto)
+          done
+      qed auto
+    qed
+  qed
+  show ?thesis
+  proof (cases "S = {}")
+    case True then show ?thesis by auto
+  next
+    case False
+    then obtain a where "a \<in> S" by auto
+    show ?thesis
+      using nf [of "(\<lambda>x. -a + x) ` S"]
+      by (metis \<open>a \<in> S\<close> add.left_inverse assms convex_translation_eq frontier_translation
+                image_eqI negligible_translation_rev)
+  qed
+qed
+
+corollary negligible_sphere: "negligible (sphere a e)"
+  using frontier_cball negligible_convex_frontier convex_cball
+  by (blast intro: negligible_subset)
+
+
+lemma non_negligible_UNIV [simp]: "\<not> negligible UNIV"
+  unfolding negligible_iff_null_sets by (auto simp: null_sets_def emeasure_lborel_UNIV)
+
+lemma negligible_interval:
+  "negligible (cbox a b) \<longleftrightarrow> box a b = {}" "negligible (box a b) \<longleftrightarrow> box a b = {}"
+   by (auto simp: negligible_iff_null_sets null_sets_def setprod_nonneg inner_diff_left box_eq_empty
+                  not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq
+            intro: eq_refl antisym less_imp_le)
+
 lemma set_integral_norm_bound:
   fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
   shows "set_integrable M k f \<Longrightarrow> norm (LINT x:k|M. f x) \<le> LINT x:k|M. norm (f x)"
@@ -1209,8 +1389,8 @@
       have sum_p': "(\<Sum>(x, k)\<in>p'. norm (integral k f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))"
         apply (subst setsum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
         unfolding norm_eq_zero
-        apply (rule integral_null)
-        apply assumption
+         apply (rule integral_null)
+        apply (simp add: content_eq_0_interior)
         apply rule
         done
       note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]]
@@ -1657,7 +1837,7 @@
           show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
             using partial_division_of_tagged_division[of p "cbox a b"] p(1)
             apply (subst setsum.over_tagged_division_lemma[OF p(1)])
-            apply (simp add: integral_null)
+            apply (simp add: content_eq_0_interior)
             apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
             apply (auto simp: tagged_partial_division_of_def)
             done
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -6,179 +6,49 @@
 
 theory Henstock_Kurzweil_Integration
 imports
-  Lebesgue_Measure
+  Lebesgue_Measure Tagged_Division
 begin
 
-lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
-  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
-  scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
-
-
-subsection \<open>Sundries\<close>
-
-lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
-lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto
-lemma conjunctD4: assumes "a \<and> b \<and> c \<and> d" shows a b c d using assms by auto
-
-lemma cond_cases: "(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)"
+(* BEGIN MOVE *)
+lemma swap_continuous:
+  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
+    shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)"
+proof -
+  have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap"
+    by auto
+  then show ?thesis
+    apply (rule ssubst)
+    apply (rule continuous_on_compose)
+    apply (simp add: split_def)
+    apply (rule continuous_intros | simp add: assms)+
+    done
+qed
+
+
+lemma norm_minus2: "norm (x1-x2, y1-y2) = norm (x2-x1, y2-y1)"
+  by (simp add: norm_minus_eqI)
+
+lemma norm_diff2: "\<lbrakk>y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \<le> e1; norm(y2 - x2) \<le> e2\<rbrakk>
+  \<Longrightarrow> norm(y - x) \<le> e"
+  using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"]
+  by (simp add: add_diff_add)
+
+lemma setcomp_dot1: "{z. P (z \<bullet> (i,0))} = {(x,y). P(x \<bullet> i)}"
   by auto
 
-declare norm_triangle_ineq4[intro]
-
-lemma transitive_stepwise_le:
-  assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" and "\<And>n. R n (Suc n)"
-  shows "\<forall>n\<ge>m. R m n"
-proof (intro allI impI)
-  show "m \<le> n \<Longrightarrow> R m n" for n
-    by (induction rule: dec_induct)
-       (use assms in blast)+
-qed
-
-subsection \<open>Some useful lemmas about intervals.\<close>
+lemma setcomp_dot2: "{z. P (z \<bullet> (0,i))} = {(x,y). P(y \<bullet> i)}"
+  by auto
+
+lemma Sigma_Int_Paircomp1: "(Sigma A B) \<inter> {(x, y). P x} = Sigma (A \<inter> {x. P x}) B"
+  by blast
+
+lemma Sigma_Int_Paircomp2: "(Sigma A B) \<inter> {(x, y). P y} = Sigma A (\<lambda>z. B z \<inter> {y. P y})"
+  by blast
 
 lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)"
   using nonempty_Basis
   by (fastforce simp add: set_eq_iff mem_box)
-
-lemma interior_subset_union_intervals:
-  assumes "i = cbox a b"
-    and "j = cbox c d"
-    and "interior j \<noteq> {}"
-    and "i \<subseteq> j \<union> s"
-    and "interior i \<inter> interior j = {}"
-  shows "interior i \<subseteq> interior s"
-proof -
-  have "box a b \<inter> cbox c d = {}"
-     using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
-     unfolding assms(1,2) interior_cbox by auto
-  moreover
-  have "box a b \<subseteq> cbox c d \<union> s"
-    apply (rule order_trans,rule box_subset_cbox)
-    using assms(4) unfolding assms(1,2)
-    apply auto
-    done
-  ultimately
-  show ?thesis
-    unfolding assms interior_cbox
-      by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
-qed
-
-lemma interior_Union_subset_cbox:
-  assumes "finite f"
-  assumes f: "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a b" "\<And>s. s \<in> f \<Longrightarrow> interior s \<subseteq> t"
-    and t: "closed t"
-  shows "interior (\<Union>f) \<subseteq> t"
-proof -
-  have [simp]: "s \<in> f \<Longrightarrow> closed s" for s
-    using f by auto
-  define E where "E = {s\<in>f. interior s = {}}"
-  then have "finite E" "E \<subseteq> {s\<in>f. interior s = {}}"
-    using \<open>finite f\<close> by auto
-  then have "interior (\<Union>f) = interior (\<Union>(f - E))"
-  proof (induction E rule: finite_subset_induct')
-    case (insert s f')
-    have "interior (\<Union>(f - insert s f') \<union> s) = interior (\<Union>(f - insert s f'))"
-      using insert.hyps \<open>finite f\<close> by (intro interior_closed_Un_empty_interior) auto
-    also have "\<Union>(f - insert s f') \<union> s = \<Union>(f - f')"
-      using insert.hyps by auto
-    finally show ?case
-      by (simp add: insert.IH)
-  qed simp
-  also have "\<dots> \<subseteq> \<Union>(f - E)"
-    by (rule interior_subset)
-  also have "\<dots> \<subseteq> t"
-  proof (rule Union_least)
-    fix s assume "s \<in> f - E"
-    with f[of s] obtain a b where s: "s \<in> f" "s = cbox a b" "box a b \<noteq> {}"
-      by (fastforce simp: E_def)
-    have "closure (interior s) \<subseteq> closure t"
-      by (intro closure_mono f \<open>s \<in> f\<close>)
-    with s \<open>closed t\<close> show "s \<subseteq> t"
-      by (simp add: closure_box)
-  qed
-  finally show ?thesis .
-qed
-
-lemma inter_interior_unions_intervals:
-    "finite f \<Longrightarrow> open s \<Longrightarrow> \<forall>t\<in>f. \<exists>a b. t = cbox a b \<Longrightarrow> \<forall>t\<in>f. s \<inter> (interior t) = {} \<Longrightarrow> s \<inter> interior (\<Union>f) = {}"
-  using interior_Union_subset_cbox[of f "UNIV - s"] by auto
-
-lemma interval_split:
-  fixes a :: "'a::euclidean_space"
-  assumes "k \<in> Basis"
-  shows
-    "cbox a b \<inter> {x. x\<bullet>k \<le> c} = cbox a (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)"
-    "cbox a b \<inter> {x. x\<bullet>k \<ge> c} = cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) b"
-  apply (rule_tac[!] set_eqI)
-  unfolding Int_iff mem_box mem_Collect_eq
-  using assms
-  apply auto
-  done
-
-lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}"
-  by (simp add: box_ne_empty)
-
-subsection \<open>Bounds on intervals where they exist.\<close>
-
-definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
-  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
-
-definition interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
-  where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
-
-lemma interval_upperbound[simp]:
-  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
-    interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
-  unfolding interval_upperbound_def euclidean_representation_setsum cbox_def
-  by (safe intro!: cSup_eq) auto
-
-lemma interval_lowerbound[simp]:
-  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
-    interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
-  unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def
-  by (safe intro!: cInf_eq) auto
-
-lemmas interval_bounds = interval_upperbound interval_lowerbound
-
-lemma
-  fixes X::"real set"
-  shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
-    and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
-  by (auto simp: interval_upperbound_def interval_lowerbound_def)
-
-lemma interval_bounds'[simp]:
-  assumes "cbox a b \<noteq> {}"
-  shows "interval_upperbound (cbox a b) = b"
-    and "interval_lowerbound (cbox a b) = a"
-  using assms unfolding box_ne_empty by auto
-
-lemma interval_upperbound_Times:
-  assumes "A \<noteq> {}" and "B \<noteq> {}"
-  shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)"
-proof-
-  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:A. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
-  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
-  ultimately show ?thesis unfolding interval_upperbound_def
-      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
-qed
-
-lemma interval_lowerbound_Times:
-  assumes "A \<noteq> {}" and "B \<noteq> {}"
-  shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)"
-proof-
-  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:A. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
-  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
-  ultimately show ?thesis unfolding interval_lowerbound_def
-      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
-qed
+(* END MOVE *)
 
 subsection \<open>Content (length, area, volume...) of an interval.\<close>
 
@@ -300,1508 +170,34 @@
     by (auto simp: not_le)
 qed
 
-subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
-
-definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))"
-
-lemma gaugeI:
-  assumes "\<And>x. x \<in> g x"
-    and "\<And>x. open (g x)"
-  shows "gauge g"
-  using assms unfolding gauge_def by auto
-
-lemma gaugeD[dest]:
-  assumes "gauge d"
-  shows "x \<in> d x"
-    and "open (d x)"
-  using assms unfolding gauge_def by auto
-
-lemma gauge_ball_dependent: "\<forall>x. 0 < e x \<Longrightarrow> gauge (\<lambda>x. ball x (e x))"
-  unfolding gauge_def by auto
-
-lemma gauge_ball[intro]: "0 < e \<Longrightarrow> gauge (\<lambda>x. ball x e)"
-  unfolding gauge_def by auto
-
-lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)"
-  by (rule gauge_ball) auto
-
-lemma gauge_inter[intro]: "gauge d1 \<Longrightarrow> gauge d2 \<Longrightarrow> gauge (\<lambda>x. d1 x \<inter> d2 x)"
-  unfolding gauge_def by auto
-
-lemma gauge_inters:
-  assumes "finite s"
-    and "\<forall>d\<in>s. gauge (f d)"
-  shows "gauge (\<lambda>x. \<Inter>{f d x | d. d \<in> s})"
-proof -
-  have *: "\<And>x. {f d x |d. d \<in> s} = (\<lambda>d. f d x) ` s"
-    by auto
-  show ?thesis
-    unfolding gauge_def unfolding *
-    using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
-qed
-
-lemma gauge_existence_lemma:
-  "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)"
-  by (metis zero_less_one)
-
-
-subsection \<open>Divisions.\<close>
-
-definition division_of (infixl "division'_of" 40)
-where
-  "s division_of i \<longleftrightarrow>
-    finite s \<and>
-    (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = cbox a b)) \<and>
-    (\<forall>k1\<in>s. \<forall>k2\<in>s. k1 \<noteq> k2 \<longrightarrow> interior(k1) \<inter> interior(k2) = {}) \<and>
-    (\<Union>s = i)"
-
-lemma division_ofD[dest]:
-  assumes "s division_of i"
-  shows "finite s"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
-    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}"
-    and "\<Union>s = i"
-  using assms unfolding division_of_def by auto
-
-lemma division_ofI:
-  assumes "finite s"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
-    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-    and "\<Union>s = i"
-  shows "s division_of i"
-  using assms unfolding division_of_def by auto
-
-lemma division_of_finite: "s division_of i \<Longrightarrow> finite s"
-  unfolding division_of_def by auto
-
-lemma division_of_self[intro]: "cbox a b \<noteq> {} \<Longrightarrow> {cbox a b} division_of (cbox a b)"
-  unfolding division_of_def by auto
-
-lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}"
-  unfolding division_of_def by auto
-
-lemma division_of_sing[simp]:
-  "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}"
-  (is "?l = ?r")
-proof
-  assume ?r
-  moreover
-  { fix k
-    assume "s = {{a}}" "k\<in>s"
-    then have "\<exists>x y. k = cbox x y"
-      apply (rule_tac x=a in exI)+
-      apply (force simp: cbox_sing)
-      done
-  }
-  ultimately show ?l
-    unfolding division_of_def cbox_sing by auto
-next
-  assume ?l
-  note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]]
-  {
-    fix x
-    assume x: "x \<in> s" have "x = {a}"
-      using *(2)[rule_format,OF x] by auto
-  }
-  moreover have "s \<noteq> {}"
-    using *(4) by auto
-  ultimately show ?r
-    unfolding cbox_sing by auto
-qed
-
-lemma elementary_empty: obtains p where "p division_of {}"
-  unfolding division_of_trivial by auto
-
-lemma elementary_interval: obtains p where "p division_of (cbox a b)"
-  by (metis division_of_trivial division_of_self)
-
-lemma division_contains: "s division_of i \<Longrightarrow> \<forall>x\<in>i. \<exists>k\<in>s. x \<in> k"
-  unfolding division_of_def by auto
-
-lemma forall_in_division:
-  "d division_of i \<Longrightarrow> (\<forall>x\<in>d. P x) \<longleftrightarrow> (\<forall>a b. cbox a b \<in> d \<longrightarrow> P (cbox a b))"
-  unfolding division_of_def by fastforce
-
-lemma division_of_subset:
-  assumes "p division_of (\<Union>p)"
-    and "q \<subseteq> p"
-  shows "q division_of (\<Union>q)"
-proof (rule division_ofI)
-  note * = division_ofD[OF assms(1)]
-  show "finite q"
-    using "*"(1) assms(2) infinite_super by auto
-  {
-    fix k
-    assume "k \<in> q"
-    then have kp: "k \<in> p"
-      using assms(2) by auto
-    show "k \<subseteq> \<Union>q"
-      using \<open>k \<in> q\<close> by auto
-    show "\<exists>a b. k = cbox a b"
-      using *(4)[OF kp] by auto
-    show "k \<noteq> {}"
-      using *(3)[OF kp] by auto
-  }
-  fix k1 k2
-  assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2"
-  then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2"
-    using assms(2) by auto
-  show "interior k1 \<inter> interior k2 = {}"
-    using *(5)[OF **] by auto
-qed auto
-
-lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)"
-  unfolding division_of_def by auto
-
 lemma division_of_content_0:
   assumes "content (cbox a b) = 0" "d division_of (cbox a b)"
   shows "\<forall>k\<in>d. content k = 0"
   unfolding forall_in_division[OF assms(2)]
   by (metis antisym_conv assms content_pos_le content_subset division_ofD(2))
 
-lemma division_inter:
-  fixes s1 s2 :: "'a::euclidean_space set"
-  assumes "p1 division_of s1"
-    and "p2 division_of s2"
-  shows "{k1 \<inter> k2 | k1 k2. k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)"
-  (is "?A' division_of _")
-proof -
-  let ?A = "{s. s \<in>  (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}"
-  have *: "?A' = ?A" by auto
-  show ?thesis
-    unfolding *
-  proof (rule division_ofI)
-    have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)"
-      by auto
-    moreover have "finite (p1 \<times> p2)"
-      using assms unfolding division_of_def by auto
-    ultimately show "finite ?A" by auto
-    have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s"
-      by auto
-    show "\<Union>?A = s1 \<inter> s2"
-      apply (rule set_eqI)
-      unfolding * and UN_iff
-      using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)]
-      apply auto
-      done
-    {
-      fix k
-      assume "k \<in> ?A"
-      then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}"
-        by auto
-      then show "k \<noteq> {}"
-        by auto
-      show "k \<subseteq> s1 \<inter> s2"
-        using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)]
-        unfolding k by auto
-      obtain a1 b1 where k1: "k1 = cbox a1 b1"
-        using division_ofD(4)[OF assms(1) k(2)] by blast
-      obtain a2 b2 where k2: "k2 = cbox a2 b2"
-        using division_ofD(4)[OF assms(2) k(3)] by blast
-      show "\<exists>a b. k = cbox a b"
-        unfolding k k1 k2 unfolding Int_interval by auto
-    }
-    fix k1 k2
-    assume "k1 \<in> ?A"
-    then obtain x1 y1 where k1: "k1 = x1 \<inter> y1" "x1 \<in> p1" "y1 \<in> p2" "k1 \<noteq> {}"
-      by auto
-    assume "k2 \<in> ?A"
-    then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}"
-      by auto
-    assume "k1 \<noteq> k2"
-    then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2"
-      unfolding k1 k2 by auto
-    have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow>
-      interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow>
-      interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow>
-      interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto
-    show "interior k1 \<inter> interior k2 = {}"
-      unfolding k1 k2
-      apply (rule *)
-      using assms division_ofD(5) k1 k2(2) k2(3) th apply auto
-      done
-  qed
-qed
-
-lemma division_inter_1:
-  assumes "d division_of i"
-    and "cbox a (b::'a::euclidean_space) \<subseteq> i"
-  shows "{cbox a b \<inter> k | k. k \<in> d \<and> cbox a b \<inter> k \<noteq> {}} division_of (cbox a b)"
-proof (cases "cbox a b = {}")
-  case True
-  show ?thesis
-    unfolding True and division_of_trivial by auto
-next
-  case False
-  have *: "cbox a b \<inter> i = cbox a b" using assms(2) by auto
-  show ?thesis
-    using division_inter[OF division_of_self[OF False] assms(1)]
-    unfolding * by auto
-qed
-
-lemma elementary_inter:
-  fixes s t :: "'a::euclidean_space set"
-  assumes "p1 division_of s"
-    and "p2 division_of t"
-  shows "\<exists>p. p division_of (s \<inter> t)"
-using assms division_inter by blast
-
-lemma elementary_inters:
-  assumes "finite f"
-    and "f \<noteq> {}"
-    and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)"
-  shows "\<exists>p. p division_of (\<Inter>f)"
-  using assms
-proof (induct f rule: finite_induct)
-  case (insert x f)
-  show ?case
-  proof (cases "f = {}")
-    case True
-    then show ?thesis
-      unfolding True using insert by auto
-  next
-    case False
-    obtain p where "p division_of \<Inter>f"
-      using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
-    moreover obtain px where "px division_of x"
-      using insert(5)[rule_format,OF insertI1] ..
-    ultimately show ?thesis
-      by (simp add: elementary_inter Inter_insert)
-  qed
-qed auto
-
-lemma division_disjoint_union:
-  assumes "p1 division_of s1"
-    and "p2 division_of s2"
-    and "interior s1 \<inter> interior s2 = {}"
-  shows "(p1 \<union> p2) division_of (s1 \<union> s2)"
-proof (rule division_ofI)
-  note d1 = division_ofD[OF assms(1)]
-  note d2 = division_ofD[OF assms(2)]
-  show "finite (p1 \<union> p2)"
-    using d1(1) d2(1) by auto
-  show "\<Union>(p1 \<union> p2) = s1 \<union> s2"
-    using d1(6) d2(6) by auto
-  {
-    fix k1 k2
-    assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
-    moreover
-    let ?g="interior k1 \<inter> interior k2 = {}"
-    {
-      assume as: "k1\<in>p1" "k2\<in>p2"
-      have ?g
-        using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]]
-        using assms(3) by blast
-    }
-    moreover
-    {
-      assume as: "k1\<in>p2" "k2\<in>p1"
-      have ?g
-        using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]]
-        using assms(3) by blast
-    }
-    ultimately show ?g
-      using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto
-  }
-  fix k
-  assume k: "k \<in> p1 \<union> p2"
-  show "k \<subseteq> s1 \<union> s2"
-    using k d1(2) d2(2) by auto
-  show "k \<noteq> {}"
-    using k d1(3) d2(3) by auto
-  show "\<exists>a b. k = cbox a b"
-    using k d1(4) d2(4) by auto
-qed
-
-lemma partial_division_extend_1:
-  fixes a b c d :: "'a::euclidean_space"
-  assumes incl: "cbox c d \<subseteq> cbox a b"
-    and nonempty: "cbox c d \<noteq> {}"
-  obtains p where "p division_of (cbox a b)" "cbox c d \<in> p"
-proof
-  let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a.
-    cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)"
-  define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
-
-  show "cbox c d \<in> p"
-    unfolding p_def
-    by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
-  {
-    fix i :: 'a
-    assume "i \<in> Basis"
-    with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i"
-      unfolding box_eq_empty subset_box by (auto simp: not_le)
-  }
-  note ord = this
-
-  show "p division_of (cbox a b)"
-  proof (rule division_ofI)
-    show "finite p"
-      unfolding p_def by (auto intro!: finite_PiE)
-    {
-      fix k
-      assume "k \<in> p"
-      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
-        by (auto simp: p_def)
-      then show "\<exists>a b. k = cbox a b"
-        by auto
-      have "k \<subseteq> cbox a b \<and> k \<noteq> {}"
-      proof (simp add: k box_eq_empty subset_box not_less, safe)
-        fix i :: 'a
-        assume i: "i \<in> Basis"
-        with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
-          by (auto simp: PiE_iff)
-        with i ord[of i]
-        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
-          by auto
-      qed
-      then show "k \<noteq> {}" "k \<subseteq> cbox a b"
-        by auto
-      {
-        fix l
-        assume "l \<in> p"
-        then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
-          by (auto simp: p_def)
-        assume "l \<noteq> k"
-        have "\<exists>i\<in>Basis. f i \<noteq> g i"
-        proof (rule ccontr)
-          assume "\<not> ?thesis"
-          with f g have "f = g"
-            by (auto simp: PiE_iff extensional_def intro!: ext)
-          with \<open>l \<noteq> k\<close> show False
-            by (simp add: l k)
-        qed
-        then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
-        then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
-                  "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
-          using f g by (auto simp: PiE_iff)
-        with * ord[of i] show "interior l \<inter> interior k = {}"
-          by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i])
-      }
-      note \<open>k \<subseteq> cbox a b\<close>
-    }
-    moreover
-    {
-      fix x assume x: "x \<in> cbox a b"
-      have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
-      proof
-        fix i :: 'a
-        assume "i \<in> Basis"
-        with x ord[of i]
-        have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
-            (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
-          by (auto simp: cbox_def)
-        then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
-          by auto
-      qed
-      then obtain f where
-        f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}"
-        unfolding bchoice_iff ..
-      moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
-        by auto
-      moreover from f have "x \<in> ?B (restrict f Basis)"
-        by (auto simp: mem_box)
-      ultimately have "\<exists>k\<in>p. x \<in> k"
-        unfolding p_def by blast
-    }
-    ultimately show "\<Union>p = cbox a b"
-      by auto
-  qed
-qed
-
-lemma partial_division_extend_interval:
-  assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b"
-  obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
-proof (cases "p = {}")
-  case True
-  obtain q where "q division_of (cbox a b)"
-    by (rule elementary_interval)
-  then show ?thesis
-    using True that by blast
-next
-  case False
-  note p = division_ofD[OF assms(1)]
-  have div_cbox: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
-  proof
-    fix k
-    assume kp: "k \<in> p"
-    obtain c d where k: "k = cbox c d"
-      using p(4)[OF kp] by blast
-    have *: "cbox c d \<subseteq> cbox a b" "cbox c d \<noteq> {}"
-      using p(2,3)[OF kp, unfolded k] using assms(2)
-      by (blast intro: order.trans)+
-    obtain q where "q division_of cbox a b" "cbox c d \<in> q"
-      by (rule partial_division_extend_1[OF *])
-    then show "\<exists>q. q division_of cbox a b \<and> k \<in> q"
-      unfolding k by auto
-  qed
-  obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
-    using bchoice[OF div_cbox] by blast
-  { fix x
-    assume x: "x \<in> p"
-    have "q x division_of \<Union>q x"
-      apply (rule division_ofI)
-      using division_ofD[OF q(1)[OF x]]
-      apply auto
-      done }
-  then have "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
-    by (meson Diff_subset division_of_subset)
-  then have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
-    apply -
-    apply (rule elementary_inters [OF finite_imageI[OF p(1)]])
-    apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]])
-    done
-  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" ..
-  have "d \<union> p division_of cbox a b"
-  proof -
-    have te: "\<And>s f t. s \<noteq> {} \<Longrightarrow> \<forall>i\<in>s. f i \<union> i = t \<Longrightarrow> t = \<Inter>(f ` s) \<union> \<Union>s" by auto
-    have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
-    proof (rule te[OF False], clarify)
-      fix i
-      assume i: "i \<in> p"
-      show "\<Union>(q i - {i}) \<union> i = cbox a b"
-        using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto
-    qed
-    { fix k
-      assume k: "k \<in> p"
-      have *: "\<And>u t s. t \<inter> s = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<inter> t = {}"
-        by auto
-      have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior k = {}"
-      proof (rule *[OF inter_interior_unions_intervals])
-        note qk=division_ofD[OF q(1)[OF k]]
-        show "finite (q k - {k})" "open (interior k)" "\<forall>t\<in>q k - {k}. \<exists>a b. t = cbox a b"
-          using qk by auto
-        show "\<forall>t\<in>q k - {k}. interior k \<inter> interior t = {}"
-          using qk(5) using q(2)[OF k] by auto
-        show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q k - {k}))"
-          apply (rule interior_mono)+
-          using k
-          apply auto
-          done
-      qed } note [simp] = this
-    show "d \<union> p division_of (cbox a b)"
-      unfolding cbox_eq
-      apply (rule division_disjoint_union[OF d assms(1)])
-      apply (rule inter_interior_unions_intervals)
-      apply (rule p open_interior ballI)+
-      apply simp_all
-      done
-  qed
-  then show ?thesis
-    by (meson Un_upper2 that)
-qed
-
-lemma elementary_bounded[dest]:
-  fixes s :: "'a::euclidean_space set"
-  shows "p division_of s \<Longrightarrow> bounded s"
-  unfolding division_of_def by (metis bounded_Union bounded_cbox)
-
-lemma elementary_subset_cbox:
-  "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> cbox a (b::'a::euclidean_space)"
-  by (meson elementary_bounded bounded_subset_cbox)
-
-lemma division_union_intervals_exists:
-  fixes a b :: "'a::euclidean_space"
-  assumes "cbox a b \<noteq> {}"
-  obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
-proof (cases "cbox c d = {}")
-  case True
-  show ?thesis
-    apply (rule that[of "{}"])
-    unfolding True
-    using assms
-    apply auto
-    done
-next
-  case False
-  show ?thesis
-  proof (cases "cbox a b \<inter> cbox c d = {}")
-    case True
-    then show ?thesis
-      by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty)
-  next
-    case False
-    obtain u v where uv: "cbox a b \<inter> cbox c d = cbox u v"
-      unfolding Int_interval by auto
-    have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto
-    obtain p where "p division_of cbox c d" "cbox u v \<in> p"
-      by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
-    note p = this division_ofD[OF this(1)]
-    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))"
-      apply (rule arg_cong[of _ _ interior])
-      using p(8) uv by auto
-    also have "\<dots> = {}"
-      unfolding interior_Int
-      apply (rule inter_interior_unions_intervals)
-      using p(6) p(7)[OF p(2)] p(3)
-      apply auto
-      done
-    finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp
-    have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(p - {cbox u v})"
-      using p(8) unfolding uv[symmetric] by auto
-    have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
-    proof -
-      have "{cbox a b} division_of cbox a b"
-        by (simp add: assms division_of_self)
-      then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
-        by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> division_disjoint_union division_of_subset insert_is_Un p(1) p(8))
-    qed
-    with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe)
-  qed
-qed
-
-lemma division_of_unions:
-  assumes "finite f"
-    and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)"
-    and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-  shows "\<Union>f division_of \<Union>\<Union>f"
-  using assms
-  by (auto intro!: division_ofI)
-
-lemma elementary_union_interval:
-  fixes a b :: "'a::euclidean_space"
-  assumes "p division_of \<Union>p"
-  obtains q where "q division_of (cbox a b \<union> \<Union>p)"
-proof -
-  note assm = division_ofD[OF assms]
-  have lem1: "\<And>f s. \<Union>\<Union>(f ` s) = \<Union>((\<lambda>x. \<Union>(f x)) ` s)"
-    by auto
-  have lem2: "\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f"
+lemma setsum_content_null:
+  assumes "content (cbox a b) = 0"
+    and "p tagged_division_of (cbox a b)"
+  shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
+proof (rule setsum.neutral, rule)
+  fix y
+  assume y: "y \<in> p"
+  obtain x k where xk: "y = (x, k)"
+    using surj_pair[of y] by blast
+  note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]]
+  from this(2) obtain c d where k: "k = cbox c d" by blast
+  have "(\<lambda>(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x"
+    unfolding xk by auto
+  also have "\<dots> = 0"
+    using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d]
+    unfolding assms(1) k
     by auto
-  {
-    presume "p = {} \<Longrightarrow> thesis"
-      "cbox a b = {} \<Longrightarrow> thesis"
-      "cbox a b \<noteq> {} \<Longrightarrow> interior (cbox a b) = {} \<Longrightarrow> thesis"
-      "p \<noteq> {} \<Longrightarrow> interior (cbox a b)\<noteq>{} \<Longrightarrow> cbox a b \<noteq> {} \<Longrightarrow> thesis"
-    then show thesis by auto
-  next
-    assume as: "p = {}"
-    obtain p where "p division_of (cbox a b)"
-      by (rule elementary_interval)
-    then show thesis
-      using as that by auto
-  next
-    assume as: "cbox a b = {}"
-    show thesis
-      using as assms that by auto
-  next
-    assume as: "interior (cbox a b) = {}" "cbox a b \<noteq> {}"
-    show thesis
-      apply (rule that[of "insert (cbox a b) p"],rule division_ofI)
-      unfolding finite_insert
-      apply (rule assm(1)) unfolding Union_insert
-      using assm(2-4) as
-      apply -
-      apply (fast dest: assm(5))+
-      done
-  next
-    assume as: "p \<noteq> {}" "interior (cbox a b) \<noteq> {}" "cbox a b \<noteq> {}"
-    have "\<forall>k\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
-    proof
-      fix k
-      assume kp: "k \<in> p"
-      from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast
-      then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
-        by (meson as(3) division_union_intervals_exists)
-    qed
-    from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" ..
-    note q = division_ofD[OF this[rule_format]]
-    let ?D = "\<Union>{insert (cbox a b) (q k) | k. k \<in> p}"
-    show thesis
-    proof (rule that[OF division_ofI])
-      have *: "{insert (cbox a b) (q k) |k. k \<in> p} = (\<lambda>k. insert (cbox a b) (q k)) ` p"
-        by auto
-      show "finite ?D"
-        using "*" assm(1) q(1) by auto
-      show "\<Union>?D = cbox a b \<union> \<Union>p"
-        unfolding * lem1
-        unfolding lem2[OF as(1), of "cbox a b", symmetric]
-        using q(6)
-        by auto
-      fix k
-      assume k: "k \<in> ?D"
-      then show "k \<subseteq> cbox a b \<union> \<Union>p"
-        using q(2) by auto
-      show "k \<noteq> {}"
-        using q(3) k by auto
-      show "\<exists>a b. k = cbox a b"
-        using q(4) k by auto
-      fix k'
-      assume k': "k' \<in> ?D" "k \<noteq> k'"
-      obtain x where x: "k \<in> insert (cbox a b) (q x)" "x\<in>p"
-        using k by auto
-      obtain x' where x': "k'\<in>insert (cbox a b) (q x')" "x'\<in>p"
-        using k' by auto
-      show "interior k \<inter> interior k' = {}"
-      proof (cases "x = x'")
-        case True
-        show ?thesis
-          using True k' q(5) x' x by auto
-      next
-        case False
-        {
-          presume "k = cbox a b \<Longrightarrow> ?thesis"
-            and "k' = cbox a b \<Longrightarrow> ?thesis"
-            and "k \<noteq> cbox a b \<Longrightarrow> k' \<noteq> cbox a b \<Longrightarrow> ?thesis"
-          then show ?thesis by linarith
-        next
-          assume as': "k  = cbox a b"
-          show ?thesis
-            using as' k' q(5) x' by blast
-        next
-          assume as': "k' = cbox a b"
-          show ?thesis
-            using as' k'(2) q(5) x by blast
-        }
-        assume as': "k \<noteq> cbox a b" "k' \<noteq> cbox a b"
-        obtain c d where k: "k = cbox c d"
-          using q(4)[OF x(2,1)] by blast
-        have "interior k \<inter> interior (cbox a b) = {}"
-          using as' k'(2) q(5) x by blast
-        then have "interior k \<subseteq> interior x"
-        using interior_subset_union_intervals
-          by (metis as(2) k q(2) x interior_subset_union_intervals)
-        moreover
-        obtain c d where c_d: "k' = cbox c d"
-          using q(4)[OF x'(2,1)] by blast
-        have "interior k' \<inter> interior (cbox a b) = {}"
-          using as'(2) q(5) x' by blast
-        then have "interior k' \<subseteq> interior x'"
-          by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2))
-        ultimately show ?thesis
-          using assm(5)[OF x(2) x'(2) False] by auto
-      qed
-    qed
-  }
-qed
-
-lemma elementary_unions_intervals:
-  assumes fin: "finite f"
-    and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
-  obtains p where "p division_of (\<Union>f)"
-proof -
-  have "\<exists>p. p division_of (\<Union>f)"
-  proof (induct_tac f rule:finite_subset_induct)
-    show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
-  next
-    fix x F
-    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
-    from this(3) obtain p where p: "p division_of \<Union>F" ..
-    from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast
-    have *: "\<Union>F = \<Union>p"
-      using division_ofD[OF p] by auto
-    show "\<exists>p. p division_of \<Union>insert x F"
-      using elementary_union_interval[OF p[unfolded *], of a b]
-      unfolding Union_insert x * by metis
-  qed (insert assms, auto)
-  then show ?thesis
-    using that by auto
-qed
-
-lemma elementary_union:
-  fixes s t :: "'a::euclidean_space set"
-  assumes "ps division_of s" "pt division_of t"
-  obtains p where "p division_of (s \<union> t)"
-proof -
-  have *: "s \<union> t = \<Union>ps \<union> \<Union>pt"
-    using assms unfolding division_of_def by auto
-  show ?thesis
-    apply (rule elementary_unions_intervals[of "ps \<union> pt"])
-    using assms apply auto
-    by (simp add: * that)
-qed
-
-lemma partial_division_extend:
-  fixes t :: "'a::euclidean_space set"
-  assumes "p division_of s"
-    and "q division_of t"
-    and "s \<subseteq> t"
-  obtains r where "p \<subseteq> r" and "r division_of t"
-proof -
-  note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
-  obtain a b where ab: "t \<subseteq> cbox a b"
-    using elementary_subset_cbox[OF assms(2)] by auto
-  obtain r1 where "p \<subseteq> r1" "r1 division_of (cbox a b)"
-    using assms
-    by (metis ab dual_order.trans partial_division_extend_interval divp(6))
-  note r1 = this division_ofD[OF this(2)]
-  obtain p' where "p' division_of \<Union>(r1 - p)"
-    apply (rule elementary_unions_intervals[of "r1 - p"])
-    using r1(3,6)
-    apply auto
-    done
-  then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
-    by (metis assms(2) divq(6) elementary_inter)
-  {
-    fix x
-    assume x: "x \<in> t" "x \<notin> s"
-    then have "x\<in>\<Union>r1"
-      unfolding r1 using ab by auto
-    then obtain r where r: "r \<in> r1" "x \<in> r"
-      unfolding Union_iff ..
-    moreover
-    have "r \<notin> p"
-    proof
-      assume "r \<in> p"
-      then have "x \<in> s" using divp(2) r by auto
-      then show False using x by auto
-    qed
-    ultimately have "x\<in>\<Union>(r1 - p)" by auto
-  }
-  then have *: "t = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
-    unfolding divp divq using assms(3) by auto
-  show ?thesis
-    apply (rule that[of "p \<union> r2"])
-    unfolding *
-    defer
-    apply (rule division_disjoint_union)
-    unfolding divp(6)
-    apply(rule assms r2)+
-  proof -
-    have "interior s \<inter> interior (\<Union>(r1-p)) = {}"
-    proof (rule inter_interior_unions_intervals)
-      show "finite (r1 - p)" and "open (interior s)" and "\<forall>t\<in>r1-p. \<exists>a b. t = cbox a b"
-        using r1 by auto
-      have *: "\<And>s. (\<And>x. x \<in> s \<Longrightarrow> False) \<Longrightarrow> s = {}"
-        by auto
-      show "\<forall>t\<in>r1-p. interior s \<inter> interior t = {}"
-      proof
-        fix m x
-        assume as: "m \<in> r1 - p"
-        have "interior m \<inter> interior (\<Union>p) = {}"
-        proof (rule inter_interior_unions_intervals)
-          show "finite p" and "open (interior m)" and "\<forall>t\<in>p. \<exists>a b. t = cbox a b"
-            using divp by auto
-          show "\<forall>t\<in>p. interior m \<inter> interior t = {}"
-            by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp)
-        qed
-        then show "interior s \<inter> interior m = {}"
-          unfolding divp by auto
-      qed
-    qed
-    then show "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}"
-      using interior_subset by auto
-  qed auto
-qed
-
-lemma division_split_left_inj:
-  fixes type :: "'a::euclidean_space"
-  assumes "d division_of i"
-    and "k1 \<in> d"
-    and "k2 \<in> d"
-    and "k1 \<noteq> k2"
-    and "k1 \<inter> {x::'a. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
-    and k: "k\<in>Basis"
-  shows "content(k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
-proof -
-  note d=division_ofD[OF assms(1)]
-  have *: "\<And>(a::'a) b c. content (cbox a b \<inter> {x. x\<bullet>k \<le> c}) = 0 \<longleftrightarrow>
-    interior(cbox a b \<inter> {x. x\<bullet>k \<le> c}) = {}"
-    unfolding  interval_split[OF k] content_eq_0_interior by auto
-  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
-  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
-  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
-    by auto
-  show ?thesis
-    unfolding uv1 uv2 *
-    apply (rule **[OF d(5)[OF assms(2-4)]])
-    apply (simp add: uv1)
-    using assms(5) uv1 by auto
-qed
-
-lemma division_split_right_inj:
-  fixes type :: "'a::euclidean_space"
-  assumes "d division_of i"
-    and "k1 \<in> d"
-    and "k2 \<in> d"
-    and "k1 \<noteq> k2"
-    and "k1 \<inter> {x::'a. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
-    and k: "k \<in> Basis"
-  shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
-proof -
-  note d=division_ofD[OF assms(1)]
-  have *: "\<And>a b::'a. \<And>c. content(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = 0 \<longleftrightarrow>
-    interior(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = {}"
-    unfolding interval_split[OF k] content_eq_0_interior by auto
-  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
-  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
-  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
-    by auto
-  show ?thesis
-    unfolding uv1 uv2 *
-    apply (rule **[OF d(5)[OF assms(2-4)]])
-    apply (simp add: uv1)
-    using assms(5) uv1 by auto
+  finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
 qed
 
-
-lemma division_split:
-  fixes a :: "'a::euclidean_space"
-  assumes "p division_of (cbox a b)"
-    and k: "k\<in>Basis"
-  shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})"
-      (is "?p1 division_of ?I1")
-    and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-      (is "?p2 division_of ?I2")
-proof (rule_tac[!] division_ofI)
-  note p = division_ofD[OF assms(1)]
-  show "finite ?p1" "finite ?p2"
-    using p(1) by auto
-  show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2"
-    unfolding p(6)[symmetric] by auto
-  {
-    fix k
-    assume "k \<in> ?p1"
-    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
-    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
-    show "k \<subseteq> ?I1"
-      using l p(2) uv by force
-    show  "k \<noteq> {}"
-      by (simp add: l)
-    show  "\<exists>a b. k = cbox a b"
-      apply (simp add: l uv p(2-3)[OF l(2)])
-      apply (subst interval_split[OF k])
-      apply (auto intro: order.trans)
-      done
-    fix k'
-    assume "k' \<in> ?p1"
-    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
-    assume "k \<noteq> k'"
-    then show "interior k \<inter> interior k' = {}"
-      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
-  }
-  {
-    fix k
-    assume "k \<in> ?p2"
-    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
-    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
-    show "k \<subseteq> ?I2"
-      using l p(2) uv by force
-    show  "k \<noteq> {}"
-      by (simp add: l)
-    show  "\<exists>a b. k = cbox a b"
-      apply (simp add: l uv p(2-3)[OF l(2)])
-      apply (subst interval_split[OF k])
-      apply (auto intro: order.trans)
-      done
-    fix k'
-    assume "k' \<in> ?p2"
-    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
-    assume "k \<noteq> k'"
-    then show "interior k \<inter> interior k' = {}"
-      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
-  }
-qed
-
-subsection \<open>Tagged (partial) divisions.\<close>
-
-definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
-  where "s tagged_partial_division_of i \<longleftrightarrow>
-    finite s \<and>
-    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
-    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
-      interior k1 \<inter> interior k2 = {})"
-
-lemma tagged_partial_division_ofD[dest]:
-  assumes "s tagged_partial_division_of i"
-  shows "finite s"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow>
-      (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-  using assms unfolding tagged_partial_division_of_def by blast+
-
-definition tagged_division_of (infixr "tagged'_division'_of" 40)
-  where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-
-lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
-  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
-
-lemma tagged_division_of:
-  "s tagged_division_of i \<longleftrightarrow>
-    finite s \<and>
-    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
-    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
-      interior k1 \<inter> interior k2 = {}) \<and>
-    (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
-
-lemma tagged_division_ofI:
-  assumes "finite s"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
-      interior k1 \<inter> interior k2 = {}"
-    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  shows "s tagged_division_of i"
-  unfolding tagged_division_of
-  using assms
-  apply auto
-  apply fastforce+
-  done
-
-lemma tagged_division_ofD[dest]:  (*FIXME USE A LOCALE*)
-  assumes "s tagged_division_of i"
-  shows "finite s"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>x1 k1 x2 k2. (x1, k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
-      interior k1 \<inter> interior k2 = {}"
-    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  using assms unfolding tagged_division_of by blast+
-
-lemma division_of_tagged_division:
-  assumes "s tagged_division_of i"
-  shows "(snd ` s) division_of i"
-proof (rule division_ofI)
-  note assm = tagged_division_ofD[OF assms]
-  show "\<Union>(snd ` s) = i" "finite (snd ` s)"
-    using assm by auto
-  fix k
-  assume k: "k \<in> snd ` s"
-  then obtain xk where xk: "(xk, k) \<in> s"
-    by auto
-  then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b"
-    using assm by fastforce+
-  fix k'
-  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
-  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
-    by auto
-  then show "interior k \<inter> interior k' = {}"
-    using assm(5) k'(2) xk by blast
-qed
-
-lemma partial_division_of_tagged_division:
-  assumes "s tagged_partial_division_of i"
-  shows "(snd ` s) division_of \<Union>(snd ` s)"
-proof (rule division_ofI)
-  note assm = tagged_partial_division_ofD[OF assms]
-  show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
-    using assm by auto
-  fix k
-  assume k: "k \<in> snd ` s"
-  then obtain xk where xk: "(xk, k) \<in> s"
-    by auto
-  then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)"
-    using assm by auto
-  fix k'
-  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
-  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
-    by auto
-  then show "interior k \<inter> interior k' = {}"
-    using assm(5) k'(2) xk by auto
-qed
-
-lemma tagged_partial_division_subset:
-  assumes "s tagged_partial_division_of i"
-    and "t \<subseteq> s"
-  shows "t tagged_partial_division_of i"
-  using assms
-  unfolding tagged_partial_division_of_def
-  using finite_subset[OF assms(2)]
-  by blast
-
-lemma (in comm_monoid_set) over_tagged_division_lemma:
-  assumes "p tagged_division_of i"
-    and "\<And>u v. cbox u v \<noteq> {} \<Longrightarrow> content (cbox u v) = 0 \<Longrightarrow> d (cbox u v) = \<^bold>1"
-  shows "F (\<lambda>(x,k). d k) p = F d (snd ` p)"
-proof -
-  have *: "(\<lambda>(x,k). d k) = d \<circ> snd"
-    unfolding o_def by (rule ext) auto
-  note assm = tagged_division_ofD[OF assms(1)]
-  show ?thesis
-    unfolding *
-  proof (rule reindex_nontrivial[symmetric])
-    show "finite p"
-      using assm by auto
-    fix x y
-    assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
-    obtain a b where ab: "snd x = cbox a b"
-      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
-    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
-      by (metis prod.collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+
-    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
-      by (intro assm(5)[of "fst x" _ "fst y"]) auto
-    then have "content (cbox a b) = 0"
-      unfolding \<open>snd x = snd y\<close>[symmetric] ab content_eq_0_interior by auto
-    then have "d (cbox a b) = \<^bold>1"
-      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
-    then show "d (snd x) = \<^bold>1"
-      unfolding ab by auto
-  qed
-qed
-
-lemma tag_in_interval: "p tagged_division_of i \<Longrightarrow> (x, k) \<in> p \<Longrightarrow> x \<in> i"
-  by auto
-
-lemma tagged_division_of_empty: "{} tagged_division_of {}"
-  unfolding tagged_division_of by auto
-
-lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \<longleftrightarrow> p = {}"
-  unfolding tagged_partial_division_of_def by auto
-
-lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \<longleftrightarrow> p = {}"
-  unfolding tagged_division_of by auto
-
-lemma tagged_division_of_self: "x \<in> cbox a b \<Longrightarrow> {(x,cbox a b)} tagged_division_of (cbox a b)"
-  by (rule tagged_division_ofI) auto
-
-lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}"
-  unfolding box_real[symmetric]
-  by (rule tagged_division_of_self)
-
-lemma tagged_division_union:
-  assumes "p1 tagged_division_of s1"
-    and "p2 tagged_division_of s2"
-    and "interior s1 \<inter> interior s2 = {}"
-  shows "(p1 \<union> p2) tagged_division_of (s1 \<union> s2)"
-proof (rule tagged_division_ofI)
-  note p1 = tagged_division_ofD[OF assms(1)]
-  note p2 = tagged_division_ofD[OF assms(2)]
-  show "finite (p1 \<union> p2)"
-    using p1(1) p2(1) by auto
-  show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2"
-    using p1(6) p2(6) by blast
-  fix x k
-  assume xk: "(x, k) \<in> p1 \<union> p2"
-  show "x \<in> k" "\<exists>a b. k = cbox a b"
-    using xk p1(2,4) p2(2,4) by auto
-  show "k \<subseteq> s1 \<union> s2"
-    using xk p1(3) p2(3) by blast
-  fix x' k'
-  assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')"
-  have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}"
-    using assms(3) interior_mono by blast
-  show "interior k \<inter> interior k' = {}"
-    apply (cases "(x, k) \<in> p1")
-    apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2))
-    by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
-qed
-
-lemma tagged_division_unions:
-  assumes "finite iset"
-    and "\<forall>i\<in>iset. pfn i tagged_division_of i"
-    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior(i1) \<inter> interior(i2) = {}"
-  shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)"
-proof (rule tagged_division_ofI)
-  note assm = tagged_division_ofD[OF assms(2)[rule_format]]
-  show "finite (\<Union>(pfn ` iset))"
-    using assms by auto
-  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)"
-    by blast
-  also have "\<dots> = \<Union>iset"
-    using assm(6) by auto
-  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" .
-  fix x k
-  assume xk: "(x, k) \<in> \<Union>(pfn ` iset)"
-  then obtain i where i: "i \<in> iset" "(x, k) \<in> pfn i"
-    by auto
-  show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>iset"
-    using assm(2-4)[OF i] using i(1) by auto
-  fix x' k'
-  assume xk': "(x', k') \<in> \<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')"
-  then obtain i' where i': "i' \<in> iset" "(x', k') \<in> pfn i'"
-    by auto
-  have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}"
-    using i(1) i'(1)
-    using assms(3)[rule_format] interior_mono
-    by blast
-  show "interior k \<inter> interior k' = {}"
-    apply (cases "i = i'")
-    using assm(5) i' i(2) xk'(2) apply blast
-    using "*" assm(3) i' i by auto
-qed
-
-lemma tagged_partial_division_of_union_self:
-  assumes "p tagged_partial_division_of s"
-  shows "p tagged_division_of (\<Union>(snd ` p))"
-  apply (rule tagged_division_ofI)
-  using tagged_partial_division_ofD[OF assms]
-  apply auto
-  done
-
-lemma tagged_division_of_union_self:
-  assumes "p tagged_division_of s"
-  shows "p tagged_division_of (\<Union>(snd ` p))"
-  apply (rule tagged_division_ofI)
-  using tagged_division_ofD[OF assms]
-  apply auto
-  done
-
-subsection \<open>Functions closed on boxes: morphisms from boxes to monoids\<close>
-
-text \<open>This auxiliary structure is used to sum up over the elements of a division. Main theorem is
-  @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and
-  @{typ bool}.\<close>
-
-lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}"
-  using content_empty unfolding empty_as_interval by auto
-
-paragraph \<open>Using additivity of lifted function to encode definedness.\<close>
-
-definition lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
-where
-  "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
-
-lemma lift_option_simps[simp]:
-  "lift_option f (Some a) (Some b) = Some (f a b)"
-  "lift_option f None b' = None"
-  "lift_option f a' None = None"
-  by (auto simp: lift_option_def)
-
-lemma comm_monoid_lift_option:
-  assumes "comm_monoid f z"
-  shows "comm_monoid (lift_option f) (Some z)"
-proof -
-  from assms interpret comm_monoid f z .
-  show ?thesis
-    by standard (auto simp: lift_option_def ac_simps split: bind_split)
-qed
-
-lemma comm_monoid_and: "comm_monoid HOL.conj True"
-  by standard auto
-
-lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True"
-  by (rule comm_monoid_set.intro) (fact comm_monoid_and)
-
-paragraph \<open>Operative\<close>
-
-definition (in comm_monoid) operative :: "('b::euclidean_space set \<Rightarrow> 'a) \<Rightarrow> bool"
-  where "operative g \<longleftrightarrow>
-    (\<forall>a b. content (cbox a b) = 0 \<longrightarrow> g (cbox a b) = \<^bold>1) \<and>
-    (\<forall>a b c. \<forall>k\<in>Basis. g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c}))"
-
-lemma (in comm_monoid) operativeD[dest]:
-  assumes "operative g"
-  shows "\<And>a b. content (cbox a b) = 0 \<Longrightarrow> g (cbox a b) = \<^bold>1"
-    and "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-  using assms unfolding operative_def by auto
-
-lemma (in comm_monoid) operative_empty: "operative g \<Longrightarrow> g {} = \<^bold>1"
-  unfolding operative_def by (rule property_empty_interval) auto
-
 lemma operative_content[intro]: "add.operative content"
-  by (force simp add: add.operative_def content_split[symmetric])
-
-definition "division_points (k::('a::euclidean_space) set) d =
-   {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
-     (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
-
-lemma division_points_finite:
-  fixes i :: "'a::euclidean_space set"
-  assumes "d division_of i"
-  shows "finite (division_points i d)"
-proof -
-  note assm = division_ofD[OF assms]
-  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
-    (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
-  have *: "division_points i d = \<Union>(?M ` Basis)"
-    unfolding division_points_def by auto
-  show ?thesis
-    unfolding * using assm by auto
-qed
-
-lemma division_points_subset:
-  fixes a :: "'a::euclidean_space"
-  assumes "d division_of (cbox a b)"
-    and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
-    and k: "k \<in> Basis"
-  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq>
-      division_points (cbox a b) d" (is ?t1)
-    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
-      division_points (cbox a b) d" (is ?t2)
-proof -
-  note assm = division_ofD[OF assms(1)]
-  have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-    "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else  b \<bullet> i) *\<^sub>R i) \<bullet> i"
-    "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i"
-    "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c"
-    using assms using less_imp_le by auto
-  show ?t1 (*FIXME a horrible mess*)
-    unfolding division_points_def interval_split[OF k, of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
-    unfolding *
-    apply (rule subsetI)
-    unfolding mem_Collect_eq split_beta
-    apply (erule bexE conjE)+
-    apply (simp add: )
-    apply (erule exE conjE)+
-  proof
-    fix i l x
-    assume as:
-      "a \<bullet> fst x < snd x" "snd x < (if fst x = k then c else b \<bullet> fst x)"
-      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
-      and fstx: "fst x \<in> Basis"
-    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
-    have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
-      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
-    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
-      using l using as(6) unfolding box_ne_empty[symmetric] by auto
-    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
-      using as(1-3,5) fstx
-      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
-      apply (auto split: if_split_asm)
-      done
-    show "snd x < b \<bullet> fst x"
-      using as(2) \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
-  qed
-  show ?t2
-    unfolding division_points_def interval_split[OF k, of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
-    unfolding *
-    unfolding subset_eq
-    apply rule
-    unfolding mem_Collect_eq split_beta
-    apply (erule bexE conjE)+
-    apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
-    apply (erule exE conjE)+
-  proof
-    fix i l x
-    assume as:
-      "(if fst x = k then c else a \<bullet> fst x) < snd x" "snd x < b \<bullet> fst x"
-      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
-      and fstx: "fst x \<in> Basis"
-    from assm(4)[OF this(5)] guess u v by (elim exE) note l=this
-    have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
-      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
-    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
-      using l using as(6) unfolding box_ne_empty[symmetric] by auto
-    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
-      using as(1-3,5) fstx
-      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
-      apply (auto split: if_split_asm)
-      done
-    show "a \<bullet> fst x < snd x"
-      using as(1) \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm)
-   qed
-qed
-
-lemma division_points_psubset:
-  fixes a :: "'a::euclidean_space"
-  assumes "d division_of (cbox a b)"
-      and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
-      and "l \<in> d"
-      and "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c"
-      and k: "k \<in> Basis"
-  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subset>
-         division_points (cbox a b) d" (is "?D1 \<subset> ?D")
-    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset>
-         division_points (cbox a b) d" (is "?D2 \<subset> ?D")
-proof -
-  have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-    using assms(2) by (auto intro!:less_imp_le)
-  guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this
-  have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i"
-    using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty
-    using subset_box(1)
-    apply auto
-    apply blast+
-    done
-  have *: "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
-          "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
-    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
-    using uv[rule_format, of k] ab k
-    by auto
-  have "\<exists>x. x \<in> ?D - ?D1"
-    using assms(3-)
-    unfolding division_points_def interval_bounds[OF ab]
-    apply -
-    apply (erule disjE)
-    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
-    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
-    done
-  moreover have "?D1 \<subseteq> ?D"
-    by (auto simp add: assms division_points_subset)
-  ultimately show "?D1 \<subset> ?D"
-    by blast
-  have *: "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
-    "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
-    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
-    using uv[rule_format, of k] ab k
-    by auto
-  have "\<exists>x. x \<in> ?D - ?D2"
-    using assms(3-)
-    unfolding division_points_def interval_bounds[OF ab]
-    apply -
-    apply (erule disjE)
-    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
-    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
-    done
-  moreover have "?D2 \<subseteq> ?D"
-    by (auto simp add: assms division_points_subset)
-  ultimately show "?D2 \<subset> ?D"
-    by blast
-qed
-
-lemma (in comm_monoid_set) operative_division:
-  fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
-  assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)"
-proof -
-  define C where [abs_def]: "C = card (division_points (cbox a b) d)"
-  then show ?thesis
-    using d
-  proof (induction C arbitrary: a b d rule: less_induct)
-    case (less a b d)
-    show ?case
-    proof cases
-      show "content (cbox a b) = 0 \<Longrightarrow> F g d = g (cbox a b)"
-        using division_of_content_0[OF _ less.prems] operativeD(1)[OF  g] division_ofD(4)[OF less.prems]
-        by (fastforce intro!: neutral)
-    next
-      assume "content (cbox a b) \<noteq> 0"
-      note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq]
-      then have ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-        by (auto intro!: less_imp_le)
-      show "F g d = g (cbox a b)"
-      proof (cases "division_points (cbox a b) d = {}")
-        case True
-        { fix u v and j :: 'b
-          assume j: "j \<in> Basis" and as: "cbox u v \<in> d"
-          then have "cbox u v \<noteq> {}"
-            using less.prems by blast
-          then have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j"
-            using j unfolding box_ne_empty by auto
-          have *: "\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> Q (cbox u v)"
-            using as j by auto
-          have "(j, u\<bullet>j) \<notin> division_points (cbox a b) d"
-               "(j, v\<bullet>j) \<notin> division_points (cbox a b) d" using True by auto
-          note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps]
-          note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
-          moreover
-          have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j"
-            using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as]
-            apply (metis j subset_box(1) uv(1))
-            by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1))
-          ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j"
-            unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force }
-        then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and>
-          (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
-          unfolding forall_in_division[OF less.prems] by blast
-        have "(1/2) *\<^sub>R (a+b) \<in> cbox a b"
-          unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps)
-        note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff]
-        then guess i .. note i=this
-        guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this
-        have "cbox a b \<in> d"
-        proof -
-          have "u = a" "v = b"
-            unfolding euclidean_eq_iff[where 'a='b]
-          proof safe
-            fix j :: 'b
-            assume j: "j \<in> Basis"
-            note i(2)[unfolded uv mem_box,rule_format,of j]
-            then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> j"
-              using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
-          qed
-          then have "i = cbox a b" using uv by auto
-          then show ?thesis using i by auto
-        qed
-        then have deq: "d = insert (cbox a b) (d - {cbox a b})"
-          by auto
-        have "F g (d - {cbox a b}) = \<^bold>1"
-        proof (intro neutral ballI)
-          fix x
-          assume x: "x \<in> d - {cbox a b}"
-          then have "x\<in>d"
-            by auto note d'[rule_format,OF this]
-          then guess u v by (elim exE conjE) note uv=this
-          have "u \<noteq> a \<or> v \<noteq> b"
-            using x[unfolded uv] by auto
-          then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j: "j \<in> Basis"
-            unfolding euclidean_eq_iff[where 'a='b] by auto
-          then have "u\<bullet>j = v\<bullet>j"
-            using uv(2)[rule_format,OF j] by auto
-          then have "content (cbox u v) = 0"
-            unfolding content_eq_0 using j
-            by force
-          then show "g x = \<^bold>1"
-            unfolding uv(1) by (rule operativeD(1)[OF g])
-        qed
-        then show "F g d = g (cbox a b)"
-          using division_ofD[OF less.prems]
-          apply (subst deq)
-          apply (subst insert)
-          apply auto
-          done
-      next
-        case False
-        then have "\<exists>x. x \<in> division_points (cbox a b) d"
-          by auto
-        then guess k c
-          unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv
-          apply (elim exE conjE)
-          done
-        note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']]
-        from this(3) guess j .. note j=this
-        define d1 where "d1 = {l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
-        define d2 where "d2 = {l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
-        define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)"
-        define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)"
-        note division_points_psubset[OF \<open>d division_of cbox a b\<close> ab kc(1-2) j]
-        note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
-        then have *: "F g d1 = g (cbox a b \<inter> {x. x\<bullet>k \<le> c})" "F g d2 = g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-          unfolding interval_split[OF kc(4)]
-          apply (rule_tac[!] "less.hyps"[rule_format])
-          using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c]
-          apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>])
-          done
-        { fix l y
-          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} = y \<inter> {x. x \<bullet> k \<le> c}" "l \<noteq> y"
-          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
-          have "g (l \<inter> {x. x \<bullet> k \<le> c}) = \<^bold>1"
-            unfolding leq interval_split[OF kc(4)]
-            apply (rule operativeD[OF g])
-            unfolding interval_split[symmetric, OF kc(4)]
-            using division_split_left_inj less as kc leq by blast
-        } note fxk_le = this
-        { fix l y
-          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} = y \<inter> {x. c \<le> x \<bullet> k}" "l \<noteq> y"
-          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
-          have "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
-            unfolding leq interval_split[OF kc(4)]
-            apply (rule operativeD(1)[OF g])
-            unfolding interval_split[symmetric,OF kc(4)]
-            using division_split_right_inj less leq as kc by blast
-        } note fxk_ge = this
-        have d1_alt: "d1 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<le> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
-          using d1_def by auto
-        have d2_alt: "d2 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<ge> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
-          using d2_def by auto
-        have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
-          unfolding * using g kc(4) by blast
-        also have "F g d1 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d"
-          unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
-          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
-        also have "F g d2 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d"
-          unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
-          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
-        also have *: "\<forall>x\<in>d. g x = g (x \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (x \<inter> {x. c \<le> x \<bullet> k})"
-          unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
-          using g kc(4) by blast
-        have "F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d \<^bold>* F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d = F g d"
-          using * by (simp add: distrib)
-        finally show ?thesis by auto
-      qed
-    qed
-  qed
-qed
-
-lemma (in comm_monoid_set) operative_tagged_division:
-  assumes f: "operative g" and d: "d tagged_division_of (cbox a b)"
-  shows "F (\<lambda>(x, l). g l) d = g (cbox a b)"
-  unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric]
-  by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d])
+  by (force simp add: add.operative_def content_split[symmetric] content_eq_0_interior)
 
 lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> setsum content d = content (cbox a b)"
   by (metis operative_content setsum.operative_division)
@@ -1813,537 +209,8 @@
 lemma content_real_eq_0: "content {a .. b::real} = 0 \<longleftrightarrow> a \<ge> b"
   by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0)
 
-lemma interval_real_split:
-  "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
-  "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
-  apply (metis Int_atLeastAtMostL1 atMost_def)
-  apply (metis Int_atLeastAtMostL2 atLeast_def)
-  done
-
-lemma (in comm_monoid) operative_1_lt:
-  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
-    ((\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1) \<and> (\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
-  apply (simp add: operative_def content_real_eq_0 atMost_def[symmetric] atLeast_def[symmetric]
-              del: content_real_if)
-proof safe
-  fix a b c :: real
-  assume *: "\<forall>a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
-  assume "a < c" "c < b"
-  with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    by (simp add: less_imp_le min.absorb2 max.absorb2)
-next
-  fix a b c :: real
-  assume as: "\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1"
-    "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-  from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2)
-  have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
-    "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    by auto
-  show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
-    by (auto simp: min_def max_def le_less)
-qed
-
-lemma (in comm_monoid) operative_1_le:
-  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
-    ((\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1) \<and> (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
-  unfolding operative_1_lt
-proof safe
-  fix a b c :: real
-  assume as: "\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b"
-  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    apply (rule as(1)[rule_format])
-    using as(2-)
-    apply auto
-    done
-next
-  fix a b c :: real
-  assume "\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1"
-    and "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    and "a \<le> c"
-    and "c \<le> b"
-  note as = this[rule_format]
-  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
-  proof (cases "c = a \<or> c = b")
-    case False
-    then show ?thesis
-      apply -
-      apply (subst as(2))
-      using as(3-)
-      apply auto
-      done
-  next
-    case True
-    then show ?thesis
-    proof
-      assume *: "c = a"
-      then have "g {a .. c} = \<^bold>1"
-        apply -
-        apply (rule as(1)[rule_format])
-        apply auto
-        done
-      then show ?thesis
-        unfolding * by auto
-    next
-      assume *: "c = b"
-      then have "g {c .. b} = \<^bold>1"
-        apply -
-        apply (rule as(1)[rule_format])
-        apply auto
-        done
-      then show ?thesis
-        unfolding * by auto
-    qed
-  qed
-qed
-
-subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close>
-
-definition fine  (infixr "fine" 46)
-  where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
-
-lemma fineI:
-  assumes "\<And>x k. (x, k) \<in> s \<Longrightarrow> k \<subseteq> d x"
-  shows "d fine s"
-  using assms unfolding fine_def by auto
-
-lemma fineD[dest]:
-  assumes "d fine s"
-  shows "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> d x"
-  using assms unfolding fine_def by auto
-
-lemma fine_inter: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p"
-  unfolding fine_def by auto
-
-lemma fine_inters:
- "(\<lambda>x. \<Inter>{f d x | d.  d \<in> s}) fine p \<longleftrightarrow> (\<forall>d\<in>s. (f d) fine p)"
-  unfolding fine_def by blast
-
-lemma fine_union: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
-  unfolding fine_def by blast
-
-lemma fine_unions: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
-  unfolding fine_def by auto
-
-lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
-  unfolding fine_def by blast
-
-subsection \<open>Some basic combining lemmas.\<close>
-
-lemma tagged_division_unions_exists:
-  assumes "finite iset"
-    and "\<forall>i\<in>iset. \<exists>p. p tagged_division_of i \<and> d fine p"
-    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
-    and "\<Union>iset = i"
-   obtains p where "p tagged_division_of i" and "d fine p"
-proof -
-  obtain pfn where pfn:
-    "\<And>x. x \<in> iset \<Longrightarrow> pfn x tagged_division_of x"
-    "\<And>x. x \<in> iset \<Longrightarrow> d fine pfn x"
-    using bchoice[OF assms(2)] by auto
-  show thesis
-    apply (rule_tac p="\<Union>(pfn ` iset)" in that)
-    using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force
-    by (metis (mono_tags, lifting) fine_unions imageE pfn(2))
-qed
-
-
-subsection \<open>The set we're concerned with must be closed.\<close>
-
-lemma division_of_closed:
-  fixes i :: "'n::euclidean_space set"
-  shows "s division_of i \<Longrightarrow> closed i"
-  unfolding division_of_def by fastforce
-
-subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close>
-
-lemma interval_bisection_step:
-  fixes type :: "'a::euclidean_space"
-  assumes "P {}"
-    and "\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P (s \<union> t)"
-    and "\<not> P (cbox a (b::'a))"
-  obtains c d where "\<not> P (cbox c d)"
-    and "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
-proof -
-  have "cbox a b \<noteq> {}"
-    using assms(1,3) by metis
-  then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
-    by (force simp: mem_box)
-  { fix f
-    have "\<lbrakk>finite f;
-           \<And>s. s\<in>f \<Longrightarrow> P s;
-           \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b;
-           \<And>s t. s\<in>f \<Longrightarrow> t\<in>f \<Longrightarrow> s \<noteq> t \<Longrightarrow> interior s \<inter> interior t = {}\<rbrakk> \<Longrightarrow> P (\<Union>f)"
-    proof (induct f rule: finite_induct)
-      case empty
-      show ?case
-        using assms(1) by auto
-    next
-      case (insert x f)
-      show ?case
-        unfolding Union_insert
-        apply (rule assms(2)[rule_format])
-        using inter_interior_unions_intervals [of f "interior x"]
-        apply (auto simp: insert)
-        by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff)
-    qed
-  } note UN_cases = this
-  let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or>
-    (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}"
-  let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
-  {
-    presume "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d) \<Longrightarrow> False"
-    then show thesis
-      unfolding atomize_not not_all
-      by (blast intro: that)
-  }
-  assume as: "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d)"
-  have "P (\<Union>?A)"
-  proof (rule UN_cases)
-    let ?B = "(\<lambda>s. cbox (\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i::'a)
-      (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)) ` {s. s \<subseteq> Basis}"
-    have "?A \<subseteq> ?B"
-    proof
-      fix x
-      assume "x \<in> ?A"
-      then obtain c d
-        where x:  "x = cbox c d"
-                  "\<And>i. i \<in> Basis \<Longrightarrow>
-                        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-                        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" by blast
-      show "x \<in> ?B"
-        unfolding image_iff x
-        apply (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI)
-        apply (rule arg_cong2 [where f = cbox])
-        using x(2) ab
-        apply (auto simp add: euclidean_eq_iff[where 'a='a])
-        by fastforce
-    qed
-    then show "finite ?A"
-      by (rule finite_subset) auto
-  next
-    fix s
-    assume "s \<in> ?A"
-    then obtain c d
-      where s: "s = cbox c d"
-               "\<And>i. i \<in> Basis \<Longrightarrow>
-                     c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-                     c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
-      by blast
-    show "P s"
-      unfolding s
-      apply (rule as[rule_format])
-      using ab s(2) by force
-    show "\<exists>a b. s = cbox a b"
-      unfolding s by auto
-    fix t
-    assume "t \<in> ?A"
-    then obtain e f where t:
-      "t = cbox e f"
-      "\<And>i. i \<in> Basis \<Longrightarrow>
-        e \<bullet> i = a \<bullet> i \<and> f \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-        e \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> f \<bullet> i = b \<bullet> i"
-      by blast
-    assume "s \<noteq> t"
-    then have "\<not> (c = e \<and> d = f)"
-      unfolding s t by auto
-    then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis"
-      unfolding euclidean_eq_iff[where 'a='a] by auto
-    then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
-      using s(2) t(2) apply fastforce
-      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
-    have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
-      by auto
-    show "interior s \<inter> interior t = {}"
-      unfolding s t interior_cbox
-    proof (rule *)
-      fix x
-      assume "x \<in> box c d" "x \<in> box e f"
-      then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
-        unfolding mem_box using i'
-        by force+
-      show False  using s(2)[OF i']
-      proof safe
-        assume as: "c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2"
-        show False
-          using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps)
-      next
-        assume as: "c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i"
-        show False
-          using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
-      qed
-    qed
-  qed
-  also have "\<Union>?A = cbox a b"
-  proof (rule set_eqI,rule)
-    fix x
-    assume "x \<in> \<Union>?A"
-    then obtain c d where x:
-      "x \<in> cbox c d"
-      "\<And>i. i \<in> Basis \<Longrightarrow>
-        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
-      by blast
-    show "x\<in>cbox a b"
-      unfolding mem_box
-    proof safe
-      fix i :: 'a
-      assume i: "i \<in> Basis"
-      then show "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i"
-        using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto
-    qed
-  next
-    fix x
-    assume x: "x \<in> cbox a b"
-    have "\<forall>i\<in>Basis.
-      \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d"
-      (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d")
-      unfolding mem_box
-    proof
-      fix i :: 'a
-      assume i: "i \<in> Basis"
-      have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)"
-        using x[unfolded mem_box,THEN bspec, OF i] by auto
-      then show "\<exists>c d. ?P i c d"
-        by blast
-    qed
-    then show "x\<in>\<Union>?A"
-      unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff
-      apply auto
-      apply (rule_tac x="cbox xa xaa" in exI)
-      unfolding mem_box
-      apply auto
-      done
-  qed
-  finally show False
-    using assms by auto
-qed
-
-lemma interval_bisection:
-  fixes type :: "'a::euclidean_space"
-  assumes "P {}"
-    and "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))"
-    and "\<not> P (cbox a (b::'a))"
-  obtains x where "x \<in> cbox a b"
-    and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
-proof -
-  have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and>
-    (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
-       2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>x. ?P x")
-  proof
-    show "?P x" for x
-    proof (cases "P (cbox (fst x) (snd x))")
-      case True
-      then show ?thesis by auto
-    next
-      case as: False
-      obtain c d where "\<not> P (cbox c d)"
-        "\<forall>i\<in>Basis.
-           fst x \<bullet> i \<le> c \<bullet> i \<and>
-           c \<bullet> i \<le> d \<bullet> i \<and>
-           d \<bullet> i \<le> snd x \<bullet> i \<and>
-           2 * (d \<bullet> i - c \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i"
-        by (rule interval_bisection_step[of P, OF assms(1-2) as])
-      then show ?thesis
-        apply -
-        apply (rule_tac x="(c,d)" in exI)
-        apply auto
-        done
-    qed
-  qed
-  then obtain f where f:
-    "\<forall>x.
-      \<not> P (cbox (fst x) (snd x)) \<longrightarrow>
-      \<not> P (cbox (fst (f x)) (snd (f x))) \<and>
-        (\<forall>i\<in>Basis.
-            fst x \<bullet> i \<le> fst (f x) \<bullet> i \<and>
-            fst (f x) \<bullet> i \<le> snd (f x) \<bullet> i \<and>
-            snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
-            2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)"
-    apply -
-    apply (drule choice)
-    apply blast
-    done
-  define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
-  have "A 0 = a" "B 0 = b" "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
-    (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
-    2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
-  proof -
-    show "A 0 = a" "B 0 = b"
-      unfolding ab_def by auto
-    note S = ab_def funpow.simps o_def id_apply
-    show "?P n" for n
-    proof (induct n)
-      case 0
-      then show ?case
-        unfolding S
-        apply (rule f[rule_format]) using assms(3)
-        apply auto
-        done
-    next
-      case (Suc n)
-      show ?case
-        unfolding S
-        apply (rule f[rule_format])
-        using Suc
-        unfolding S
-        apply auto
-        done
-    qed
-  qed
-  note AB = this(1-2) conjunctD2[OF this(3),rule_format]
-
-  have interv: "\<exists>n. \<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e"
-    if e: "0 < e" for e
-  proof -
-    obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n"
-      using real_arch_pow[of 2 "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
-    show ?thesis
-    proof (rule exI [where x=n], clarify)
-      fix x y
-      assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)"
-      have "dist x y \<le> setsum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
-        unfolding dist_norm by(rule norm_le_l1)
-      also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
-      proof (rule setsum_mono)
-        fix i :: 'a
-        assume i: "i \<in> Basis"
-        show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
-          using xy[unfolded mem_box,THEN bspec, OF i]
-          by (auto simp: inner_diff_left)
-      qed
-      also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
-        unfolding setsum_divide_distrib
-      proof (rule setsum_mono)
-        show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i
-        proof (induct n)
-          case 0
-          then show ?case
-            unfolding AB by auto
-        next
-          case (Suc n)
-          have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
-            using AB(4)[of i n] using i by auto
-          also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n"
-            using Suc by (auto simp add: field_simps)
-          finally show ?case .
-        qed
-      qed
-      also have "\<dots> < e"
-        using n using e by (auto simp add: field_simps)
-      finally show "dist x y < e" .
-    qed
-  qed
-  {
-    fix n m :: nat
-    assume "m \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)"
-    proof (induction rule: inc_induct)
-      case (step i)
-      show ?case
-        using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto
-    qed simp
-  } note ABsubset = this
-  have "\<exists>a. \<forall>n. a\<in> cbox (A n) (B n)"
-    by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv])
-      (metis nat.exhaust AB(1-3) assms(1,3))
-  then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
-    by blast
-  show thesis
-  proof (rule that[rule_format, of x0])
-    show "x0\<in>cbox a b"
-      using x0[of 0] unfolding AB .
-    fix e :: real
-    assume "e > 0"
-    from interv[OF this] obtain n
-      where n: "\<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" ..
-    have "\<not> P (cbox (A n) (B n))"
-      apply (cases "0 < n")
-      using AB(3)[of "n - 1"] assms(3) AB(1-2)
-      apply auto
-      done
-    moreover have "cbox (A n) (B n) \<subseteq> ball x0 e"
-      using n using x0[of n] by auto
-    moreover have "cbox (A n) (B n) \<subseteq> cbox a b"
-      unfolding AB(1-2)[symmetric] by (rule ABsubset) auto
-    ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
-      apply (rule_tac x="A n" in exI)
-      apply (rule_tac x="B n" in exI)
-      apply (auto simp: x0)
-      done
-  qed
-qed
-
-
-subsection \<open>Cousin's lemma.\<close>
-
-lemma fine_division_exists:
-  fixes a b :: "'a::euclidean_space"
-  assumes "gauge g"
-  obtains p where "p tagged_division_of (cbox a b)" "g fine p"
-proof -
-  presume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p) \<Longrightarrow> False"
-  then obtain p where "p tagged_division_of (cbox a b)" "g fine p"
-    by blast
-  then show thesis ..
-next
-  assume as: "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)"
-  obtain x where x:
-      "x \<in> (cbox a b)"
-      "\<And>e. 0 < e \<Longrightarrow>
-        \<exists>c d.
-          x \<in> cbox c d \<and>
-          cbox c d \<subseteq> ball x e \<and>
-          cbox c d \<subseteq> (cbox a b) \<and>
-          \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
-    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ as])
-    apply (simp add: fine_def)
-    apply (metis tagged_division_union fine_union)
-    apply (auto simp: )
-    done
-  obtain e where e: "e > 0" "ball x e \<subseteq> g x"
-    using gaugeD[OF assms, of x] unfolding open_contains_ball by auto
-  from x(2)[OF e(1)]
-  obtain c d where c_d: "x \<in> cbox c d"
-                        "cbox c d \<subseteq> ball x e"
-                        "cbox c d \<subseteq> cbox a b"
-                        "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
-    by blast
-  have "g fine {(x, cbox c d)}"
-    unfolding fine_def using e using c_d(2) by auto
-  then show False
-    using tagged_division_of_self[OF c_d(1)] using c_d by auto
-qed
-
-lemma fine_division_exists_real:
-  fixes a b :: real
-  assumes "gauge g"
-  obtains p where "p tagged_division_of {a .. b}" "g fine p"
-  by (metis assms box_real(2) fine_division_exists)
-
-subsection \<open>Division filter\<close>
-
-text \<open>Divisions over all gauges towards finer divisions.\<close>
-
-definition division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
-  where "division_filter s = (INF g:{g. gauge g}. principal {p. p tagged_division_of s \<and> g fine p})"
-
-lemma eventually_division_filter:
-  "(\<forall>\<^sub>F p in division_filter s. P p) \<longleftrightarrow>
-    (\<exists>g. gauge g \<and> (\<forall>p. p tagged_division_of s \<and> g fine p \<longrightarrow> P p))"
-  unfolding division_filter_def
-proof (subst eventually_INF_base; clarsimp)
-  fix g1 g2 :: "'a \<Rightarrow> 'a set" show "gauge g1 \<Longrightarrow> gauge g2 \<Longrightarrow> \<exists>x. gauge x \<and>
-    {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g1 fine p} \<and>
-    {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g2 fine p}"
-    by (intro exI[of _ "\<lambda>x. g1 x \<inter> g2 x"]) (auto simp: fine_inter)
-qed (auto simp: eventually_principal)
-
-lemma division_filter_not_empty: "division_filter (cbox a b) \<noteq> bot"
-  unfolding trivial_limit_def eventually_division_filter
-  by (auto elim: fine_division_exists)
-
-lemma eventually_division_filter_tagged_division:
-  "eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
-  unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto
+lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}"
+  using content_empty unfolding empty_as_interval by auto
 
 subsection \<open>Gauge integral\<close>
 
@@ -2421,26 +288,6 @@
 lemma has_integral_integral: "f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s"
   by auto
 
-lemma setsum_content_null:
-  assumes "content (cbox a b) = 0"
-    and "p tagged_division_of (cbox a b)"
-  shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
-proof (rule setsum.neutral, rule)
-  fix y
-  assume y: "y \<in> p"
-  obtain x k where xk: "y = (x, k)"
-    using surj_pair[of y] by blast
-  note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]]
-  from this(2) obtain c d where k: "k = cbox c d" by blast
-  have "(\<lambda>(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x"
-    unfolding xk by auto
-  also have "\<dots> = 0"
-    using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d]
-    unfolding assms(1) k
-    by auto
-  finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
-qed
-
 subsection \<open>Basic theorems about integrals.\<close>
 
 lemma has_integral_unique:
@@ -3153,38 +1000,32 @@
 
 subsection \<open>Additivity of integral on abutting intervals.\<close>
 
-lemma tagged_division_split_left_inj:
-  fixes x1 :: "'a::euclidean_space"
+lemma tagged_division_split_left_inj_content:
   assumes d: "d tagged_division_of i"
-    and k12: "(x1, k1) \<in> d"
-             "(x2, k2) \<in> d"
-             "k1 \<noteq> k2"
-             "k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
-             "k \<in> Basis"
+    and "(x1, k1) \<in> d" "(x2, k2) \<in> d" "k1 \<noteq> k2" "k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}" "k \<in> Basis"
   shows "content (k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
 proof -
-  have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c"
-    by force
+  from tagged_division_ofD(4)[OF d \<open>(x1, k1) \<in> d\<close>] obtain a b where k1: "k1 = cbox a b"
+    by auto
   show ?thesis
-    using k12
-    by (fastforce intro!:  division_split_left_inj[OF division_of_tagged_division[OF d]] *)
+    unfolding k1 interval_split[OF \<open>k \<in> Basis\<close>]
+    unfolding content_eq_0_interior
+    unfolding interval_split[OF \<open>k \<in> Basis\<close>, symmetric] k1[symmetric]
+    by (rule tagged_division_split_left_inj[OF assms])
 qed
 
-lemma tagged_division_split_right_inj:
-  fixes x1 :: "'a::euclidean_space"
+lemma tagged_division_split_right_inj_content:
   assumes d: "d tagged_division_of i"
-    and k12: "(x1, k1) \<in> d"
-             "(x2, k2) \<in> d"
-             "k1 \<noteq> k2"
-             "k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
-             "k \<in> Basis"
+    and "(x1, k1) \<in> d" "(x2, k2) \<in> d" "k1 \<noteq> k2" "k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}" "k \<in> Basis"
   shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
 proof -
-  have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c"
-    by force
+  from tagged_division_ofD(4)[OF d \<open>(x1, k1) \<in> d\<close>] obtain a b where k1: "k1 = cbox a b"
+    by auto
   show ?thesis
-    using k12
-    by (fastforce intro!:  division_split_right_inj[OF division_of_tagged_division[OF d]] *)
+    unfolding k1 interval_split[OF \<open>k \<in> Basis\<close>]
+    unfolding content_eq_0_interior
+    unfolding interval_split[OF \<open>k \<in> Basis\<close>, symmetric] k1[symmetric]
+    by (rule tagged_division_split_right_inj[OF assms])
 qed
 
 lemma has_integral_split:
@@ -3262,7 +1103,8 @@
     have lem1: "\<And>f P Q. (\<forall>x k. (x, k) \<in> {(x, f k) | x k. P x k} \<longrightarrow> Q x k) \<longleftrightarrow>
                          (\<forall>x k. P x k \<longrightarrow> Q x (f k))"
       by auto
-    have fin_finite: "finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}" if "finite s" for f s P
+    have fin_finite: "finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}"
+      if "finite s" for s and f :: "'a set \<Rightarrow> 'a set" and P :: "'a \<Rightarrow> 'a set \<Rightarrow> bool"
     proof -
       from that have "finite ((\<lambda>(x, k). (x, f k)) ` s)"
         by auto
@@ -3376,8 +1218,9 @@
       also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
         (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
         unfolding lem3[OF p(3)]
-        by (subst setsum.reindex_nontrivial[OF p(3)], auto intro!: k eq0 tagged_division_split_left_inj[OF p(1)] tagged_division_split_right_inj[OF p(1)]
-              simp: cont_eq)+
+        by (subst (1 2) setsum.reindex_nontrivial[OF p(3)])
+           (auto intro!: k eq0 tagged_division_split_left_inj_content[OF p(1)] tagged_division_split_right_inj_content[OF p(1)]
+                 simp: cont_eq)+
       also note setsum.distrib[symmetric]
       also have "\<And>x. x \<in> p \<Longrightarrow>
                     (\<lambda>(x,ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x +
@@ -3407,34 +1250,6 @@
 
 subsection \<open>A sort of converse, integrability on subintervals.\<close>
 
-lemma tagged_division_union_interval:
-  fixes a :: "'a::euclidean_space"
-  assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
-    and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-    and k: "k \<in> Basis"
-  shows "(p1 \<union> p2) tagged_division_of (cbox a b)"
-proof -
-  have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-    by auto
-  show ?thesis
-    apply (subst *)
-    apply (rule tagged_division_union[OF assms(1-2)])
-    unfolding interval_split[OF k] interior_cbox
-    using k
-    apply (auto simp add: box_def elim!: ballE[where x=k])
-    done
-qed
-
-lemma tagged_division_union_interval_real:
-  fixes a :: real
-  assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})"
-    and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
-    and k: "k \<in> Basis"
-  shows "(p1 \<union> p2) tagged_division_of {a .. b}"
-  using assms
-  unfolding box_real[symmetric]
-  by (rule tagged_division_union_interval)
-
 lemma has_integral_separate_sides:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
   assumes "(f has_integral i) (cbox a b)"
@@ -3599,10 +1414,10 @@
     qed
   next
     fix a b :: 'a
-    assume "content (cbox a b) = 0"
+    assume "box a b = {}"
     then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0"
       using has_integral_null_eq
-      by (auto simp: integrable_on_null)
+      by (auto simp: integrable_on_null content_eq_0_interior)
   qed
 qed
 
@@ -4047,47 +1862,6 @@
 
 subsection \<open>Negligibility of hyperplane.\<close>
 
-lemma interval_doublesplit:
-  fixes a :: "'a::euclidean_space"
-  assumes "k \<in> Basis"
-  shows "cbox a b \<inter> {x . \<bar>x\<bullet>k - c\<bar> \<le> (e::real)} =
-    cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i)
-     (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)"
-proof -
-  have *: "\<And>x c e::real. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
-    by auto
-  have **: "\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}"
-    by blast
-  show ?thesis
-    unfolding * ** interval_split[OF assms] by (rule refl)
-qed
-
-lemma division_doublesplit:
-  fixes a :: "'a::euclidean_space"
-  assumes "p division_of (cbox a b)"
-    and k: "k \<in> Basis"
-  shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
-         division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
-proof -
-  have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
-    by auto
-  have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
-    by auto
-  note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
-  note division_split(2)[OF this, where c="c-e" and k=k,OF k]
-  then show ?thesis
-    apply (rule **)
-    subgoal
-      apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric])
-      apply (rule equalityI)
-      apply blast
-      apply clarsimp
-      apply (rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
-      apply auto
-      done
-    by (simp add: interval_split k interval_doublesplit)
-qed
-
 lemma content_doublesplit:
   fixes a :: "'a::euclidean_space"
   assumes "0 < e"
@@ -4240,6 +2014,8 @@
       also have "\<dots> < e"
       proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases)
         case prems: (1 u v)
+        then have *: "content (cbox u v) = 0"
+          unfolding content_eq_0_interior by simp
         have "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<le> content (cbox u v)"
           unfolding interval_doublesplit[OF k]
           apply (rule content_subset)
@@ -4247,7 +2023,7 @@
           apply auto
           done
         then show ?case
-          unfolding prems interval_doublesplit[OF k]
+          unfolding * interval_doublesplit[OF k]
           by (blast intro: antisym)
       next
         have "(\<Sum>l\<in>snd ` p. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) =
@@ -4281,188 +2057,9 @@
 qed
 
 
-subsection \<open>A technical lemma about "refinement" of division.\<close>
-
-lemma tagged_division_finer:
-  fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
-  assumes "p tagged_division_of (cbox a b)"
-    and "gauge d"
-  obtains q where "q tagged_division_of (cbox a b)"
-    and "d fine q"
-    and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
-proof -
-  let ?P = "\<lambda>p. p tagged_partial_division_of (cbox a b) \<longrightarrow> gauge d \<longrightarrow>
-    (\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and>
-      (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))"
-  {
-    have *: "finite p" "p tagged_partial_division_of (cbox a b)"
-      using assms(1)
-      unfolding tagged_division_of_def
-      by auto
-    presume "\<And>p. finite p \<Longrightarrow> ?P p"
-    from this[rule_format,OF * assms(2)] guess q .. note q=this
-    then show ?thesis
-      apply -
-      apply (rule that[of q])
-      unfolding tagged_division_ofD[OF assms(1)]
-      apply auto
-      done
-  }
-  fix p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
-  assume as: "finite p"
-  show "?P p"
-    apply rule
-    apply rule
-    using as
-  proof (induct p)
-    case empty
-    show ?case
-      apply (rule_tac x="{}" in exI)
-      unfolding fine_def
-      apply auto
-      done
-  next
-    case (insert xk p)
-    guess x k using surj_pair[of xk] by (elim exE) note xk=this
-    note tagged_partial_division_subset[OF insert(4) subset_insertI]
-    from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this]
-    have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}"
-      unfolding xk by auto
-    note p = tagged_partial_division_ofD[OF insert(4)]
-    from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this
-
-    have "finite {k. \<exists>x. (x, k) \<in> p}"
-      apply (rule finite_subset[of _ "snd ` p"])
-      using p
-      apply safe
-      apply (metis image_iff snd_conv)
-      apply auto
-      done
-    then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
-      apply (rule inter_interior_unions_intervals)
-      apply (rule open_interior)
-      apply (rule_tac[!] ballI)
-      unfolding mem_Collect_eq
-      apply (erule_tac[!] exE)
-      apply (drule p(4)[OF insertI2])
-      apply assumption
-      apply (rule p(5))
-      unfolding uv xk
-      apply (rule insertI1)
-      apply (rule insertI2)
-      apply assumption
-      using insert(2)
-      unfolding uv xk
-      apply auto
-      done
-    show ?case
-    proof (cases "cbox u v \<subseteq> d x")
-      case True
-      then show ?thesis
-        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
-        apply rule
-        unfolding * uv
-        apply (rule tagged_division_union)
-        apply (rule tagged_division_of_self)
-        apply (rule p[unfolded xk uv] insertI1)+
-        apply (rule q1)
-        apply (rule int)
-        apply rule
-        apply (rule fine_union)
-        apply (subst fine_def)
-        defer
-        apply (rule q1)
-        unfolding Ball_def split_paired_All split_conv
-        apply rule
-        apply rule
-        apply rule
-        apply rule
-        apply (erule insertE)
-        apply (simp add: uv xk)
-        apply (rule UnI2)
-        apply (drule q1(3)[rule_format])
-        unfolding xk uv
-        apply auto
-        done
-    next
-      case False
-      from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this
-      show ?thesis
-        apply (rule_tac x="q2 \<union> q1" in exI)
-        apply rule
-        unfolding * uv
-        apply (rule tagged_division_union q2 q1 int fine_union)+
-        unfolding Ball_def split_paired_All split_conv
-        apply rule
-        apply (rule fine_union)
-        apply (rule q1 q2)+
-        apply rule
-        apply rule
-        apply rule
-        apply rule
-        apply (erule insertE)
-        apply (rule UnI2)
-        apply (simp add: False uv xk)
-        apply (drule q1(3)[rule_format])
-        using False
-        unfolding xk uv
-        apply auto
-        done
-    qed
-  qed
-qed
-
 
 subsection \<open>Hence the main theorem about negligible sets.\<close>
 
-lemma finite_product_dependent:
-  assumes "finite s"
-    and "\<And>x. x \<in> s \<Longrightarrow> finite (t x)"
-  shows "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
-  using assms
-proof induct
-  case (insert x s)
-  have *: "{(i, j) |i j. i \<in> insert x s \<and> j \<in> t i} =
-    (\<lambda>y. (x,y)) ` (t x) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
-  show ?case
-    unfolding *
-    apply (rule finite_UnI)
-    using insert
-    apply auto
-    done
-qed auto
-
-lemma sum_sum_product:
-  assumes "finite s"
-    and "\<forall>i\<in>s. finite (t i)"
-  shows "setsum (\<lambda>i. setsum (x i) (t i)::real) s =
-    setsum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}"
-  using assms
-proof induct
-  case (insert a s)
-  have *: "{(i, j) |i j. i \<in> insert a s \<and> j \<in> t i} =
-    (\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
-  show ?case
-    unfolding *
-    apply (subst setsum.union_disjoint)
-    unfolding setsum.insert[OF insert(1-2)]
-    prefer 4
-    apply (subst insert(3))
-    unfolding add_right_cancel
-  proof -
-    show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
-      apply (subst setsum.reindex)
-      unfolding inj_on_def
-      apply auto
-      done
-    show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
-      apply (rule finite_product_dependent)
-      using insert
-      apply auto
-      done
-  qed (insert insert, auto)
-qed auto
-
 lemma has_integral_negligible:
   fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
   assumes "negligible s"
@@ -4767,7 +2364,7 @@
 qed
 
 lemma negligible_Un_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> negligible s \<and> negligible t"
-  using negligible_Un negligible_subset by blast 
+  using negligible_Un negligible_subset by blast
 
 lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}"
   using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] negligible_subset by blast
@@ -4915,10 +2512,10 @@
 proof safe
   fix a b :: 'b
   show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
-    if "content (cbox a b) = 0"
+    if "box a b = {}"
     apply (rule_tac x=f in exI)
     using assms that
-    apply (auto intro!: integrable_on_null)
+    apply (auto simp: content_eq_0_interior)
     done
   {
     fix c g
@@ -5037,33 +2634,6 @@
 
 subsection \<open>Specialization of additivity to one dimension.\<close>
 
-subsection \<open>Special case of additivity we need for the FTC.\<close>
-
-lemma additive_tagged_division_1:
-  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
-  assumes "a \<le> b"
-    and "p tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
-proof -
-  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
-  have ***: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
-    using assms by auto
-  have *: "add.operative ?f"
-    unfolding add.operative_1_lt box_eq_empty
-    by auto
-  have **: "cbox a b \<noteq> {}"
-    using assms(1) by auto
-  note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]]
-  note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric]
-  show ?thesis
-    unfolding *
-    apply (rule setsum.cong)
-    unfolding split_paired_all split_conv
-    using assms(2)
-    apply auto
-    done
-qed
-
 
 subsection \<open>A useful lemma allowing us to factor out the content size.\<close>
 
@@ -5351,20 +2921,6 @@
 qed
 
 
-subsection \<open>Attempt a systematic general set of "offset" results for components.\<close>
-
-lemma gauge_modify:
-  assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
-  shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
-  using assms
-  unfolding gauge_def
-  apply safe
-  defer
-  apply (erule_tac x="f x" in allE)
-  apply (erule_tac x="d (f x)" in allE)
-  apply auto
-  done
-
 
 subsection \<open>Only need trivial subintervals if the interval itself is trivial.\<close>
 
@@ -5502,10 +3058,12 @@
   shows "comm_monoid.operative op \<and> True (\<lambda>i. f integrable_on i)"
   unfolding comm_monoid.operative_def[OF comm_monoid_and]
   apply safe
-  apply (subst integrable_on_def)
-  unfolding has_integral_null_eq
-  apply (rule, rule refl)
-  apply (rule, assumption, assumption)+
+     apply (subst integrable_on_def)
+     apply rule
+     apply (rule has_integral_null_eq[where i=0, THEN iffD2])
+      apply (simp add: content_eq_0_interior)
+     apply rule
+    apply (rule, assumption, assumption)+
   unfolding integrable_on_def
   by (auto intro!: has_integral_split)
 
@@ -6099,18 +3657,6 @@
 
 subsection \<open>Stronger form of FCT; quite a tedious proof.\<close>
 
-lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
-  by (meson zero_less_one)
-
-lemma additive_tagged_division_1':
-  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
-  assumes "a \<le> b"
-    and "p tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
-  using additive_tagged_division_1[OF _ assms(2), of f]
-  using assms(1)
-  by auto
-
 lemma split_minus[simp]: "(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x"
   by (simp add: split_def)
 
@@ -7695,39 +5241,6 @@
   shows "f integrable_on s \<longleftrightarrow> f integrable_on t"
 by (blast intro: integrable_spike_set assms negligible_subset)
 
-(*lemma integral_spike_set:
- "\<forall>f:real^M->real^N g s t.
-        negligible(s DIFF t \<union> t DIFF s)
-        \<longrightarrow> integral s f = integral t f"
-qed  REPEAT STRIP_TAC THEN REWRITE_TAC[integral] THEN
-  AP_TERM_TAC THEN ABS_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
-  ASM_MESON_TAC[]);;
-
-lemma has_integral_interior:
- "\<forall>f:real^M->real^N y s.
-        negligible(frontier s)
-        \<longrightarrow> ((f has_integral y) (interior s) \<longleftrightarrow> (f has_integral y) s)"
-qed  REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
-  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ]
-    NEGLIGIBLE_SUBSET)) THEN
-  REWRITE_TAC[frontier] THEN
-  MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN
-  MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN
-  SET_TAC[]);;
-
-lemma has_integral_closure:
- "\<forall>f:real^M->real^N y s.
-        negligible(frontier s)
-        \<longrightarrow> ((f has_integral y) (closure s) \<longleftrightarrow> (f has_integral y) s)"
-qed  REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
-  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ]
-    NEGLIGIBLE_SUBSET)) THEN
-  REWRITE_TAC[frontier] THEN
-  MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN
-  MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN
-  SET_TAC[]);;*)
-
-
 subsection \<open>More lemmas that are useful later\<close>
 
 lemma has_integral_subset_component_le:
@@ -8412,34 +5925,26 @@
 
 subsection \<open>Also tagged divisions\<close>
 
+lemma has_integral_iff: "(f has_integral i) s \<longleftrightarrow> (f integrable_on s \<and> integral s f = i)"
+  by blast
+
 lemma has_integral_combine_tagged_division:
   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
   assumes "p tagged_division_of s"
     and "\<forall>(x,k) \<in> p. (f has_integral (i k)) k"
-  shows "(f has_integral (setsum (\<lambda>(x,k). i k) p)) s"
+  shows "(f has_integral (\<Sum>(x,k)\<in>p. i k)) s"
 proof -
-  have *: "(f has_integral (setsum (\<lambda>k. integral k f) (snd ` p))) s"
-    apply (rule has_integral_combine_division)
-    apply (rule division_of_tagged_division[OF assms(1)])
+  have *: "(f has_integral (\<Sum>k\<in>snd`p. integral k f)) s"
     using assms(2)
-    unfolding has_integral_integral[symmetric]
-    apply safe
+    apply (intro has_integral_combine_division)
+    apply (auto simp: has_integral_integral[symmetric] intro: division_of_tagged_division[OF assms(1)])
     apply auto
     done
-  then show ?thesis
-    apply -
-    apply (rule subst[where P="\<lambda>i. (f has_integral i) s"])
-    defer
-    apply assumption
-    apply (rule trans[of _ "setsum (\<lambda>(x,k). integral k f) p"])
-    apply (subst eq_commute)
-    apply (rule setsum.over_tagged_division_lemma[OF assms(1)])
-    apply (rule integral_null)
-    apply assumption
-    apply (rule setsum.cong)
-    using assms(2)
-    apply auto
-    done
+  also have "(\<Sum>k\<in>snd`p. integral k f) = (\<Sum>(x, k)\<in>p. integral k f)"
+    by (intro setsum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null)
+       (simp add: content_eq_0_interior)
+  finally show ?thesis
+    using assms by (auto simp add: has_integral_iff intro!: setsum.cong)
 qed
 
 lemma integral_combine_tagged_division_bottomup:
@@ -9594,84 +7099,6 @@
 
 subsection \<open>differentiation under the integral sign\<close>
 
-lemma tube_lemma:
-  assumes "compact K"
-  assumes "open W"
-  assumes "{x0} \<times> K \<subseteq> W"
-  shows "\<exists>X0. x0 \<in> X0 \<and> open X0 \<and> X0 \<times> K \<subseteq> W"
-proof -
-  {
-    fix y assume "y \<in> K"
-    then have "(x0, y) \<in> W" using assms by auto
-    with \<open>open W\<close>
-    have "\<exists>X0 Y. open X0 \<and> open Y \<and> x0 \<in> X0 \<and> y \<in> Y \<and> X0 \<times> Y \<subseteq> W"
-      by (rule open_prod_elim) blast
-  }
-  then obtain X0 Y where
-    *: "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W"
-    by metis
-  from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
-  with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
-    by (rule compactE)
-  then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
-    by (force intro!: choice)
-  with * CC show ?thesis
-    by (force intro!: exI[where x="\<Inter>C\<in>CC. X0 (c C)"]) (* SLOW *)
-qed
-
-lemma continuous_on_prod_compactE:
-  fixes fx::"'a::topological_space \<times> 'b::topological_space \<Rightarrow> 'c::metric_space"
-    and e::real
-  assumes cont_fx: "continuous_on (U \<times> C) fx"
-  assumes "compact C"
-  assumes [intro]: "x0 \<in> U"
-  notes [continuous_intros] = continuous_on_compose2[OF cont_fx]
-  assumes "e > 0"
-  obtains X0 where "x0 \<in> X0" "open X0"
-    "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
-proof -
-  define psi where "psi = (\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t)))"
-  define W0 where "W0 = {(x, t) \<in> U \<times> C. psi (x, t) < e}"
-  have W0_eq: "W0 = psi -` {..<e} \<inter> U \<times> C"
-    by (auto simp: vimage_def W0_def)
-  have "open {..<e}" by simp
-  have "continuous_on (U \<times> C) psi"
-    by (auto intro!: continuous_intros simp: psi_def split_beta')
-  from this[unfolded continuous_on_open_invariant, rule_format, OF \<open>open {..<e}\<close>]
-  obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C"
-    unfolding W0_eq by blast
-  have "{x0} \<times> C \<subseteq> W \<inter> U \<times> C"
-    unfolding W
-    by (auto simp: W0_def psi_def \<open>0 < e\<close>)
-  then have "{x0} \<times> C \<subseteq> W" by blast
-  from tube_lemma[OF \<open>compact C\<close> \<open>open W\<close> this]
-  obtain X0 where X0: "x0 \<in> X0" "open X0" "X0 \<times> C \<subseteq> W"
-    by blast
-
-  have "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
-  proof safe
-    fix x assume x: "x \<in> X0" "x \<in> U"
-    fix t assume t: "t \<in> C"
-    have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)"
-      by (auto simp: psi_def)
-    also
-    {
-      have "(x, t) \<in> X0 \<times> C"
-        using t x
-        by auto
-      also note \<open>\<dots> \<subseteq> W\<close>
-      finally have "(x, t) \<in> W" .
-      with t x have "(x, t) \<in> W \<inter> U \<times> C"
-        by blast
-      also note \<open>W \<inter> U \<times> C = W0 \<inter> U \<times> C\<close>
-      finally  have "psi (x, t) < e"
-        by (auto simp: W0_def)
-    }
-    finally show "dist (fx (x, t)) (fx (x0, t)) \<le> e" by simp
-  qed
-  from X0(1,2) this show ?thesis ..
-qed
-
 lemma integral_continuous_on_param:
   fixes f::"'a::topological_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
   assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). f x t)"
@@ -9721,26 +7148,6 @@
   qed
 qed (auto intro!: continuous_on_const)
 
-lemma eventually_closed_segment:
-  fixes x0::"'a::real_normed_vector"
-  assumes "open X0" "x0 \<in> X0"
-  shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0"
-proof -
-  from openE[OF assms]
-  obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" .
-  then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e"
-    by (auto simp: dist_commute eventually_at)
-  then show ?thesis
-  proof eventually_elim
-    case (elim x)
-    have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp
-    from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim]
-    have "closed_segment x0 x \<subseteq> ball x0 e" .
-    also note \<open>\<dots> \<subseteq> X0\<close>
-    finally show ?case .
-  qed
-qed
-
 lemma leibniz_rule:
   fixes f::"'a::banach \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
   assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
@@ -9834,8 +7241,7 @@
   qed (rule blinfun.bounded_linear_right)
 qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps)
 
-lemma
-  has_vector_derivative_eq_has_derivative_blinfun:
+lemma has_vector_derivative_eq_has_derivative_blinfun:
   "(f has_vector_derivative f') (at x within U) \<longleftrightarrow>
     (f has_derivative blinfun_scaleR_left f') (at x within U)"
   by (simp add: has_vector_derivative_def)
@@ -9867,8 +7273,7 @@
     done
 qed
 
-lemma
-  has_field_derivative_eq_has_derivative_blinfun:
+lemma has_field_derivative_eq_has_derivative_blinfun:
   "(f has_field_derivative f') (at x within U) \<longleftrightarrow> (f has_derivative blinfun_mult_right f') (at x within U)"
   by (simp add: has_field_derivative_def)
 
@@ -9900,8 +7305,7 @@
 
 subsection \<open>Exchange uniform limit and integral\<close>
 
-lemma
-  uniform_limit_integral:
+lemma uniform_limit_integral:
   fixes f::"'a \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
   assumes u: "uniform_limit (cbox a b) f g F"
   assumes c: "\<And>n. continuous_on (cbox a b) (f n)"
@@ -10095,18 +7499,6 @@
 
 subsection \<open>Compute a double integral using iterated integrals and switching the order of integration\<close>
 
-lemma setcomp_dot1: "{z. P (z \<bullet> (i,0))} = {(x,y). P(x \<bullet> i)}"
-  by auto
-
-lemma setcomp_dot2: "{z. P (z \<bullet> (0,i))} = {(x,y). P(y \<bullet> i)}"
-  by auto
-
-lemma Sigma_Int_Paircomp1: "(Sigma A B) \<inter> {(x, y). P x} = Sigma (A \<inter> {x. P x}) B"
-  by blast
-
-lemma Sigma_Int_Paircomp2: "(Sigma A B) \<inter> {(x, y). P y} = Sigma A (\<lambda>z. B z \<inter> {y. P y})"
-  by blast
-
 lemma continuous_on_imp_integrable_on_Pair1:
   fixes f :: "_ \<Rightarrow> 'b::banach"
   assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \<in> cbox a b"
@@ -10158,11 +7550,6 @@
     done
 qed
 
-lemma norm_diff2: "\<lbrakk>y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \<le> e1; norm(y2 - x2) \<le> e2\<rbrakk>
-            \<Longrightarrow> norm(y - x) \<le> e"
-using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"]
-  by (simp add: add_diff_add)
-
 lemma integral_split:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
   assumes f: "f integrable_on (cbox a b)"
@@ -10186,9 +7573,11 @@
                     \<le> e * content (cbox (a,c) (b,d)))"
 proof (auto simp: comm_monoid.operative_def[OF comm_monoid_and])
   fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b
-  assume c0: "content (cbox (a, c) (b, d)) = 0"
+  assume *: "box (a, c) (b, d) = {}"
      and cb1: "cbox (u, w) (v, z) \<subseteq> cbox (a, c) (b, d)"
      and cb2: "cbox (u, w) (v, z) \<subseteq> s"
+  then have c0: "content (cbox (a, c) (b, d)) = 0"
+    using * unfolding content_eq_0_interior by simp
   have c0': "content (cbox (u, w) (v, z)) = 0"
     by (fact content_0_subset [OF c0 cb1])
   show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))))
@@ -10274,9 +7663,6 @@
      integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. k))"
   by (simp add: content_Pair)
 
-lemma norm_minus2: "norm (x1-x2, y1-y2) = norm (x2-x1, y2-y1)"
-  by (simp add: norm_minus_eqI)
-
 lemma integral_prod_continuous:
   fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
   assumes "continuous_on (cbox (a,c) (b,d)) f" (is "continuous_on ?CBOX f")
@@ -10392,20 +7778,6 @@
     by simp
 qed
 
-lemma swap_continuous:
-  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
-    shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)"
-proof -
-  have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap"
-    by auto
-  then show ?thesis
-    apply (rule ssubst)
-    apply (rule continuous_on_compose)
-    apply (simp add: split_def)
-    apply (rule continuous_intros | simp add: assms)+
-    done
-qed
-
 lemma integral_swap_2dim:
   fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
   assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
--- a/src/HOL/Analysis/Lebesgue_Measure.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -8,7 +8,7 @@
 section \<open>Lebesgue measure\<close>
 
 theory Lebesgue_Measure
-  imports Finite_Product_Measure Bochner_Integration Caratheodory
+  imports Finite_Product_Measure Bochner_Integration Caratheodory Complete_Measure Summation_Tests
 begin
 
 subsection \<open>Every right continuous and nondecreasing function gives rise to a measure\<close>
@@ -356,6 +356,12 @@
 definition lborel :: "('a :: euclidean_space) measure" where
   "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. interval_measure (\<lambda>x. x)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
 
+abbreviation lebesgue :: "'a::euclidean_space measure"
+  where "lebesgue \<equiv> completion lborel"
+
+abbreviation lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
+  where "lebesgue_on \<Omega> \<equiv> restrict_space (completion lborel) \<Omega>"
+
 lemma
   shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel"
     and space_lborel[simp]: "space lborel = space borel"
@@ -659,6 +665,105 @@
     by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq)
 qed
 
+lemma
+  fixes c :: "'a::euclidean_space \<Rightarrow> real" and t
+  assumes c: "\<And>j. j \<in> Basis \<Longrightarrow> c j \<noteq> 0"
+  defines "T == (\<lambda>x. t + (\<Sum>j\<in>Basis. (c j * (x \<bullet> j)) *\<^sub>R j))"
+  shows lebesgue_affine_euclidean: "lebesgue = density (distr lebesgue lebesgue T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))" (is "_ = ?D")
+    and lebesgue_affine_measurable: "T \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
+proof -
+  have T_borel[measurable]: "T \<in> borel \<rightarrow>\<^sub>M borel"
+    by (auto simp: T_def[abs_def])
+  { fix A :: "'a set" assume A: "A \<in> sets borel"
+    then have "emeasure lborel A = 0 \<longleftrightarrow> emeasure (density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))) A = 0"
+      unfolding T_def using c by (subst lborel_affine_euclidean[symmetric]) auto
+    also have "\<dots> \<longleftrightarrow> emeasure (distr lebesgue lborel T) A = 0"
+      using A c by (simp add: distr_completion emeasure_density nn_integral_cmult setprod_nonneg cong: distr_cong)
+    finally have "emeasure lborel A = 0 \<longleftrightarrow> emeasure (distr lebesgue lborel T) A = 0" . }
+  then have eq: "null_sets lborel = null_sets (distr lebesgue lborel T)"
+    by (auto simp: null_sets_def)
+
+  show "T \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
+    by (rule completion.measurable_completion2) (auto simp: eq measurable_completion)
+
+  have "lebesgue = completion (density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>)))"
+    using c by (subst lborel_affine_euclidean[of c t]) (simp_all add: T_def[abs_def])
+  also have "\<dots> = density (completion (distr lebesgue lborel T)) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))"
+    using c by (auto intro!: always_eventually setprod_pos completion_density_eq simp: distr_completion cong: distr_cong)
+  also have "\<dots> = density (distr lebesgue lebesgue T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))"
+    by (subst completion.completion_distr_eq) (auto simp: eq measurable_completion)
+  finally show "lebesgue = density (distr lebesgue lebesgue T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))" .
+qed
+
+lemma lebesgue_measurable_scaling[measurable]: "op *\<^sub>R x \<in> lebesgue \<rightarrow>\<^sub>M lebesgue"
+proof cases
+  assume "x = 0"
+  then have "op *\<^sub>R x = (\<lambda>x. 0::'a)"
+    by (auto simp: fun_eq_iff)
+  then show ?thesis by auto
+next
+  assume "x \<noteq> 0" then show ?thesis
+    using lebesgue_affine_measurable[of "\<lambda>_. x" 0]
+    unfolding scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric] euclidean_representation
+    by (auto simp add: ac_simps)
+qed
+
+lemma
+  fixes m :: real and \<delta> :: "'a::euclidean_space"
+  defines "T r d x \<equiv> r *\<^sub>R x + d"
+  shows emeasure_lebesgue_affine: "emeasure lebesgue (T m \<delta> ` S) = \<bar>m\<bar> ^ DIM('a) * emeasure lebesgue S" (is ?e)
+    and measure_lebesgue_affine: "measure lebesgue (T m \<delta> ` S) = \<bar>m\<bar> ^ DIM('a) * measure lebesgue S" (is ?m)
+proof -
+  show ?e
+  proof cases
+    assume "m = 0" then show ?thesis
+      by (simp add: image_constant_conv T_def[abs_def])
+  next
+    let ?T = "T m \<delta>" and ?T' = "T (1 / m) (- ((1/m) *\<^sub>R \<delta>))"
+    assume "m \<noteq> 0"
+    then have s_comp_s: "?T' \<circ> ?T = id" "?T \<circ> ?T' = id"
+      by (auto simp: T_def[abs_def] fun_eq_iff scaleR_add_right scaleR_diff_right)
+    then have "inv ?T' = ?T" "bij ?T'"
+      by (auto intro: inv_unique_comp o_bij)
+    then have eq: "T m \<delta> ` S = T (1 / m) ((-1/m) *\<^sub>R \<delta>) -` S \<inter> space lebesgue"
+      using bij_vimage_eq_inv_image[OF \<open>bij ?T'\<close>, of S] by auto
+
+    have trans_eq_T: "(\<lambda>x. \<delta> + (\<Sum>j\<in>Basis. (m * (x \<bullet> j)) *\<^sub>R j)) = T m \<delta>" for m \<delta>
+      unfolding T_def[abs_def] scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric]
+      by (auto simp add: euclidean_representation ac_simps)
+
+    have T[measurable]: "T r d \<in> lebesgue \<rightarrow>\<^sub>M lebesgue" for r d
+      using lebesgue_affine_measurable[of "\<lambda>_. r" d]
+      by (cases "r = 0") (auto simp: trans_eq_T T_def[abs_def])
+
+    show ?thesis
+    proof cases
+      assume "S \<in> sets lebesgue" with \<open>m \<noteq> 0\<close> show ?thesis
+        unfolding eq
+        apply (subst lebesgue_affine_euclidean[of "\<lambda>_. m" \<delta>])
+        apply (simp_all add: emeasure_density trans_eq_T nn_integral_cmult emeasure_distr
+                        del: space_completion emeasure_completion)
+        apply (simp add: vimage_comp s_comp_s setprod_constant)
+        done
+    next
+      assume "S \<notin> sets lebesgue"
+      moreover have "?T ` S \<notin> sets lebesgue"
+      proof
+        assume "?T ` S \<in> sets lebesgue"
+        then have "?T -` (?T ` S) \<inter> space lebesgue \<in> sets lebesgue"
+          by (rule measurable_sets[OF T])
+        also have "?T -` (?T ` S) \<inter> space lebesgue = S"
+          by (simp add: vimage_comp s_comp_s eq)
+        finally show False using \<open>S \<notin> sets lebesgue\<close> by auto
+      qed
+      ultimately show ?thesis
+        by (simp add: emeasure_notin_sets)
+    qed
+  qed
+  show ?m
+    unfolding measure_def \<open>?e\<close> by (simp add: enn2real_mult setprod_nonneg)
+qed
+
 lemma divideR_right:
   fixes x y :: "'a::real_normed_vector"
   shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x"
@@ -780,4 +885,105 @@
     by (auto simp: mult.commute)
 qed
 
+abbreviation lmeasurable :: "'a::euclidean_space set set"
+where
+  "lmeasurable \<equiv> fmeasurable lebesgue"
+
+lemma lmeasurable_iff_integrable:
+  "S \<in> lmeasurable \<longleftrightarrow> integrable lebesgue (indicator S :: 'a::euclidean_space \<Rightarrow> real)"
+  by (auto simp: fmeasurable_def integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator)
+
+lemma lmeasurable_cbox [iff]: "cbox a b \<in> lmeasurable"
+  and lmeasurable_box [iff]: "box a b \<in> lmeasurable"
+  by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq)
+
+lemma lmeasurable_compact: "compact S \<Longrightarrow> S \<in> lmeasurable"
+  using emeasure_compact_finite[of S] by (intro fmeasurableI) (auto simp: borel_compact)
+
+lemma lmeasurable_open: "bounded S \<Longrightarrow> open S \<Longrightarrow> S \<in> lmeasurable"
+  using emeasure_bounded_finite[of S] by (intro fmeasurableI) (auto simp: borel_open)
+
+lemma lmeasurable_ball: "ball a r \<in> lmeasurable"
+  by (simp add: lmeasurable_open)
+
+lemma lmeasurable_interior: "bounded S \<Longrightarrow> interior S \<in> lmeasurable"
+  by (simp add: bounded_interior lmeasurable_open)
+
+lemma null_sets_cbox_Diff_box: "cbox a b - box a b \<in> null_sets lborel"
+proof -
+  have "emeasure lborel (cbox a b - box a b) = 0"
+    by (subst emeasure_Diff) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq box_subset_cbox)
+  then have "cbox a b - box a b \<in> null_sets lborel"
+    by (auto simp: null_sets_def)
+  then show ?thesis
+    by (auto dest!: AE_not_in)
+qed
+subsection\<open> A nice lemma for negligibility proofs.\<close>
+
+lemma summable_iff_suminf_neq_top: "(\<And>n. f n \<ge> 0) \<Longrightarrow> \<not> summable f \<Longrightarrow> (\<Sum>i. ennreal (f i)) = top"
+  by (metis summable_suminf_not_top)
+
+proposition starlike_negligible_bounded_gmeasurable:
+  fixes S :: "'a :: euclidean_space set"
+  assumes S: "S \<in> sets lebesgue" and "bounded S"
+      and eq1: "\<And>c x. \<lbrakk>(c *\<^sub>R x) \<in> S; 0 \<le> c; x \<in> S\<rbrakk> \<Longrightarrow> c = 1"
+    shows "S \<in> null_sets lebesgue"
+proof -
+  obtain M where "0 < M" "S \<subseteq> ball 0 M"
+    using \<open>bounded S\<close> by (auto dest: bounded_subset_ballD)
+
+  let ?f = "\<lambda>n. root DIM('a) (Suc n)"
+
+  have vimage_eq_image: "op *\<^sub>R (?f n) -` S = op *\<^sub>R (1 / ?f n) ` S" for n
+    apply safe
+    subgoal for x by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto
+    subgoal by auto
+    done
+
+  have eq: "(1 / ?f n) ^ DIM('a) = 1 / Suc n" for n
+    by (simp add: field_simps)
+
+  { fix n x assume x: "root DIM('a) (1 + real n) *\<^sub>R x \<in> S"
+    have "1 * norm x \<le> root DIM('a) (1 + real n) * norm x"
+      by (rule mult_mono) auto
+    also have "\<dots> < M"
+      using x \<open>S \<subseteq> ball 0 M\<close> by auto
+    finally have "norm x < M" by simp }
+  note less_M = this
+
+  have "(\<Sum>n. ennreal (1 / Suc n)) = top"
+    using not_summable_harmonic[where 'a=real] summable_Suc_iff[where f="\<lambda>n. 1 / (real n)"]
+    by (intro summable_iff_suminf_neq_top) (auto simp add: inverse_eq_divide)
+  then have "top * emeasure lebesgue S = (\<Sum>n. (1 / ?f n)^DIM('a) * emeasure lebesgue S)"
+    unfolding ennreal_suminf_multc eq by simp
+  also have "\<dots> = (\<Sum>n. emeasure lebesgue (op *\<^sub>R (?f n) -` S))"
+    unfolding vimage_eq_image using emeasure_lebesgue_affine[of "1 / ?f n" 0 S for n] by simp
+  also have "\<dots> = emeasure lebesgue (\<Union>n. op *\<^sub>R (?f n) -` S)"
+  proof (intro suminf_emeasure)
+    show "disjoint_family (\<lambda>n. op *\<^sub>R (?f n) -` S)"
+      unfolding disjoint_family_on_def
+    proof safe
+      fix m n :: nat and x assume "m \<noteq> n" "?f m *\<^sub>R x \<in> S" "?f n *\<^sub>R x \<in> S"
+      with eq1[of "?f m / ?f n" "?f n *\<^sub>R x"] show "x \<in> {}"
+        by auto
+    qed
+    have "op *\<^sub>R (?f i) -` S \<in> sets lebesgue" for i
+      using measurable_sets[OF lebesgue_measurable_scaling[of "?f i"] S] by auto
+    then show "range (\<lambda>i. op *\<^sub>R (?f i) -` S) \<subseteq> sets lebesgue"
+      by auto
+  qed
+  also have "\<dots> \<le> emeasure lebesgue (ball 0 M :: 'a set)"
+    using less_M by (intro emeasure_mono) auto
+  also have "\<dots> < top"
+    using lmeasurable_ball by (auto simp: fmeasurable_def)
+  finally have "emeasure lebesgue S = 0"
+    by (simp add: ennreal_top_mult split: if_split_asm)
+  then show "S \<in> null_sets lebesgue"
+    unfolding null_sets_def using \<open>S \<in> sets lebesgue\<close> by auto
+qed
+
+corollary starlike_negligible_compact:
+  "compact S \<Longrightarrow> (\<And>c x. \<lbrakk>(c *\<^sub>R x) \<in> S; 0 \<le> c; x \<in> S\<rbrakk> \<Longrightarrow> c = 1) \<Longrightarrow> S \<in> null_sets lebesgue"
+  using starlike_negligible_bounded_gmeasurable[of S] by (auto simp: compact_eq_bounded_closed)
+
 end
--- a/src/HOL/Analysis/Measure_Space.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -965,6 +965,16 @@
 translations
   "AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)"
 
+abbreviation
+  "set_almost_everywhere A M P \<equiv> AE x in M. x \<in> A \<longrightarrow> P x"
+
+syntax
+  "_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
+  ("AE _\<in>_ in _./ _" [0,0,0,10] 10)
+
+translations
+  "AE x\<in>A in M. P" \<rightleftharpoons> "CONST set_almost_everywhere A M (\<lambda>x. P)"
+
 lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
   unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)
 
@@ -1117,6 +1127,12 @@
     unfolding eventually_ae_filter by auto
 qed auto
 
+lemma pairwise_alt: "pairwise R S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S-{x}. R x y)"
+  by (auto simp add: pairwise_def)
+
+lemma AE_pairwise: "countable F \<Longrightarrow> pairwise (\<lambda>A B. AE x in M. R x A B) F \<longleftrightarrow> (AE x in M. pairwise (R x) F)"
+  unfolding pairwise_alt by (simp add: AE_ball_countable)
+
 lemma AE_discrete_difference:
   assumes X: "countable X"
   assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
@@ -1443,6 +1459,12 @@
   by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top
            del: real_of_ereal_enn2ereal)
 
+lemma measure_eq_AE:
+  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "measure M A = measure M B"
+  using assms emeasure_eq_AE[OF assms] by (simp add: measure_def)
+
 lemma measure_Union:
   "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M B \<noteq> \<infinity> \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow>
     measure M (A \<union> B) = measure M A + measure M B"
@@ -1544,6 +1566,12 @@
     done
 qed
 
+lemma measure_Un_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A \<union> B) = measure M A"
+  by (simp add: measure_def emeasure_Un_null_set)
+
+lemma measure_Diff_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A - B) = measure M A"
+  by (simp add: measure_def emeasure_Diff_null_set)
+
 lemma measure_eq_setsum_singleton:
   "finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
     measure M S = (\<Sum>x\<in>S. measure M {x})"
@@ -1576,6 +1604,151 @@
     using fin A by (auto intro!: Lim_emeasure_decseq)
 qed auto
 
+subsection \<open>Set of measurable sets with finite measure\<close>
+
+definition fmeasurable :: "'a measure \<Rightarrow> 'a set set"
+where
+  "fmeasurable M = {A\<in>sets M. emeasure M A < \<infinity>}"
+
+lemma fmeasurableD[dest, measurable_dest]: "A \<in> fmeasurable M \<Longrightarrow> A \<in> sets M"
+  by (auto simp: fmeasurable_def)
+
+lemma fmeasurableD2: "A \<in> fmeasurable M \<Longrightarrow> emeasure M A \<noteq> top"
+  by (auto simp: fmeasurable_def)
+
+lemma fmeasurableI: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> A \<in> fmeasurable M"
+  by (auto simp: fmeasurable_def)
+
+lemma fmeasurableI_null_sets: "A \<in> null_sets M \<Longrightarrow> A \<in> fmeasurable M"
+  by (auto simp: fmeasurable_def)
+
+lemma fmeasurableI2: "A \<in> fmeasurable M \<Longrightarrow> B \<subseteq> A \<Longrightarrow> B \<in> sets M \<Longrightarrow> B \<in> fmeasurable M"
+  using emeasure_mono[of B A M] by (auto simp: fmeasurable_def)
+
+lemma measure_mono_fmeasurable:
+  "A \<subseteq> B \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> measure M A \<le> measure M B"
+  by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono)
+
+lemma emeasure_eq_measure2: "A \<in> fmeasurable M \<Longrightarrow> emeasure M A = measure M A"
+  by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top)
+
+interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M"
+proof (rule ring_of_setsI)
+  show "fmeasurable M \<subseteq> Pow (space M)" "{} \<in> fmeasurable M"
+    by (auto simp: fmeasurable_def dest: sets.sets_into_space)
+  fix a b assume *: "a \<in> fmeasurable M" "b \<in> fmeasurable M"
+  then have "emeasure M (a \<union> b) \<le> emeasure M a + emeasure M b"
+    by (intro emeasure_subadditive) auto
+  also have "\<dots> < top"
+    using * by (auto simp: fmeasurable_def)
+  finally show  "a \<union> b \<in> fmeasurable M"
+    using * by (auto intro: fmeasurableI)
+  show "a - b \<in> fmeasurable M"
+    using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def Diff_subset)
+qed
+
+lemma fmeasurable_Diff: "A \<in> fmeasurable M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A - B \<in> fmeasurable M"
+  using fmeasurableI2[of A M "A - B"] by auto
+
+lemma fmeasurable_UN:
+  assumes "countable I" "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> A" "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" "A \<in> fmeasurable M"
+  shows "(\<Union>i\<in>I. F i) \<in> fmeasurable M"
+proof (rule fmeasurableI2)
+  show "A \<in> fmeasurable M" "(\<Union>i\<in>I. F i) \<subseteq> A" using assms by auto
+  show "(\<Union>i\<in>I. F i) \<in> sets M"
+    using assms by (intro sets.countable_UN') auto
+qed
+
+lemma fmeasurable_INT:
+  assumes "countable I" "i \<in> I" "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" "F i \<in> fmeasurable M"
+  shows "(\<Inter>i\<in>I. F i) \<in> fmeasurable M"
+proof (rule fmeasurableI2)
+  show "F i \<in> fmeasurable M" "(\<Inter>i\<in>I. F i) \<subseteq> F i"
+    using assms by auto
+  show "(\<Inter>i\<in>I. F i) \<in> sets M"
+    using assms by (intro sets.countable_INT') auto
+qed
+
+lemma measure_Un2:
+  "A \<in> fmeasurable M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)"
+  using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff)
+
+lemma measure_Un3:
+  assumes "A \<in> fmeasurable M" "B \<in> fmeasurable M"
+  shows "measure M (A \<union> B) = measure M A + measure M B - measure M (A \<inter> B)"
+proof -
+  have "measure M (A \<union> B) = measure M A + measure M (B - A)"
+    using assms by (rule measure_Un2)
+  also have "B - A = B - (A \<inter> B)"
+    by auto
+  also have "measure M (B - (A \<inter> B)) = measure M B - measure M (A \<inter> B)"
+    using assms by (intro measure_Diff) (auto simp: fmeasurable_def)
+  finally show ?thesis
+    by simp
+qed
+
+lemma measure_Un_AE:
+  "AE x in M. x \<notin> A \<or> x \<notin> B \<Longrightarrow> A \<in> fmeasurable M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow>
+  measure M (A \<union> B) = measure M A + measure M B"
+  by (subst measure_Un2) (auto intro!: measure_eq_AE)
+
+lemma measure_UNION_AE:
+  assumes I: "finite I"
+  shows "(\<And>i. i \<in> I \<Longrightarrow> F i \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>i j. AE x in M. x \<notin> F i \<or> x \<notin> F j) I \<Longrightarrow>
+    measure M (\<Union>i\<in>I. F i) = (\<Sum>i\<in>I. measure M (F i))"
+  unfolding AE_pairwise[OF countable_finite, OF I]
+  using I
+  apply (induction I rule: finite_induct)
+   apply simp
+  apply (simp add: pairwise_insert)
+  apply (subst measure_Un_AE)
+  apply auto
+  done
+
+lemma measure_UNION':
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>i j. disjnt (F i) (F j)) I \<Longrightarrow>
+    measure M (\<Union>i\<in>I. F i) = (\<Sum>i\<in>I. measure M (F i))"
+  by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually)
+
+lemma measure_Union_AE:
+  "finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>S T. AE x in M. x \<notin> S \<or> x \<notin> T) F \<Longrightarrow>
+    measure M (\<Union>F) = (\<Sum>S\<in>F. measure M S)"
+  using measure_UNION_AE[of F "\<lambda>x. x" M] by simp
+
+lemma measure_Union':
+  "finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> fmeasurable M) \<Longrightarrow> pairwise disjnt F \<Longrightarrow> measure M (\<Union>F) = (\<Sum>S\<in>F. measure M S)"
+  using measure_UNION'[of F "\<lambda>x. x" M] by simp
+
+lemma measure_Un_le:
+  assumes "A \<in> sets M" "B \<in> sets M" shows "measure M (A \<union> B) \<le> measure M A + measure M B"
+proof cases
+  assume "A \<in> fmeasurable M \<and> B \<in> fmeasurable M"
+  with measure_subadditive[of A M B] assms show ?thesis
+    by (auto simp: fmeasurableD2)
+next
+  assume "\<not> (A \<in> fmeasurable M \<and> B \<in> fmeasurable M)"
+  then have "A \<union> B \<notin> fmeasurable M"
+    using fmeasurableI2[of "A \<union> B" M A] fmeasurableI2[of "A \<union> B" M B] assms by auto
+  with assms show ?thesis
+    by (auto simp: fmeasurable_def measure_def less_top[symmetric])
+qed
+
+lemma measure_UNION_le:
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M) \<Longrightarrow> measure M (\<Union>i\<in>I. F i) \<le> (\<Sum>i\<in>I. measure M (F i))"
+proof (induction I rule: finite_induct)
+  case (insert i I)
+  then have "measure M (\<Union>i\<in>insert i I. F i) \<le> measure M (F i) + measure M (\<Union>i\<in>I. F i)"
+    by (auto intro!: measure_Un_le)
+  also have "measure M (\<Union>i\<in>I. F i) \<le> (\<Sum>i\<in>I. measure M (F i))"
+    using insert by auto
+  finally show ?case
+    using insert by simp
+qed simp
+
+lemma measure_Union_le:
+  "finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> sets M) \<Longrightarrow> measure M (\<Union>F) \<le> (\<Sum>S\<in>F. measure M S)"
+  using measure_UNION_le[of F "\<lambda>x. x" M] by simp
+
 subsection \<open>Measure spaces with @{term "emeasure M (space M) < \<infinity>"}\<close>
 
 locale finite_measure = sigma_finite_measure M for M +
@@ -1588,6 +1761,9 @@
 lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> top"
   using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique)
 
+lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M"
+  by (auto simp: fmeasurable_def less_top[symmetric])
+
 lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)"
   by (intro emeasure_eq_ennreal_measure) simp
 
@@ -3135,7 +3311,7 @@
 proof -
   { fix a m assume "a \<in> sigma_sets \<Omega> m" "m \<in> M"
     then have "a \<in> sigma_sets \<Omega> (\<Union>M)"
-     by induction (auto intro: sigma_sets.intros) }
+     by induction (auto intro: sigma_sets.intros(2-)) }
   then show "sets (SUP m:M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
     apply (subst sets_Sup_eq[where X="\<Omega>"])
     apply (auto simp add: M) []
@@ -3317,7 +3493,7 @@
   assume "?M \<noteq> 0"
   then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
     using reals_Archimedean[of "?m x / ?M" for x]
-    by (auto simp: field_simps not_le[symmetric] measure_nonneg divide_le_0_iff measure_le_0_iff)
+    by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff)
   have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
   proof (rule ccontr)
     fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
--- a/src/HOL/Analysis/Path_Connected.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -2315,10 +2315,10 @@
    "outside s \<inter> s = {}"
   by (auto simp: outside_def)
 
-lemma inside_inter_outside [simp]: "inside s \<inter> outside s = {}"
+lemma inside_Int_outside [simp]: "inside s \<inter> outside s = {}"
   by (auto simp: inside_def outside_def)
 
-lemma inside_union_outside [simp]: "inside s \<union> outside s = (- s)"
+lemma inside_Un_outside [simp]: "inside s \<union> outside s = (- s)"
   by (auto simp: inside_def outside_def)
 
 lemma inside_eq_outside:
@@ -2606,7 +2606,7 @@
   by (simp add: inside_def connected_component_UNIV)
 
 lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)"
-using inside_empty inside_union_outside by blast
+using inside_empty inside_Un_outside by blast
 
 lemma inside_same_component:
    "\<lbrakk>connected_component (- s) x y; x \<in> inside s\<rbrakk> \<Longrightarrow> y \<in> inside s"
@@ -2666,7 +2666,7 @@
   fixes s :: "'a :: {real_normed_vector, perfect_space} set"
   assumes "convex s"
     shows "outside s = - s"
-  by (metis ComplD assms convex_in_outside equalityI inside_union_outside subsetI sup.cobounded2)
+  by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2)
 
 lemma inside_convex:
   fixes s :: "'a :: {real_normed_vector, perfect_space} set"
@@ -2761,7 +2761,7 @@
   have "closure (inside s) \<inter> - inside s = closure (inside s) - interior (inside s)"
     by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside)
   moreover have "- inside s \<inter> - outside s = s"
-    by (metis (no_types) compl_sup double_compl inside_union_outside)
+    by (metis (no_types) compl_sup double_compl inside_Un_outside)
   moreover have "closure (inside s) \<subseteq> - outside s"
     by (metis (no_types) assms closure_inside_subset union_with_inside)
   ultimately have "closure (inside s) - interior (inside s) \<subseteq> s"
@@ -5460,21 +5460,6 @@
 
 subsection\<open>Components, continuity, openin, closedin\<close>
 
-lemma continuous_openin_preimage_eq:
-   "continuous_on S f \<longleftrightarrow>
-    (\<forall>t. open t \<longrightarrow> openin (subtopology euclidean S) {x. x \<in> S \<and> f x \<in> t})"
-apply (auto simp: continuous_openin_preimage_gen)
-apply (fastforce simp add: continuous_on_open openin_open)
-done
-
-lemma continuous_closedin_preimage_eq:
-   "continuous_on S f \<longleftrightarrow>
-    (\<forall>t. closed t \<longrightarrow> closedin (subtopology euclidean S) {x. x \<in> S \<and> f x \<in> t})"
-apply safe
-apply (simp add: continuous_closedin_preimage)
-apply (fastforce simp add: continuous_on_closed closedin_closed)
-done
-
 lemma continuous_on_components_gen:
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
   assumes "\<And>c. c \<in> components S \<Longrightarrow>
@@ -6108,7 +6093,6 @@
 apply (metis assms homotopy_eqv_homotopic_triviality_imp)
 by (metis (no_types) assms homotopy_eqv_homotopic_triviality_imp homotopy_eqv_sym)
 
-
 lemma homotopy_eqv_cohomotopic_triviality_null_imp:
   fixes S :: "'a::real_normed_vector set"
     and T :: "'b::real_normed_vector set"
@@ -6156,7 +6140,6 @@
 apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp)
 by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_eqv_sym)
 
-
 lemma homotopy_eqv_contractible_sets:
   fixes S :: "'a::real_normed_vector set"
     and T :: "'b::real_normed_vector set"
@@ -6215,6 +6198,50 @@
 lemma homeomorphic_contractible:
   fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
   shows "\<lbrakk>contractible S; S homeomorphic T\<rbrakk> \<Longrightarrow> contractible T"
-by (metis homeomorphic_contractible_eq)
+  by (metis homeomorphic_contractible_eq)
+
+subsection\<open>Misc other results\<close>
+
+lemma bounded_connected_Compl_real:
+  fixes S :: "real set"
+  assumes "bounded S" and conn: "connected(- S)"
+    shows "S = {}"
+proof -
+  obtain a b where "S \<subseteq> box a b"
+    by (meson assms bounded_subset_open_interval)
+  then have "a \<notin> S" "b \<notin> S"
+    by auto
+  then have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> - S"
+    by (meson Compl_iff conn connected_iff_interval)
+  then show ?thesis
+    using \<open>S \<subseteq> box a b\<close> by auto
+qed
+
+lemma bounded_connected_Compl_1:
+  fixes S :: "'a::{euclidean_space} set"
+  assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1"
+    shows "S = {}"
+proof -
+  have "DIM('a) = DIM(real)"
+    by (simp add: "1")
+  then obtain f::"'a \<Rightarrow> real" and g
+  where "linear f" "\<And>x. norm(f x) = norm x" "\<And>x. g(f x) = x" "\<And>y. f(g y) = y"
+    by (rule isomorphisms_UNIV_UNIV) blast
+  with \<open>bounded S\<close> have "bounded (f ` S)"
+    using bounded_linear_image linear_linear by blast
+  have "connected (f ` (-S))"
+    using connected_linear_image assms \<open>linear f\<close> by blast
+  moreover have "f ` (-S) = - (f ` S)"
+    apply (rule bij_image_Compl_eq)
+    apply (auto simp: bij_def)
+     apply (metis \<open>\<And>x. g (f x) = x\<close> injI)
+    by (metis UNIV_I \<open>\<And>y. f (g y) = y\<close> image_iff)
+  finally have "connected (- (f ` S))"
+    by simp
+  then have "f ` S = {}"
+    using \<open>bounded (f ` S)\<close> bounded_connected_Compl_real by blast
+  then show ?thesis
+    by blast
+qed
 
 end
--- a/src/HOL/Analysis/Set_Integral.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Set_Integral.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -27,16 +27,6 @@
 translations
 "LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
 
-abbreviation
-  "set_almost_everywhere A M P \<equiv> AE x in M. x \<in> A \<longrightarrow> P x"
-
-syntax
-  "_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
-("AE _\<in>_ in _./ _" [0,0,0,10] 10)
-
-translations
-  "AE x\<in>A in M. P" == "CONST set_almost_everywhere A M (\<lambda>x. P)"
-
 (*
     Notation for integration wrt lebesgue measure on the reals:
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/Tagged_Division.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,2817 @@
+(*  Title:      HOL/Analysis/Tagged_Division.thy
+    Author:     John Harrison
+    Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP
+*)
+
+section \<open>Tagged divisions used for the Henstock-Kurzweil gauge integration\<close>
+
+theory Tagged_Division
+imports
+  Topology_Euclidean_Space
+begin
+
+lemma finite_product_dependent:
+  assumes "finite s"
+    and "\<And>x. x \<in> s \<Longrightarrow> finite (t x)"
+  shows "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
+  using assms
+proof induct
+  case (insert x s)
+  have *: "{(i, j) |i j. i \<in> insert x s \<and> j \<in> t i} =
+    (\<lambda>y. (x,y)) ` (t x) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
+  show ?case
+    unfolding *
+    apply (rule finite_UnI)
+    using insert
+    apply auto
+    done
+qed auto
+
+lemma sum_sum_product:
+  assumes "finite s"
+    and "\<forall>i\<in>s. finite (t i)"
+  shows "setsum (\<lambda>i. setsum (x i) (t i)::real) s =
+    setsum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}"
+  using assms
+proof induct
+  case (insert a s)
+  have *: "{(i, j) |i j. i \<in> insert a s \<and> j \<in> t i} =
+    (\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
+  show ?case
+    unfolding *
+    apply (subst setsum.union_disjoint)
+    unfolding setsum.insert[OF insert(1-2)]
+    prefer 4
+    apply (subst insert(3))
+    unfolding add_right_cancel
+  proof -
+    show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
+      apply (subst setsum.reindex)
+      unfolding inj_on_def
+      apply auto
+      done
+    show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
+      apply (rule finite_product_dependent)
+      using insert
+      apply auto
+      done
+  qed (insert insert, auto)
+qed auto
+
+lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
+  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
+  scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
+
+
+subsection \<open>Sundries\<close>
+
+
+text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
+lemma wf_finite_segments:
+  assumes "irrefl r" and "trans r" and "\<And>x. finite {y. (y, x) \<in> r}"
+  shows "wf (r)"
+  apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms)
+  using acyclic_def assms irrefl_def trans_Restr by fastforce
+
+text\<open>For creating values between @{term u} and @{term v}.\<close>
+lemma scaling_mono:
+  fixes u::"'a::linordered_field"
+  assumes "u \<le> v" "0 \<le> r" "r \<le> s"
+    shows "u + r * (v - u) / s \<le> v"
+proof -
+  have "r/s \<le> 1" using assms
+    using divide_le_eq_1 by fastforce
+  then have "(r/s) * (v - u) \<le> 1 * (v - u)"
+    by (meson diff_ge_0_iff_ge mult_right_mono \<open>u \<le> v\<close>)
+  then show ?thesis
+    by (simp add: field_simps)
+qed
+
+lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
+lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto
+lemma conjunctD4: assumes "a \<and> b \<and> c \<and> d" shows a b c d using assms by auto
+
+lemma cond_cases: "(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)"
+  by auto
+
+declare norm_triangle_ineq4[intro]
+
+lemma transitive_stepwise_le:
+  assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" and "\<And>n. R n (Suc n)"
+  shows "\<forall>n\<ge>m. R m n"
+proof (intro allI impI)
+  show "m \<le> n \<Longrightarrow> R m n" for n
+    by (induction rule: dec_induct)
+       (use assms in blast)+
+qed
+
+subsection \<open>Some useful lemmas about intervals.\<close>
+
+lemma interior_subset_union_intervals:
+  assumes "i = cbox a b"
+    and "j = cbox c d"
+    and "interior j \<noteq> {}"
+    and "i \<subseteq> j \<union> s"
+    and "interior i \<inter> interior j = {}"
+  shows "interior i \<subseteq> interior s"
+proof -
+  have "box a b \<inter> cbox c d = {}"
+     using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
+     unfolding assms(1,2) interior_cbox by auto
+  moreover
+  have "box a b \<subseteq> cbox c d \<union> s"
+    apply (rule order_trans,rule box_subset_cbox)
+    using assms(4) unfolding assms(1,2)
+    apply auto
+    done
+  ultimately
+  show ?thesis
+    unfolding assms interior_cbox
+      by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
+qed
+
+lemma interior_Union_subset_cbox:
+  assumes "finite f"
+  assumes f: "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a b" "\<And>s. s \<in> f \<Longrightarrow> interior s \<subseteq> t"
+    and t: "closed t"
+  shows "interior (\<Union>f) \<subseteq> t"
+proof -
+  have [simp]: "s \<in> f \<Longrightarrow> closed s" for s
+    using f by auto
+  define E where "E = {s\<in>f. interior s = {}}"
+  then have "finite E" "E \<subseteq> {s\<in>f. interior s = {}}"
+    using \<open>finite f\<close> by auto
+  then have "interior (\<Union>f) = interior (\<Union>(f - E))"
+  proof (induction E rule: finite_subset_induct')
+    case (insert s f')
+    have "interior (\<Union>(f - insert s f') \<union> s) = interior (\<Union>(f - insert s f'))"
+      using insert.hyps \<open>finite f\<close> by (intro interior_closed_Un_empty_interior) auto
+    also have "\<Union>(f - insert s f') \<union> s = \<Union>(f - f')"
+      using insert.hyps by auto
+    finally show ?case
+      by (simp add: insert.IH)
+  qed simp
+  also have "\<dots> \<subseteq> \<Union>(f - E)"
+    by (rule interior_subset)
+  also have "\<dots> \<subseteq> t"
+  proof (rule Union_least)
+    fix s assume "s \<in> f - E"
+    with f[of s] obtain a b where s: "s \<in> f" "s = cbox a b" "box a b \<noteq> {}"
+      by (fastforce simp: E_def)
+    have "closure (interior s) \<subseteq> closure t"
+      by (intro closure_mono f \<open>s \<in> f\<close>)
+    with s \<open>closed t\<close> show "s \<subseteq> t"
+      by simp
+  qed
+  finally show ?thesis .
+qed
+
+lemma inter_interior_unions_intervals:
+    "finite f \<Longrightarrow> open s \<Longrightarrow> \<forall>t\<in>f. \<exists>a b. t = cbox a b \<Longrightarrow> \<forall>t\<in>f. s \<inter> (interior t) = {} \<Longrightarrow> s \<inter> interior (\<Union>f) = {}"
+  using interior_Union_subset_cbox[of f "UNIV - s"] by auto
+
+lemma interval_split:
+  fixes a :: "'a::euclidean_space"
+  assumes "k \<in> Basis"
+  shows
+    "cbox a b \<inter> {x. x\<bullet>k \<le> c} = cbox a (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)"
+    "cbox a b \<inter> {x. x\<bullet>k \<ge> c} = cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) b"
+  apply (rule_tac[!] set_eqI)
+  unfolding Int_iff mem_box mem_Collect_eq
+  using assms
+  apply auto
+  done
+
+lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}"
+  by (simp add: box_ne_empty)
+
+subsection \<open>Bounds on intervals where they exist.\<close>
+
+definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
+  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
+
+definition interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
+  where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
+
+lemma interval_upperbound[simp]:
+  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
+    interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
+  unfolding interval_upperbound_def euclidean_representation_setsum cbox_def
+  by (safe intro!: cSup_eq) auto
+
+lemma interval_lowerbound[simp]:
+  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
+    interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
+  unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def
+  by (safe intro!: cInf_eq) auto
+
+lemmas interval_bounds = interval_upperbound interval_lowerbound
+
+lemma
+  fixes X::"real set"
+  shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
+    and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
+  by (auto simp: interval_upperbound_def interval_lowerbound_def)
+
+lemma interval_bounds'[simp]:
+  assumes "cbox a b \<noteq> {}"
+  shows "interval_upperbound (cbox a b) = b"
+    and "interval_lowerbound (cbox a b) = a"
+  using assms unfolding box_ne_empty by auto
+
+lemma interval_upperbound_Times:
+  assumes "A \<noteq> {}" and "B \<noteq> {}"
+  shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)"
+proof-
+  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:A. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
+  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
+  ultimately show ?thesis unfolding interval_upperbound_def
+      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
+qed
+
+lemma interval_lowerbound_Times:
+  assumes "A \<noteq> {}" and "B \<noteq> {}"
+  shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)"
+proof-
+  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:A. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
+  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
+  ultimately show ?thesis unfolding interval_lowerbound_def
+      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
+qed
+
+subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
+
+definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))"
+
+lemma gaugeI:
+  assumes "\<And>x. x \<in> g x"
+    and "\<And>x. open (g x)"
+  shows "gauge g"
+  using assms unfolding gauge_def by auto
+
+lemma gaugeD[dest]:
+  assumes "gauge d"
+  shows "x \<in> d x"
+    and "open (d x)"
+  using assms unfolding gauge_def by auto
+
+lemma gauge_ball_dependent: "\<forall>x. 0 < e x \<Longrightarrow> gauge (\<lambda>x. ball x (e x))"
+  unfolding gauge_def by auto
+
+lemma gauge_ball[intro]: "0 < e \<Longrightarrow> gauge (\<lambda>x. ball x e)"
+  unfolding gauge_def by auto
+
+lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)"
+  by (rule gauge_ball) auto
+
+lemma gauge_inter[intro]: "gauge d1 \<Longrightarrow> gauge d2 \<Longrightarrow> gauge (\<lambda>x. d1 x \<inter> d2 x)"
+  unfolding gauge_def by auto
+
+lemma gauge_inters:
+  assumes "finite s"
+    and "\<forall>d\<in>s. gauge (f d)"
+  shows "gauge (\<lambda>x. \<Inter>{f d x | d. d \<in> s})"
+proof -
+  have *: "\<And>x. {f d x |d. d \<in> s} = (\<lambda>d. f d x) ` s"
+    by auto
+  show ?thesis
+    unfolding gauge_def unfolding *
+    using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
+qed
+
+lemma gauge_existence_lemma:
+  "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)"
+  by (metis zero_less_one)
+
+subsection \<open>Attempt a systematic general set of "offset" results for components.\<close>
+
+lemma gauge_modify:
+  assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
+  shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
+  using assms
+  unfolding gauge_def
+  apply safe
+  defer
+  apply (erule_tac x="f x" in allE)
+  apply (erule_tac x="d (f x)" in allE)
+  apply auto
+  done
+
+subsection \<open>Divisions.\<close>
+
+definition division_of (infixl "division'_of" 40)
+where
+  "s division_of i \<longleftrightarrow>
+    finite s \<and>
+    (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = cbox a b)) \<and>
+    (\<forall>k1\<in>s. \<forall>k2\<in>s. k1 \<noteq> k2 \<longrightarrow> interior(k1) \<inter> interior(k2) = {}) \<and>
+    (\<Union>s = i)"
+
+lemma division_ofD[dest]:
+  assumes "s division_of i"
+  shows "finite s"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
+    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}"
+    and "\<Union>s = i"
+  using assms unfolding division_of_def by auto
+
+lemma division_ofI:
+  assumes "finite s"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
+    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+    and "\<Union>s = i"
+  shows "s division_of i"
+  using assms unfolding division_of_def by auto
+
+lemma division_of_finite: "s division_of i \<Longrightarrow> finite s"
+  unfolding division_of_def by auto
+
+lemma division_of_self[intro]: "cbox a b \<noteq> {} \<Longrightarrow> {cbox a b} division_of (cbox a b)"
+  unfolding division_of_def by auto
+
+lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}"
+  unfolding division_of_def by auto
+
+lemma division_of_sing[simp]:
+  "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}"
+  (is "?l = ?r")
+proof
+  assume ?r
+  moreover
+  { fix k
+    assume "s = {{a}}" "k\<in>s"
+    then have "\<exists>x y. k = cbox x y"
+      apply (rule_tac x=a in exI)+
+      apply (force simp: cbox_sing)
+      done
+  }
+  ultimately show ?l
+    unfolding division_of_def cbox_sing by auto
+next
+  assume ?l
+  note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]]
+  {
+    fix x
+    assume x: "x \<in> s" have "x = {a}"
+      using *(2)[rule_format,OF x] by auto
+  }
+  moreover have "s \<noteq> {}"
+    using *(4) by auto
+  ultimately show ?r
+    unfolding cbox_sing by auto
+qed
+
+lemma elementary_empty: obtains p where "p division_of {}"
+  unfolding division_of_trivial by auto
+
+lemma elementary_interval: obtains p where "p division_of (cbox a b)"
+  by (metis division_of_trivial division_of_self)
+
+lemma division_contains: "s division_of i \<Longrightarrow> \<forall>x\<in>i. \<exists>k\<in>s. x \<in> k"
+  unfolding division_of_def by auto
+
+lemma forall_in_division:
+  "d division_of i \<Longrightarrow> (\<forall>x\<in>d. P x) \<longleftrightarrow> (\<forall>a b. cbox a b \<in> d \<longrightarrow> P (cbox a b))"
+  unfolding division_of_def by fastforce
+
+lemma division_of_subset:
+  assumes "p division_of (\<Union>p)"
+    and "q \<subseteq> p"
+  shows "q division_of (\<Union>q)"
+proof (rule division_ofI)
+  note * = division_ofD[OF assms(1)]
+  show "finite q"
+    using "*"(1) assms(2) infinite_super by auto
+  {
+    fix k
+    assume "k \<in> q"
+    then have kp: "k \<in> p"
+      using assms(2) by auto
+    show "k \<subseteq> \<Union>q"
+      using \<open>k \<in> q\<close> by auto
+    show "\<exists>a b. k = cbox a b"
+      using *(4)[OF kp] by auto
+    show "k \<noteq> {}"
+      using *(3)[OF kp] by auto
+  }
+  fix k1 k2
+  assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2"
+  then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2"
+    using assms(2) by auto
+  show "interior k1 \<inter> interior k2 = {}"
+    using *(5)[OF **] by auto
+qed auto
+
+lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)"
+  unfolding division_of_def by auto
+
+lemma division_inter:
+  fixes s1 s2 :: "'a::euclidean_space set"
+  assumes "p1 division_of s1"
+    and "p2 division_of s2"
+  shows "{k1 \<inter> k2 | k1 k2. k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)"
+  (is "?A' division_of _")
+proof -
+  let ?A = "{s. s \<in>  (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}"
+  have *: "?A' = ?A" by auto
+  show ?thesis
+    unfolding *
+  proof (rule division_ofI)
+    have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)"
+      by auto
+    moreover have "finite (p1 \<times> p2)"
+      using assms unfolding division_of_def by auto
+    ultimately show "finite ?A" by auto
+    have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s"
+      by auto
+    show "\<Union>?A = s1 \<inter> s2"
+      apply (rule set_eqI)
+      unfolding * and UN_iff
+      using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)]
+      apply auto
+      done
+    {
+      fix k
+      assume "k \<in> ?A"
+      then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}"
+        by auto
+      then show "k \<noteq> {}"
+        by auto
+      show "k \<subseteq> s1 \<inter> s2"
+        using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)]
+        unfolding k by auto
+      obtain a1 b1 where k1: "k1 = cbox a1 b1"
+        using division_ofD(4)[OF assms(1) k(2)] by blast
+      obtain a2 b2 where k2: "k2 = cbox a2 b2"
+        using division_ofD(4)[OF assms(2) k(3)] by blast
+      show "\<exists>a b. k = cbox a b"
+        unfolding k k1 k2 unfolding Int_interval by auto
+    }
+    fix k1 k2
+    assume "k1 \<in> ?A"
+    then obtain x1 y1 where k1: "k1 = x1 \<inter> y1" "x1 \<in> p1" "y1 \<in> p2" "k1 \<noteq> {}"
+      by auto
+    assume "k2 \<in> ?A"
+    then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}"
+      by auto
+    assume "k1 \<noteq> k2"
+    then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2"
+      unfolding k1 k2 by auto
+    have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow>
+      interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow>
+      interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow>
+      interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto
+    show "interior k1 \<inter> interior k2 = {}"
+      unfolding k1 k2
+      apply (rule *)
+      using assms division_ofD(5) k1 k2(2) k2(3) th apply auto
+      done
+  qed
+qed
+
+lemma division_inter_1:
+  assumes "d division_of i"
+    and "cbox a (b::'a::euclidean_space) \<subseteq> i"
+  shows "{cbox a b \<inter> k | k. k \<in> d \<and> cbox a b \<inter> k \<noteq> {}} division_of (cbox a b)"
+proof (cases "cbox a b = {}")
+  case True
+  show ?thesis
+    unfolding True and division_of_trivial by auto
+next
+  case False
+  have *: "cbox a b \<inter> i = cbox a b" using assms(2) by auto
+  show ?thesis
+    using division_inter[OF division_of_self[OF False] assms(1)]
+    unfolding * by auto
+qed
+
+lemma elementary_inter:
+  fixes s t :: "'a::euclidean_space set"
+  assumes "p1 division_of s"
+    and "p2 division_of t"
+  shows "\<exists>p. p division_of (s \<inter> t)"
+using assms division_inter by blast
+
+lemma elementary_inters:
+  assumes "finite f"
+    and "f \<noteq> {}"
+    and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)"
+  shows "\<exists>p. p division_of (\<Inter>f)"
+  using assms
+proof (induct f rule: finite_induct)
+  case (insert x f)
+  show ?case
+  proof (cases "f = {}")
+    case True
+    then show ?thesis
+      unfolding True using insert by auto
+  next
+    case False
+    obtain p where "p division_of \<Inter>f"
+      using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
+    moreover obtain px where "px division_of x"
+      using insert(5)[rule_format,OF insertI1] ..
+    ultimately show ?thesis
+      by (simp add: elementary_inter Inter_insert)
+  qed
+qed auto
+
+lemma division_disjoint_union:
+  assumes "p1 division_of s1"
+    and "p2 division_of s2"
+    and "interior s1 \<inter> interior s2 = {}"
+  shows "(p1 \<union> p2) division_of (s1 \<union> s2)"
+proof (rule division_ofI)
+  note d1 = division_ofD[OF assms(1)]
+  note d2 = division_ofD[OF assms(2)]
+  show "finite (p1 \<union> p2)"
+    using d1(1) d2(1) by auto
+  show "\<Union>(p1 \<union> p2) = s1 \<union> s2"
+    using d1(6) d2(6) by auto
+  {
+    fix k1 k2
+    assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
+    moreover
+    let ?g="interior k1 \<inter> interior k2 = {}"
+    {
+      assume as: "k1\<in>p1" "k2\<in>p2"
+      have ?g
+        using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]]
+        using assms(3) by blast
+    }
+    moreover
+    {
+      assume as: "k1\<in>p2" "k2\<in>p1"
+      have ?g
+        using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]]
+        using assms(3) by blast
+    }
+    ultimately show ?g
+      using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto
+  }
+  fix k
+  assume k: "k \<in> p1 \<union> p2"
+  show "k \<subseteq> s1 \<union> s2"
+    using k d1(2) d2(2) by auto
+  show "k \<noteq> {}"
+    using k d1(3) d2(3) by auto
+  show "\<exists>a b. k = cbox a b"
+    using k d1(4) d2(4) by auto
+qed
+
+lemma partial_division_extend_1:
+  fixes a b c d :: "'a::euclidean_space"
+  assumes incl: "cbox c d \<subseteq> cbox a b"
+    and nonempty: "cbox c d \<noteq> {}"
+  obtains p where "p division_of (cbox a b)" "cbox c d \<in> p"
+proof
+  let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a.
+    cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)"
+  define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
+
+  show "cbox c d \<in> p"
+    unfolding p_def
+    by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
+  {
+    fix i :: 'a
+    assume "i \<in> Basis"
+    with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i"
+      unfolding box_eq_empty subset_box by (auto simp: not_le)
+  }
+  note ord = this
+
+  show "p division_of (cbox a b)"
+  proof (rule division_ofI)
+    show "finite p"
+      unfolding p_def by (auto intro!: finite_PiE)
+    {
+      fix k
+      assume "k \<in> p"
+      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
+        by (auto simp: p_def)
+      then show "\<exists>a b. k = cbox a b"
+        by auto
+      have "k \<subseteq> cbox a b \<and> k \<noteq> {}"
+      proof (simp add: k box_eq_empty subset_box not_less, safe)
+        fix i :: 'a
+        assume i: "i \<in> Basis"
+        with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+          by (auto simp: PiE_iff)
+        with i ord[of i]
+        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
+          by auto
+      qed
+      then show "k \<noteq> {}" "k \<subseteq> cbox a b"
+        by auto
+      {
+        fix l
+        assume "l \<in> p"
+        then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
+          by (auto simp: p_def)
+        assume "l \<noteq> k"
+        have "\<exists>i\<in>Basis. f i \<noteq> g i"
+        proof (rule ccontr)
+          assume "\<not> ?thesis"
+          with f g have "f = g"
+            by (auto simp: PiE_iff extensional_def fun_eq_iff)
+          with \<open>l \<noteq> k\<close> show False
+            by (simp add: l k)
+        qed
+        then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
+        then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+                  "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
+          using f g by (auto simp: PiE_iff)
+        with * ord[of i] show "interior l \<inter> interior k = {}"
+          by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i])
+      }
+      note \<open>k \<subseteq> cbox a b\<close>
+    }
+    moreover
+    {
+      fix x assume x: "x \<in> cbox a b"
+      have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
+      proof
+        fix i :: 'a
+        assume "i \<in> Basis"
+        with x ord[of i]
+        have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
+            (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
+          by (auto simp: cbox_def)
+        then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
+          by auto
+      qed
+      then obtain f where
+        f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}"
+        unfolding bchoice_iff ..
+      moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
+        by auto
+      moreover from f have "x \<in> ?B (restrict f Basis)"
+        by (auto simp: mem_box)
+      ultimately have "\<exists>k\<in>p. x \<in> k"
+        unfolding p_def by blast
+    }
+    ultimately show "\<Union>p = cbox a b"
+      by auto
+  qed
+qed
+
+lemma partial_division_extend_interval:
+  assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b"
+  obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
+proof (cases "p = {}")
+  case True
+  obtain q where "q division_of (cbox a b)"
+    by (rule elementary_interval)
+  then show ?thesis
+    using True that by blast
+next
+  case False
+  note p = division_ofD[OF assms(1)]
+  have div_cbox: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
+  proof
+    fix k
+    assume kp: "k \<in> p"
+    obtain c d where k: "k = cbox c d"
+      using p(4)[OF kp] by blast
+    have *: "cbox c d \<subseteq> cbox a b" "cbox c d \<noteq> {}"
+      using p(2,3)[OF kp, unfolded k] using assms(2)
+      by (blast intro: order.trans)+
+    obtain q where "q division_of cbox a b" "cbox c d \<in> q"
+      by (rule partial_division_extend_1[OF *])
+    then show "\<exists>q. q division_of cbox a b \<and> k \<in> q"
+      unfolding k by auto
+  qed
+  obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
+    using bchoice[OF div_cbox] by blast
+  { fix x
+    assume x: "x \<in> p"
+    have "q x division_of \<Union>q x"
+      apply (rule division_ofI)
+      using division_ofD[OF q(1)[OF x]]
+      apply auto
+      done }
+  then have "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
+    by (meson Diff_subset division_of_subset)
+  then have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
+    apply -
+    apply (rule elementary_inters [OF finite_imageI[OF p(1)]])
+    apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]])
+    done
+  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" ..
+  have "d \<union> p division_of cbox a b"
+  proof -
+    have te: "\<And>s f t. s \<noteq> {} \<Longrightarrow> \<forall>i\<in>s. f i \<union> i = t \<Longrightarrow> t = \<Inter>(f ` s) \<union> \<Union>s" by auto
+    have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
+    proof (rule te[OF False], clarify)
+      fix i
+      assume i: "i \<in> p"
+      show "\<Union>(q i - {i}) \<union> i = cbox a b"
+        using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto
+    qed
+    { fix k
+      assume k: "k \<in> p"
+      have *: "\<And>u t s. t \<inter> s = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<inter> t = {}"
+        by auto
+      have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior k = {}"
+      proof (rule *[OF inter_interior_unions_intervals])
+        note qk=division_ofD[OF q(1)[OF k]]
+        show "finite (q k - {k})" "open (interior k)" "\<forall>t\<in>q k - {k}. \<exists>a b. t = cbox a b"
+          using qk by auto
+        show "\<forall>t\<in>q k - {k}. interior k \<inter> interior t = {}"
+          using qk(5) using q(2)[OF k] by auto
+        show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q k - {k}))"
+          apply (rule interior_mono)+
+          using k
+          apply auto
+          done
+      qed } note [simp] = this
+    show "d \<union> p division_of (cbox a b)"
+      unfolding cbox_eq
+      apply (rule division_disjoint_union[OF d assms(1)])
+      apply (rule inter_interior_unions_intervals)
+      apply (rule p open_interior ballI)+
+      apply simp_all
+      done
+  qed
+  then show ?thesis
+    by (meson Un_upper2 that)
+qed
+
+lemma elementary_bounded[dest]:
+  fixes s :: "'a::euclidean_space set"
+  shows "p division_of s \<Longrightarrow> bounded s"
+  unfolding division_of_def by (metis bounded_Union bounded_cbox)
+
+lemma elementary_subset_cbox:
+  "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> cbox a (b::'a::euclidean_space)"
+  by (meson elementary_bounded bounded_subset_cbox)
+
+lemma division_union_intervals_exists:
+  fixes a b :: "'a::euclidean_space"
+  assumes "cbox a b \<noteq> {}"
+  obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
+proof (cases "cbox c d = {}")
+  case True
+  show ?thesis
+    apply (rule that[of "{}"])
+    unfolding True
+    using assms
+    apply auto
+    done
+next
+  case False
+  show ?thesis
+  proof (cases "cbox a b \<inter> cbox c d = {}")
+    case True
+    then show ?thesis
+      by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty)
+  next
+    case False
+    obtain u v where uv: "cbox a b \<inter> cbox c d = cbox u v"
+      unfolding Int_interval by auto
+    have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto
+    obtain p where "p division_of cbox c d" "cbox u v \<in> p"
+      by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
+    note p = this division_ofD[OF this(1)]
+    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))"
+      apply (rule arg_cong[of _ _ interior])
+      using p(8) uv by auto
+    also have "\<dots> = {}"
+      unfolding interior_Int
+      apply (rule inter_interior_unions_intervals)
+      using p(6) p(7)[OF p(2)] p(3)
+      apply auto
+      done
+    finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp
+    have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(p - {cbox u v})"
+      using p(8) unfolding uv[symmetric] by auto
+    have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
+    proof -
+      have "{cbox a b} division_of cbox a b"
+        by (simp add: assms division_of_self)
+      then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
+        by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> division_disjoint_union division_of_subset insert_is_Un p(1) p(8))
+    qed
+    with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe)
+  qed
+qed
+
+lemma division_of_unions:
+  assumes "finite f"
+    and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)"
+    and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+  shows "\<Union>f division_of \<Union>\<Union>f"
+  using assms
+  by (auto intro!: division_ofI)
+
+lemma elementary_union_interval:
+  fixes a b :: "'a::euclidean_space"
+  assumes "p division_of \<Union>p"
+  obtains q where "q division_of (cbox a b \<union> \<Union>p)"
+proof -
+  note assm = division_ofD[OF assms]
+  have lem1: "\<And>f s. \<Union>\<Union>(f ` s) = \<Union>((\<lambda>x. \<Union>(f x)) ` s)"
+    by auto
+  have lem2: "\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f"
+    by auto
+  {
+    presume "p = {} \<Longrightarrow> thesis"
+      "cbox a b = {} \<Longrightarrow> thesis"
+      "cbox a b \<noteq> {} \<Longrightarrow> interior (cbox a b) = {} \<Longrightarrow> thesis"
+      "p \<noteq> {} \<Longrightarrow> interior (cbox a b)\<noteq>{} \<Longrightarrow> cbox a b \<noteq> {} \<Longrightarrow> thesis"
+    then show thesis by auto
+  next
+    assume as: "p = {}"
+    obtain p where "p division_of (cbox a b)"
+      by (rule elementary_interval)
+    then show thesis
+      using as that by auto
+  next
+    assume as: "cbox a b = {}"
+    show thesis
+      using as assms that by auto
+  next
+    assume as: "interior (cbox a b) = {}" "cbox a b \<noteq> {}"
+    show thesis
+      apply (rule that[of "insert (cbox a b) p"],rule division_ofI)
+      unfolding finite_insert
+      apply (rule assm(1)) unfolding Union_insert
+      using assm(2-4) as
+      apply -
+      apply (fast dest: assm(5))+
+      done
+  next
+    assume as: "p \<noteq> {}" "interior (cbox a b) \<noteq> {}" "cbox a b \<noteq> {}"
+    have "\<forall>k\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
+    proof
+      fix k
+      assume kp: "k \<in> p"
+      from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast
+      then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
+        by (meson as(3) division_union_intervals_exists)
+    qed
+    from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" ..
+    note q = division_ofD[OF this[rule_format]]
+    let ?D = "\<Union>{insert (cbox a b) (q k) | k. k \<in> p}"
+    show thesis
+    proof (rule that[OF division_ofI])
+      have *: "{insert (cbox a b) (q k) |k. k \<in> p} = (\<lambda>k. insert (cbox a b) (q k)) ` p"
+        by auto
+      show "finite ?D"
+        using "*" assm(1) q(1) by auto
+      show "\<Union>?D = cbox a b \<union> \<Union>p"
+        unfolding * lem1
+        unfolding lem2[OF as(1), of "cbox a b", symmetric]
+        using q(6)
+        by auto
+      fix k
+      assume k: "k \<in> ?D"
+      then show "k \<subseteq> cbox a b \<union> \<Union>p"
+        using q(2) by auto
+      show "k \<noteq> {}"
+        using q(3) k by auto
+      show "\<exists>a b. k = cbox a b"
+        using q(4) k by auto
+      fix k'
+      assume k': "k' \<in> ?D" "k \<noteq> k'"
+      obtain x where x: "k \<in> insert (cbox a b) (q x)" "x\<in>p"
+        using k by auto
+      obtain x' where x': "k'\<in>insert (cbox a b) (q x')" "x'\<in>p"
+        using k' by auto
+      show "interior k \<inter> interior k' = {}"
+      proof (cases "x = x'")
+        case True
+        show ?thesis
+          using True k' q(5) x' x by auto
+      next
+        case False
+        {
+          presume "k = cbox a b \<Longrightarrow> ?thesis"
+            and "k' = cbox a b \<Longrightarrow> ?thesis"
+            and "k \<noteq> cbox a b \<Longrightarrow> k' \<noteq> cbox a b \<Longrightarrow> ?thesis"
+          then show ?thesis by linarith
+        next
+          assume as': "k  = cbox a b"
+          show ?thesis
+            using as' k' q(5) x' by blast
+        next
+          assume as': "k' = cbox a b"
+          show ?thesis
+            using as' k'(2) q(5) x by blast
+        }
+        assume as': "k \<noteq> cbox a b" "k' \<noteq> cbox a b"
+        obtain c d where k: "k = cbox c d"
+          using q(4)[OF x(2,1)] by blast
+        have "interior k \<inter> interior (cbox a b) = {}"
+          using as' k'(2) q(5) x by blast
+        then have "interior k \<subseteq> interior x"
+        using interior_subset_union_intervals
+          by (metis as(2) k q(2) x interior_subset_union_intervals)
+        moreover
+        obtain c d where c_d: "k' = cbox c d"
+          using q(4)[OF x'(2,1)] by blast
+        have "interior k' \<inter> interior (cbox a b) = {}"
+          using as'(2) q(5) x' by blast
+        then have "interior k' \<subseteq> interior x'"
+          by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2))
+        ultimately show ?thesis
+          using assm(5)[OF x(2) x'(2) False] by auto
+      qed
+    qed
+  }
+qed
+
+lemma elementary_unions_intervals:
+  assumes fin: "finite f"
+    and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
+  obtains p where "p division_of (\<Union>f)"
+proof -
+  have "\<exists>p. p division_of (\<Union>f)"
+  proof (induct_tac f rule:finite_subset_induct)
+    show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
+  next
+    fix x F
+    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
+    from this(3) obtain p where p: "p division_of \<Union>F" ..
+    from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast
+    have *: "\<Union>F = \<Union>p"
+      using division_ofD[OF p] by auto
+    show "\<exists>p. p division_of \<Union>insert x F"
+      using elementary_union_interval[OF p[unfolded *], of a b]
+      unfolding Union_insert x * by metis
+  qed (insert assms, auto)
+  then show ?thesis
+    using that by auto
+qed
+
+lemma elementary_union:
+  fixes s t :: "'a::euclidean_space set"
+  assumes "ps division_of s" "pt division_of t"
+  obtains p where "p division_of (s \<union> t)"
+proof -
+  have *: "s \<union> t = \<Union>ps \<union> \<Union>pt"
+    using assms unfolding division_of_def by auto
+  show ?thesis
+    apply (rule elementary_unions_intervals[of "ps \<union> pt"])
+    using assms apply auto
+    by (simp add: * that)
+qed
+
+lemma partial_division_extend:
+  fixes t :: "'a::euclidean_space set"
+  assumes "p division_of s"
+    and "q division_of t"
+    and "s \<subseteq> t"
+  obtains r where "p \<subseteq> r" and "r division_of t"
+proof -
+  note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
+  obtain a b where ab: "t \<subseteq> cbox a b"
+    using elementary_subset_cbox[OF assms(2)] by auto
+  obtain r1 where "p \<subseteq> r1" "r1 division_of (cbox a b)"
+    using assms
+    by (metis ab dual_order.trans partial_division_extend_interval divp(6))
+  note r1 = this division_ofD[OF this(2)]
+  obtain p' where "p' division_of \<Union>(r1 - p)"
+    apply (rule elementary_unions_intervals[of "r1 - p"])
+    using r1(3,6)
+    apply auto
+    done
+  then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
+    by (metis assms(2) divq(6) elementary_inter)
+  {
+    fix x
+    assume x: "x \<in> t" "x \<notin> s"
+    then have "x\<in>\<Union>r1"
+      unfolding r1 using ab by auto
+    then obtain r where r: "r \<in> r1" "x \<in> r"
+      unfolding Union_iff ..
+    moreover
+    have "r \<notin> p"
+    proof
+      assume "r \<in> p"
+      then have "x \<in> s" using divp(2) r by auto
+      then show False using x by auto
+    qed
+    ultimately have "x\<in>\<Union>(r1 - p)" by auto
+  }
+  then have *: "t = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
+    unfolding divp divq using assms(3) by auto
+  show ?thesis
+    apply (rule that[of "p \<union> r2"])
+    unfolding *
+    defer
+    apply (rule division_disjoint_union)
+    unfolding divp(6)
+    apply(rule assms r2)+
+  proof -
+    have "interior s \<inter> interior (\<Union>(r1-p)) = {}"
+    proof (rule inter_interior_unions_intervals)
+      show "finite (r1 - p)" and "open (interior s)" and "\<forall>t\<in>r1-p. \<exists>a b. t = cbox a b"
+        using r1 by auto
+      have *: "\<And>s. (\<And>x. x \<in> s \<Longrightarrow> False) \<Longrightarrow> s = {}"
+        by auto
+      show "\<forall>t\<in>r1-p. interior s \<inter> interior t = {}"
+      proof
+        fix m x
+        assume as: "m \<in> r1 - p"
+        have "interior m \<inter> interior (\<Union>p) = {}"
+        proof (rule inter_interior_unions_intervals)
+          show "finite p" and "open (interior m)" and "\<forall>t\<in>p. \<exists>a b. t = cbox a b"
+            using divp by auto
+          show "\<forall>t\<in>p. interior m \<inter> interior t = {}"
+            by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp)
+        qed
+        then show "interior s \<inter> interior m = {}"
+          unfolding divp by auto
+      qed
+    qed
+    then show "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}"
+      using interior_subset by auto
+  qed auto
+qed
+
+
+lemma division_split:
+  fixes a :: "'a::euclidean_space"
+  assumes "p division_of (cbox a b)"
+    and k: "k\<in>Basis"
+  shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})"
+      (is "?p1 division_of ?I1")
+    and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+      (is "?p2 division_of ?I2")
+proof (rule_tac[!] division_ofI)
+  note p = division_ofD[OF assms(1)]
+  show "finite ?p1" "finite ?p2"
+    using p(1) by auto
+  show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2"
+    unfolding p(6)[symmetric] by auto
+  {
+    fix k
+    assume "k \<in> ?p1"
+    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
+    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
+    show "k \<subseteq> ?I1"
+      using l p(2) uv by force
+    show  "k \<noteq> {}"
+      by (simp add: l)
+    show  "\<exists>a b. k = cbox a b"
+      apply (simp add: l uv p(2-3)[OF l(2)])
+      apply (subst interval_split[OF k])
+      apply (auto intro: order.trans)
+      done
+    fix k'
+    assume "k' \<in> ?p1"
+    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
+    assume "k \<noteq> k'"
+    then show "interior k \<inter> interior k' = {}"
+      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
+  }
+  {
+    fix k
+    assume "k \<in> ?p2"
+    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
+    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
+    show "k \<subseteq> ?I2"
+      using l p(2) uv by force
+    show  "k \<noteq> {}"
+      by (simp add: l)
+    show  "\<exists>a b. k = cbox a b"
+      apply (simp add: l uv p(2-3)[OF l(2)])
+      apply (subst interval_split[OF k])
+      apply (auto intro: order.trans)
+      done
+    fix k'
+    assume "k' \<in> ?p2"
+    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
+    assume "k \<noteq> k'"
+    then show "interior k \<inter> interior k' = {}"
+      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
+  }
+qed
+
+subsection \<open>Tagged (partial) divisions.\<close>
+
+definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
+  where "s tagged_partial_division_of i \<longleftrightarrow>
+    finite s \<and>
+    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
+    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
+      interior k1 \<inter> interior k2 = {})"
+
+lemma tagged_partial_division_ofD[dest]:
+  assumes "s tagged_partial_division_of i"
+  shows "finite s"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow>
+      (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+  using assms unfolding tagged_partial_division_of_def by blast+
+
+definition tagged_division_of (infixr "tagged'_division'_of" 40)
+  where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+
+lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
+  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
+
+lemma tagged_division_of:
+  "s tagged_division_of i \<longleftrightarrow>
+    finite s \<and>
+    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
+    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
+      interior k1 \<inter> interior k2 = {}) \<and>
+    (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
+
+lemma tagged_division_ofI:
+  assumes "finite s"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
+      interior k1 \<inter> interior k2 = {}"
+    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  shows "s tagged_division_of i"
+  unfolding tagged_division_of
+  using assms
+  apply auto
+  apply fastforce+
+  done
+
+lemma tagged_division_ofD[dest]:  (*FIXME USE A LOCALE*)
+  assumes "s tagged_division_of i"
+  shows "finite s"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>x1 k1 x2 k2. (x1, k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
+      interior k1 \<inter> interior k2 = {}"
+    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  using assms unfolding tagged_division_of by blast+
+
+lemma division_of_tagged_division:
+  assumes "s tagged_division_of i"
+  shows "(snd ` s) division_of i"
+proof (rule division_ofI)
+  note assm = tagged_division_ofD[OF assms]
+  show "\<Union>(snd ` s) = i" "finite (snd ` s)"
+    using assm by auto
+  fix k
+  assume k: "k \<in> snd ` s"
+  then obtain xk where xk: "(xk, k) \<in> s"
+    by auto
+  then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b"
+    using assm by fastforce+
+  fix k'
+  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
+  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
+    by auto
+  then show "interior k \<inter> interior k' = {}"
+    using assm(5) k'(2) xk by blast
+qed
+
+lemma partial_division_of_tagged_division:
+  assumes "s tagged_partial_division_of i"
+  shows "(snd ` s) division_of \<Union>(snd ` s)"
+proof (rule division_ofI)
+  note assm = tagged_partial_division_ofD[OF assms]
+  show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
+    using assm by auto
+  fix k
+  assume k: "k \<in> snd ` s"
+  then obtain xk where xk: "(xk, k) \<in> s"
+    by auto
+  then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)"
+    using assm by auto
+  fix k'
+  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
+  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
+    by auto
+  then show "interior k \<inter> interior k' = {}"
+    using assm(5) k'(2) xk by auto
+qed
+
+lemma tagged_partial_division_subset:
+  assumes "s tagged_partial_division_of i"
+    and "t \<subseteq> s"
+  shows "t tagged_partial_division_of i"
+  using assms
+  unfolding tagged_partial_division_of_def
+  using finite_subset[OF assms(2)]
+  by blast
+
+lemma tag_in_interval: "p tagged_division_of i \<Longrightarrow> (x, k) \<in> p \<Longrightarrow> x \<in> i"
+  by auto
+
+lemma tagged_division_of_empty: "{} tagged_division_of {}"
+  unfolding tagged_division_of by auto
+
+lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \<longleftrightarrow> p = {}"
+  unfolding tagged_partial_division_of_def by auto
+
+lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \<longleftrightarrow> p = {}"
+  unfolding tagged_division_of by auto
+
+lemma tagged_division_of_self: "x \<in> cbox a b \<Longrightarrow> {(x,cbox a b)} tagged_division_of (cbox a b)"
+  by (rule tagged_division_ofI) auto
+
+lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}"
+  unfolding box_real[symmetric]
+  by (rule tagged_division_of_self)
+
+lemma tagged_division_union:
+  assumes "p1 tagged_division_of s1"
+    and "p2 tagged_division_of s2"
+    and "interior s1 \<inter> interior s2 = {}"
+  shows "(p1 \<union> p2) tagged_division_of (s1 \<union> s2)"
+proof (rule tagged_division_ofI)
+  note p1 = tagged_division_ofD[OF assms(1)]
+  note p2 = tagged_division_ofD[OF assms(2)]
+  show "finite (p1 \<union> p2)"
+    using p1(1) p2(1) by auto
+  show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2"
+    using p1(6) p2(6) by blast
+  fix x k
+  assume xk: "(x, k) \<in> p1 \<union> p2"
+  show "x \<in> k" "\<exists>a b. k = cbox a b"
+    using xk p1(2,4) p2(2,4) by auto
+  show "k \<subseteq> s1 \<union> s2"
+    using xk p1(3) p2(3) by blast
+  fix x' k'
+  assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')"
+  have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}"
+    using assms(3) interior_mono by blast
+  show "interior k \<inter> interior k' = {}"
+    apply (cases "(x, k) \<in> p1")
+    apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2))
+    by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
+qed
+
+lemma tagged_division_unions:
+  assumes "finite iset"
+    and "\<forall>i\<in>iset. pfn i tagged_division_of i"
+    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior(i1) \<inter> interior(i2) = {}"
+  shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)"
+proof (rule tagged_division_ofI)
+  note assm = tagged_division_ofD[OF assms(2)[rule_format]]
+  show "finite (\<Union>(pfn ` iset))"
+    using assms by auto
+  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)"
+    by blast
+  also have "\<dots> = \<Union>iset"
+    using assm(6) by auto
+  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" .
+  fix x k
+  assume xk: "(x, k) \<in> \<Union>(pfn ` iset)"
+  then obtain i where i: "i \<in> iset" "(x, k) \<in> pfn i"
+    by auto
+  show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>iset"
+    using assm(2-4)[OF i] using i(1) by auto
+  fix x' k'
+  assume xk': "(x', k') \<in> \<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')"
+  then obtain i' where i': "i' \<in> iset" "(x', k') \<in> pfn i'"
+    by auto
+  have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}"
+    using i(1) i'(1)
+    using assms(3)[rule_format] interior_mono
+    by blast
+  show "interior k \<inter> interior k' = {}"
+    apply (cases "i = i'")
+    using assm(5) i' i(2) xk'(2) apply blast
+    using "*" assm(3) i' i by auto
+qed
+
+lemma tagged_partial_division_of_union_self:
+  assumes "p tagged_partial_division_of s"
+  shows "p tagged_division_of (\<Union>(snd ` p))"
+  apply (rule tagged_division_ofI)
+  using tagged_partial_division_ofD[OF assms]
+  apply auto
+  done
+
+lemma tagged_division_of_union_self:
+  assumes "p tagged_division_of s"
+  shows "p tagged_division_of (\<Union>(snd ` p))"
+  apply (rule tagged_division_ofI)
+  using tagged_division_ofD[OF assms]
+  apply auto
+  done
+
+subsection \<open>Functions closed on boxes: morphisms from boxes to monoids\<close>
+
+text \<open>This auxiliary structure is used to sum up over the elements of a division. Main theorem is
+  @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and
+  @{typ bool}.\<close>
+
+paragraph \<open>Using additivity of lifted function to encode definedness.\<close>
+
+definition lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
+where
+  "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
+
+lemma lift_option_simps[simp]:
+  "lift_option f (Some a) (Some b) = Some (f a b)"
+  "lift_option f None b' = None"
+  "lift_option f a' None = None"
+  by (auto simp: lift_option_def)
+
+lemma comm_monoid_lift_option:
+  assumes "comm_monoid f z"
+  shows "comm_monoid (lift_option f) (Some z)"
+proof -
+  from assms interpret comm_monoid f z .
+  show ?thesis
+    by standard (auto simp: lift_option_def ac_simps split: bind_split)
+qed
+
+lemma comm_monoid_and: "comm_monoid HOL.conj True"
+  by standard auto
+
+lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True"
+  by (rule comm_monoid_set.intro) (fact comm_monoid_and)
+
+paragraph \<open>Operative\<close>
+
+definition (in comm_monoid) operative :: "('b::euclidean_space set \<Rightarrow> 'a) \<Rightarrow> bool"
+  where "operative g \<longleftrightarrow>
+    (\<forall>a b. box a b = {} \<longrightarrow> g (cbox a b) = \<^bold>1) \<and>
+    (\<forall>a b c. \<forall>k\<in>Basis. g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c}))"
+
+lemma (in comm_monoid) operativeD[dest]:
+  assumes "operative g"
+  shows "\<And>a b. box a b = {} \<Longrightarrow> g (cbox a b) = \<^bold>1"
+    and "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+  using assms unfolding operative_def by auto
+
+lemma (in comm_monoid) operative_empty:
+  assumes g: "operative g" shows "g {} = \<^bold>1"
+proof -
+  have *: "cbox One (-One) = ({}::'b set)"
+    by (auto simp: box_eq_empty inner_setsum_left inner_Basis setsum.If_cases ex_in_conv)
+  moreover have "box One (-One) = ({}::'b set)"
+    using box_subset_cbox[of One "-One"] by (auto simp: *)
+  ultimately show ?thesis
+    using operativeD(1)[OF g, of One "-One"] by simp
+qed
+
+definition "division_points (k::('a::euclidean_space) set) d =
+   {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
+     (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
+
+lemma division_points_finite:
+  fixes i :: "'a::euclidean_space set"
+  assumes "d division_of i"
+  shows "finite (division_points i d)"
+proof -
+  note assm = division_ofD[OF assms]
+  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
+    (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
+  have *: "division_points i d = \<Union>(?M ` Basis)"
+    unfolding division_points_def by auto
+  show ?thesis
+    unfolding * using assm by auto
+qed
+
+lemma division_points_subset:
+  fixes a :: "'a::euclidean_space"
+  assumes "d division_of (cbox a b)"
+    and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
+    and k: "k \<in> Basis"
+  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq>
+      division_points (cbox a b) d" (is ?t1)
+    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
+      division_points (cbox a b) d" (is ?t2)
+proof -
+  note assm = division_ofD[OF assms(1)]
+  have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+    "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else  b \<bullet> i) *\<^sub>R i) \<bullet> i"
+    "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i"
+    "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c"
+    using assms using less_imp_le by auto
+  show ?t1 (*FIXME a horrible mess*)
+    unfolding division_points_def interval_split[OF k, of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
+    unfolding *
+    apply (rule subsetI)
+    unfolding mem_Collect_eq split_beta
+    apply (erule bexE conjE)+
+    apply (simp add: )
+    apply (erule exE conjE)+
+  proof
+    fix i l x
+    assume as:
+      "a \<bullet> fst x < snd x" "snd x < (if fst x = k then c else b \<bullet> fst x)"
+      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
+      and fstx: "fst x \<in> Basis"
+    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
+    have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
+      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
+    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
+      using l using as(6) unfolding box_ne_empty[symmetric] by auto
+    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
+      using as(1-3,5) fstx
+      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
+      apply (auto split: if_split_asm)
+      done
+    show "snd x < b \<bullet> fst x"
+      using as(2) \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
+  qed
+  show ?t2
+    unfolding division_points_def interval_split[OF k, of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
+    unfolding *
+    unfolding subset_eq
+    apply rule
+    unfolding mem_Collect_eq split_beta
+    apply (erule bexE conjE)+
+    apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
+    apply (erule exE conjE)+
+  proof
+    fix i l x
+    assume as:
+      "(if fst x = k then c else a \<bullet> fst x) < snd x" "snd x < b \<bullet> fst x"
+      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
+      and fstx: "fst x \<in> Basis"
+    from assm(4)[OF this(5)] guess u v by (elim exE) note l=this
+    have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
+      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
+    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
+      using l using as(6) unfolding box_ne_empty[symmetric] by auto
+    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
+      using as(1-3,5) fstx
+      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
+      apply (auto split: if_split_asm)
+      done
+    show "a \<bullet> fst x < snd x"
+      using as(1) \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm)
+   qed
+qed
+
+lemma division_points_psubset:
+  fixes a :: "'a::euclidean_space"
+  assumes "d division_of (cbox a b)"
+      and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
+      and "l \<in> d"
+      and "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c"
+      and k: "k \<in> Basis"
+  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subset>
+         division_points (cbox a b) d" (is "?D1 \<subset> ?D")
+    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset>
+         division_points (cbox a b) d" (is "?D2 \<subset> ?D")
+proof -
+  have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+    using assms(2) by (auto intro!:less_imp_le)
+  guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this
+  have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i"
+    using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty
+    using subset_box(1)
+    apply auto
+    apply blast+
+    done
+  have *: "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
+          "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
+    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
+    using uv[rule_format, of k] ab k
+    by auto
+  have "\<exists>x. x \<in> ?D - ?D1"
+    using assms(3-)
+    unfolding division_points_def interval_bounds[OF ab]
+    apply -
+    apply (erule disjE)
+    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
+    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
+    done
+  moreover have "?D1 \<subseteq> ?D"
+    by (auto simp add: assms division_points_subset)
+  ultimately show "?D1 \<subset> ?D"
+    by blast
+  have *: "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
+    "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
+    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
+    using uv[rule_format, of k] ab k
+    by auto
+  have "\<exists>x. x \<in> ?D - ?D2"
+    using assms(3-)
+    unfolding division_points_def interval_bounds[OF ab]
+    apply -
+    apply (erule disjE)
+    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
+    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
+    done
+  moreover have "?D2 \<subseteq> ?D"
+    by (auto simp add: assms division_points_subset)
+  ultimately show "?D2 \<subset> ?D"
+    by blast
+qed
+
+lemma division_split_left_inj:
+  fixes type :: "'a::euclidean_space"
+  assumes "d division_of i"
+    and "k1 \<in> d"
+    and "k2 \<in> d"
+    and "k1 \<noteq> k2"
+    and "k1 \<inter> {x::'a. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
+    and k: "k\<in>Basis"
+  shows "interior (k1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
+proof -
+  note d=division_ofD[OF assms(1)]
+  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
+  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
+  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
+    by auto
+  show ?thesis
+    unfolding uv1 uv2
+    apply (rule **[OF d(5)[OF assms(2-4)]])
+    apply (simp add: uv1)
+    using assms(5) uv1 by auto
+qed
+
+lemma division_split_right_inj:
+  fixes type :: "'a::euclidean_space"
+  assumes "d division_of i"
+    and "k1 \<in> d"
+    and "k2 \<in> d"
+    and "k1 \<noteq> k2"
+    and "k1 \<inter> {x::'a. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
+    and k: "k \<in> Basis"
+  shows "interior (k1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
+proof -
+  note d=division_ofD[OF assms(1)]
+  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
+  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
+  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
+    by auto
+  show ?thesis
+    unfolding uv1 uv2
+    apply (rule **[OF d(5)[OF assms(2-4)]])
+    apply (simp add: uv1)
+    using assms(5) uv1 by auto
+qed
+
+lemma interval_doublesplit:
+  fixes a :: "'a::euclidean_space"
+  assumes "k \<in> Basis"
+  shows "cbox a b \<inter> {x . \<bar>x\<bullet>k - c\<bar> \<le> (e::real)} =
+    cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i)
+     (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)"
+proof -
+  have *: "\<And>x c e::real. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
+    by auto
+  have **: "\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}"
+    by blast
+  show ?thesis
+    unfolding * ** interval_split[OF assms] by (rule refl)
+qed
+
+lemma division_doublesplit:
+  fixes a :: "'a::euclidean_space"
+  assumes "p division_of (cbox a b)"
+    and k: "k \<in> Basis"
+  shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
+         division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
+proof -
+  have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
+    by auto
+  have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
+    by auto
+  note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
+  note division_split(2)[OF this, where c="c-e" and k=k,OF k]
+  then show ?thesis
+    apply (rule **)
+    subgoal
+      apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric])
+      apply (rule equalityI)
+      apply blast
+      apply clarsimp
+      apply (rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
+      apply auto
+      done
+    by (simp add: interval_split k interval_doublesplit)
+qed
+
+lemma (in comm_monoid_set) operative_division:
+  fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
+  assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)"
+proof -
+  define C where [abs_def]: "C = card (division_points (cbox a b) d)"
+  then show ?thesis
+    using d
+  proof (induction C arbitrary: a b d rule: less_induct)
+    case (less a b d)
+    show ?case
+    proof cases
+      assume "box a b = {}"
+      { fix k assume "k\<in>d"
+        then obtain a' b' where k: "k = cbox a' b'"
+          using division_ofD(4)[OF less.prems] by blast
+        with \<open>k\<in>d\<close> division_ofD(2)[OF less.prems] have "cbox a' b' \<subseteq> cbox a b"
+          by auto
+        then have "box a' b' \<subseteq> box a b"
+          unfolding subset_box by auto
+        then have "g k = \<^bold>1"
+          using operativeD(1)[OF g, of a' b'] k by (simp add: \<open>box a b = {}\<close>) }
+      then show "box a b = {} \<Longrightarrow> F g d = g (cbox a b)"
+        by (auto intro!: neutral simp: operativeD(1)[OF g])
+    next
+      assume "box a b \<noteq> {}"
+      then have ab: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" and ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+        by (auto simp: box_ne_empty)
+      show "F g d = g (cbox a b)"
+      proof (cases "division_points (cbox a b) d = {}")
+        case True
+        { fix u v and j :: 'b
+          assume j: "j \<in> Basis" and as: "cbox u v \<in> d"
+          then have "cbox u v \<noteq> {}"
+            using less.prems by blast
+          then have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j"
+            using j unfolding box_ne_empty by auto
+          have *: "\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> Q (cbox u v)"
+            using as j by auto
+          have "(j, u\<bullet>j) \<notin> division_points (cbox a b) d"
+               "(j, v\<bullet>j) \<notin> division_points (cbox a b) d" using True by auto
+          note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps]
+          note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
+          moreover
+          have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j"
+            using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as]
+            apply (metis j subset_box(1) uv(1))
+            by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1))
+          ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j"
+            unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force }
+        then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and>
+          (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
+          unfolding forall_in_division[OF less.prems] by blast
+        have "(1/2) *\<^sub>R (a+b) \<in> cbox a b"
+          unfolding mem_box using ab by (auto simp: inner_simps)
+        note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff]
+        then guess i .. note i=this
+        guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this
+        have "cbox a b \<in> d"
+        proof -
+          have "u = a" "v = b"
+            unfolding euclidean_eq_iff[where 'a='b]
+          proof safe
+            fix j :: 'b
+            assume j: "j \<in> Basis"
+            note i(2)[unfolded uv mem_box,rule_format,of j]
+            then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> j"
+              using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
+          qed
+          then have "i = cbox a b" using uv by auto
+          then show ?thesis using i by auto
+        qed
+        then have deq: "d = insert (cbox a b) (d - {cbox a b})"
+          by auto
+        have "F g (d - {cbox a b}) = \<^bold>1"
+        proof (intro neutral ballI)
+          fix x
+          assume x: "x \<in> d - {cbox a b}"
+          then have "x\<in>d"
+            by auto note d'[rule_format,OF this]
+          then guess u v by (elim exE conjE) note uv=this
+          have "u \<noteq> a \<or> v \<noteq> b"
+            using x[unfolded uv] by auto
+          then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j: "j \<in> Basis"
+            unfolding euclidean_eq_iff[where 'a='b] by auto
+          then have "u\<bullet>j = v\<bullet>j"
+            using uv(2)[rule_format,OF j] by auto
+          then have "box u v = {}"
+            using j unfolding box_eq_empty by (auto intro!: bexI[of _ j])
+          then show "g x = \<^bold>1"
+            unfolding uv(1) by (rule operativeD(1)[OF g])
+        qed
+        then show "F g d = g (cbox a b)"
+          using division_ofD[OF less.prems]
+          apply (subst deq)
+          apply (subst insert)
+          apply auto
+          done
+      next
+        case False
+        then have "\<exists>x. x \<in> division_points (cbox a b) d"
+          by auto
+        then guess k c
+          unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv
+          apply (elim exE conjE)
+          done
+        note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']]
+        from this(3) guess j .. note j=this
+        define d1 where "d1 = {l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+        define d2 where "d2 = {l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+        define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)"
+        define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)"
+        note division_points_psubset[OF \<open>d division_of cbox a b\<close> ab kc(1-2) j]
+        note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
+        then have *: "F g d1 = g (cbox a b \<inter> {x. x\<bullet>k \<le> c})" "F g d2 = g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+          unfolding interval_split[OF kc(4)]
+          apply (rule_tac[!] "less.hyps"[rule_format])
+          using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c]
+          apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>])
+          done
+        { fix l y
+          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} = y \<inter> {x. x \<bullet> k \<le> c}" "l \<noteq> y"
+          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
+          have "g (l \<inter> {x. x \<bullet> k \<le> c}) = \<^bold>1"
+            unfolding leq interval_split[OF kc(4)]
+            apply (rule operativeD[OF g])
+            unfolding interior_cbox[symmetric] interval_split[symmetric, OF kc(4)]
+            using division_split_left_inj less as kc leq by blast
+        } note fxk_le = this
+        { fix l y
+          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} = y \<inter> {x. c \<le> x \<bullet> k}" "l \<noteq> y"
+          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
+          have "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
+            unfolding leq interval_split[OF kc(4)]
+            apply (rule operativeD(1)[OF g])
+            unfolding interior_cbox[symmetric] interval_split[symmetric,OF kc(4)]
+            using division_split_right_inj less leq as kc by blast
+        } note fxk_ge = this
+        have d1_alt: "d1 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<le> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+          using d1_def by auto
+        have d2_alt: "d2 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<ge> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+          using d2_def by auto
+        have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
+          unfolding * using g kc(4) by blast
+        also have "F g d1 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d"
+          unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
+          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
+        also have "F g d2 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d"
+          unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
+          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
+        also have *: "\<forall>x\<in>d. g x = g (x \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (x \<inter> {x. c \<le> x \<bullet> k})"
+          unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
+          using g kc(4) by blast
+        have "F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d \<^bold>* F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d = F g d"
+          using * by (simp add: distrib)
+        finally show ?thesis by auto
+      qed
+    qed
+  qed
+qed
+
+lemma (in comm_monoid_set) over_tagged_division_lemma:
+  assumes "p tagged_division_of i"
+    and "\<And>u v. cbox u v \<noteq> {} \<Longrightarrow> box u v = {} \<Longrightarrow> d (cbox u v) = \<^bold>1"
+  shows "F (\<lambda>(x,k). d k) p = F d (snd ` p)"
+proof -
+  have *: "(\<lambda>(x,k). d k) = d \<circ> snd"
+    unfolding o_def by (rule ext) auto
+  note assm = tagged_division_ofD[OF assms(1)]
+  show ?thesis
+    unfolding *
+  proof (rule reindex_nontrivial[symmetric])
+    show "finite p"
+      using assm by auto
+    fix x y
+    assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
+    obtain a b where ab: "snd x = cbox a b"
+      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
+    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
+      by (metis prod.collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+
+    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
+      by (intro assm(5)[of "fst x" _ "fst y"]) auto
+    then have "box a b = {}"
+      unfolding \<open>snd x = snd y\<close>[symmetric] ab by auto
+    then have "d (cbox a b) = \<^bold>1"
+      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
+    then show "d (snd x) = \<^bold>1"
+      unfolding ab by auto
+  qed
+qed
+
+lemma (in comm_monoid_set) operative_tagged_division:
+  assumes f: "operative g" and d: "d tagged_division_of (cbox a b)"
+  shows "F (\<lambda>(x, l). g l) d = g (cbox a b)"
+  unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric]
+  by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d])
+
+lemma interval_real_split:
+  "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
+  "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
+  apply (metis Int_atLeastAtMostL1 atMost_def)
+  apply (metis Int_atLeastAtMostL2 atLeast_def)
+  done
+
+lemma (in comm_monoid) operative_1_lt:
+  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
+    ((\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1) \<and> (\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
+  apply (simp add: operative_def atMost_def[symmetric] atLeast_def[symmetric])
+proof safe
+  fix a b c :: real
+  assume *: "\<forall>a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
+  assume "a < c" "c < b"
+  with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    by (simp add: less_imp_le min.absorb2 max.absorb2)
+next
+  fix a b c :: real
+  assume as: "\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1"
+    "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+  from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2)
+  have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
+    "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    by auto
+  show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
+    by (auto simp: min_def max_def le_less)
+qed
+
+lemma (in comm_monoid) operative_1_le:
+  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
+    ((\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1) \<and> (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
+  unfolding operative_1_lt
+proof safe
+  fix a b c :: real
+  assume as: "\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b"
+  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    apply (rule as(1)[rule_format])
+    using as(2-)
+    apply auto
+    done
+next
+  fix a b c :: real
+  assume "\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1"
+    and "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    and "a \<le> c"
+    and "c \<le> b"
+  note as = this[rule_format]
+  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
+  proof (cases "c = a \<or> c = b")
+    case False
+    then show ?thesis
+      apply -
+      apply (subst as(2))
+      using as(3-)
+      apply auto
+      done
+  next
+    case True
+    then show ?thesis
+    proof
+      assume *: "c = a"
+      then have "g {a .. c} = \<^bold>1"
+        apply -
+        apply (rule as(1)[rule_format])
+        apply auto
+        done
+      then show ?thesis
+        unfolding * by auto
+    next
+      assume *: "c = b"
+      then have "g {c .. b} = \<^bold>1"
+        apply -
+        apply (rule as(1)[rule_format])
+        apply auto
+        done
+      then show ?thesis
+        unfolding * by auto
+    qed
+  qed
+qed
+
+lemma tagged_division_union_interval:
+  fixes a :: "'a::euclidean_space"
+  assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
+    and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+    and k: "k \<in> Basis"
+  shows "(p1 \<union> p2) tagged_division_of (cbox a b)"
+proof -
+  have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+    by auto
+  show ?thesis
+    apply (subst *)
+    apply (rule tagged_division_union[OF assms(1-2)])
+    unfolding interval_split[OF k] interior_cbox
+    using k
+    apply (auto simp add: box_def elim!: ballE[where x=k])
+    done
+qed
+
+lemma tagged_division_union_interval_real:
+  fixes a :: real
+  assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})"
+    and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
+    and k: "k \<in> Basis"
+  shows "(p1 \<union> p2) tagged_division_of {a .. b}"
+  using assms
+  unfolding box_real[symmetric]
+  by (rule tagged_division_union_interval)
+
+lemma tagged_division_split_left_inj:
+  "d tagged_division_of i \<Longrightarrow> (x1, k1) \<in> d \<Longrightarrow> (x2, k2) \<in> d \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow>
+    k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c} \<Longrightarrow> k \<in> Basis \<Longrightarrow>
+    interior (k1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
+  by (intro division_split_left_inj[of "snd`d" i k1 k2, OF division_of_tagged_division])
+     (auto simp add: snd_def[abs_def] image_iff split: prod.split )
+
+lemma tagged_division_split_right_inj:
+  "d tagged_division_of i \<Longrightarrow> (x1, k1) \<in> d \<Longrightarrow> (x2, k2) \<in> d \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow>
+    k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c} \<Longrightarrow> k \<in> Basis \<Longrightarrow>
+    interior (k1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
+  by (intro division_split_right_inj[of "snd`d" i k1 k2, OF division_of_tagged_division])
+     (auto simp add: snd_def[abs_def] image_iff split: prod.split )
+
+subsection \<open>Special case of additivity we need for the FTC.\<close>
+
+lemma additive_tagged_division_1:
+  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
+  assumes "a \<le> b"
+    and "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
+proof -
+  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
+  have ***: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
+    using assms by auto
+  have *: "add.operative ?f"
+    unfolding add.operative_1_lt box_eq_empty
+    by auto
+  have **: "cbox a b \<noteq> {}"
+    using assms(1) by auto
+  note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]]
+  note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric]
+  show ?thesis
+    unfolding *
+    apply (rule setsum.cong)
+    unfolding split_paired_all split_conv
+    using assms(2)
+    apply auto
+    done
+qed
+
+lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
+  by (meson zero_less_one)
+
+lemma additive_tagged_division_1':
+  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
+  assumes "a \<le> b"
+    and "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
+  using additive_tagged_division_1[OF _ assms(2), of f]
+  using assms(1)
+  by auto
+
+subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close>
+
+definition fine  (infixr "fine" 46)
+  where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
+
+lemma fineI:
+  assumes "\<And>x k. (x, k) \<in> s \<Longrightarrow> k \<subseteq> d x"
+  shows "d fine s"
+  using assms unfolding fine_def by auto
+
+lemma fineD[dest]:
+  assumes "d fine s"
+  shows "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> d x"
+  using assms unfolding fine_def by auto
+
+lemma fine_inter: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p"
+  unfolding fine_def by auto
+
+lemma fine_inters:
+ "(\<lambda>x. \<Inter>{f d x | d.  d \<in> s}) fine p \<longleftrightarrow> (\<forall>d\<in>s. (f d) fine p)"
+  unfolding fine_def by blast
+
+lemma fine_union: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
+  unfolding fine_def by blast
+
+lemma fine_unions: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
+  unfolding fine_def by auto
+
+lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
+  unfolding fine_def by blast
+
+subsection \<open>Some basic combining lemmas.\<close>
+
+lemma tagged_division_unions_exists:
+  assumes "finite iset"
+    and "\<forall>i\<in>iset. \<exists>p. p tagged_division_of i \<and> d fine p"
+    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
+    and "\<Union>iset = i"
+   obtains p where "p tagged_division_of i" and "d fine p"
+proof -
+  obtain pfn where pfn:
+    "\<And>x. x \<in> iset \<Longrightarrow> pfn x tagged_division_of x"
+    "\<And>x. x \<in> iset \<Longrightarrow> d fine pfn x"
+    using bchoice[OF assms(2)] by auto
+  show thesis
+    apply (rule_tac p="\<Union>(pfn ` iset)" in that)
+    using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force
+    by (metis (mono_tags, lifting) fine_unions imageE pfn(2))
+qed
+
+
+subsection \<open>The set we're concerned with must be closed.\<close>
+
+lemma division_of_closed:
+  fixes i :: "'n::euclidean_space set"
+  shows "s division_of i \<Longrightarrow> closed i"
+  unfolding division_of_def by fastforce
+
+subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close>
+
+lemma interval_bisection_step:
+  fixes type :: "'a::euclidean_space"
+  assumes "P {}"
+    and "\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P (s \<union> t)"
+    and "\<not> P (cbox a (b::'a))"
+  obtains c d where "\<not> P (cbox c d)"
+    and "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
+proof -
+  have "cbox a b \<noteq> {}"
+    using assms(1,3) by metis
+  then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
+    by (force simp: mem_box)
+  { fix f
+    have "\<lbrakk>finite f;
+           \<And>s. s\<in>f \<Longrightarrow> P s;
+           \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b;
+           \<And>s t. s\<in>f \<Longrightarrow> t\<in>f \<Longrightarrow> s \<noteq> t \<Longrightarrow> interior s \<inter> interior t = {}\<rbrakk> \<Longrightarrow> P (\<Union>f)"
+    proof (induct f rule: finite_induct)
+      case empty
+      show ?case
+        using assms(1) by auto
+    next
+      case (insert x f)
+      show ?case
+        unfolding Union_insert
+        apply (rule assms(2)[rule_format])
+        using inter_interior_unions_intervals [of f "interior x"]
+        apply (auto simp: insert)
+        by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff)
+    qed
+  } note UN_cases = this
+  let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or>
+    (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}"
+  let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
+  {
+    presume "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d) \<Longrightarrow> False"
+    then show thesis
+      unfolding atomize_not not_all
+      by (blast intro: that)
+  }
+  assume as: "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d)"
+  have "P (\<Union>?A)"
+  proof (rule UN_cases)
+    let ?B = "(\<lambda>s. cbox (\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i::'a)
+      (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)) ` {s. s \<subseteq> Basis}"
+    have "?A \<subseteq> ?B"
+    proof
+      fix x
+      assume "x \<in> ?A"
+      then obtain c d
+        where x:  "x = cbox c d"
+                  "\<And>i. i \<in> Basis \<Longrightarrow>
+                        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+                        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" by blast
+      show "x \<in> ?B"
+        unfolding image_iff x
+        apply (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI)
+        apply (rule arg_cong2 [where f = cbox])
+        using x(2) ab
+        apply (auto simp add: euclidean_eq_iff[where 'a='a])
+        by fastforce
+    qed
+    then show "finite ?A"
+      by (rule finite_subset) auto
+  next
+    fix s
+    assume "s \<in> ?A"
+    then obtain c d
+      where s: "s = cbox c d"
+               "\<And>i. i \<in> Basis \<Longrightarrow>
+                     c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+                     c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
+      by blast
+    show "P s"
+      unfolding s
+      apply (rule as[rule_format])
+      using ab s(2) by force
+    show "\<exists>a b. s = cbox a b"
+      unfolding s by auto
+    fix t
+    assume "t \<in> ?A"
+    then obtain e f where t:
+      "t = cbox e f"
+      "\<And>i. i \<in> Basis \<Longrightarrow>
+        e \<bullet> i = a \<bullet> i \<and> f \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+        e \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> f \<bullet> i = b \<bullet> i"
+      by blast
+    assume "s \<noteq> t"
+    then have "\<not> (c = e \<and> d = f)"
+      unfolding s t by auto
+    then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis"
+      unfolding euclidean_eq_iff[where 'a='a] by auto
+    then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
+      using s(2) t(2) apply fastforce
+      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
+    have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
+      by auto
+    show "interior s \<inter> interior t = {}"
+      unfolding s t interior_cbox
+    proof (rule *)
+      fix x
+      assume "x \<in> box c d" "x \<in> box e f"
+      then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
+        unfolding mem_box using i'
+        by force+
+      show False  using s(2)[OF i']
+      proof safe
+        assume as: "c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2"
+        show False
+          using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps)
+      next
+        assume as: "c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i"
+        show False
+          using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
+      qed
+    qed
+  qed
+  also have "\<Union>?A = cbox a b"
+  proof (rule set_eqI,rule)
+    fix x
+    assume "x \<in> \<Union>?A"
+    then obtain c d where x:
+      "x \<in> cbox c d"
+      "\<And>i. i \<in> Basis \<Longrightarrow>
+        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
+      by blast
+    show "x\<in>cbox a b"
+      unfolding mem_box
+    proof safe
+      fix i :: 'a
+      assume i: "i \<in> Basis"
+      then show "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i"
+        using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto
+    qed
+  next
+    fix x
+    assume x: "x \<in> cbox a b"
+    have "\<forall>i\<in>Basis.
+      \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d"
+      (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d")
+      unfolding mem_box
+    proof
+      fix i :: 'a
+      assume i: "i \<in> Basis"
+      have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)"
+        using x[unfolded mem_box,THEN bspec, OF i] by auto
+      then show "\<exists>c d. ?P i c d"
+        by blast
+    qed
+    then show "x\<in>\<Union>?A"
+      unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff
+      apply auto
+      apply (rule_tac x="cbox xa xaa" in exI)
+      unfolding mem_box
+      apply auto
+      done
+  qed
+  finally show False
+    using assms by auto
+qed
+
+lemma interval_bisection:
+  fixes type :: "'a::euclidean_space"
+  assumes "P {}"
+    and "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))"
+    and "\<not> P (cbox a (b::'a))"
+  obtains x where "x \<in> cbox a b"
+    and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
+proof -
+  have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and>
+    (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
+       2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>x. ?P x")
+  proof
+    show "?P x" for x
+    proof (cases "P (cbox (fst x) (snd x))")
+      case True
+      then show ?thesis by auto
+    next
+      case as: False
+      obtain c d where "\<not> P (cbox c d)"
+        "\<forall>i\<in>Basis.
+           fst x \<bullet> i \<le> c \<bullet> i \<and>
+           c \<bullet> i \<le> d \<bullet> i \<and>
+           d \<bullet> i \<le> snd x \<bullet> i \<and>
+           2 * (d \<bullet> i - c \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i"
+        by (rule interval_bisection_step[of P, OF assms(1-2) as])
+      then show ?thesis
+        apply -
+        apply (rule_tac x="(c,d)" in exI)
+        apply auto
+        done
+    qed
+  qed
+  then obtain f where f:
+    "\<forall>x.
+      \<not> P (cbox (fst x) (snd x)) \<longrightarrow>
+      \<not> P (cbox (fst (f x)) (snd (f x))) \<and>
+        (\<forall>i\<in>Basis.
+            fst x \<bullet> i \<le> fst (f x) \<bullet> i \<and>
+            fst (f x) \<bullet> i \<le> snd (f x) \<bullet> i \<and>
+            snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
+            2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)"
+    apply -
+    apply (drule choice)
+    apply blast
+    done
+  define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
+  have "A 0 = a" "B 0 = b" "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
+    (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
+    2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
+  proof -
+    show "A 0 = a" "B 0 = b"
+      unfolding ab_def by auto
+    note S = ab_def funpow.simps o_def id_apply
+    show "?P n" for n
+    proof (induct n)
+      case 0
+      then show ?case
+        unfolding S
+        apply (rule f[rule_format]) using assms(3)
+        apply auto
+        done
+    next
+      case (Suc n)
+      show ?case
+        unfolding S
+        apply (rule f[rule_format])
+        using Suc
+        unfolding S
+        apply auto
+        done
+    qed
+  qed
+  note AB = this(1-2) conjunctD2[OF this(3),rule_format]
+
+  have interv: "\<exists>n. \<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e"
+    if e: "0 < e" for e
+  proof -
+    obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n"
+      using real_arch_pow[of 2 "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
+    show ?thesis
+    proof (rule exI [where x=n], clarify)
+      fix x y
+      assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)"
+      have "dist x y \<le> setsum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
+        unfolding dist_norm by(rule norm_le_l1)
+      also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
+      proof (rule setsum_mono)
+        fix i :: 'a
+        assume i: "i \<in> Basis"
+        show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
+          using xy[unfolded mem_box,THEN bspec, OF i]
+          by (auto simp: inner_diff_left)
+      qed
+      also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
+        unfolding setsum_divide_distrib
+      proof (rule setsum_mono)
+        show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i
+        proof (induct n)
+          case 0
+          then show ?case
+            unfolding AB by auto
+        next
+          case (Suc n)
+          have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
+            using AB(4)[of i n] using i by auto
+          also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n"
+            using Suc by (auto simp add: field_simps)
+          finally show ?case .
+        qed
+      qed
+      also have "\<dots> < e"
+        using n using e by (auto simp add: field_simps)
+      finally show "dist x y < e" .
+    qed
+  qed
+  {
+    fix n m :: nat
+    assume "m \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)"
+    proof (induction rule: inc_induct)
+      case (step i)
+      show ?case
+        using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto
+    qed simp
+  } note ABsubset = this
+  have "\<exists>a. \<forall>n. a\<in> cbox (A n) (B n)"
+    by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv])
+      (metis nat.exhaust AB(1-3) assms(1,3))
+  then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
+    by blast
+  show thesis
+  proof (rule that[rule_format, of x0])
+    show "x0\<in>cbox a b"
+      using x0[of 0] unfolding AB .
+    fix e :: real
+    assume "e > 0"
+    from interv[OF this] obtain n
+      where n: "\<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" ..
+    have "\<not> P (cbox (A n) (B n))"
+      apply (cases "0 < n")
+      using AB(3)[of "n - 1"] assms(3) AB(1-2)
+      apply auto
+      done
+    moreover have "cbox (A n) (B n) \<subseteq> ball x0 e"
+      using n using x0[of n] by auto
+    moreover have "cbox (A n) (B n) \<subseteq> cbox a b"
+      unfolding AB(1-2)[symmetric] by (rule ABsubset) auto
+    ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
+      apply (rule_tac x="A n" in exI)
+      apply (rule_tac x="B n" in exI)
+      apply (auto simp: x0)
+      done
+  qed
+qed
+
+
+subsection \<open>Cousin's lemma.\<close>
+
+lemma fine_division_exists:
+  fixes a b :: "'a::euclidean_space"
+  assumes "gauge g"
+  obtains p where "p tagged_division_of (cbox a b)" "g fine p"
+proof -
+  presume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p) \<Longrightarrow> False"
+  then obtain p where "p tagged_division_of (cbox a b)" "g fine p"
+    by blast
+  then show thesis ..
+next
+  assume as: "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)"
+  obtain x where x:
+      "x \<in> (cbox a b)"
+      "\<And>e. 0 < e \<Longrightarrow>
+        \<exists>c d.
+          x \<in> cbox c d \<and>
+          cbox c d \<subseteq> ball x e \<and>
+          cbox c d \<subseteq> (cbox a b) \<and>
+          \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
+    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ as])
+    apply (simp add: fine_def)
+    apply (metis tagged_division_union fine_union)
+    apply (auto simp: )
+    done
+  obtain e where e: "e > 0" "ball x e \<subseteq> g x"
+    using gaugeD[OF assms, of x] unfolding open_contains_ball by auto
+  from x(2)[OF e(1)]
+  obtain c d where c_d: "x \<in> cbox c d"
+                        "cbox c d \<subseteq> ball x e"
+                        "cbox c d \<subseteq> cbox a b"
+                        "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
+    by blast
+  have "g fine {(x, cbox c d)}"
+    unfolding fine_def using e using c_d(2) by auto
+  then show False
+    using tagged_division_of_self[OF c_d(1)] using c_d by auto
+qed
+
+lemma fine_division_exists_real:
+  fixes a b :: real
+  assumes "gauge g"
+  obtains p where "p tagged_division_of {a .. b}" "g fine p"
+  by (metis assms box_real(2) fine_division_exists)
+
+subsection \<open>A technical lemma about "refinement" of division.\<close>
+
+lemma tagged_division_finer:
+  fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
+  assumes "p tagged_division_of (cbox a b)"
+    and "gauge d"
+  obtains q where "q tagged_division_of (cbox a b)"
+    and "d fine q"
+    and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
+proof -
+  let ?P = "\<lambda>p. p tagged_partial_division_of (cbox a b) \<longrightarrow> gauge d \<longrightarrow>
+    (\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and>
+      (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))"
+  {
+    have *: "finite p" "p tagged_partial_division_of (cbox a b)"
+      using assms(1)
+      unfolding tagged_division_of_def
+      by auto
+    presume "\<And>p. finite p \<Longrightarrow> ?P p"
+    from this[rule_format,OF * assms(2)] guess q .. note q=this
+    then show ?thesis
+      apply -
+      apply (rule that[of q])
+      unfolding tagged_division_ofD[OF assms(1)]
+      apply auto
+      done
+  }
+  fix p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
+  assume as: "finite p"
+  show "?P p"
+    apply rule
+    apply rule
+    using as
+  proof (induct p)
+    case empty
+    show ?case
+      apply (rule_tac x="{}" in exI)
+      unfolding fine_def
+      apply auto
+      done
+  next
+    case (insert xk p)
+    guess x k using surj_pair[of xk] by (elim exE) note xk=this
+    note tagged_partial_division_subset[OF insert(4) subset_insertI]
+    from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this]
+    have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}"
+      unfolding xk by auto
+    note p = tagged_partial_division_ofD[OF insert(4)]
+    from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this
+
+    have "finite {k. \<exists>x. (x, k) \<in> p}"
+      apply (rule finite_subset[of _ "snd ` p"])
+      using p
+      apply safe
+      apply (metis image_iff snd_conv)
+      apply auto
+      done
+    then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
+      apply (rule inter_interior_unions_intervals)
+      apply (rule open_interior)
+      apply (rule_tac[!] ballI)
+      unfolding mem_Collect_eq
+      apply (erule_tac[!] exE)
+      apply (drule p(4)[OF insertI2])
+      apply assumption
+      apply (rule p(5))
+      unfolding uv xk
+      apply (rule insertI1)
+      apply (rule insertI2)
+      apply assumption
+      using insert(2)
+      unfolding uv xk
+      apply auto
+      done
+    show ?case
+    proof (cases "cbox u v \<subseteq> d x")
+      case True
+      then show ?thesis
+        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
+        apply rule
+        unfolding * uv
+        apply (rule tagged_division_union)
+        apply (rule tagged_division_of_self)
+        apply (rule p[unfolded xk uv] insertI1)+
+        apply (rule q1)
+        apply (rule int)
+        apply rule
+        apply (rule fine_union)
+        apply (subst fine_def)
+        defer
+        apply (rule q1)
+        unfolding Ball_def split_paired_All split_conv
+        apply rule
+        apply rule
+        apply rule
+        apply rule
+        apply (erule insertE)
+        apply (simp add: uv xk)
+        apply (rule UnI2)
+        apply (drule q1(3)[rule_format])
+        unfolding xk uv
+        apply auto
+        done
+    next
+      case False
+      from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this
+      show ?thesis
+        apply (rule_tac x="q2 \<union> q1" in exI)
+        apply rule
+        unfolding * uv
+        apply (rule tagged_division_union q2 q1 int fine_union)+
+        unfolding Ball_def split_paired_All split_conv
+        apply rule
+        apply (rule fine_union)
+        apply (rule q1 q2)+
+        apply rule
+        apply rule
+        apply rule
+        apply rule
+        apply (erule insertE)
+        apply (rule UnI2)
+        apply (simp add: False uv xk)
+        apply (drule q1(3)[rule_format])
+        using False
+        unfolding xk uv
+        apply auto
+        done
+    qed
+  qed
+qed
+
+subsubsection \<open>Covering lemma\<close>
+
+text\<open> Some technical lemmas used in the approximation results that follow. Proof of the covering
+  lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's
+  "Introduction to Gauge Integrals". \<close>
+
+proposition covering_lemma:
+  assumes "S \<subseteq> cbox a b" "box a b \<noteq> {}" "gauge g"
+  obtains \<D> where
+    "countable \<D>"  "\<Union>\<D> \<subseteq> cbox a b"
+    "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
+    "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
+    "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
+    "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
+    "S \<subseteq> \<Union>\<D>"
+proof -
+  have aibi: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" and normab: "0 < norm(b - a)"
+    using \<open>box a b \<noteq> {}\<close> box_eq_empty box_sing by fastforce+
+  let ?K0 = "\<lambda>(n, f::'a\<Rightarrow>nat).
+                    cbox (\<Sum>i \<in> Basis. (a \<bullet> i + (f i / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
+                         (\<Sum>i \<in> Basis. (a \<bullet> i + ((f i + 1) / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)"
+  let ?D0 = "?K0 ` (SIGMA n:UNIV. PiE Basis (\<lambda>i::'a. lessThan (2^n)))"
+  obtain \<D>0 where count: "countable \<D>0"
+             and sub: "\<Union>\<D>0 \<subseteq> cbox a b"
+             and int:  "\<And>K. K \<in> \<D>0 \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
+             and intdj: "\<And>A B. \<lbrakk>A \<in> \<D>0; B \<in> \<D>0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
+             and SK: "\<And>x. x \<in> S \<Longrightarrow> \<exists>K \<in> \<D>0. x \<in> K \<and> K \<subseteq> g x"
+             and cbox: "\<And>u v. cbox u v \<in> \<D>0 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
+             and fin: "\<And>K. K \<in> \<D>0 \<Longrightarrow> finite {L \<in> \<D>0. K \<subseteq> L}"
+  proof
+    show "countable ?D0"
+      by (simp add: countable_PiE)
+  next
+    show "\<Union>?D0 \<subseteq> cbox a b"
+      apply (simp add: UN_subset_iff)
+      apply (intro conjI allI ballI subset_box_imp)
+       apply (simp add: divide_simps zero_le_mult_iff aibi)
+      apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem)
+      done
+  next
+    show "\<And>K. K \<in> ?D0 \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
+      using \<open>box a b \<noteq> {}\<close>
+      by (clarsimp simp: box_eq_empty) (fastforce simp add: divide_simps dest: PiE_mem)
+  next
+    have realff: "(real w) * 2^m < (real v) * 2^n \<longleftrightarrow> w * 2^m < v * 2^n" for m n v w
+      using of_nat_less_iff less_imp_of_nat_less by fastforce
+    have *: "\<forall>v w. ?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
+      for m n --\<open>The symmetry argument requires a single HOL formula\<close>
+    proof (rule linorder_wlog [where a=m and b=n], intro allI impI)
+      fix v w m and n::nat
+      assume "m \<le> n" --\<open>WLOG we can assume @{term"m \<le> n"}, when the first disjunct becomes impossible\<close>
+      have "?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
+        apply (simp add: subset_box disjoint_interval)
+        apply (rule ccontr)
+        apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le)
+        apply (drule_tac x=i in bspec, assumption)
+        using \<open>m\<le>n\<close> realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"]
+        apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff)
+        apply (simp add: power_add, metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc)+
+        done
+      then show "?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
+        by meson
+    qed auto
+    show "\<And>A B. \<lbrakk>A \<in> ?D0; B \<in> ?D0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
+      apply (erule imageE SigmaE)+
+      using * by simp
+  next
+    show "\<exists>K \<in> ?D0. x \<in> K \<and> K \<subseteq> g x" if "x \<in> S" for x
+    proof (simp only: bex_simps split_paired_Bex_Sigma)
+      show "\<exists>n. \<exists>f \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ n}. x \<in> ?K0(n,f) \<and> ?K0(n,f) \<subseteq> g x"
+      proof -
+        obtain e where "0 < e"
+                   and e: "\<And>y. (\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> e) \<Longrightarrow> y \<in> g x"
+        proof -
+          have "x \<in> g x" "open (g x)"
+            using \<open>gauge g\<close> by (auto simp: gauge_def)
+          then obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "ball x \<epsilon> \<subseteq> g x"
+            using openE by blast
+          have "norm (x - y) < \<epsilon>"
+               if "(\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> \<epsilon> / (2 * real DIM('a)))" for y
+          proof -
+            have "norm (x - y) \<le> (\<Sum>i\<in>Basis. \<bar>x \<bullet> i - y \<bullet> i\<bar>)"
+              by (metis (no_types, lifting) inner_diff_left norm_le_l1 setsum.cong)
+            also have "... \<le> DIM('a) * (\<epsilon> / (2 * real DIM('a)))"
+              by (meson setsum_bounded_above that)
+            also have "... = \<epsilon> / 2"
+              by (simp add: divide_simps)
+            also have "... < \<epsilon>"
+              by (simp add: \<open>0 < \<epsilon>\<close>)
+            finally show ?thesis .
+          qed
+          then show ?thesis
+            by (rule_tac e = "\<epsilon> / 2 / DIM('a)" in that) (simp_all add:  \<open>0 < \<epsilon>\<close> dist_norm subsetD [OF \<epsilon>])
+        qed
+        have xab: "x \<in> cbox a b"
+          using \<open>x \<in> S\<close> \<open>S \<subseteq> cbox a b\<close> by blast
+        obtain n where n: "norm (b - a) / 2^n < e"
+          using real_arch_pow_inv [of "e / norm(b - a)" "1/2"] normab \<open>0 < e\<close>
+          by (auto simp: divide_simps)
+        then have "norm (b - a) < e * 2^n"
+          by (auto simp: divide_simps)
+        then have bai: "b \<bullet> i - a \<bullet> i < e * 2 ^ n" if "i \<in> Basis" for i
+        proof -
+          have "b \<bullet> i - a \<bullet> i \<le> norm (b - a)"
+            by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that)
+          also have "... < e * 2 ^ n"
+            using \<open>norm (b - a) < e * 2 ^ n\<close> by blast
+          finally show ?thesis .
+        qed
+        have D: "(a + n \<le> x \<and> x \<le> a + m) \<Longrightarrow> (a + n \<le> y \<and> y \<le> a + m) \<Longrightarrow> abs(x - y) \<le> m - n"
+                 for a m n x and y::real
+          by auto
+        have "\<forall>i\<in>Basis. \<exists>k<2 ^ n. (a \<bullet> i + real k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i \<and>
+               x \<bullet> i \<le> a \<bullet> i + (real k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n)"
+        proof
+          fix i::'a assume "i \<in> Basis"
+          consider "x \<bullet> i = b \<bullet> i" | "x \<bullet> i < b \<bullet> i"
+            using \<open>i \<in> Basis\<close> mem_box(2) xab by force
+          then show "\<exists>k<2 ^ n. (a \<bullet> i + real k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i \<and>
+                          x \<bullet> i \<le> a \<bullet> i + (real k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n)"
+          proof cases
+            case 1 then show ?thesis
+              by (rule_tac x = "2^n - 1" in exI) (auto simp: algebra_simps divide_simps of_nat_diff \<open>i \<in> Basis\<close> aibi)
+          next
+            case 2
+            then have abi_less: "a \<bullet> i < b \<bullet> i"
+              using \<open>i \<in> Basis\<close> xab by (auto simp: mem_box)
+            let ?k = "nat \<lfloor>2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)\<rfloor>"
+            show ?thesis
+            proof (intro exI conjI)
+              show "?k < 2 ^ n"
+                using aibi xab \<open>i \<in> Basis\<close>
+                by (force simp: nat_less_iff floor_less_iff divide_simps 2 mem_box)
+            next
+              have "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le>
+                    a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
+                apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
+                using aibi [OF \<open>i \<in> Basis\<close>] xab 2
+                  apply (simp_all add: \<open>i \<in> Basis\<close> mem_box divide_simps)
+                done
+              also have "... = x \<bullet> i"
+                using abi_less by (simp add: divide_simps)
+              finally show "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i" .
+            next
+              have "x \<bullet> i \<le> a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
+                using abi_less by (simp add: divide_simps algebra_simps)
+              also have "... \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
+                apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
+                using aibi [OF \<open>i \<in> Basis\<close>] xab
+                  apply (auto simp: \<open>i \<in> Basis\<close> mem_box divide_simps)
+                done
+              finally show "x \<bullet> i \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n" .
+            qed
+          qed
+        qed
+        then have "\<exists>f\<in>Basis \<rightarrow>\<^sub>E {..<2 ^ n}. x \<in> ?K0(n,f)"
+          apply (simp add: mem_box Bex_def)
+          apply (clarify dest!: bchoice)
+          apply (rule_tac x="restrict f Basis" in exI, simp)
+          done
+        moreover have "\<And>f. x \<in> ?K0(n,f) \<Longrightarrow> ?K0(n,f) \<subseteq> g x"
+          apply (clarsimp simp add: mem_box)
+          apply (rule e)
+          apply (drule bspec D, assumption)+
+          apply (erule order_trans)
+          apply (simp add: divide_simps)
+          using bai by (force simp: algebra_simps)
+        ultimately show ?thesis by auto
+      qed
+    qed
+  next
+    show "\<And>u v. cbox u v \<in> ?D0 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
+      by (force simp: eq_cbox box_eq_empty field_simps dest!: aibi)
+  next
+    obtain j::'a where "j \<in> Basis"
+      using nonempty_Basis by blast
+    have "finite {L \<in> ?D0. ?K0(n,f) \<subseteq> L}" if "f \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ n}" for n f
+    proof (rule finite_subset)
+      let ?B = "(\<lambda>(n, f::'a\<Rightarrow>nat). cbox (\<Sum>i\<in>Basis. (a \<bullet> i + (f i) / 2^n * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
+                                        (\<Sum>i\<in>Basis. (a \<bullet> i + ((f i) + 1) / 2^n * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i))
+                ` (SIGMA m:atMost n. PiE Basis (\<lambda>i::'a. lessThan (2^m)))"
+      have "?K0(m,g) \<in> ?B" if "g \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ m}" "?K0(n,f) \<subseteq> ?K0(m,g)" for m g
+      proof -
+        have dd: "w / m \<le> v / n \<and> (v+1) / n \<le> (w+1) / m
+                  \<Longrightarrow> inverse n \<le> inverse m" for w m v n::real
+          by (auto simp: divide_simps algebra_simps)
+        have bjaj: "b \<bullet> j - a \<bullet> j > 0"
+          using \<open>j \<in> Basis\<close> \<open>box a b \<noteq> {}\<close> box_eq_empty(1) by fastforce
+        have "((g j) / 2 ^ m) * (b \<bullet> j - a \<bullet> j) \<le> ((f j) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<and>
+              (((f j) + 1) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<le> (((g j) + 1) / 2 ^ m) * (b \<bullet> j - a \<bullet> j)"
+          using that \<open>j \<in> Basis\<close> by (simp add: subset_box algebra_simps divide_simps aibi)
+        then have "((g j) / 2 ^ m) \<le> ((f j) / 2 ^ n) \<and>
+          ((real(f j) + 1) / 2 ^ n) \<le> ((real(g j) + 1) / 2 ^ m)"
+          by (metis bjaj mult.commute of_nat_1 of_nat_add real_mult_le_cancel_iff2)
+        then have "inverse (2^n) \<le> (inverse (2^m) :: real)"
+          by (rule dd)
+        then have "m \<le> n"
+          by auto
+        show ?thesis
+          by (rule imageI) (simp add: \<open>m \<le> n\<close> that)
+      qed
+      then show "{L \<in> ?D0. ?K0(n,f) \<subseteq> L} \<subseteq> ?B"
+        by auto
+      show "finite ?B"
+        by (intro finite_imageI finite_SigmaI finite_atMost finite_lessThan finite_PiE finite_Basis)
+    qed
+    then show "finite {L \<in> ?D0. K \<subseteq> L}" if "K \<in> ?D0" for K
+      using that by auto
+  qed
+  let ?D1 = "{K \<in> \<D>0. \<exists>x \<in> S \<inter> K. K \<subseteq> g x}"
+  obtain \<D> where count: "countable \<D>"
+             and sub: "\<Union>\<D> \<subseteq> cbox a b"  "S \<subseteq> \<Union>\<D>"
+             and int:  "\<And>K. K \<in> \<D> \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
+             and intdj: "\<And>A B. \<lbrakk>A \<in> \<D>; B \<in> \<D>\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
+             and SK: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x. x \<in> S \<inter> K \<and> K \<subseteq> g x"
+             and cbox: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
+             and fin: "\<And>K. K \<in> \<D> \<Longrightarrow> finite {L. L \<in> \<D> \<and> K \<subseteq> L}"
+  proof
+    show "countable ?D1" using count countable_subset
+      by (simp add: count countable_subset)
+    show "\<Union>?D1 \<subseteq> cbox a b"
+      using sub by blast
+    show "S \<subseteq> \<Union>?D1"
+      using SK by (force simp:)
+    show "\<And>K. K \<in> ?D1 \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
+      using int by blast
+    show "\<And>A B. \<lbrakk>A \<in> ?D1; B \<in> ?D1\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
+      using intdj by blast
+    show "\<And>K. K \<in> ?D1 \<Longrightarrow> \<exists>x. x \<in> S \<inter> K \<and> K \<subseteq> g x"
+      by auto
+    show "\<And>u v. cbox u v \<in> ?D1 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
+      using cbox by blast
+    show "\<And>K. K \<in> ?D1 \<Longrightarrow> finite {L. L \<in> ?D1 \<and> K \<subseteq> L}"
+      using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset)
+  qed
+  let ?\<D> = "{K \<in> \<D>. \<forall>K'. K' \<in> \<D> \<and> K \<noteq> K' \<longrightarrow> ~(K \<subseteq> K')}"
+  show ?thesis
+  proof (rule that)
+    show "countable ?\<D>"
+      by (blast intro: countable_subset [OF _ count])
+    show "\<Union>?\<D> \<subseteq> cbox a b"
+      using sub by blast
+    show "S \<subseteq> \<Union>?\<D>"
+    proof clarsimp
+      fix x
+      assume "x \<in> S"
+      then obtain X where "x \<in> X" "X \<in> \<D>" using \<open>S \<subseteq> \<Union>\<D>\<close> by blast
+      let ?R = "{(K,L). K \<in> \<D> \<and> L \<in> \<D> \<and> L \<subset> K}"
+      have irrR: "irrefl ?R" by (force simp: irrefl_def)
+      have traR: "trans ?R" by (force simp: trans_def)
+      have finR: "\<And>x. finite {y. (y, x) \<in> ?R}"
+        by simp (metis (mono_tags, lifting) fin \<open>X \<in> \<D>\<close> finite_subset mem_Collect_eq psubset_imp_subset subsetI)
+      have "{X \<in> \<D>. x \<in> X} \<noteq> {}"
+        using \<open>X \<in> \<D>\<close> \<open>x \<in> X\<close> by blast
+      then obtain Y where "Y \<in> {X \<in> \<D>. x \<in> X}" "\<And>Y'. (Y', Y) \<in> ?R \<Longrightarrow> Y' \<notin> {X \<in> \<D>. x \<in> X}"
+        by (rule wfE_min' [OF wf_finite_segments [OF irrR traR finR]]) blast
+      then show "\<exists>Y. Y \<in> \<D> \<and> (\<forall>K'. K' \<in> \<D> \<and> Y \<noteq> K' \<longrightarrow> \<not> Y \<subseteq> K') \<and> x \<in> Y"
+        by blast
+    qed
+    show "\<And>K. K \<in> ?\<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
+      by (blast intro: dest: int)
+    show "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) ?\<D>"
+      using intdj by (simp add: pairwise_def) metis
+    show "\<And>K. K \<in> ?\<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
+      using SK by force
+    show "\<And>u v. cbox u v \<in> ?\<D> \<Longrightarrow> \<exists>n. \<forall>i\<in>Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
+      using cbox by force
+    qed
+qed
+
+subsection \<open>Division filter\<close>
+
+text \<open>Divisions over all gauges towards finer divisions.\<close>
+
+definition division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
+  where "division_filter s = (INF g:{g. gauge g}. principal {p. p tagged_division_of s \<and> g fine p})"
+
+lemma eventually_division_filter:
+  "(\<forall>\<^sub>F p in division_filter s. P p) \<longleftrightarrow>
+    (\<exists>g. gauge g \<and> (\<forall>p. p tagged_division_of s \<and> g fine p \<longrightarrow> P p))"
+  unfolding division_filter_def
+proof (subst eventually_INF_base; clarsimp)
+  fix g1 g2 :: "'a \<Rightarrow> 'a set" show "gauge g1 \<Longrightarrow> gauge g2 \<Longrightarrow> \<exists>x. gauge x \<and>
+    {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g1 fine p} \<and>
+    {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g2 fine p}"
+    by (intro exI[of _ "\<lambda>x. g1 x \<inter> g2 x"]) (auto simp: fine_inter)
+qed (auto simp: eventually_principal)
+
+lemma division_filter_not_empty: "division_filter (cbox a b) \<noteq> bot"
+  unfolding trivial_limit_def eventually_division_filter
+  by (auto elim: fine_division_exists)
+
+lemma eventually_division_filter_tagged_division:
+  "eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
+  unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto
+
+end
\ No newline at end of file
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -550,6 +550,11 @@
 lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union>{A - s|s. s\<in>S}"
   by auto
 
+lemma closedin_Union:
+  assumes "finite S" "\<And>T. T \<in> S \<Longrightarrow> closedin U T"
+    shows "closedin U (\<Union>S)"
+  using assms by induction auto
+
 lemma closedin_Inter[intro]:
   assumes Ke: "K \<noteq> {}"
     and Kc: "\<And>S. S \<in>K \<Longrightarrow> closedin U S"
@@ -1409,7 +1414,7 @@
   then have "cbox a b \<subseteq> box c d" "box c d \<subseteq>cbox a b"
     by auto
   then show ?rhs
-    apply (simp add: subset_box) 
+    apply (simp add: subset_box)
     using \<open>cbox a b = box c d\<close> box_ne_empty box_sing
     apply (fastforce simp add:)
     done
@@ -1430,7 +1435,7 @@
     by auto
   then show ?rhs
     apply (simp add: subset_box)
-    using box_ne_empty(2) \<open>box a b = box c d\<close> 
+    using box_ne_empty(2) \<open>box a b = box c d\<close>
     apply auto
      apply (meson euclidean_eqI less_eq_real_def not_less)+
     done
@@ -6164,6 +6169,22 @@
     using * by auto
 qed
 
+lemma continuous_openin_preimage_eq:
+   "continuous_on S f \<longleftrightarrow>
+    (\<forall>t. open t \<longrightarrow> openin (subtopology euclidean S) {x. x \<in> S \<and> f x \<in> t})"
+apply safe
+apply (simp add: continuous_openin_preimage_gen)
+apply (fastforce simp add: continuous_on_open openin_open)
+done
+
+lemma continuous_closedin_preimage_eq:
+   "continuous_on S f \<longleftrightarrow>
+    (\<forall>t. closed t \<longrightarrow> closedin (subtopology euclidean S) {x. x \<in> S \<and> f x \<in> t})"
+apply safe
+apply (simp add: continuous_closedin_preimage)
+apply (fastforce simp add: continuous_on_closed closedin_closed)
+done
+
 lemma continuous_open_preimage:
   assumes "continuous_on s f"
     and "open s"
@@ -9893,6 +9914,178 @@
     by (force simp: bounded_linear_def bounded_linear_axioms_def \<open>linear f'\<close>)
 qed
 
+subsection\<open>Pasting functions together\<close>
+
+subsubsection\<open>on open sets\<close>
+
+lemma pasting_lemma:
+  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)"
+      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
+      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+      and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
+    shows "continuous_on S g"
+proof (clarsimp simp: continuous_openin_preimage_eq)
+  fix U :: "'b set"
+  assume "open U"
+  have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S"
+    using clo openin_imp_subset by blast
+  have *: "{x \<in> S. g x \<in> U} = \<Union>{{x. x \<in> (T i) \<and> (f i x) \<in> U} |i. i \<in> I}"
+    apply (auto simp: dest: S)
+      apply (metis (no_types, lifting) g mem_Collect_eq)
+    using clo f g openin_imp_subset by fastforce
+  show "openin (subtopology euclidean S) {x \<in> S. g x \<in> U}"
+    apply (subst *)
+    apply (rule openin_Union, clarify)
+    apply (metis (full_types) \<open>open U\<close> cont clo openin_trans continuous_openin_preimage_gen)
+    done
+qed 
+
+lemma pasting_lemma_exists:
+  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes S: "S \<subseteq> (\<Union>i \<in> I. T i)"
+      and clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)"
+      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
+      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+    obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
+proof
+  show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
+    apply (rule pasting_lemma [OF clo cont])
+     apply (blast intro: f)+
+    apply (metis (mono_tags, lifting) S UN_iff subsetCE someI)
+    done
+next
+  fix x i
+  assume "i \<in> I" "x \<in> S \<inter> T i"
+  then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
+    by (metis (no_types, lifting) IntD2 IntI f someI_ex)
+qed
+
+subsubsection\<open>Likewise on closed sets, with a finiteness assumption\<close>
+
+lemma pasting_lemma_closed:
+  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes "finite I"
+      and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)"
+      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
+      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+      and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
+    shows "continuous_on S g"
+proof (clarsimp simp: continuous_closedin_preimage_eq)
+  fix U :: "'b set"
+  assume "closed U"
+  have *: "{x \<in> S. g x \<in> U} = \<Union>{{x. x \<in> (T i) \<and> (f i x) \<in> U} |i. i \<in> I}"
+    apply auto
+    apply (metis (no_types, lifting) g mem_Collect_eq)
+    using clo closedin_closed apply blast
+    apply (metis Int_iff f g clo closedin_limpt inf.absorb_iff2)
+    done
+  show "closedin (subtopology euclidean S) {x \<in> S. g x \<in> U}"
+    apply (subst *)
+    apply (rule closedin_Union)
+    using \<open>finite I\<close> apply simp
+    apply (blast intro: \<open>closed U\<close> continuous_closedin_preimage cont clo closedin_trans)
+    done
+qed
+
+lemma pasting_lemma_exists_closed:
+  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes "finite I"
+      and S: "S \<subseteq> (\<Union>i \<in> I. T i)"
+      and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)"
+      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
+      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+    obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
+proof
+  show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
+    apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont])
+     apply (blast intro: f)+
+    apply (metis (mono_tags, lifting) S UN_iff subsetCE someI)
+    done
+next
+  fix x i
+  assume "i \<in> I" "x \<in> S \<inter> T i"
+  then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
+    by (metis (no_types, lifting) IntD2 IntI f someI_ex)
+qed
+
+lemma tube_lemma:
+  assumes "compact K"
+  assumes "open W"
+  assumes "{x0} \<times> K \<subseteq> W"
+  shows "\<exists>X0. x0 \<in> X0 \<and> open X0 \<and> X0 \<times> K \<subseteq> W"
+proof -
+  {
+    fix y assume "y \<in> K"
+    then have "(x0, y) \<in> W" using assms by auto
+    with \<open>open W\<close>
+    have "\<exists>X0 Y. open X0 \<and> open Y \<and> x0 \<in> X0 \<and> y \<in> Y \<and> X0 \<times> Y \<subseteq> W"
+      by (rule open_prod_elim) blast
+  }
+  then obtain X0 Y where
+    *: "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W"
+    by metis
+  from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
+  with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
+    by (rule compactE)
+  then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
+    by (force intro!: choice)
+  with * CC show ?thesis
+    by (force intro!: exI[where x="\<Inter>C\<in>CC. X0 (c C)"]) (* SLOW *)
+qed
+
+lemma continuous_on_prod_compactE:
+  fixes fx::"'a::topological_space \<times> 'b::topological_space \<Rightarrow> 'c::metric_space"
+    and e::real
+  assumes cont_fx: "continuous_on (U \<times> C) fx"
+  assumes "compact C"
+  assumes [intro]: "x0 \<in> U"
+  notes [continuous_intros] = continuous_on_compose2[OF cont_fx]
+  assumes "e > 0"
+  obtains X0 where "x0 \<in> X0" "open X0"
+    "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
+proof -
+  define psi where "psi = (\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t)))"
+  define W0 where "W0 = {(x, t) \<in> U \<times> C. psi (x, t) < e}"
+  have W0_eq: "W0 = psi -` {..<e} \<inter> U \<times> C"
+    by (auto simp: vimage_def W0_def)
+  have "open {..<e}" by simp
+  have "continuous_on (U \<times> C) psi"
+    by (auto intro!: continuous_intros simp: psi_def split_beta')
+  from this[unfolded continuous_on_open_invariant, rule_format, OF \<open>open {..<e}\<close>]
+  obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C"
+    unfolding W0_eq by blast
+  have "{x0} \<times> C \<subseteq> W \<inter> U \<times> C"
+    unfolding W
+    by (auto simp: W0_def psi_def \<open>0 < e\<close>)
+  then have "{x0} \<times> C \<subseteq> W" by blast
+  from tube_lemma[OF \<open>compact C\<close> \<open>open W\<close> this]
+  obtain X0 where X0: "x0 \<in> X0" "open X0" "X0 \<times> C \<subseteq> W"
+    by blast
+
+  have "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
+  proof safe
+    fix x assume x: "x \<in> X0" "x \<in> U"
+    fix t assume t: "t \<in> C"
+    have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)"
+      by (auto simp: psi_def)
+    also
+    {
+      have "(x, t) \<in> X0 \<times> C"
+        using t x
+        by auto
+      also note \<open>\<dots> \<subseteq> W\<close>
+      finally have "(x, t) \<in> W" .
+      with t x have "(x, t) \<in> W \<inter> U \<times> C"
+        by blast
+      also note \<open>W \<inter> U \<times> C = W0 \<inter> U \<times> C\<close>
+      finally  have "psi (x, t) < e"
+        by (auto simp: W0_def)
+    }
+    finally show "dist (fx (x, t)) (fx (x0, t)) \<le> e" by simp
+  qed
+  from X0(1,2) this show ?thesis ..
+qed
 
 no_notation
   eucl_less (infix "<e" 50)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Argo.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,27 @@
+(*  Title:      HOL/Argo.thy
+    Author:     Sascha Boehme
+*)
+
+theory Argo
+imports HOL
+begin
+
+ML_file "~~/src/Tools/Argo/argo_expr.ML"
+ML_file "~~/src/Tools/Argo/argo_term.ML"
+ML_file "~~/src/Tools/Argo/argo_lit.ML"
+ML_file "~~/src/Tools/Argo/argo_proof.ML"
+ML_file "~~/src/Tools/Argo/argo_rewr.ML"
+ML_file "~~/src/Tools/Argo/argo_cls.ML"
+ML_file "~~/src/Tools/Argo/argo_common.ML"
+ML_file "~~/src/Tools/Argo/argo_cc.ML"
+ML_file "~~/src/Tools/Argo/argo_simplex.ML"
+ML_file "~~/src/Tools/Argo/argo_thy.ML"
+ML_file "~~/src/Tools/Argo/argo_heap.ML"
+ML_file "~~/src/Tools/Argo/argo_cdcl.ML"
+ML_file "~~/src/Tools/Argo/argo_core.ML"
+ML_file "~~/src/Tools/Argo/argo_clausify.ML"
+ML_file "~~/src/Tools/Argo/argo_solver.ML"
+
+ML_file "Tools/Argo/argo_tactic.ML"
+
+end
--- a/src/HOL/Library/Indicator_Function.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Library/Indicator_Function.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -28,6 +28,9 @@
 lemma indicator_eq_1_iff: "indicator A x = (1::'a::zero_neq_one) \<longleftrightarrow> x \<in> A"
   by (auto simp: indicator_def)
 
+lemma indicator_UNIV [simp]: "indicator UNIV = (\<lambda>x. 1)"
+  by auto
+
 lemma indicator_leI:
   "(x \<in> A \<Longrightarrow> y \<in> B) \<Longrightarrow> (indicator A x :: 'a::linordered_nonzero_semiring) \<le> indicator B y"
   by (auto simp: indicator_def)
--- a/src/HOL/ROOT	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/ROOT	Fri Sep 30 10:00:49 2016 +0200
@@ -626,6 +626,7 @@
     Sudoku
     Code_Timing
     Perm_Fragments
+    Argo_Examples
   theories [skip_proofs = false]
     Meson_Test
   document_files "root.bib" "root.tex"
--- a/src/HOL/Real.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Real.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -1807,4 +1807,9 @@
   for x y :: real
   by auto
 
+
+subsection \<open>Setup for Argo\<close>
+
+ML_file "Tools/Argo/argo_real.ML"
+
 end
--- a/src/HOL/SAT.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/SAT.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -8,12 +8,13 @@
 section \<open>Reconstructing external resolution proofs for propositional logic\<close>
 
 theory SAT
-imports HOL
+imports Argo
 begin
 
 ML_file "Tools/prop_logic.ML"
 ML_file "Tools/sat_solver.ML"
 ML_file "Tools/sat.ML"
+ML_file "Tools/Argo/argo_sat_solver.ML"
 
 method_setup sat = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac)\<close>
   "SAT solver"
--- a/src/HOL/Set.thy	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Set.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -1847,7 +1847,7 @@
 text \<open>Misc\<close>
 
 definition pairwise :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
-  where "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x \<noteq> y \<longrightarrow> R x y)"
+  where "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. x \<noteq> y \<longrightarrow> R x y)"
 
 lemma pairwise_subset: "pairwise P S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> pairwise P T"
   by (force simp: pairwise_def)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Argo/argo_real.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,297 @@
+(*  Title:      HOL/Tools/argo_real.ML
+    Author:     Sascha Boehme
+
+Extension of the Argo tactic for the reals.
+*)
+
+structure Argo_Real: sig end =
+struct
+
+(* translating input terms *)
+
+fun trans_type _ @{typ Real.real} tcx = SOME (Argo_Expr.Real, tcx)
+  | trans_type _ _ _ = NONE
+
+fun trans_term f (@{const Groups.uminus_class.uminus (real)} $ t) tcx =
+      tcx |> f t |>> Argo_Expr.mk_neg |> SOME
+  | trans_term f (@{const Groups.plus_class.plus (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_add2 |> SOME
+  | trans_term f (@{const Groups.minus_class.minus (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_sub |> SOME
+  | trans_term f (@{const Groups.times_class.times (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_mul |> SOME
+  | trans_term f (@{const Rings.divide_class.divide (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_div |> SOME
+  | trans_term f (@{const Orderings.ord_class.min (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_min |> SOME
+  | trans_term f (@{const Orderings.ord_class.max (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_max |> SOME
+  | trans_term f (@{const Groups.abs_class.abs (real)} $ t) tcx =
+      tcx |> f t |>> Argo_Expr.mk_abs |> SOME
+  | trans_term f (@{const Orderings.ord_class.less (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_lt |> SOME
+  | trans_term f (@{const Orderings.ord_class.less_eq (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_le |> SOME
+  | trans_term _ t tcx =
+      (case try HOLogic.dest_number t of
+        SOME (@{typ Real.real}, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx)
+      | _ => NONE)
+
+
+(* reverse translation *)
+
+fun mk_plus t1 t2 = @{const Groups.plus_class.plus (real)} $ t1 $ t2
+fun mk_sum ts = uncurry (fold_rev mk_plus) (split_last ts)
+fun mk_times t1 t2 = @{const Groups.times_class.times (real)} $ t1 $ t2
+fun mk_divide t1 t2 = @{const Rings.divide_class.divide (real)} $ t1 $ t2
+fun mk_le t1 t2 = @{const Orderings.ord_class.less_eq (real)} $ t1 $ t2
+fun mk_lt t1 t2 = @{const Orderings.ord_class.less (real)} $ t1 $ t2
+
+fun mk_real_num i = HOLogic.mk_number @{typ Real.real} i
+
+fun mk_number n =
+  let val (p, q) = Rat.dest n
+  in if q = 1 then mk_real_num p else mk_divide (mk_real_num p) (mk_real_num q) end
+
+fun term_of _ (Argo_Expr.E (Argo_Expr.Num n, _)) = SOME (mk_number n)
+  | term_of f (Argo_Expr.E (Argo_Expr.Neg, [e])) =
+      SOME (@{const Groups.uminus_class.uminus (real)} $ f e)
+  | term_of f (Argo_Expr.E (Argo_Expr.Add, es)) = SOME (mk_sum (map f es))
+  | term_of f (Argo_Expr.E (Argo_Expr.Sub, [e1, e2])) =
+      SOME (@{const Groups.minus_class.minus (real)} $ f e1 $ f e2)
+  | term_of f (Argo_Expr.E (Argo_Expr.Mul, [e1, e2])) = SOME (mk_times (f e1) (f e2))
+  | term_of f (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) = SOME (mk_divide (f e1) (f e2))
+  | term_of f (Argo_Expr.E (Argo_Expr.Min, [e1, e2])) =
+      SOME (@{const Orderings.ord_class.min (real)} $ f e1 $ f e2)
+  | term_of f (Argo_Expr.E (Argo_Expr.Max, [e1, e2])) =
+      SOME (@{const Orderings.ord_class.max (real)} $ f e1 $ f e2)
+  | term_of f (Argo_Expr.E (Argo_Expr.Abs, [e])) = SOME (@{const Groups.abs_class.abs (real)} $ f e)
+  | term_of f (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (mk_le (f e1) (f e2))
+  | term_of f (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (mk_lt (f e1) (f e2))
+  | term_of _ _ = NONE
+
+
+(* proof replay for rewrite steps *)
+
+fun mk_rewr thm = thm RS @{thm eq_reflection}
+
+fun by_simp ctxt t = 
+  let fun prove {context, ...} = HEADGOAL (Simplifier.simp_tac context)
+  in Goal.prove ctxt [] [] (HOLogic.mk_Trueprop t) prove end
+
+fun prove_num_pred ctxt n =
+  by_simp ctxt (uncurry mk_lt (apply2 mk_number (if @0 < n then (@0, n) else (n, @0))))
+
+fun simp_conv ctxt t = Conv.rewr_conv (mk_rewr (by_simp ctxt t))
+
+fun nums_conv mk f ctxt n m =
+  simp_conv ctxt (HOLogic.mk_eq (mk (mk_number n) (mk_number m), mk_number (f (n, m))))
+
+val add_nums_conv = nums_conv mk_plus (op +)
+val mul_nums_conv = nums_conv mk_times (op *)
+val div_nums_conv = nums_conv mk_divide (op /)
+
+fun cmp_nums_conv ctxt b ct =
+  let val t = if b then @{const HOL.True} else @{const HOL.False}
+  in simp_conv ctxt (HOLogic.mk_eq (Thm.term_of ct, t)) ct end
+
+local
+
+fun is_add2 (@{const Groups.plus_class.plus (real)} $ _ $ _) = true
+  | is_add2 _ = false
+
+fun is_add3 (@{const Groups.plus_class.plus (real)} $ _ $ t) = is_add2 t
+  | is_add3 _ = false
+
+val flatten_thm = mk_rewr @{lemma "(a::real) + b + c = a + (b + c)" by simp}
+  
+fun flatten_conv ct =
+  if is_add2 (Thm.term_of ct) then Argo_Tactic.flatten_conv flatten_conv flatten_thm ct
+  else Conv.all_conv ct
+
+val swap_conv = Conv.rewrs_conv (map mk_rewr @{lemma 
+  "(a::real) + (b + c) = b + (a + c)"
+  "(a::real) + b = b + a"
+  by simp_all})
+
+val assoc_conv = Conv.rewr_conv (mk_rewr @{lemma "(a::real) + (b + c) = (a + b) + c" by simp})
+
+val norm_monom_thm = mk_rewr @{lemma "1 * (a::real) = a" by simp}
+fun norm_monom_conv n = if n = @1 then Conv.rewr_conv norm_monom_thm else Conv.all_conv
+
+val add2_thms = map mk_rewr @{lemma
+  "n * (a::real) + m * a = (n + m) * a"
+  "n * (a::real) + a = (n + 1) * a"
+  "(a::real) + m * a = (1 + m) * a"
+  "(a::real) + a = (1 + 1) * a"
+  by algebra+}
+
+val add3_thms = map mk_rewr @{lemma
+  "n * (a::real) + (m * a + b) = (n + m) * a + b"
+  "n * (a::real) + (a + b) = (n + 1) * a + b"
+  "(a::real) + (m * a + b) = (1 + m) * a + b"
+  "(a::real) + (a + b) = (1 + 1) * a + b"
+  by algebra+}
+
+fun choose_conv cv2 cv3 ct = if is_add3 (Thm.term_of ct) then cv3 ct else cv2 ct
+
+fun join_num_conv ctxt n m =
+  let val conv = add_nums_conv ctxt n m
+  in choose_conv conv (assoc_conv then_conv Conv.arg1_conv conv) end
+
+fun join_monom_conv ctxt n m =
+  let
+    val conv = Conv.arg1_conv (add_nums_conv ctxt n m) then_conv norm_monom_conv (n + m)
+    fun seq_conv thms cv = Conv.rewrs_conv thms then_conv cv
+  in choose_conv (seq_conv add2_thms conv) (seq_conv add3_thms (Conv.arg1_conv conv)) end
+
+fun join_conv NONE = join_num_conv
+  | join_conv (SOME _) = join_monom_conv
+
+fun bubble_down_conv _ _ [] ct = Conv.no_conv ct
+  | bubble_down_conv _ _ [_] ct = Conv.all_conv ct
+  | bubble_down_conv ctxt i ((m1 as (n1, i1)) :: (m2 as (n2, i2)) :: ms) ct =
+      if i1 = i then
+        if i2 = i then
+          (join_conv i ctxt n1 n2 then_conv bubble_down_conv ctxt i ((n1 + n2, i) :: ms)) ct
+        else (swap_conv then_conv Conv.arg_conv (bubble_down_conv ctxt i (m1 :: ms))) ct
+      else Conv.arg_conv (bubble_down_conv ctxt i (m2 :: ms)) ct
+
+fun drop_var i ms = filter_out (fn (_, i') => i' = i) ms
+
+fun permute_conv _ [] [] ct = Conv.all_conv ct
+  | permute_conv ctxt (ms as ((_, i) :: _)) [] ct =
+      (bubble_down_conv ctxt i ms then_conv permute_conv ctxt (drop_var i ms) []) ct
+  | permute_conv ctxt ms1 ms2 ct =
+      let val (ms2', (_, i)) = split_last ms2
+      in (bubble_down_conv ctxt i ms1 then_conv permute_conv ctxt (drop_var i ms1) ms2') ct end
+
+val no_monom_conv = Conv.rewr_conv (mk_rewr @{lemma "0 * (a::real) = 0" by simp})
+
+val norm_sum_conv = Conv.rewrs_conv (map mk_rewr @{lemma
+  "0 * (a::real) + b = b"
+  "(a::real) + 0 * b = a"
+  "0 + (a::real) = a"
+  "(a::real) + 0 = a"
+  by simp_all})
+
+fun drop0_conv ct =
+  if is_add2 (Thm.term_of ct) then
+    ((norm_sum_conv then_conv drop0_conv) else_conv Conv.arg_conv drop0_conv) ct
+  else Conv.try_conv no_monom_conv ct
+
+fun full_add_conv ctxt ms1 ms2 =
+  if eq_list (op =) (ms1, ms2) then flatten_conv
+  else flatten_conv then_conv permute_conv ctxt ms1 ms2 then_conv drop0_conv
+
+in
+
+fun add_conv ctxt (ms1, ms2 as [(n, NONE)]) =
+      if n = @0 then full_add_conv ctxt ms1 [] else full_add_conv ctxt ms1 ms2
+  | add_conv ctxt (ms1, ms2) = full_add_conv ctxt ms1 ms2
+
+end
+
+val mul_sum_thm = mk_rewr @{lemma "(x::real) * (y + z) = x * y + x * z" by (rule ring_distribs)}
+fun mul_sum_conv ct =
+  Conv.try_conv (Conv.rewr_conv mul_sum_thm then_conv Conv.binop_conv mul_sum_conv) ct
+
+fun var_of_add_cmp (_ $ _ $ (_ $ _ $ (_ $ _ $ Var v))) = v
+  | var_of_add_cmp t = raise TERM ("var_of_add_cmp", [t])
+
+fun add_cmp_conv ctxt n thm =
+  let val v = var_of_add_cmp (Thm.prop_of thm)
+  in Conv.rewr_conv (Thm.instantiate ([], [(v, Thm.cterm_of ctxt (mk_number n))]) thm) end
+
+fun mul_cmp_conv ctxt n pos_thm neg_thm =
+  let val thm = if @0 < n then pos_thm else neg_thm
+  in Conv.rewr_conv (prove_num_pred ctxt n RS thm) end
+
+val neg_thm = mk_rewr @{lemma "-(x::real) = -1 * x" by simp}
+val sub_thm = mk_rewr @{lemma "(x::real) - y = x + -1 * y" by simp}
+val mul_zero_thm = mk_rewr @{lemma "0 * (x::real) = 0" by (rule mult_zero_left)}
+val mul_one_thm = mk_rewr @{lemma "1 * (x::real) = x" by (rule mult_1)}
+val mul_comm_thm = mk_rewr @{lemma "(x::real) * y = y * x" by simp}
+val mul_assoc_thm = mk_rewr @{lemma "(x::real) * (y * z) = (x * y) * z" by simp}
+val div_zero_thm = mk_rewr @{lemma "0 / (x::real) = 0" by simp}
+val div_one_thm = mk_rewr @{lemma "(x::real) / 1 = x" by simp}
+val div_mul_thm = mk_rewr @{lemma "(x::real) / y = x * (1 / y)" by simp}
+val div_inv_thm = mk_rewr @{lemma "(x::real) / y = (1 / y) * x" by simp}
+val div_left_thm = mk_rewr @{lemma "((x::real) * y) / z = x * (y / z)" by simp}
+val div_right_thm = mk_rewr @{lemma "(x::real) / (y * z) = (1 / y) * (x / z)" by simp}
+val min_thm = mk_rewr @{lemma "min (x::real) y = (if x <= y then x else y)" by (rule min_def)}
+val max_thm = mk_rewr @{lemma "max (x::real) y = (if x <= y then y else x)" by (rule max_def)}
+val abs_thm = mk_rewr @{lemma "abs (x::real) = (if 0 <= x then x else -x)" by simp}
+val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x <= y) & (y <= x))" by (rule order_eq_iff)}
+val add_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x + n <= y + n)" by simp}
+val add_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x + n < y + n)" by simp}
+val sub_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x - y <= 0)" by simp}
+val sub_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x - y < 0)" by simp}
+val pos_mul_le_thm = mk_rewr @{lemma "0 < n ==> ((x::real) <= y) = (n * x <= n * y)" by simp}
+val neg_mul_le_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) <= y) = (n * y <= n * x)" by simp}
+val pos_mul_lt_thm = mk_rewr @{lemma "0 < n ==> ((x::real) < y) = (n * x < n * y)" by simp}
+val neg_mul_lt_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) < y) = (n * y < n * x)" by simp}
+val not_le_thm = mk_rewr @{lemma "(~((x::real) <= y)) = (y < x)" by auto}
+val not_lt_thm = mk_rewr @{lemma "(~((x::real) < y)) = (y <= x)" by auto}
+
+fun replay_rewr _ Argo_Proof.Rewr_Neg = Conv.rewr_conv neg_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Add ps) = add_conv ctxt ps
+  | replay_rewr _ Argo_Proof.Rewr_Sub = Conv.rewr_conv sub_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Zero = Conv.rewr_conv mul_zero_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_One = Conv.rewr_conv mul_one_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Mul_Nums (n, m)) = mul_nums_conv ctxt n m
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Comm = Conv.rewr_conv mul_comm_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Assoc = Conv.rewr_conv mul_assoc_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Sum = mul_sum_conv
+  | replay_rewr ctxt (Argo_Proof.Rewr_Div_Nums (n, m)) = div_nums_conv ctxt n m
+  | replay_rewr _ Argo_Proof.Rewr_Div_Zero = Conv.rewr_conv div_zero_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_One = Conv.rewr_conv div_one_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Mul = Conv.rewr_conv div_mul_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Inv = Conv.rewr_conv div_inv_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Left = Conv.rewr_conv div_left_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Right = Conv.rewr_conv div_right_thm
+  | replay_rewr _ Argo_Proof.Rewr_Min = Conv.rewr_conv min_thm
+  | replay_rewr _ Argo_Proof.Rewr_Max = Conv.rewr_conv max_thm
+  | replay_rewr _ Argo_Proof.Rewr_Abs = Conv.rewr_conv abs_thm
+  | replay_rewr _ Argo_Proof.Rewr_Eq_Le = Conv.rewr_conv eq_le_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Nums (_, b)) = cmp_nums_conv ctxt b
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Le, n)) =
+      add_cmp_conv ctxt n add_le_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Lt, n)) =
+      add_cmp_conv ctxt n add_lt_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Le) = Conv.rewr_conv sub_le_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Lt) = Conv.rewr_conv sub_lt_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Le, n)) =
+      mul_cmp_conv ctxt n pos_mul_le_thm neg_mul_le_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Lt, n)) =
+      mul_cmp_conv ctxt n pos_mul_lt_thm neg_mul_lt_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) = Conv.rewr_conv not_le_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) = Conv.rewr_conv not_lt_thm
+  | replay_rewr _ _ = Conv.no_conv
+
+
+(* proof replay *)
+
+val combine_thms = @{lemma
+  "(a::real) < b ==> c < d ==> a + c < b + d"
+  "(a::real) < b ==> c <= d ==> a + c < b + d"
+  "(a::real) <= b ==> c < d ==> a + c < b + d"
+  "(a::real) <= b ==> c <= d ==> a + c <= b + d"
+  by auto}
+
+fun combine thm1 thm2 = hd (Argo_Tactic.discharges thm2 (Argo_Tactic.discharges thm1 combine_thms))
+
+fun replay _ _ Argo_Proof.Linear_Comb prems = SOME (uncurry (fold_rev combine) (split_last prems))
+  | replay _ _ _ _ = NONE
+
+
+(* real extension of the Argo solver *)
+
+val _ = Theory.setup (Argo_Tactic.add_extension {
+  trans_type = trans_type,
+  trans_term = trans_term,
+  term_of = term_of,
+  replay_rewr = replay_rewr,
+  replay = replay})
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Argo/argo_sat_solver.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,27 @@
+(*  Title:      HOL/Tools/argo_sat_solver.ML
+    Author:     Sascha Boehme
+
+A SAT solver based on the Argo solver.
+This SAT solver produces models. For proofs use Argo_Tactic.prove instead.
+*)
+
+structure Argo_SAT_Solver: sig end =
+struct
+
+fun con_of i = (string_of_int i, Argo_Expr.Bool)
+
+fun expr_of Prop_Logic.True = Argo_Expr.true_expr
+  | expr_of Prop_Logic.False = Argo_Expr.false_expr
+  | expr_of (Prop_Logic.Not p) = Argo_Expr.mk_not (expr_of p)
+  | expr_of (Prop_Logic.Or (p1, p2)) = Argo_Expr.mk_or2 (expr_of p1) (expr_of p2)
+  | expr_of (Prop_Logic.And (p1, p2)) = Argo_Expr.mk_and2 (expr_of p1) (expr_of p2)
+  | expr_of (Prop_Logic.BoolVar i) = Argo_Expr.mk_con (con_of i)
+
+fun argo_solver p =
+  let val argo = Argo_Solver.assert [expr_of p] Argo_Solver.context
+  in SAT_Solver.SATISFIABLE (Argo_Solver.model_of argo o con_of) end
+  handle Argo_Proof.UNSAT _ => SAT_Solver.UNSATISFIABLE NONE
+
+val () = SAT_Solver.add_solver ("argo", argo_solver)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Argo/argo_tactic.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,730 @@
+(*  Title:      HOL/Tools/argo_tactic.ML
+    Author:     Sascha Boehme
+
+HOL method and tactic for the Argo solver.
+*)
+
+signature ARGO_TACTIC =
+sig
+  val trace: string Config.T
+  val timeout: real Config.T
+
+  (* extending the tactic *)
+  type trans_context =
+    Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table
+  type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context
+  type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option
+  type extension = {
+    trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans',
+    trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans',
+    term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option,
+    replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv,
+    replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option}
+  val add_extension: extension -> theory -> theory
+
+  (* proof utilities *)
+  val discharges: thm -> thm list -> thm list
+  val flatten_conv: conv -> thm -> conv
+
+  (* interface to the tactic as well as the underlying checker and prover *)
+  datatype result = Satisfiable of term -> bool option | Unsatisfiable
+  val check: term list -> Proof.context -> result * Proof.context
+  val prove: thm list -> Proof.context -> thm option * Proof.context
+  val argo_tac: Proof.context -> thm list -> int -> tactic
+end
+
+structure Argo_Tactic: ARGO_TACTIC =
+struct
+
+(* readable fresh names for terms *)
+
+fun fresh_name n = Name.variant (case Long_Name.base_name n of "" => "x" | n' => n')
+
+fun fresh_type_name (Type (n, _)) = fresh_name n
+  | fresh_type_name (TFree (n, _)) = fresh_name n
+  | fresh_type_name (TVar ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i)
+
+fun fresh_term_name (Const (n, _)) = fresh_name n
+  | fresh_term_name (Free (n, _)) = fresh_name n
+  | fresh_term_name (Var ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i)
+  | fresh_term_name _ = fresh_name ""
+
+
+(* tracing *)
+
+datatype mode = None | Basic | Full
+
+fun string_of_mode None = "none"
+  | string_of_mode Basic = "basic"
+  | string_of_mode Full = "full"
+
+fun requires_mode None = []
+  | requires_mode Basic = [Basic, Full]
+  | requires_mode Full = [Full]
+
+val trace = Attrib.setup_config_string @{binding argo_trace} (K (string_of_mode None))
+
+fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode
+
+fun output mode ctxt msg = if allows_mode ctxt mode then Output.tracing ("Argo: " ^ msg) else ()
+val tracing = output Basic
+val full_tracing = output Full
+
+fun with_mode mode ctxt f = if allows_mode ctxt mode then f ctxt else ()
+val with_tracing = with_mode Basic
+val with_full_tracing = with_mode Full
+
+
+(* timeout *)
+
+val timeout = Attrib.setup_config_real @{binding argo_timeout} (K 10.0)
+
+fun time_of_timeout ctxt = Time.fromReal (Config.get ctxt timeout)
+
+fun with_timeout ctxt f x = Timeout.apply (time_of_timeout ctxt) f x
+
+
+(* extending the tactic *)
+
+type trans_context =
+  Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table
+
+type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context
+type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option
+
+type extension = {
+  trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans',
+  trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans',
+  term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option,
+  replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv,
+  replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option}
+
+fun eq_extension ((serial1, _), (serial2, _)) = (serial1 = serial2)
+
+structure Extensions = Theory_Data
+(
+  type T = (serial * extension) list
+  val empty = []
+  val extend = I
+  val merge = merge eq_extension 
+)
+
+fun add_extension ext = Extensions.map (insert eq_extension (serial (), ext))
+fun get_extensions ctxt = Extensions.get (Proof_Context.theory_of ctxt)
+fun apply_first ctxt f = get_first (fn (_, e) => f e) (get_extensions ctxt)
+
+fun ext_trans sel ctxt f x tcx = apply_first ctxt (fn ext => sel ext f x tcx)
+
+val ext_trans_type = ext_trans (fn {trans_type, ...}: extension => trans_type)
+val ext_trans_term = ext_trans (fn {trans_term, ...}: extension => trans_term)
+
+fun ext_term_of ctxt f e = apply_first ctxt (fn {term_of, ...}: extension => term_of f e)
+
+fun ext_replay_rewr ctxt r = 
+  get_extensions ctxt
+  |> map (fn (_, {replay_rewr, ...}: extension) => replay_rewr ctxt r)
+  |> Conv.first_conv
+
+fun ext_replay cprop_of ctxt rule prems =
+  (case apply_first ctxt (fn {replay, ...}: extension => replay cprop_of ctxt rule prems) of
+    SOME thm => thm
+  | NONE => raise THM ("failed to replay " ^ quote (Argo_Proof.string_of_rule rule), 0, []))
+
+
+(* translating input terms *)
+
+fun add_new_type T (names, types, terms) =
+  let
+    val (n, names') = fresh_type_name T names
+    val ty = Argo_Expr.Type n
+  in (ty, (names', Typtab.update (T, ty) types, terms)) end
+
+fun add_type T (tcx as (_, types, _)) =
+  (case Typtab.lookup types T of
+    SOME ty => (ty, tcx)
+  | NONE => add_new_type T tcx)
+
+fun trans_type _ @{typ HOL.bool} = pair Argo_Expr.Bool
+  | trans_type ctxt (Type (@{type_name "fun"}, [T1, T2])) =
+      trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func
+  | trans_type ctxt T = (fn tcx =>
+      (case ext_trans_type ctxt (trans_type ctxt) T tcx of
+        SOME result => result
+      | NONE => add_type T tcx))
+
+fun add_new_term ctxt t T tcx =
+  let
+    val (ty, (names, types, terms)) = trans_type ctxt T tcx
+    val (n, names') = fresh_term_name t names
+    val c = (n, ty)
+  in (Argo_Expr.mk_con c, (names', types, Termtab.update (t, c) terms)) end
+
+fun add_term ctxt t (tcx as (_, _, terms)) =
+  (case Termtab.lookup terms t of
+    SOME c => (Argo_Expr.mk_con c, tcx)
+  | NONE => add_new_term ctxt t (Term.fastype_of t) tcx)
+
+fun mk_eq @{typ HOL.bool} = Argo_Expr.mk_iff
+  | mk_eq _ = Argo_Expr.mk_eq
+
+fun trans_term _ @{const HOL.True} = pair Argo_Expr.true_expr
+  | trans_term _ @{const HOL.False} = pair Argo_Expr.false_expr
+  | trans_term ctxt (@{const HOL.Not} $ t) = trans_term ctxt t #>> Argo_Expr.mk_not
+  | trans_term ctxt (@{const HOL.conj} $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2
+  | trans_term ctxt (@{const HOL.disj} $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2
+  | trans_term ctxt (@{const HOL.implies} $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp
+  | trans_term ctxt (Const (@{const_name HOL.If}, _) $ t1 $ t2 $ t3) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>>
+      (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3)
+  | trans_term ctxt (Const (@{const_name HOL.eq}, T) $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq (Term.domain_type T))
+  | trans_term ctxt (t as (t1 $ t2)) = (fn tcx =>
+      (case ext_trans_term ctxt (trans_term ctxt) t tcx of
+        SOME result => result
+      | NONE => tcx |> trans_term ctxt t1 ||>> trans_term ctxt t2 |>> uncurry Argo_Expr.mk_app))
+  | trans_term ctxt t = (fn tcx =>
+      (case ext_trans_term ctxt (trans_term ctxt) t tcx of
+        SOME result => result
+      | NONE => add_term ctxt t tcx))
+
+fun translate ctxt prop = trans_term ctxt (HOLogic.dest_Trueprop prop)
+
+
+(* invoking the solver *)
+
+type data = {
+  names: Name.context,
+  types: Argo_Expr.typ Typtab.table,
+  terms: (string * Argo_Expr.typ) Termtab.table,
+  argo: Argo_Solver.context}
+
+fun mk_data names types terms argo: data = {names=names, types=types, terms=terms, argo=argo}
+val empty_data = mk_data Name.context Typtab.empty Termtab.empty Argo_Solver.context
+
+structure Solver_Data = Proof_Data
+(
+  type T = data option
+  fun init _ = SOME empty_data
+)
+
+datatype ('m, 'p) solver_result = Model of 'm | Proof of 'p
+
+fun raw_solve es argo = Model (Argo_Solver.assert es argo)
+  handle Argo_Proof.UNSAT proof => Proof proof
+
+fun value_of terms model t =
+  (case Termtab.lookup terms t of
+    SOME c => model c
+  | _ => NONE)
+
+fun trace_props props ctxt =
+  tracing ctxt (Pretty.string_of (Pretty.big_list "using these propositions:"
+    (map (Syntax.pretty_term ctxt) props)))
+
+fun trace_result ctxt ({elapsed, ...}: Timing.timing) msg =
+  tracing ctxt ("found a " ^ msg ^ " in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms")
+
+fun solve props ctxt =
+  (case Solver_Data.get ctxt of
+    NONE => error "bad Argo solver context"
+  | SOME {names, types, terms, argo} =>
+      let
+        val _ = with_tracing ctxt (trace_props props)
+        val (es, (names', types', terms')) = fold_map (translate ctxt) props (names, types, terms)
+        val _ = tracing ctxt "starting the prover"
+      in
+        (case Timing.timing (raw_solve es) argo of
+          (time, Proof proof) =>
+            let val _ = trace_result ctxt time "proof"
+            in (Proof (terms', proof), Solver_Data.put NONE ctxt) end
+        | (time, Model argo') =>
+            let
+              val _ = trace_result ctxt time "model"
+              val model = value_of terms' (Argo_Solver.model_of argo')
+            in (Model model, Solver_Data.put (SOME (mk_data names' types' terms' argo')) ctxt) end)
+      end)
+
+
+(* reverse translation *)
+
+structure Contab = Table(type key = string * Argo_Expr.typ val ord = Argo_Expr.con_ord)
+
+fun mk_nary f ts = uncurry (fold_rev (curry f)) (split_last ts)
+
+fun mk_nary' d _ [] = d
+  | mk_nary' _ f ts = mk_nary f ts
+
+fun mk_ite t1 t2 t3 =
+  let
+    val T = Term.fastype_of t2
+    val ite = Const (@{const_name HOL.If}, [@{typ HOL.bool}, T, T] ---> T)
+  in ite $ t1 $ t2 $ t3 end
+
+fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = @{const HOL.True}
+  | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = @{const HOL.False}
+  | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e)
+  | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) =
+      mk_nary' @{const HOL.True} HOLogic.mk_conj (map (term_of cx) es)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) =
+      mk_nary' @{const HOL.False} HOLogic.mk_disj (map (term_of cx) es)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) =
+      HOLogic.mk_imp (term_of cx e1, term_of cx e2)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) =
+      HOLogic.mk_eq (term_of cx e1, term_of cx e2)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) =
+      mk_ite (term_of cx e1) (term_of cx e2) (term_of cx e3)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Eq, [e1, e2])) =
+      HOLogic.mk_eq (term_of cx e1, term_of cx e2)
+  | term_of cx (Argo_Expr.E (Argo_Expr.App, [e1, e2])) =
+      term_of cx e1 $ term_of cx e2
+  | term_of (_, cons) (Argo_Expr.E (Argo_Expr.Con (c as (n, _)), _)) =
+      (case Contab.lookup cons c of
+        SOME t => t
+      | NONE => error ("Unexpected expression named " ^ quote n))
+  | term_of (cx as (ctxt, _)) e =
+      (case ext_term_of ctxt (term_of cx) e of
+        SOME t => t
+      | NONE => raise Fail "bad expression")
+
+fun as_prop ct = Thm.apply @{cterm HOL.Trueprop} ct
+
+fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e)
+fun cprop_of ctxt cons e = as_prop (cterm_of ctxt cons e)
+
+
+(* generic proof tools *)
+
+fun discharge thm rule = thm INCR_COMP rule
+fun discharge2 thm1 thm2 rule = discharge thm2 (discharge thm1 rule)
+fun discharges thm rules = [thm] RL rules
+
+fun under_assumption f ct =
+  let val cprop = as_prop ct
+  in Thm.implies_intr cprop (f (Thm.assume cprop)) end
+
+fun instantiate cv ct = Thm.instantiate ([], [(Term.dest_Var (Thm.term_of cv), ct)])
+
+
+(* proof replay for tautologies *)
+
+fun prove_taut ctxt ns t = Goal.prove ctxt ns [] (HOLogic.mk_Trueprop t)
+  (fn {context, ...} => HEADGOAL (Classical.fast_tac context))
+
+fun with_frees ctxt n mk =
+  let
+    val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1))
+    val ts = map (Free o rpair @{typ bool}) ns
+    val t = mk_nary HOLogic.mk_disj (mk ts)
+  in prove_taut ctxt ns t end
+
+fun taut_and1_term ts = mk_nary HOLogic.mk_conj ts :: map HOLogic.mk_not ts
+fun taut_and2_term i ts = [HOLogic.mk_not (mk_nary HOLogic.mk_conj ts), nth ts i]
+fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)]
+fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts
+
+val iff_1_taut = @{lemma "P = Q | P | Q" by fast}
+val iff_2_taut = @{lemma "P = Q | (~P) | (~Q)" by fast}
+val iff_3_taut = @{lemma "~(P = Q) | (~P) | Q" by fast}
+val iff_4_taut = @{lemma "~(P = Q) | P | (~Q)" by fast}
+val ite_then_taut = @{lemma "~P | (if P then t else u) = t" by auto}
+val ite_else_taut = @{lemma "P | (if P then t else u) = u" by auto}
+
+fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term
+  | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i)
+  | taut_rule_of ctxt (Argo_Proof.Taut_Or_1 (i, n)) = with_frees ctxt n (taut_or1_term i)
+  | taut_rule_of ctxt (Argo_Proof.Taut_Or_2 n) = with_frees ctxt n taut_or2_term
+  | taut_rule_of _ Argo_Proof.Taut_Iff_1 = iff_1_taut
+  | taut_rule_of _ Argo_Proof.Taut_Iff_2 = iff_2_taut
+  | taut_rule_of _ Argo_Proof.Taut_Iff_3 = iff_3_taut
+  | taut_rule_of _ Argo_Proof.Taut_Iff_4 = iff_4_taut
+  | taut_rule_of _ Argo_Proof.Taut_Ite_Then = ite_then_taut
+  | taut_rule_of _ Argo_Proof.Taut_Ite_Else = ite_else_taut
+
+fun replay_taut ctxt k ct =
+  let val rule = taut_rule_of ctxt k
+  in Thm.instantiate (Thm.match (Thm.cprop_of rule, ct)) rule end
+
+
+(* proof replay for conjunct extraction *)
+
+fun replay_conjunct 0 1 thm = thm
+  | replay_conjunct 0 _ thm = discharge thm @{thm conjunct1}
+  | replay_conjunct 1 2 thm = discharge thm @{thm conjunct2}
+  | replay_conjunct i n thm = replay_conjunct (i - 1) (n - 1) (discharge thm @{thm conjunct2})
+
+
+(* proof replay for rewrite steps *)
+
+fun mk_rewr thm = thm RS @{thm eq_reflection}
+
+fun not_nary_conv rule i ct =
+  if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct
+  else Conv.all_conv ct
+
+val flatten_and_thm = @{lemma "(P1 & P2) & P3 == P1 & P2 & P3" by simp}
+val flatten_or_thm = @{lemma "(P1 | P2) | P3 == P1 | P2 | P3" by simp}
+
+fun flatten_conv cv rule ct = (
+  Conv.try_conv (Conv.arg_conv cv) then_conv
+  Conv.try_conv (Conv.rewr_conv rule then_conv cv)) ct
+
+fun flat_conj_conv ct =
+  (case Thm.term_of ct of
+    @{const HOL.conj} $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct
+  | _ => Conv.all_conv ct)
+
+fun flat_disj_conv ct =
+  (case Thm.term_of ct of
+    @{const HOL.disj} $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct
+  | _ => Conv.all_conv ct)
+
+fun explode rule1 rule2 thm =
+  explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm]
+val explode_conj = explode @{thm conjunct1} @{thm conjunct2}
+val explode_ndis = explode @{lemma "~(P | Q) ==> ~P" by auto} @{lemma "~(P | Q) ==> ~Q" by auto}
+
+fun pick_false i thms = nth thms i
+
+fun pick_dual rule (i1, i2) thms =
+  rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1]
+val pick_dual_conj = pick_dual @{lemma "~P ==> P ==> False" by auto}
+val pick_dual_ndis = pick_dual @{lemma "~P ==> P ==> ~True" by auto}
+
+fun join thm0 rule is thms =
+  let
+    val l = length thms
+    val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is []
+  in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end
+
+val join_conj = join @{lemma "True" by auto} @{lemma "P ==> Q ==> P & Q" by auto}
+val join_ndis = join @{lemma "~False" by auto} @{lemma "~P ==> ~Q ==> ~(P | Q)" by auto}
+
+val false_thm = @{lemma "False ==> P" by auto}
+val ntrue_thm = @{lemma "~True ==> P" by auto}
+val iff_conj_thm = @{lemma "(P ==> Q) ==> (Q ==> P) ==> P = Q" by auto}
+val iff_ndis_thm = @{lemma "(~P ==> ~Q) ==> (~Q ==> ~P) ==> P = Q" by auto}
+
+fun iff_intro rule lf rf ct =
+  let
+    val lhs = under_assumption lf ct
+    val rhs = rf (Thm.dest_arg (snd (Thm.dest_implies (Thm.cprop_of lhs))))
+  in mk_rewr (discharge2 lhs rhs rule) end
+
+fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct
+fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply @{cterm HOL.Not} ct)
+
+fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1))
+fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g))
+val sort_conj = sort_nary with_conj join_conj explode_conj
+val sort_ndis = sort_nary with_ndis join_ndis explode_ndis 
+
+val not_true_thm = mk_rewr @{lemma "(~True) = False" by auto}
+val not_false_thm = mk_rewr @{lemma "(~False) = True" by auto}
+val not_not_thm = mk_rewr @{lemma "(~~P) = P" by auto}
+val not_and_thm = mk_rewr @{lemma "(~(P & Q)) = (~P | ~Q)" by auto}
+val not_or_thm = mk_rewr @{lemma "(~(P | Q)) = (~P & ~Q)" by auto}
+val not_iff_thms = map mk_rewr
+  @{lemma "(~((~P) = Q)) = (P = Q)" "(~(P = (~Q))) = (P = Q)" "(~(P = Q)) = ((~P) = Q)" by auto}
+val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto}
+val iff_false_thms = map mk_rewr @{lemma "(False = P) = (~P)" "(P = False) = (~P)" by auto}
+val iff_not_not_thm = mk_rewr @{lemma "((~P) = (~Q)) = (P = Q)" by auto}
+val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto}
+val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto}
+val iff_dual_thms = map mk_rewr @{lemma "(P = (~P)) = False" "((~P) = P) = False" by auto}
+val imp_thm = mk_rewr @{lemma "(P --> Q) = (~P | Q)" by auto}
+val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((~P | Q) & (P | R) & (Q | R))" by auto}
+val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto}
+val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto}
+val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto}
+val eq_refl_thm = mk_rewr @{lemma "(t = t) = True" by auto}
+val eq_symm_thm = mk_rewr @{lemma "(t1 = t2) = (t2 = t1)" by auto}
+
+fun replay_rewr _ Argo_Proof.Rewr_Not_True = Conv.rewr_conv not_true_thm
+  | replay_rewr _ Argo_Proof.Rewr_Not_False = Conv.rewr_conv not_false_thm
+  | replay_rewr _ Argo_Proof.Rewr_Not_Not = Conv.rewr_conv not_not_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Not_And i) = not_nary_conv not_and_thm i
+  | replay_rewr _ (Argo_Proof.Rewr_Not_Or i) = not_nary_conv not_or_thm i
+  | replay_rewr _ Argo_Proof.Rewr_Not_Iff = Conv.rewrs_conv not_iff_thms
+  | replay_rewr _ (Argo_Proof.Rewr_And_False i) = with_conj (pick_false i) (K false_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_And_Dual ip) = with_conj (pick_dual_conj ip) (K false_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_And_Sort is) = flat_conj_conv then_conv sort_conj is
+  | replay_rewr _ (Argo_Proof.Rewr_Or_True i) = with_ndis (pick_false i) (K ntrue_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_Or_Dual ip) = with_ndis (pick_dual_ndis ip) (K ntrue_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_Or_Sort is) = flat_disj_conv then_conv sort_ndis is
+  | replay_rewr _ Argo_Proof.Rewr_Iff_True = Conv.rewrs_conv iff_true_thms
+  | replay_rewr _ Argo_Proof.Rewr_Iff_False = Conv.rewrs_conv iff_false_thms
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Not_Not = Conv.rewr_conv iff_not_not_thm
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Refl = Conv.rewr_conv iff_refl_thm
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Symm = Conv.rewr_conv iff_symm_thm
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Dual  = Conv.rewrs_conv iff_dual_thms
+  | replay_rewr _ Argo_Proof.Rewr_Imp = Conv.rewr_conv imp_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_Prop = Conv.rewr_conv ite_prop_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_True = Conv.rewr_conv ite_true_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_False = Conv.rewr_conv ite_false_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_Eq = Conv.rewr_conv ite_eq_thm
+  | replay_rewr _ Argo_Proof.Rewr_Eq_Refl = Conv.rewr_conv eq_refl_thm
+  | replay_rewr _ Argo_Proof.Rewr_Eq_Symm = Conv.rewr_conv eq_symm_thm
+  | replay_rewr ctxt r = ext_replay_rewr ctxt r
+
+fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
+
+fun replay_conv _ Argo_Proof.Keep_Conv ct = Conv.all_conv ct
+  | replay_conv ctxt (Argo_Proof.Then_Conv (c1, c2)) ct = 
+      (replay_conv ctxt c1 then_conv replay_conv ctxt c2) ct
+  | replay_conv ctxt (Argo_Proof.Args_Conv cs) ct = replay_args_conv ctxt cs ct
+  | replay_conv ctxt (Argo_Proof.Rewr_Conv r) ct = replay_rewr ctxt r ct
+
+and replay_args_conv _ [] ct = Conv.all_conv ct
+  | replay_args_conv ctxt [c] ct = Conv.arg_conv (replay_conv ctxt c) ct
+  | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct
+  | replay_args_conv ctxt (c :: cs) ct =
+      (case Term.head_of (Thm.term_of ct) of
+        Const (@{const_name HOL.If}, _) =>
+          let val (cs', c') = split_last cs
+          in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end
+      | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct)
+
+fun replay_rewrite ctxt c thm = Conv.fconv_rule (HOLogic.Trueprop_conv (replay_conv ctxt c)) thm
+
+
+(* proof replay for clauses *)
+
+val prep_clause_rule = @{lemma "P ==> ~P ==> False" by fast}
+val extract_lit_rule = @{lemma "(~(P | Q) ==> False) ==> ~P ==> ~Q ==> False" by fast}
+
+fun add_lit i thm lits =
+  let val ct = Thm.cprem_of thm 1
+  in (Thm.implies_elim thm (Thm.assume ct), (i, ct) :: lits) end
+
+fun extract_lits [] _ = error "Bad clause"
+  | extract_lits [i] (thm, lits) = add_lit i thm lits
+  | extract_lits (i :: is) (thm, lits) =
+      extract_lits is (add_lit i (discharge thm extract_lit_rule) lits)
+
+fun lit_ord ((l1, _), (l2, _)) = int_ord (abs l1, abs l2)
+
+fun replay_with_lits [] thm lits = (thm, lits)
+  | replay_with_lits is thm lits =
+      extract_lits is (discharge thm prep_clause_rule, lits)
+      ||> Ord_List.make lit_ord
+
+fun replay_clause is thm = replay_with_lits is thm []
+
+
+(* proof replay for unit resolution *)
+
+val unit_rule = @{lemma "(P ==> False) ==> (~P ==> False) ==> False" by fast}
+val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1))
+val bogus_ct = @{cterm HOL.True}
+
+fun replay_unit_res lit (pthm, plits) (nthm, nlits) =
+  let
+    val plit = the (AList.lookup (op =) plits lit)
+    val nlit = the (AList.lookup (op =) nlits (~lit))
+    val prune = Ord_List.remove lit_ord (lit, bogus_ct)
+  in
+    unit_rule
+    |> instantiate unit_rule_var (Thm.dest_arg plit)
+    |> Thm.elim_implies (Thm.implies_intr plit pthm)
+    |> Thm.elim_implies (Thm.implies_intr nlit nthm)
+    |> rpair (Ord_List.union lit_ord (prune nlits) (prune plits))
+  end
+
+
+(* proof replay for hypothesis *)
+
+val dneg_rule = @{lemma "~~P ==> P" by auto}
+
+fun replay_hyp i ct =
+  if i < 0 then (Thm.assume ct, [(~i, ct)])
+  else
+    let val cu = as_prop (Thm.apply @{cterm HOL.Not} (Thm.apply @{cterm HOL.Not} (Thm.dest_arg ct)))
+    in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end
+
+
+(* proof replay for lemma *)
+
+fun replay_lemma is (thm, lits) = replay_with_lits is thm lits
+
+
+(* proof replay for reflexivity *)
+
+val refl_rule = @{thm refl}
+val refl_rule_var = Thm.dest_arg1 (Thm.dest_arg (Thm.cprop_of refl_rule))
+
+fun replay_refl ct = Thm.instantiate (Thm.match (refl_rule_var, ct)) refl_rule
+
+
+(* proof replay for symmetry *)
+
+val symm_rules = @{lemma "a = b ==> b = a" "~(a = b) ==> ~(b = a)" by simp_all}
+
+fun replay_symm thm = hd (discharges thm symm_rules)
+
+
+(* proof replay for transitivity *)
+
+val trans_rules = @{lemma
+  "~(a = b) ==> b = c ==> ~(a = c)"
+  "a = b ==> ~(b = c) ==> ~(a = c)"
+  "a = b ==> b = c ==> a = c"
+  by simp_all}
+
+fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules))
+
+
+(* proof replay for congruence *)
+
+fun replay_cong thm1 thm2 = discharge thm2 (discharge thm1 @{thm cong})
+
+
+(* proof replay for substitution *)
+
+val subst_rule1 = @{lemma "~(p a) ==> p = q ==> a = b ==> ~(q b)" by simp}
+val subst_rule2 = @{lemma "p a ==> p = q ==> a = b ==> q b" by simp}
+
+fun replay_subst thm1 thm2 thm3 =
+  subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3]
+
+
+(* proof replay *)
+
+structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord)
+
+val unclausify_rule1 = @{lemma "(~P ==> False) ==> P" by auto}
+val unclausify_rule2 = @{lemma "(P ==> False) ==> ~P" by auto}
+
+fun unclausify (thm, lits) ls =
+  (case (Thm.prop_of thm, lits) of
+    (@{const HOL.Trueprop} $ @{const HOL.False}, [(_, ct)]) =>
+      let val thm = Thm.implies_intr ct thm
+      in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end
+  | _ => (thm, Ord_List.union lit_ord lits ls))
+
+fun with_thms f tps = fold_map unclausify tps [] |>> f
+
+fun bad_premises () = raise Fail "bad number of premises"
+fun with_thms1 f = with_thms (fn [thm] => f thm | _ => bad_premises ())
+fun with_thms2 f = with_thms (fn [thm1, thm2] => f thm1 thm2 | _ => bad_premises ())
+fun with_thms3 f = with_thms (fn [thm1, thm2, thm3] => f thm1 thm2 thm3 | _ => bad_premises ())
+
+fun replay_rule (ctxt, cons, facts) prems rule =
+  (case rule of
+    Argo_Proof.Axiom i => (nth facts i, [])
+  | Argo_Proof.Taut (k, concl) => (replay_taut ctxt k (cprop_of ctxt cons concl), [])
+  | Argo_Proof.Conjunct (i, n) => with_thms1 (replay_conjunct i n) prems
+  | Argo_Proof.Rewrite c => with_thms1 (replay_rewrite ctxt c) prems
+  | Argo_Proof.Clause is => replay_clause is (fst (hd prems))
+  | Argo_Proof.Unit_Res i => replay_unit_res i (hd prems) (hd (tl prems))
+  | Argo_Proof.Hyp (i, concl) => replay_hyp i (cprop_of ctxt cons concl)
+  | Argo_Proof.Lemma is => replay_lemma is (hd prems)
+  | Argo_Proof.Refl concl => (replay_refl (cterm_of ctxt cons concl), [])
+  | Argo_Proof.Symm => with_thms1 replay_symm prems
+  | Argo_Proof.Trans => with_thms2 replay_trans prems
+  | Argo_Proof.Cong => with_thms2 replay_cong prems
+  | Argo_Proof.Subst => with_thms3 replay_subst prems
+  | _ => with_thms (ext_replay (cprop_of ctxt cons) ctxt rule) prems)
+
+fun with_cache f proof thm_cache =
+  (case Thm_Cache.lookup thm_cache (Argo_Proof.id_of proof) of
+    SOME thm => (thm, thm_cache)
+  | NONE =>
+      let val (thm, thm_cache') = f proof thm_cache
+      in (thm, Thm_Cache.update (Argo_Proof.id_of proof, thm) thm_cache') end)
+
+fun trace_step ctxt proof_id rule proofs = with_full_tracing ctxt (fn ctxt' =>
+  let
+    val id = Argo_Proof.string_of_proof_id proof_id
+    val ids = map (Argo_Proof.string_of_proof_id o Argo_Proof.id_of) proofs
+    val rule_string = Argo_Proof.string_of_rule rule
+  in full_tracing ctxt' ("  " ^ id ^ " <- " ^ space_implode " " ids ^ " . " ^ rule_string) end)
+
+fun replay_bottom_up (env as (ctxt, _, _)) proof thm_cache =
+  let
+    val (proof_id, rule, proofs) = Argo_Proof.dest proof
+    val (prems, thm_cache) = fold_map (with_cache (replay_bottom_up env)) proofs thm_cache
+    val _ = trace_step ctxt proof_id rule proofs
+  in (replay_rule env prems rule, thm_cache) end
+
+fun replay_proof env proof = with_cache (replay_bottom_up env) proof Thm_Cache.empty
+
+fun replay ctxt terms facts proof =
+  let
+    val env = (ctxt, Termtab.fold (Contab.update o swap) terms Contab.empty, facts)
+    val _ = tracing ctxt "replaying the proof"
+    val ({elapsed=t, ...}, ((thm, _), _)) = Timing.timing (replay_proof env) proof
+    val _ = tracing ctxt ("replayed the proof in " ^ string_of_int (Time.toMilliseconds t) ^ " ms")
+  in thm end
+
+
+(* normalizing goals *)
+
+fun instantiate_elim_rule thm =
+  let val ct = Drule.strip_imp_concl (Thm.cprop_of thm)
+  in
+    (case Thm.term_of ct of
+      @{const HOL.Trueprop} $ Var (_, @{typ HOL.bool}) =>
+        instantiate (Thm.dest_arg ct) @{cterm HOL.False} thm
+    | Var _ => instantiate ct @{cprop HOL.False} thm
+    | _ => thm)
+  end
+
+fun atomize_conv ctxt ct =
+  (case Thm.term_of ct of
+    @{const HOL.Trueprop} $ _ => Conv.all_conv
+  | @{const Pure.imp} $ _ $ _ =>
+      Conv.binop_conv (atomize_conv ctxt) then_conv
+      Conv.rewr_conv @{thm atomize_imp}
+  | Const (@{const_name Pure.eq}, _) $ _ $ _ =>
+      Conv.binop_conv (atomize_conv ctxt) then_conv
+      Conv.rewr_conv @{thm atomize_eq}
+  | Const (@{const_name Pure.all}, _) $ Abs _ =>
+      Conv.binder_conv (atomize_conv o snd) ctxt then_conv
+      Conv.rewr_conv @{thm atomize_all}
+  | _ => Conv.all_conv) ct
+
+fun normalize ctxt thm =
+  thm
+  |> instantiate_elim_rule
+  |> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion)
+  |> Drule.forall_intr_vars
+  |> Conv.fconv_rule (atomize_conv ctxt)
+
+
+(* prover, tactic and method *)
+
+datatype result = Satisfiable of term -> bool option | Unsatisfiable
+
+fun check props ctxt =
+  (case solve props ctxt of
+    (Proof _, ctxt') => (Unsatisfiable, ctxt')
+  | (Model model, ctxt') => (Satisfiable model, ctxt'))
+
+fun prove thms ctxt =
+  let val thms' = map (normalize ctxt) thms
+  in
+    (case solve (map Thm.prop_of thms') ctxt of
+      (Model _, ctxt') => (NONE, ctxt')
+    | (Proof (terms, proof), ctxt') => (SOME (replay ctxt' terms thms' proof), ctxt'))
+  end
+
+fun argo_tac ctxt thms =
+  CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1
+    (Conv.try_conv (Conv.rewr_conv @{thm atomize_eq})))) ctxt)
+  THEN' Tactic.resolve_tac ctxt [@{thm ccontr}]
+  THEN' Subgoal.FOCUS (fn {context, prems, ...} =>
+    (case with_timeout context (prove (thms @ prems)) context of
+      (SOME thm, _) => Tactic.resolve_tac context [thm] 1
+    | (NONE, _) => Tactical.no_tac)) ctxt
+
+val _ =
+  Theory.setup (Method.setup @{binding argo}
+    (Scan.optional Attrib.thms [] >>
+      (fn thms => fn ctxt => METHOD (fn facts =>
+        HEADGOAL (argo_tac ctxt (thms @ facts)))))
+    "Applies the Argo prover")
+
+end
--- a/src/HOL/Tools/try0.ML	Thu Sep 29 16:49:42 2016 +0200
+++ b/src/HOL/Tools/try0.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -88,6 +88,7 @@
    ("auto", ((true, true), full_attrs)),
    ("blast", ((false, true), clas_attrs)),
    ("metis", ((false, true), metis_attrs)),
+   ("argo", ((false, true), no_attrs)),
    ("linarith", ((false, true), no_attrs)),
    ("presburger", ((false, true), no_attrs)),
    ("algebra", ((false, true), no_attrs)),
@@ -111,6 +112,7 @@
       |> Simplifier_Trace.disable
       |> Context_Position.set_visible false
       |> Config.put Unify.trace_bound (Config.get ctxt Unify.search_bound)
+      |> Config.put Argo_Tactic.trace "none"
       |> Proof_Context.background_theory (fn thy =>
           thy
           |> Context_Position.set_visible_global false
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Argo_Examples.thy	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,1328 @@
+(*  Title:      HOL/ex/Argo_Examples.thy
+    Author:     Sascha Boehme
+*)
+
+section \<open>Argo\<close>
+
+theory Argo_Examples
+imports Complex_Main
+begin
+
+text \<open>
+  This theory is intended to showcase and test different features of the \<open>argo\<close> proof method.
+
+  The \<open>argo\<close> proof method can be applied to propositional problems, problems involving equality
+  reasoning and problems of linear real arithmetic.
+
+  The \<open>argo\<close> proof method provides two options. To specify an upper limit of the proof methods
+  run time in seconds, use the option \<open>argo_timeout\<close>. To specify the amount of output, use the
+  option \<open>argo_trace\<close> with value \<open>none\<close> for no tracing output, value \<open>basic\<close> for viewing the
+  underlying propositions and some timings, and value \<open>full\<close> for additionally inspecting the
+  proof replay steps.
+\<close>
+
+declare[[argo_trace = full]]
+
+subsection \<open>Propositional logic\<close>
+
+notepad
+begin
+  have "True" by argo
+next
+  have "~False" by argo
+next
+  fix P :: bool
+  assume "False"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "~True"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "P"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "~~P"
+  then have "P" by argo
+next
+  fix P Q R :: bool
+  assume "P & Q & R"
+  then have "R & P & Q" by argo
+next
+  fix P Q R :: bool
+  assume "P & (Q & True & R) & (Q & P) & True"
+  then have "R & P & Q" by argo
+next
+  fix P Q1 Q2 Q3 Q4 Q5 :: bool
+  assume "Q1 & (Q2 & P & Q3) & (Q4 & ~P & Q5)"
+  then have "~True" by argo
+next
+  fix P Q1 Q2 Q3  :: bool
+  assume "(Q1 & False) & Q2 & Q3"
+  then have "P::bool" by argo
+next
+  fix P Q R :: bool
+  assume "P | Q | R"
+  then have "R | P | Q" by argo
+next
+  fix P Q R :: bool
+  assume "P | (Q | False | R) | (Q | P) | False"
+  then have "R | P | Q" by argo
+next
+  fix P Q1 Q2 Q3 Q4 :: bool
+  have "(Q1 & P & Q2) --> False | (Q3 | (Q4 | P) | False)" by argo
+next
+  fix Q1 Q2 Q3 Q4 :: bool
+  have "Q1 | (Q2 | True | Q3) | Q4" by argo
+next
+  fix P Q R :: bool
+  assume "(P & Q) | (P & ~Q) | (P & R) | (P & ~R)"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "P = True"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "False = P"
+  then have "~P" by argo
+next
+  fix P Q :: bool
+  assume "P = (~P)"
+  then have "Q" by argo
+next
+  fix P :: bool
+  have "(~P) = (~P)" by argo
+next
+  fix P Q :: bool
+  assume "P" and "~Q"
+  then have "P = (~Q)" by argo
+next
+  fix P Q :: bool
+  assume "((P::bool) = Q) | (Q = P)"
+  then have "(P --> Q) & (Q --> P)" by argo
+next
+  fix P Q :: bool
+  assume "(P::bool) = Q"
+  then have "Q = P" by argo
+next
+  fix P Q R :: bool
+  assume "if P then Q else R"
+  then have "Q | R" by argo
+next
+  fix P Q :: bool
+  assume "P | Q"
+     and "P | ~Q"
+     and "~P | Q"
+     and "~P | ~Q"
+  then have "False" by argo
+next
+  fix P Q R :: bool
+  assume "P | Q | R"
+     and "P | Q | ~R"
+     and "P | ~Q | R"
+     and "P | ~Q | ~R"
+     and "~P | Q | R"
+     and "~P | Q | ~R"
+     and "~P | ~Q | R"
+     and "~P | ~Q | ~R"
+  then have "False" by argo
+next
+  fix a b c d e f g h i j k l m n p q :: bool
+  assume "(a & b | c & d) & (e & f | g & h) | (i & j | k & l) & (m & n | p & q)"
+  then have "(a & b | c & d) & (e & f | g & h) | (i & j | k & l) & (m & n | p & q)" by argo
+next
+  fix P :: bool
+  have "P=P=P=P=P=P=P=P=P=P" by argo
+next
+  fix a b c d e f p q x :: bool
+  assume "a | b | c | d"
+     and "e | f | (a & d)"
+     and "~(a | (c & ~c)) | b"
+     and "~(b & (x | ~x)) | c"
+     and "~(d | False) | c"
+     and "~(c | (~p & (p | (q & ~q))))"
+  then have "False" by argo
+end
+
+
+subsection \<open>Equality, congruence and predicates\<close>
+
+notepad
+begin
+  fix t :: "'a"
+  have "t = t" by argo
+next
+  fix t u :: "'a"
+  assume "t = u"
+  then have "u = t" by argo
+next
+  fix s t u :: "'a"
+  assume "s = t" and "t = u"
+  then have "s = u" by argo
+next
+  fix s t u v :: "'a"
+  assume "s = t" and "t = u" and "u = v" and "u = s"
+  then have "s = v" by argo
+next
+  fix s t u v w :: "'a"
+  assume "s = t" and "t = u" and "s = v" and "v = w"
+  then have "w = u" by argo
+next
+  fix s t u a b c :: "'a"
+  assume "s = t" and "t = u" and "a = b" and "b = c"
+  then have "s = a --> c = u" by argo
+next
+  fix a b c d :: "'a"
+  assume "(a = b & b = c) | (a = d & d = c)"
+  then have "a = c" by argo
+next
+  fix a b1 b2 b3 b4 c d :: "'a"
+  assume "(a = b1 & ((b1 = b2 & b2 = b3) | (b1 = b4 & b4 = b3)) & b3 = c) | (a = d & d = c)"
+  then have "a = c" by argo
+next
+  fix a b :: "'a"
+  have "(if True then a else b) = a" by argo
+next
+  fix a b :: "'a"
+  have "(if False then a else b) = b" by argo
+next
+  fix a b :: "'a"
+  have "(if \<not>True then a else b) = b" by argo
+next
+  fix a b :: "'a"
+  have "(if \<not>False then a else b) = a" by argo
+next
+  fix P :: "bool"
+  fix a :: "'a"
+  have "(if P then a else a) = a" by argo
+next
+  fix P :: "bool"
+  fix a b c :: "'a"
+  assume "P" and "a = c"
+  then have "(if P then a else b) = c" by argo
+next
+  fix P :: "bool"
+  fix a b c :: "'a"
+  assume "~P" and "b = c"
+  then have "(if P then a else b) = c" by argo
+next
+  fix P Q :: "bool"
+  fix a b c d :: "'a"
+  assume "P" and "Q" and "a = d"
+  then have "(if P then (if Q then a else b) else c) = d" by argo
+next
+  fix a b c :: "'a"
+  assume "a \<noteq> b" and "b = c"
+  then have "a \<noteq> c" by argo
+next
+  fix a b c :: "'a"
+  assume "a \<noteq> b" and "a = c"
+  then have "c \<noteq> b" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "b \<noteq> d"
+  then have "a \<noteq> c" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "d \<noteq> b"
+  then have "a \<noteq> c" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "b \<noteq> d"
+  then have "c \<noteq> a" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "d \<noteq> b"
+  then have "c \<noteq> a" by argo
+next
+  fix a b c d e f :: "'a"
+  assume "a \<noteq> b" and "b = c" and "b = d" and "d = e" and "a = f"
+  then have "f \<noteq> e" by argo
+next
+  fix a b :: "'a" and f :: "'a \<Rightarrow> 'a"
+  assume "a = b"
+  then have "f a = f b" by argo
+next
+  fix a b c :: "'a" and f :: "'a \<Rightarrow> 'a"
+  assume "f a = f b" and "b = c"
+  then have "f a = f c" by argo
+next
+  fix a :: "'a" and f :: "'a \<Rightarrow> 'a"
+  assume "f a = a"
+  then have "f (f a) = a" by argo
+next
+  fix a b :: "'a" and f g :: "'a \<Rightarrow> 'a"
+  assume "a = b"
+  then have "g (f a) = g (f b)" by argo
+next
+  fix a b :: "'a" and f g :: "'a \<Rightarrow> 'a"
+  assume "f a = b" and "g b = a"
+  then have "f (g (f a)) = b" by argo
+next
+  fix a b :: "'a" and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  assume "a = b"
+  then have "g a b = g b a" by argo
+next
+  fix a b :: "'a" and f :: "'a \<Rightarrow> 'a" and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  assume "f a = b"
+  then have "g (f a) b = g b (f a)" by argo
+next
+  fix a b c d e g h :: "'a" and f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  assume "c = d" and "e = c" and "e = b" and "b = h" and "f g h = d" and "f g d = a"
+  then have "a = b" by argo
+next
+  fix a b :: "'a" and P :: "'a \<Rightarrow> bool"
+  assume "P a" and "a = b"
+  then have "P b" by argo
+next
+  fix a b :: "'a" and P :: "'a \<Rightarrow> bool"
+  assume "~ P a" and "a = b"
+  then have "~ P b" by argo
+next
+  fix a b c d :: "'a" and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assume "P a b" and "a = c" and "b = d"
+  then have "P c d" by argo
+next
+  fix a b c d :: "'a" and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assume "~ P a b" and "a = c" and "b = d"
+  then have "~ P c d" by argo
+end
+
+
+subsection \<open>Linear real arithmetic\<close>
+
+subsubsection \<open>Negation and subtraction\<close>
+
+notepad
+begin
+  fix a b :: real
+  have
+    "-a = -1 * a"
+    "-(-a) = a"
+    "a - b = a + -1 * b"
+    "a - (-b) = a + b"
+    by argo+
+end
+
+
+subsubsection \<open>Multiplication\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(2::real) * 3 = 6"
+    "0 * a = 0"
+    "a * 0 = 0"
+    "1 * a = a"
+    "a * 1 = a"
+    "2 * a = a * 2"
+    "2 * a * 3 = 6 * a"
+    "2 * a * 3 * 5 = 30 * a"
+    "a * 0 * b = 0"
+    "a * (0 * b) = 0"
+    "a * (b * c) = (a * b) * c"
+    "a * (b * (c * d)) = ((a * b) * c) * d"
+    "a * (b + c + d) = a * b + a * c + a * d"
+    by argo+
+end
+
+
+subsubsection \<open>Division\<close>
+
+notepad
+begin
+  fix a b c :: real
+  have
+    "(6::real) / 2 = 3"
+    "a / 0 = a / 0"
+    "a / 0 <= a / 0"
+    "~(a / 0 < a / 0)"
+    "0 / a = 0"
+    "a / 1 = a"
+    "a / 3 = 1/3 * a"
+    "6 * a / 2 = 3 * a"
+    "a / ((5 * b) / 2) = 2/5 * a / b"
+    "a / (5 * (b / 2)) = 2/5 * a / b"
+    "(a / 5) * (b / 2) = 1/10 * a * b"
+    "a / (3 * b) = 1/3 * a / b"
+    "(a + b) / 5 = 1/5 * a + 1/5 * b"
+    "a / (5 * 1/5) = a"
+    by argo+
+end
+
+
+subsubsection \<open>Addition\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "a + b = b + a"
+    "a + b + c = c + b + a"
+    "a + b + c + d = d + c + b + a"
+    "a + (b + (c + d)) = ((a + b) + c) + d"
+    "(5::real) + -3 = 2"
+    "(3::real) + 5 + -1 = 7"
+    "2 + a = a + 2"
+    "a + b + a = b + 2 * a"
+    "-1 + a + -1 + 2 + b + 5/3 + b + 1/3 + 5 * b + 2/3 = 8/3 + a + 7 * b"
+    "1 + b + b + 5 * b + 3 * a + 7 + a + 2 = 10 + 4 * a + 7 * b"
+    by argo+
+end
+
+
+subsubsection \<open>Minimum and maximum\<close>
+
+notepad
+begin
+  fix a b :: real
+  have
+    "min (3::real) 5 = 3"
+    "min (5::real) 3 = 3"
+    "min (3::real) (-5) = -5"
+    "min (-5::real) 3 = -5"
+    "min a a = a"
+    "a \<le> b \<longrightarrow> min a b = a"
+    "a > b \<longrightarrow> min a b = b"
+    "min a b \<le> a"
+    "min a b \<le> b"
+    "min a b = min b a"
+    by argo+
+next
+  fix a b :: real
+  have
+    "max (3::real) 5 = 5"
+    "max (5::real) 3 = 5"
+    "max (3::real) (-5) = 3"
+    "max (-5::real) 3 = 3"
+    "max a a = a"
+    "a \<le> b \<longrightarrow> max a b = b"
+    "a > b \<longrightarrow> max a b = a"
+    "a \<le> max a b"
+    "b \<le> max a b"
+    "max a b = max b a"
+    by argo+
+next
+  fix a b :: real
+  have
+    "min a b \<le> max a b"
+    "min a b + max a b = a + b"
+    "a < b \<longrightarrow> min a b < max a b"
+    by argo+
+end
+
+
+subsubsection \<open>Absolute value\<close>
+
+notepad
+begin
+  fix a :: real
+  have
+    "abs (3::real) = 3"
+    "abs (-3::real) = 3"
+    "0 \<le> abs a"
+    "a \<le> abs a"
+    "a \<ge> 0 \<longrightarrow> abs a = a"
+    "a < 0 \<longrightarrow> abs a = -a"
+    "abs (abs a) = abs a"
+    by argo+
+end
+
+
+subsubsection \<open>Equality\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(3::real) = 3"
+    "~((3::real) = 4)"
+    "~((4::real) = 3)"
+    "3 * a = 5 --> a = 5/3"
+    "-3 * a = 5 --> -5/3 = a"
+    "5 = 3 * a --> 5/3  = a "
+    "5 = -3 * a --> a = -5/3"
+    "2 + 3 * a = 4 --> a = 2/3"
+    "4 = 2 + 3 * a --> 2/3 = a"
+    "2 + 3 * a + 5 * b + c = 4 --> 3 * a + 5 * b + c = 2"
+    "4 = 2 + 3 * a + 5 * b + c --> 2 = 3 * a + 5 * b + c"
+    "-2 * a + b + -3 * c = 7 --> -7 = 2 * a + -1 * b + 3 * c"
+    "7 = -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c = -7"
+    "-2 * a + b + -3 * c + 4 * d = 7 --> -7 = 2 * a + -1 * b + 3 * c + -4 * d"
+    "7 = -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d = -7"
+    "a + 3 * b = 5 * c + b --> a + 2 * b + -5 * c = 0"
+    by argo+
+end
+
+
+subsubsection \<open>Less-equal\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(3::real) <= 3"
+    "(3::real) <= 4"
+    "~((4::real) <= 3)"
+    "3 * a <= 5 --> a <= 5/3"
+    "-3 * a <= 5 --> -5/3 <= a"
+    "5 <= 3 * a --> 5/3  <= a "
+    "5 <= -3 * a --> a <= -5/3"
+    "2 + 3 * a <= 4 --> a <= 2/3"
+    "4 <= 2 + 3 * a --> 2/3 <= a"
+    "2 + 3 * a + 5 * b + c <= 4 --> 3 * a + 5 * b + c <= 2"
+    "4 <= 2 + 3 * a + 5 * b + c --> 2 <= 3 * a + 5 * b + c"
+    "-2 * a + b + -3 * c <= 7 --> -7 <= 2 * a + -1 * b + 3 * c"
+    "7 <= -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c <= -7"
+    "-2 * a + b + -3 * c + 4 * d <= 7 --> -7 <= 2 * a + -1 * b + 3 * c + -4 * d"
+    "7 <= -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d <= -7"
+    "a + 3 * b <= 5 * c + b --> a + 2 * b + -5 * c <= 0"
+    by argo+
+end
+
+subsubsection \<open>Less\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(3::real) < 4"
+    "~((3::real) < 3)"
+    "~((4::real) < 3)"
+    "3 * a < 5 --> a < 5/3"
+    "-3 * a < 5 --> -5/3 < a"
+    "5 < 3 * a --> 5/3  < a "
+    "5 < -3 * a --> a < -5/3"
+    "2 + 3 * a < 4 --> a < 2/3"
+    "4 < 2 + 3 * a --> 2/3 < a"
+    "2 + 3 * a + 5 * b + c < 4 --> 3 * a + 5 * b + c < 2"
+    "4 < 2 + 3 * a + 5 * b + c --> 2 < 3 * a + 5 * b + c"
+    "-2 * a + b + -3 * c < 7 --> -7 < 2 * a + -1 * b + 3 * c"
+    "7 < -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c < -7"
+    "-2 * a + b + -3 * c + 4 * d < 7 --> -7 < 2 * a + -1 * b + 3 * c + -4 * d"
+    "7 < -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d < -7"
+    "a + 3 * b < 5 * c + b --> a + 2 * b + -5 * c < 0"
+    by argo+
+end
+
+
+subsubsection \<open>Other examples\<close>
+
+notepad
+begin
+  have
+    "(0::real) < 1"
+    "(47::real) + 11 < 8 * 15"
+    by argo+
+next
+  fix a :: real
+  assume "a < 3"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  assume "a <= 3"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  assume "~(3 < a)"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  assume "~(3 <= a)"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  have "a < 3 | a = 3 | a > 3" by argo
+next
+  fix a b :: real
+  assume "0 < a" and "a < b"
+  then have "0 < b" by argo
+next
+  fix a b :: real
+  assume "0 < a" and "a \<le> b"
+  then have "0 \<le> b" by argo
+next
+  fix a b :: real
+  assume "0 \<le> a" and "a < b"
+  then have "0 \<le> b" by argo
+next
+  fix a b :: real
+  assume "0 \<le> a" and "a \<le> b"
+  then have "0 \<le> b" by argo
+next
+  fix a b c :: real
+  assume "2 \<le> a" and "3 \<le> b" and "c \<le> 5"
+  then have "-2 * a + -3 * b + 5 * c < 13" by argo
+next
+  fix a b c :: real
+  assume "2 \<le> a" and "3 \<le> b" and "c \<le> 5"
+  then have "-2 * a + -3 * b + 5 * c \<le> 12" by argo
+next
+  fix a b :: real
+  assume "a = 2" and "b = 3"
+  then have "a + b > 5 \<or> a < b" by argo
+next
+  fix a b c :: real
+  assume "5 < b + c" and "a + c < 0" and "a > 0"
+  then have "b > 0" by argo
+next
+  fix a b c :: real
+  assume "a + b < 7" and "5 < b + c" and "a + c < 0" and "a > 0"
+  then have "0 < b \<and> b < 7" by argo
+next
+  fix a b c :: real
+  assume "a < b" and "b < c" and "c < a"
+  then have "False" by argo
+next
+  fix a b :: real
+  assume "a - 5 > b"
+  then have "b < a" by argo
+next
+  fix a b :: real
+  have "(a - b) - a = (a - a) - b" by argo
+next
+  fix n m n' m' :: real
+  have "
+    (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
+    (n = n' & n' < m) | (n = m & m < n') |
+    (n' < m & m < n) | (n' < m & m = n) |
+    (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) |
+    (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) |
+    (m = n & n < n') | (m = n' & n' < n) |
+    (n' = m & m = n)"
+    by argo
+end
+
+
+subsection \<open>Larger examples\<close>
+
+declare[[argo_trace = basic, argo_timeout = 60]]
+
+
+text \<open>Translated from TPTP problem library: PUZ015-2.006.dimacs\<close>
+
+lemma assumes 1: "~x0"
+  and 2: "~x30"
+  and 3: "~x29"
+  and 4: "~x59"
+  and 5: "x1 | x31 | x0"
+  and 6: "x2 | x32 | x1"
+  and 7: "x3 | x33 | x2"
+  and 8: "x4 | x34 | x3"
+  and 9: "x35 | x4"
+  and 10: "x5 | x36 | x30"
+  and 11: "x6 | x37 | x5 | x31"
+  and 12: "x7 | x38 | x6 | x32"
+  and 13: "x8 | x39 | x7 | x33"
+  and 14: "x9 | x40 | x8 | x34"
+  and 15: "x41 | x9 | x35"
+  and 16: "x10 | x42 | x36"
+  and 17: "x11 | x43 | x10 | x37"
+  and 18: "x12 | x44 | x11 | x38"
+  and 19: "x13 | x45 | x12 | x39"
+  and 20: "x14 | x46 | x13 | x40"
+  and 21: "x47 | x14 | x41"
+  and 22: "x15 | x48 | x42"
+  and 23: "x16 | x49 | x15 | x43"
+  and 24: "x17 | x50 | x16 | x44"
+  and 25: "x18 | x51 | x17 | x45"
+  and 26: "x19 | x52 | x18 | x46"
+  and 27: "x53 | x19 | x47"
+  and 28: "x20 | x54 | x48"
+  and 29: "x21 | x55 | x20 | x49"
+  and 30: "x22 | x56 | x21 | x50"
+  and 31: "x23 | x57 | x22 | x51"
+  and 32: "x24 | x58 | x23 | x52"
+  and 33: "x59 | x24 | x53"
+  and 34: "x25 | x54"
+  and 35: "x26 | x25 | x55"
+  and 36: "x27 | x26 | x56"
+  and 37: "x28 | x27 | x57"
+  and 38: "x29 | x28 | x58"
+  and 39: "~x1 | ~x31"
+  and 40: "~x1 | ~x0"
+  and 41: "~x31 | ~x0"
+  and 42: "~x2 | ~x32"
+  and 43: "~x2 | ~x1"
+  and 44: "~x32 | ~x1"
+  and 45: "~x3 | ~x33"
+  and 46: "~x3 | ~x2"
+  and 47: "~x33 | ~x2"
+  and 48: "~x4 | ~x34"
+  and 49: "~x4 | ~x3"
+  and 50: "~x34 | ~x3"
+  and 51: "~x35 | ~x4"
+  and 52: "~x5 | ~x36"
+  and 53: "~x5 | ~x30"
+  and 54: "~x36 | ~x30"
+  and 55: "~x6 | ~x37"
+  and 56: "~x6 | ~x5"
+  and 57: "~x6 | ~x31"
+  and 58: "~x37 | ~x5"
+  and 59: "~x37 | ~x31"
+  and 60: "~x5 | ~x31"
+  and 61: "~x7 | ~x38"
+  and 62: "~x7 | ~x6"
+  and 63: "~x7 | ~x32"
+  and 64: "~x38 | ~x6"
+  and 65: "~x38 | ~x32"
+  and 66: "~x6 | ~x32"
+  and 67: "~x8 | ~x39"
+  and 68: "~x8 | ~x7"
+  and 69: "~x8 | ~x33"
+  and 70: "~x39 | ~x7"
+  and 71: "~x39 | ~x33"
+  and 72: "~x7 | ~x33"
+  and 73: "~x9 | ~x40"
+  and 74: "~x9 | ~x8"
+  and 75: "~x9 | ~x34"
+  and 76: "~x40 | ~x8"
+  and 77: "~x40 | ~x34"
+  and 78: "~x8 | ~x34"
+  and 79: "~x41 | ~x9"
+  and 80: "~x41 | ~x35"
+  and 81: "~x9 | ~x35"
+  and 82: "~x10 | ~x42"
+  and 83: "~x10 | ~x36"
+  and 84: "~x42 | ~x36"
+  and 85: "~x11 | ~x43"
+  and 86: "~x11 | ~x10"
+  and 87: "~x11 | ~x37"
+  and 88: "~x43 | ~x10"
+  and 89: "~x43 | ~x37"
+  and 90: "~x10 | ~x37"
+  and 91: "~x12 | ~x44"
+  and 92: "~x12 | ~x11"
+  and 93: "~x12 | ~x38"
+  and 94: "~x44 | ~x11"
+  and 95: "~x44 | ~x38"
+  and 96: "~x11 | ~x38"
+  and 97: "~x13 | ~x45"
+  and 98: "~x13 | ~x12"
+  and 99: "~x13 | ~x39"
+  and 100: "~x45 | ~x12"
+  and 101: "~x45 | ~x39"
+  and 102: "~x12 | ~x39"
+  and 103: "~x14 | ~x46"
+  and 104: "~x14 | ~x13"
+  and 105: "~x14 | ~x40"
+  and 106: "~x46 | ~x13"
+  and 107: "~x46 | ~x40"
+  and 108: "~x13 | ~x40"
+  and 109: "~x47 | ~x14"
+  and 110: "~x47 | ~x41"
+  and 111: "~x14 | ~x41"
+  and 112: "~x15 | ~x48"
+  and 113: "~x15 | ~x42"
+  and 114: "~x48 | ~x42"
+  and 115: "~x16 | ~x49"
+  and 116: "~x16 | ~x15"
+  and 117: "~x16 | ~x43"
+  and 118: "~x49 | ~x15"
+  and 119: "~x49 | ~x43"
+  and 120: "~x15 | ~x43"
+  and 121: "~x17 | ~x50"
+  and 122: "~x17 | ~x16"
+  and 123: "~x17 | ~x44"
+  and 124: "~x50 | ~x16"
+  and 125: "~x50 | ~x44"
+  and 126: "~x16 | ~x44"
+  and 127: "~x18 | ~x51"
+  and 128: "~x18 | ~x17"
+  and 129: "~x18 | ~x45"
+  and 130: "~x51 | ~x17"
+  and 131: "~x51 | ~x45"
+  and 132: "~x17 | ~x45"
+  and 133: "~x19 | ~x52"
+  and 134: "~x19 | ~x18"
+  and 135: "~x19 | ~x46"
+  and 136: "~x52 | ~x18"
+  and 137: "~x52 | ~x46"
+  and 138: "~x18 | ~x46"
+  and 139: "~x53 | ~x19"
+  and 140: "~x53 | ~x47"
+  and 141: "~x19 | ~x47"
+  and 142: "~x20 | ~x54"
+  and 143: "~x20 | ~x48"
+  and 144: "~x54 | ~x48"
+  and 145: "~x21 | ~x55"
+  and 146: "~x21 | ~x20"
+  and 147: "~x21 | ~x49"
+  and 148: "~x55 | ~x20"
+  and 149: "~x55 | ~x49"
+  and 150: "~x20 | ~x49"
+  and 151: "~x22 | ~x56"
+  and 152: "~x22 | ~x21"
+  and 153: "~x22 | ~x50"
+  and 154: "~x56 | ~x21"
+  and 155: "~x56 | ~x50"
+  and 156: "~x21 | ~x50"
+  and 157: "~x23 | ~x57"
+  and 158: "~x23 | ~x22"
+  and 159: "~x23 | ~x51"
+  and 160: "~x57 | ~x22"
+  and 161: "~x57 | ~x51"
+  and 162: "~x22 | ~x51"
+  and 163: "~x24 | ~x58"
+  and 164: "~x24 | ~x23"
+  and 165: "~x24 | ~x52"
+  and 166: "~x58 | ~x23"
+  and 167: "~x58 | ~x52"
+  and 168: "~x23 | ~x52"
+  and 169: "~x59 | ~x24"
+  and 170: "~x59 | ~x53"
+  and 171: "~x24 | ~x53"
+  and 172: "~x25 | ~x54"
+  and 173: "~x26 | ~x25"
+  and 174: "~x26 | ~x55"
+  and 175: "~x25 | ~x55"
+  and 176: "~x27 | ~x26"
+  and 177: "~x27 | ~x56"
+  and 178: "~x26 | ~x56"
+  and 179: "~x28 | ~x27"
+  and 180: "~x28 | ~x57"
+  and 181: "~x27 | ~x57"
+  and 182: "~x29 | ~x28"
+  and 183: "~x29 | ~x58"
+  and 184: "~x28 | ~x58"
+  shows "False"
+    using assms
+    by argo
+
+
+text \<open>Translated from TPTP problem library: MSC007-1.008.dimacs\<close>
+
+lemma assumes 1: "x0 | x1 | x2 | x3 | x4 | x5 | x6"
+  and 2: "x7 | x8 | x9 | x10 | x11 | x12 | x13"
+  and 3: "x14 | x15 | x16 | x17 | x18 | x19 | x20"
+  and 4: "x21 | x22 | x23 | x24 | x25 | x26 | x27"
+  and 5: "x28 | x29 | x30 | x31 | x32 | x33 | x34"
+  and 6: "x35 | x36 | x37 | x38 | x39 | x40 | x41"
+  and 7: "x42 | x43 | x44 | x45 | x46 | x47 | x48"
+  and 8: "x49 | x50 | x51 | x52 | x53 | x54 | x55"
+  and 9: "~x0 | ~x7"
+  and 10: "~x0 | ~x14"
+  and 11: "~x0 | ~x21"
+  and 12: "~x0 | ~x28"
+  and 13: "~x0 | ~x35"
+  and 14: "~x0 | ~x42"
+  and 15: "~x0 | ~x49"
+  and 16: "~x7 | ~x14"
+  and 17: "~x7 | ~x21"
+  and 18: "~x7 | ~x28"
+  and 19: "~x7 | ~x35"
+  and 20: "~x7 | ~x42"
+  and 21: "~x7 | ~x49"
+  and 22: "~x14 | ~x21"
+  and 23: "~x14 | ~x28"
+  and 24: "~x14 | ~x35"
+  and 25: "~x14 | ~x42"
+  and 26: "~x14 | ~x49"
+  and 27: "~x21 | ~x28"
+  and 28: "~x21 | ~x35"
+  and 29: "~x21 | ~x42"
+  and 30: "~x21 | ~x49"
+  and 31: "~x28 | ~x35"
+  and 32: "~x28 | ~x42"
+  and 33: "~x28 | ~x49"
+  and 34: "~x35 | ~x42"
+  and 35: "~x35 | ~x49"
+  and 36: "~x42 | ~x49"
+  and 37: "~x1 | ~x8"
+  and 38: "~x1 | ~x15"
+  and 39: "~x1 | ~x22"
+  and 40: "~x1 | ~x29"
+  and 41: "~x1 | ~x36"
+  and 42: "~x1 | ~x43"
+  and 43: "~x1 | ~x50"
+  and 44: "~x8 | ~x15"
+  and 45: "~x8 | ~x22"
+  and 46: "~x8 | ~x29"
+  and 47: "~x8 | ~x36"
+  and 48: "~x8 | ~x43"
+  and 49: "~x8 | ~x50"
+  and 50: "~x15 | ~x22"
+  and 51: "~x15 | ~x29"
+  and 52: "~x15 | ~x36"
+  and 53: "~x15 | ~x43"
+  and 54: "~x15 | ~x50"
+  and 55: "~x22 | ~x29"
+  and 56: "~x22 | ~x36"
+  and 57: "~x22 | ~x43"
+  and 58: "~x22 | ~x50"
+  and 59: "~x29 | ~x36"
+  and 60: "~x29 | ~x43"
+  and 61: "~x29 | ~x50"
+  and 62: "~x36 | ~x43"
+  and 63: "~x36 | ~x50"
+  and 64: "~x43 | ~x50"
+  and 65: "~x2 | ~x9"
+  and 66: "~x2 | ~x16"
+  and 67: "~x2 | ~x23"
+  and 68: "~x2 | ~x30"
+  and 69: "~x2 | ~x37"
+  and 70: "~x2 | ~x44"
+  and 71: "~x2 | ~x51"
+  and 72: "~x9 | ~x16"
+  and 73: "~x9 | ~x23"
+  and 74: "~x9 | ~x30"
+  and 75: "~x9 | ~x37"
+  and 76: "~x9 | ~x44"
+  and 77: "~x9 | ~x51"
+  and 78: "~x16 | ~x23"
+  and 79: "~x16 | ~x30"
+  and 80: "~x16 | ~x37"
+  and 81: "~x16 | ~x44"
+  and 82: "~x16 | ~x51"
+  and 83: "~x23 | ~x30"
+  and 84: "~x23 | ~x37"
+  and 85: "~x23 | ~x44"
+  and 86: "~x23 | ~x51"
+  and 87: "~x30 | ~x37"
+  and 88: "~x30 | ~x44"
+  and 89: "~x30 | ~x51"
+  and 90: "~x37 | ~x44"
+  and 91: "~x37 | ~x51"
+  and 92: "~x44 | ~x51"
+  and 93: "~x3 | ~x10"
+  and 94: "~x3 | ~x17"
+  and 95: "~x3 | ~x24"
+  and 96: "~x3 | ~x31"
+  and 97: "~x3 | ~x38"
+  and 98: "~x3 | ~x45"
+  and 99: "~x3 | ~x52"
+  and 100: "~x10 | ~x17"
+  and 101: "~x10 | ~x24"
+  and 102: "~x10 | ~x31"
+  and 103: "~x10 | ~x38"
+  and 104: "~x10 | ~x45"
+  and 105: "~x10 | ~x52"
+  and 106: "~x17 | ~x24"
+  and 107: "~x17 | ~x31"
+  and 108: "~x17 | ~x38"
+  and 109: "~x17 | ~x45"
+  and 110: "~x17 | ~x52"
+  and 111: "~x24 | ~x31"
+  and 112: "~x24 | ~x38"
+  and 113: "~x24 | ~x45"
+  and 114: "~x24 | ~x52"
+  and 115: "~x31 | ~x38"
+  and 116: "~x31 | ~x45"
+  and 117: "~x31 | ~x52"
+  and 118: "~x38 | ~x45"
+  and 119: "~x38 | ~x52"
+  and 120: "~x45 | ~x52"
+  and 121: "~x4 | ~x11"
+  and 122: "~x4 | ~x18"
+  and 123: "~x4 | ~x25"
+  and 124: "~x4 | ~x32"
+  and 125: "~x4 | ~x39"
+  and 126: "~x4 | ~x46"
+  and 127: "~x4 | ~x53"
+  and 128: "~x11 | ~x18"
+  and 129: "~x11 | ~x25"
+  and 130: "~x11 | ~x32"
+  and 131: "~x11 | ~x39"
+  and 132: "~x11 | ~x46"
+  and 133: "~x11 | ~x53"
+  and 134: "~x18 | ~x25"
+  and 135: "~x18 | ~x32"
+  and 136: "~x18 | ~x39"
+  and 137: "~x18 | ~x46"
+  and 138: "~x18 | ~x53"
+  and 139: "~x25 | ~x32"
+  and 140: "~x25 | ~x39"
+  and 141: "~x25 | ~x46"
+  and 142: "~x25 | ~x53"
+  and 143: "~x32 | ~x39"
+  and 144: "~x32 | ~x46"
+  and 145: "~x32 | ~x53"
+  and 146: "~x39 | ~x46"
+  and 147: "~x39 | ~x53"
+  and 148: "~x46 | ~x53"
+  and 149: "~x5 | ~x12"
+  and 150: "~x5 | ~x19"
+  and 151: "~x5 | ~x26"
+  and 152: "~x5 | ~x33"
+  and 153: "~x5 | ~x40"
+  and 154: "~x5 | ~x47"
+  and 155: "~x5 | ~x54"
+  and 156: "~x12 | ~x19"
+  and 157: "~x12 | ~x26"
+  and 158: "~x12 | ~x33"
+  and 159: "~x12 | ~x40"
+  and 160: "~x12 | ~x47"
+  and 161: "~x12 | ~x54"
+  and 162: "~x19 | ~x26"
+  and 163: "~x19 | ~x33"
+  and 164: "~x19 | ~x40"
+  and 165: "~x19 | ~x47"
+  and 166: "~x19 | ~x54"
+  and 167: "~x26 | ~x33"
+  and 168: "~x26 | ~x40"
+  and 169: "~x26 | ~x47"
+  and 170: "~x26 | ~x54"
+  and 171: "~x33 | ~x40"
+  and 172: "~x33 | ~x47"
+  and 173: "~x33 | ~x54"
+  and 174: "~x40 | ~x47"
+  and 175: "~x40 | ~x54"
+  and 176: "~x47 | ~x54"
+  and 177: "~x6 | ~x13"
+  and 178: "~x6 | ~x20"
+  and 179: "~x6 | ~x27"
+  and 180: "~x6 | ~x34"
+  and 181: "~x6 | ~x41"
+  and 182: "~x6 | ~x48"
+  and 183: "~x6 | ~x55"
+  and 184: "~x13 | ~x20"
+  and 185: "~x13 | ~x27"
+  and 186: "~x13 | ~x34"
+  and 187: "~x13 | ~x41"
+  and 188: "~x13 | ~x48"
+  and 189: "~x13 | ~x55"
+  and 190: "~x20 | ~x27"
+  and 191: "~x20 | ~x34"
+  and 192: "~x20 | ~x41"
+  and 193: "~x20 | ~x48"
+  and 194: "~x20 | ~x55"
+  and 195: "~x27 | ~x34"
+  and 196: "~x27 | ~x41"
+  and 197: "~x27 | ~x48"
+  and 198: "~x27 | ~x55"
+  and 199: "~x34 | ~x41"
+  and 200: "~x34 | ~x48"
+  and 201: "~x34 | ~x55"
+  and 202: "~x41 | ~x48"
+  and 203: "~x41 | ~x55"
+  and 204: "~x48 | ~x55"
+  shows "False"
+    using assms
+    by argo
+
+
+lemma "0 \<le> (yc::real) \<and>
+       0 \<le> yd \<and> 0 \<le> yb \<and> 0 \<le> ya \<Longrightarrow>
+       0 \<le> yf \<and>
+       0 \<le> xh \<and> 0 \<le> ye \<and> 0 \<le> yg \<Longrightarrow>
+       0 \<le> yw \<and> 0 \<le> xs \<and> 0 \<le> yu \<Longrightarrow>
+       0 \<le> aea \<and> 0 \<le> aee \<and> 0 \<le> aed \<Longrightarrow>
+       0 \<le> zy \<and> 0 \<le> xz \<and> 0 \<le> zw \<Longrightarrow>
+       0 \<le> zb \<and>
+       0 \<le> za \<and> 0 \<le> yy \<and> 0 \<le> yz \<Longrightarrow>
+       0 \<le> zp \<and> 0 \<le> zo \<and> 0 \<le> yq \<Longrightarrow>
+       0 \<le> adp \<and> 0 \<le> aeb \<and> 0 \<le> aec \<Longrightarrow>
+       0 \<le> acm \<and> 0 \<le> aco \<and> 0 \<le> acn \<Longrightarrow>
+       0 \<le> abl \<Longrightarrow>
+       0 \<le> zr \<and> 0 \<le> zq \<and> 0 \<le> abh \<Longrightarrow>
+       0 \<le> abq \<and> 0 \<le> zd \<and> 0 \<le> abo \<Longrightarrow>
+       0 \<le> acd \<and>
+       0 \<le> acc \<and> 0 \<le> xi \<and> 0 \<le> acb \<Longrightarrow>
+       0 \<le> acp \<and> 0 \<le> acr \<and> 0 \<le> acq \<Longrightarrow>
+       0 \<le> xw \<and>
+       0 \<le> xr \<and> 0 \<le> xv \<and> 0 \<le> xu \<Longrightarrow>
+       0 \<le> zc \<and> 0 \<le> acg \<and> 0 \<le> ach \<Longrightarrow>
+       0 \<le> zt \<and> 0 \<le> zs \<and> 0 \<le> xy \<Longrightarrow>
+       0 \<le> ady \<and> 0 \<le> adw \<and> 0 \<le> zg \<Longrightarrow>
+       0 \<le> abd \<and>
+       0 \<le> abc \<and> 0 \<le> yr \<and> 0 \<le> abb \<Longrightarrow>
+       0 \<le> adi \<and>
+       0 \<le> x \<and> 0 \<le> adh \<and> 0 \<le> xa \<Longrightarrow>
+       0 \<le> aak \<and> 0 \<le> aai \<and> 0 \<le> aad \<Longrightarrow>
+       0 \<le> aba \<and> 0 \<le> zh \<and> 0 \<le> aay \<Longrightarrow>
+       0 \<le> abg \<and> 0 \<le> ys \<and> 0 \<le> abe \<Longrightarrow>
+       0 \<le> abs1 \<and>
+       0 \<le> yt \<and> 0 \<le> abr \<and> 0 \<le> zu \<Longrightarrow>
+       0 \<le> abv \<and>
+       0 \<le> zn \<and> 0 \<le> abw \<and> 0 \<le> zm \<Longrightarrow>
+       0 \<le> adl \<and> 0 \<le> adn \<Longrightarrow>
+       0 \<le> acf \<and> 0 \<le> aca \<Longrightarrow>
+       0 \<le> ads \<and> 0 \<le> aaq \<Longrightarrow>
+       0 \<le> ada \<Longrightarrow>
+       0 \<le> aaf \<and> 0 \<le> aac \<and> 0 \<le> aag \<Longrightarrow>
+       0 \<le> aal \<and>
+       0 \<le> acu \<and> 0 \<le> acs \<and> 0 \<le> act \<Longrightarrow>
+       0 \<le> aas \<and> 0 \<le> xb \<and> 0 \<le> aat \<Longrightarrow>
+       0 \<le> zk \<and> 0 \<le> zj \<and> 0 \<le> zi \<Longrightarrow>
+       0 \<le> ack \<and>
+       0 \<le> acj \<and> 0 \<le> xc \<and> 0 \<le> aci \<Longrightarrow>
+       0 \<le> aav \<and> 0 \<le> aah \<and> 0 \<le> xd \<Longrightarrow>
+       0 \<le> abt \<and>
+       0 \<le> xo \<and> 0 \<le> abu \<and> 0 \<le> xn \<Longrightarrow>
+       0 \<le> adc \<and>
+       0 \<le> abz \<and> 0 \<le> adc \<and> 0 \<le> abz \<Longrightarrow>
+       0 \<le> xt \<and>
+       0 \<le> zz \<and> 0 \<le> aab \<and> 0 \<le> aaa \<Longrightarrow>
+       0 \<le> adq \<and>
+       0 \<le> xl \<and> 0 \<le> adr \<and> 0 \<le> adb \<Longrightarrow>
+       0 \<le> zf \<and> 0 \<le> yh \<and> 0 \<le> yi \<Longrightarrow>
+       0 \<le> aao \<and> 0 \<le> aam \<and> 0 \<le> xe \<Longrightarrow>
+       0 \<le> abk \<and>
+       0 \<le> aby \<and> 0 \<le> abj \<and> 0 \<le> abx \<Longrightarrow>
+       0 \<le> yp \<Longrightarrow>
+       0 \<le> yl \<and> 0 \<le> yj \<and> 0 \<le> ym \<Longrightarrow>
+       0 \<le> acw \<Longrightarrow>
+       0 \<le> adk \<and>
+       0 \<le> adg \<and> 0 \<le> adj \<and> 0 \<le> adf \<Longrightarrow>
+       0 \<le> adv \<and> 0 \<le> xf \<and> 0 \<le> adu \<Longrightarrow>
+       yc + yd + yb + ya = 1 \<Longrightarrow>
+       yf + xh + ye + yg = 1 \<Longrightarrow>
+       yw + xs + yu = 1 \<Longrightarrow>
+       aea + aee + aed = 1 \<Longrightarrow>
+       zy + xz + zw = 1 \<Longrightarrow>
+       zb + za + yy + yz = 1 \<Longrightarrow>
+       zp + zo + yq = 1 \<Longrightarrow>
+       adp + aeb + aec = 1 \<Longrightarrow>
+       acm + aco + acn = 1 \<Longrightarrow>
+       abl + abl = 1 \<Longrightarrow>
+       zr + zq + abh = 1 \<Longrightarrow>
+       abq + zd + abo = 1 \<Longrightarrow>
+       acd + acc + xi + acb = 1 \<Longrightarrow>
+       acp + acr + acq = 1 \<Longrightarrow>
+       xw + xr + xv + xu = 1 \<Longrightarrow>
+       zc + acg + ach = 1 \<Longrightarrow>
+       zt + zs + xy = 1 \<Longrightarrow>
+       ady + adw + zg = 1 \<Longrightarrow>
+       abd + abc + yr + abb = 1 \<Longrightarrow>
+       adi + x + adh + xa = 1 \<Longrightarrow>
+       aak + aai + aad = 1 \<Longrightarrow>
+       aba + zh + aay = 1 \<Longrightarrow>
+       abg + ys + abe = 1 \<Longrightarrow>
+       abs1 + yt + abr + zu = 1 \<Longrightarrow>
+       abv + zn + abw + zm = 1 \<Longrightarrow>
+       adl + adn = 1 \<Longrightarrow>
+       acf + aca = 1 \<Longrightarrow>
+       ads + aaq = 1 \<Longrightarrow>
+       ada + ada = 1 \<Longrightarrow>
+       aaf + aac + aag = 1 \<Longrightarrow>
+       aal + acu + acs + act = 1 \<Longrightarrow>
+       aas + xb + aat = 1 \<Longrightarrow>
+       zk + zj + zi = 1 \<Longrightarrow>
+       ack + acj + xc + aci = 1 \<Longrightarrow>
+       aav + aah + xd = 1 \<Longrightarrow>
+       abt + xo + abu + xn = 1 \<Longrightarrow>
+       adc + abz + adc + abz = 1 \<Longrightarrow>
+       xt + zz + aab + aaa = 1 \<Longrightarrow>
+       adq + xl + adr + adb = 1 \<Longrightarrow>
+       zf + yh + yi = 1 \<Longrightarrow>
+       aao + aam + xe = 1 \<Longrightarrow>
+       abk + aby + abj + abx = 1 \<Longrightarrow>
+       yp + yp = 1 \<Longrightarrow>
+       yl + yj + ym = 1 \<Longrightarrow>
+       acw + acw + acw + acw = 1 \<Longrightarrow>
+       adk + adg + adj + adf = 1 \<Longrightarrow>
+       adv + xf + adu = 1 \<Longrightarrow>
+       yd = 0 \<or> yb = 0 \<Longrightarrow>
+       xh = 0 \<or> ye = 0 \<Longrightarrow>
+       yy = 0 \<or> za = 0 \<Longrightarrow>
+       acc = 0 \<or> xi = 0 \<Longrightarrow>
+       xv = 0 \<or> xr = 0 \<Longrightarrow>
+       yr = 0 \<or> abc = 0 \<Longrightarrow>
+       zn = 0 \<or> abw = 0 \<Longrightarrow>
+       xo = 0 \<or> abu = 0 \<Longrightarrow>
+       xl = 0 \<or> adr = 0 \<Longrightarrow>
+       (yr + abd < abl \<or>
+        yr + (abd + abb) < 1) \<or>
+       yr + abd = abl \<and>
+       yr + (abd + abb) = 1 \<Longrightarrow>
+       adb + adr < xn + abu \<or>
+       adb + adr = xn + abu \<Longrightarrow>
+       (abl < abt \<or> abl < abt + xo) \<or>
+       abl = abt \<and> abl = abt + xo \<Longrightarrow>
+       yd + yc < abc + abd \<or>
+       yd + yc = abc + abd \<Longrightarrow>
+       aca < abb + yr \<or> aca = abb + yr \<Longrightarrow>
+       acb + acc < xu + xv \<or>
+       acb + acc = xu + xv \<Longrightarrow>
+       (yq < xu + xr \<or>
+        yq + zp < xu + (xr + xw)) \<or>
+       yq = xu + xr \<and>
+       yq + zp = xu + (xr + xw) \<Longrightarrow>
+       (zw < xw \<or>
+        zw < xw + xv \<or>
+        zw + zy < xw + (xv + xu)) \<or>
+       zw = xw \<and>
+       zw = xw + xv \<and>
+       zw + zy = xw + (xv + xu) \<Longrightarrow>
+       xs + yw < zs + zt \<or>
+       xs + yw = zs + zt \<Longrightarrow>
+       aab + xt < ye + yf \<or>
+       aab + xt = ye + yf \<Longrightarrow>
+       (ya + yb < yg + ye \<or>
+        ya + (yb + yc) < yg + (ye + yf)) \<or>
+       ya + yb = yg + ye \<and>
+       ya + (yb + yc) = yg + (ye + yf) \<Longrightarrow>
+       (xu + xv < acb + acc \<or>
+        xu + (xv + xw) < acb + (acc + acd)) \<or>
+       xu + xv = acb + acc \<and>
+       xu + (xv + xw) = acb + (acc + acd) \<Longrightarrow>
+       (zs < xz + zy \<or>
+        zs + xy < xz + (zy + zw)) \<or>
+       zs = xz + zy \<and>
+       zs + xy = xz + (zy + zw) \<Longrightarrow>
+       (zs + zt < xz + zy \<or>
+        zs + (zt + xy) < xz + (zy + zw)) \<or>
+       zs + zt = xz + zy \<and>
+       zs + (zt + xy) = xz + (zy + zw) \<Longrightarrow>
+       yg + ye < ya + yb \<or>
+       yg + ye = ya + yb \<Longrightarrow>
+       (abd < yc \<or> abd + abc < yc + yd) \<or>
+       abd = yc \<and> abd + abc = yc + yd \<Longrightarrow>
+       (ye + yf < adr + adq \<or>
+        ye + (yf + yg) < adr + (adq + adb)) \<or>
+       ye + yf = adr + adq \<and>
+       ye + (yf + yg) = adr + (adq + adb) \<Longrightarrow>
+       yh + yi < ym + yj \<or>
+       yh + yi = ym + yj \<Longrightarrow>
+       (abq < yl \<or> abq + abo < yl + ym) \<or>
+       abq = yl \<and> abq + abo = yl + ym \<Longrightarrow>
+       (yp < zp \<or>
+        yp < zp + zo \<or> 1 < zp + (zo + yq)) \<or>
+       yp = zp \<and>
+       yp = zp + zo \<and> 1 = zp + (zo + yq) \<Longrightarrow>
+       (abb + yr < aca \<or>
+        abb + (yr + abd) < aca + acf) \<or>
+       abb + yr = aca \<and>
+       abb + (yr + abd) = aca + acf \<Longrightarrow>
+       adw + zg < abe + ys \<or>
+       adw + zg = abe + ys \<Longrightarrow>
+       zd + abq < ys + abg \<or>
+       zd + abq = ys + abg \<Longrightarrow>
+       yt + abs1 < aby + abk \<or>
+       yt + abs1 = aby + abk \<Longrightarrow>
+       (yu < abx \<or>
+        yu < abx + aby \<or>
+        yu + yw < abx + (aby + abk)) \<or>
+       yu = abx \<and>
+       yu = abx + aby \<and>
+       yu + yw = abx + (aby + abk) \<Longrightarrow>
+       aaf < adv \<or> aaf = adv \<Longrightarrow>
+       abj + abk < yy + zb \<or>
+       abj + abk = yy + zb \<Longrightarrow>
+       (abb < yz \<or>
+        abb + abc < yz + za \<or>
+        abb + (abc + abd) < yz + (za + zb)) \<or>
+       abb = yz \<and>
+       abb + abc = yz + za \<and>
+       abb + (abc + abd) = yz + (za + zb) \<Longrightarrow>
+       (acg + zc < zd + abq \<or>
+        acg + (zc + ach)
+        < zd + (abq + abo)) \<or>
+       acg + zc = zd + abq \<and>
+       acg + (zc + ach) =
+       zd + (abq + abo) \<Longrightarrow>
+       zf < acm \<or> zf = acm \<Longrightarrow>
+       (zg + ady < acn + acm \<or>
+        zg + (ady + adw)
+        < acn + (acm + aco)) \<or>
+       zg + ady = acn + acm \<and>
+       zg + (ady + adw) =
+       acn + (acm + aco) \<Longrightarrow>
+       aay + zh < zi + zj \<or>
+       aay + zh = zi + zj \<Longrightarrow>
+       zy < zk \<or> zy = zk \<Longrightarrow>
+       (adn < zm + zn \<or>
+        adn + adl < zm + (zn + abv)) \<or>
+       adn = zm + zn \<and>
+       adn + adl = zm + (zn + abv) \<Longrightarrow>
+       zo + zp < zs + zt \<or>
+       zo + zp = zs + zt \<Longrightarrow>
+       zq + zr < zs + zt \<or>
+       zq + zr = zs + zt \<Longrightarrow>
+       (aai < adi \<or> aai < adi + adh) \<or>
+       aai = adi \<and> aai = adi + adh \<Longrightarrow>
+       (abr < acj \<or>
+        abr + (abs1 + zu)
+        < acj + (aci + ack)) \<or>
+       abr = acj \<and>
+       abr + (abs1 + zu) =
+       acj + (aci + ack) \<Longrightarrow>
+       (abl < zw \<or> 1 < zw + zy) \<or>
+       abl = zw \<and> 1 = zw + zy \<Longrightarrow>
+       (zz + aaa < act + acu \<or>
+        zz + (aaa + aab)
+        < act + (acu + aal)) \<or>
+       zz + aaa = act + acu \<and>
+       zz + (aaa + aab) =
+       act + (acu + aal) \<Longrightarrow>
+       (aam < aac \<or> aam + aao < aac + aaf) \<or>
+       aam = aac \<and> aam + aao = aac + aaf \<Longrightarrow>
+       (aak < aaf \<or> aak + aad < aaf + aag) \<or>
+       aak = aaf \<and> aak + aad = aaf + aag \<Longrightarrow>
+       (aah < aai \<or> aah + aav < aai + aak) \<or>
+       aah = aai \<and> aah + aav = aai + aak \<Longrightarrow>
+       act + (acu + aal) < aam + aao \<or>
+       act + (acu + aal) = aam + aao \<Longrightarrow>
+       (ads < aat \<or> 1 < aat + aas) \<or>
+       ads = aat \<and> 1 = aat + aas \<Longrightarrow>
+       (aba < aas \<or> aba + aay < aas + aat) \<or>
+       aba = aas \<and> aba + aay = aas + aat \<Longrightarrow>
+       acm < aav \<or> acm = aav \<Longrightarrow>
+       (ada < aay \<or> 1 < aay + aba) \<or>
+       ada = aay \<and> 1 = aay + aba \<Longrightarrow>
+       abb + (abc + abd) < abe + abg \<or>
+       abb + (abc + abd) = abe + abg \<Longrightarrow>
+       (abh < abj \<or> abh < abj + abk) \<or>
+       abh = abj \<and> abh = abj + abk \<Longrightarrow>
+       1 < abo + abq \<or> 1 = abo + abq \<Longrightarrow>
+       (acj < abr \<or> acj + aci < abr + abs1) \<or>
+       acj = abr \<and> acj + aci = abr + abs1 \<Longrightarrow>
+       (abt < abv \<or> abt + abu < abv + abw) \<or>
+       abt = abv \<and> abt + abu = abv + abw \<Longrightarrow>
+       (abx < adc \<or> abx + aby < adc + abz) \<or>
+       abx = adc \<and> abx + aby = adc + abz \<Longrightarrow>
+       (acf < acd \<or>
+        acf < acd + acc \<or>
+        1 < acd + (acc + acb)) \<or>
+       acf = acd \<and>
+       acf = acd + acc \<and>
+       1 = acd + (acc + acb) \<Longrightarrow>
+       acc + acd < acf \<or> acc + acd = acf \<Longrightarrow>
+       (acg < acq \<or> acg + ach < acq + acr) \<or>
+       acg = acq \<and> acg + ach = acq + acr \<Longrightarrow>
+       aci + (acj + ack) < acr + acp \<or>
+       aci + (acj + ack) = acr + acp \<Longrightarrow>
+       (acm < acp \<or>
+        acm + acn < acp + acq \<or>
+        acm + (acn + aco)
+        < acp + (acq + acr)) \<or>
+       acm = acp \<and>
+       acm + acn = acp + acq \<and>
+       acm + (acn + aco) =
+       acp + (acq + acr) \<Longrightarrow>
+       (acs + act < acw + acw \<or>
+        acs + (act + acu)
+        < acw + (acw + acw)) \<or>
+       acs + act = acw + acw \<and>
+       acs + (act + acu) =
+       acw + (acw + acw) \<Longrightarrow>
+       (ada < adb + adr \<or>
+        1 < adb + (adr + adq)) \<or>
+       ada = adb + adr \<and>
+       1 = adb + (adr + adq) \<Longrightarrow>
+       (adc + adc < adf + adg \<or>
+        adc + (adc + abz)
+        < adf + (adg + adk)) \<or>
+       adc + adc = adf + adg \<and>
+       adc + (adc + abz) =
+       adf + (adg + adk) \<Longrightarrow>
+       adh + adi < adj + adk \<or>
+       adh + adi = adj + adk \<Longrightarrow>
+       (adl < aec \<or> 1 < aec + adp) \<or>
+       adl = aec \<and> 1 = aec + adp \<Longrightarrow>
+       (adq < ads \<or> adq + adr < ads) \<or>
+       adq = ads \<and> adq + adr = ads \<Longrightarrow>
+       adu + adv < aed + aea \<or>
+       adu + adv = aed + aea \<Longrightarrow>
+       (adw < aee \<or> adw + ady < aee + aea) \<or>
+       adw = aee \<and> adw + ady = aee + aea \<Longrightarrow>
+       (aeb < aed \<or> aeb + aec < aed + aee) \<or>
+       aeb = aed \<and> aeb + aec = aed + aee \<Longrightarrow>
+       False"
+       by argo
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_cc.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,654 @@
+(*  Title:      Tools/Argo/argo_cc.ML
+    Author:     Sascha Boehme
+    Author:     Dmitriy Traytel and Matthias Franze, TU Muenchen
+
+Equality reasoning based on congurence closure. It features:
+
+  * congruence closure for any term that participates in equalities
+  * support for predicates
+
+These features might be added:
+
+  * caching of explanations while building proofs to obtain shorter proofs
+    and faster proof checking
+  * propagating relevant merges of equivalence classes to all other theory solvers
+  * propagating new relevant negated equalities to all other theory solvers
+  * creating lemma "f ~= g | a ~= b | f a = g b" for asserted negated equalities
+    between "f a" and "g b" (dynamic ackermannization)
+
+The implementation is inspired by:
+
+  Robert Nieuwenhuis and Albert Oliveras. Fast Congruence Closure and
+  Extensions. In Information and Computation, volume 205(4),
+  pages 557-580, 2007.
+
+  Harald Ganzinger, George Hagen, Robert Nieuwenhuis, Albert Oliveras,
+  Cesare Tinelli. DPLL(T): Fast decision procedures. In Lecture Notes in
+  Computer Science, volume 3114, pages 175-188. Springer, 2004.
+*)
+
+signature ARGO_CC =
+sig
+  (* context *)
+  type context
+  val context: context
+
+  (* simplification *)
+  val simplify: Argo_Rewr.context -> Argo_Rewr.context
+  
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> Argo_Lit.literal option * context
+
+  (* main operations *)
+  val assume: Argo_Common.literal -> context -> Argo_Lit.literal Argo_Common.implied * context
+  val check: context -> Argo_Lit.literal Argo_Common.implied * context
+  val explain: Argo_Lit.literal -> context -> (Argo_Cls.clause * context) option
+  val add_level: context -> context
+  val backtrack: context -> context
+end
+
+structure Argo_Cc: ARGO_CC =
+struct
+
+(* tables indexed by pairs of terms *)
+
+val term2_ord = prod_ord Argo_Term.term_ord Argo_Term.term_ord
+
+structure Argo_Term2tab = Table(type key = Argo_Term.term * Argo_Term.term val ord = term2_ord)
+
+
+(* equality certificates *)
+
+(*
+  The solver keeps assumed equalities to produce explanations later on.
+
+  A flat equality (lp, (t1, t2)) consists of the assumed literal and its proof
+  as well as the terms t1 and t2 that are assumed to be equal. The literal expresses
+  the equality t1 = t2.
+
+  A congruence equality (t1, t2) is an equality t1 = t2 where both terms are
+  applications (f a) and (g b).
+
+  A symmetric equality eq is a marker for applying the symmetry rule to eq.
+*)
+
+datatype eq =
+  Flat of Argo_Common.literal * (Argo_Term.term * Argo_Term.term) |
+  Cong of Argo_Term.term * Argo_Term.term |
+  Symm of eq
+
+fun dest_eq (Flat (_, tp)) = tp
+  | dest_eq (Cong tp) = tp
+  | dest_eq (Symm eq) = swap (dest_eq eq)
+
+fun symm (Symm eq) = eq
+  | symm eq = Symm eq
+
+fun negate (Flat ((lit, p), tp)) = Flat ((Argo_Lit.negate lit, p), tp)
+  | negate (Cong tp) = Cong tp
+  | negate (Symm eq) = Symm (negate eq)
+
+fun dest_app (Argo_Term.T (_, Argo_Expr.App, [t1, t2])) = (t1, t2)
+  | dest_app _ = raise Fail "bad application"
+
+
+(* context *)
+
+(*
+  Each representative keeps track of the yet unimplied atoms in which this any member of
+  this representative's equivalence class occurs. An atom is either a list of equalities
+  between two terms, a list of predicates or a certificate. The certificate denotes that
+  this equivalence class contains already implied predicates, and the literal accompanying
+  the certificate specifies the polarity of these predicates.
+*)
+
+datatype atoms =
+  Eqs of (Argo_Term.term * Argo_Term.term) list |
+  Preds of Argo_Term.term list |
+  Cert of Argo_Common.literal
+
+(*
+  Each representative has an associated ritem that contains the members of the
+  equivalence class, the yet unimplied atoms and further information.
+*)
+
+type ritem = {
+  size: int, (* the size of the equivalence class *)
+  class: Argo_Term.term list, (* the equivalence class as a list of distinct terms *)
+  occs: Argo_Term.term list, (* a list of all application terms in which members of
+    the equivalence class occur either as function or as argument *)
+  neqs: (Argo_Term.term * eq) list, (* a list of terms from disjoint equivalence classes,
+    for each term of this list there is a certificate of a negated equality that is
+    required to explain why the equivalence classes are disjoint *)
+  atoms: atoms} (* the atoms of the representative *)
+
+type repr = Argo_Term.term Argo_Termtab.table
+type rdata = ritem Argo_Termtab.table
+type apps = Argo_Term.term Argo_Term2tab.table
+type trace = (Argo_Term.term * eq) Argo_Termtab.table
+
+type context = {
+  repr: repr, (* a table mapping terms to their representatives *)
+  rdata: rdata, (* a table mapping representatives to their ritems *)
+  apps: apps, (* a table mapping a function and an argument to their application *)
+  trace: trace, (* the proof forest used to trace assumed and implied equalities *)
+  prf: Argo_Proof.context, (* the proof context *)
+  back: (repr * rdata * apps * trace) list} (* backtracking information *)
+
+fun mk_context repr rdata apps trace prf back: context =
+  {repr=repr, rdata=rdata, apps=apps, trace=trace, prf=prf, back=back}
+
+val context =
+  mk_context Argo_Termtab.empty Argo_Termtab.empty Argo_Term2tab.empty Argo_Termtab.empty
+    Argo_Proof.cc_context []
+
+fun repr_of repr t = the_default t (Argo_Termtab.lookup repr t)
+fun repr_of' ({repr, ...}: context) = repr_of repr
+fun put_repr t r = Argo_Termtab.update (t, r)
+
+fun mk_ritem size class occs neqs atoms: ritem =
+  {size=size, class=class, occs=occs, neqs=neqs, atoms=atoms}
+
+fun as_ritem t = mk_ritem 1 [t] [] [] (Eqs [])
+fun as_pred_ritem t = mk_ritem 1 [t] [] [] (Preds [t])
+fun gen_ritem_of mk rdata r = the_default (mk r) (Argo_Termtab.lookup rdata r)
+fun ritem_of rdata = gen_ritem_of as_ritem rdata
+fun ritem_of_pred rdata = gen_ritem_of as_pred_ritem rdata
+fun ritem_of' ({rdata, ...}: context) = ritem_of rdata
+fun put_ritem r ri = Argo_Termtab.update (r, ri)
+
+fun add_occ r occ = Argo_Termtab.map_default (r, as_ritem r)
+  (fn {size, class, occs, neqs, atoms}: ritem => mk_ritem size class (occ :: occs) neqs atoms)
+
+fun put_atoms atoms ({size, class, occs, neqs, ...}: ritem) = mk_ritem size class occs neqs atoms
+
+fun add_eq_atom r atom = Argo_Termtab.map_default (r, as_ritem r)
+  (fn ri as {atoms=Eqs atoms, ...}: ritem => put_atoms (Eqs (atom :: atoms)) ri
+    | ri => put_atoms (Eqs [atom]) ri)
+
+fun lookup_app apps tp = Argo_Term2tab.lookup apps tp
+fun put_app tp app = Argo_Term2tab.update_new (tp, app)
+
+
+(* traces for explanations *)
+
+(*
+  Assumed and implied equalities are collected in a proof forest for being able to
+  produce explanations. For each equivalence class there is one proof tree. The
+  equality certificates are oriented towards a root term, that is not necessarily
+  the representative of the equivalence class.
+*)
+
+(*
+  Whenever two equivalence classes are merged due to an equality t1 = t2, the shorter
+  of the two paths, either from t1 to its root or t2 to its root, is re-oriented such
+  that the relevant ti becomes the new root of its tree. Then, a new edge between ti
+  and the other term of the equality t1 = t2 is added to connect the two proof trees.
+*)
+
+fun depth_of trace t =
+  (case Argo_Termtab.lookup trace t of
+    NONE => 0
+  | SOME (t', _) => 1 + depth_of trace t')
+
+fun reorient t trace =
+  (case Argo_Termtab.lookup trace t of
+    NONE => trace
+  | SOME (t', eq) => Argo_Termtab.update (t', (t, symm eq)) (reorient t' trace))
+
+fun new_edge from to eq trace = Argo_Termtab.update (from, (to, eq)) (reorient from trace)
+
+fun with_shortest f (t1, t2) eq trace =
+  (if depth_of trace t1 <= depth_of trace t2 then f t1 t2 eq else f t2 t1 (symm eq)) trace
+
+fun add_edge eq trace = with_shortest new_edge (dest_eq eq) eq trace
+
+(*
+  To produce an explanation that t1 and t2 are equal, the paths to their root are
+  extracted from the proof forest. Common ancestors in both paths are dropped.
+*)
+
+fun path_to_root trace path t =
+  (case Argo_Termtab.lookup trace t of
+    NONE => (t, path)
+  | SOME (t', _) => path_to_root trace (t :: path) t')
+
+fun drop_common root (t1 :: path1) (t2 :: path2) =
+      if Argo_Term.eq_term (t1, t2) then drop_common t1 path1 path2 else root
+  | drop_common root _ _ = root
+
+fun common_ancestor trace t1 t2 =
+  let val ((root, path1), (_, path2)) = apply2 (path_to_root trace []) (t1, t2)
+  in drop_common root path1 path2 end
+
+(*
+  The proof of an assumed literal is typically a hypothesis. If the assumed literal is
+  already known to be a unit literal, then there is already a proof for it.
+*)
+
+fun proof_of (lit, NONE) lits prf =
+      (insert Argo_Lit.eq_lit (Argo_Lit.negate lit) lits, Argo_Proof.mk_hyp lit prf)
+  | proof_of (_, SOME p) lits prf = (lits, (p, prf))
+
+(*
+  The explanation of equality between two terms t1 and t2 is computed based on the
+  paths from t1 and t2 to their common ancestor t in the proof forest. For each of
+  the two paths, a transitive proof of equality t1 = t and t = t2 is constructed,
+  such that t1 = t2 follows by transitivity.
+  
+  Each edge of the paths denotes an assumed or implied equality. Implied equalities
+  might be due to congruences (f a = g b) for which the equalities f = g and a = b
+  need to be explained recursively.
+*)
+
+fun mk_eq_proof trace t1 t2 lits prf =
+  if Argo_Term.eq_term (t1, t2) then (lits, Argo_Proof.mk_refl t1 prf)
+  else
+    let
+      val root = common_ancestor trace t1 t2
+      val (lits, (p1, prf)) = trans_proof I I trace t1 root lits prf
+      val (lits, (p2, prf)) = trans_proof swap symm trace t2 root lits prf
+    in (lits, Argo_Proof.mk_trans p1 p2 prf) end
+
+and trans_proof sw sy trace t root lits prf =
+  if Argo_Term.eq_term (t, root) then (lits, Argo_Proof.mk_refl t prf)
+  else
+    (case Argo_Termtab.lookup trace t of
+      NONE => raise Fail "bad trace"
+    | SOME (t', eq) => 
+        let
+          val (lits, (p1, prf)) = proof_step trace (sy eq) lits prf
+          val (lits, (p2, prf)) = trans_proof sw sy trace t' root lits prf
+        in (lits, uncurry Argo_Proof.mk_trans (sw (p1, p2)) prf) end)
+
+and proof_step _ (Flat (cert, _)) lits prf = proof_of cert lits prf
+  | proof_step trace (Cong tp) lits prf =
+      let
+        val ((t1, t2), (u1, u2)) = apply2 dest_app tp
+        val (lits, (p1, prf)) = mk_eq_proof trace t1 u1 lits prf
+        val (lits, (p2, prf)) = mk_eq_proof trace t2 u2 lits prf
+      in (lits, Argo_Proof.mk_cong p1 p2 prf) end
+  | proof_step trace (Symm eq) lits prf =
+      proof_step trace eq lits prf ||> uncurry Argo_Proof.mk_symm
+
+(*
+  All clauses produced by a theory solver are expected to be a lemma.
+  The lemma proof must hence be the last proof step.
+*)
+
+fun close_proof lit lits (p, prf) = (lit :: lits, Argo_Proof.mk_lemma [lit] p prf)
+
+(*
+  The explanation for the equality of t1 and t2 used the above algorithm.
+*)
+
+fun explain_eq lit t1 t2 ({repr, rdata, apps, trace, prf, back}: context) =
+  let val (lits, (p, prf)) = mk_eq_proof trace t1 t2 [] prf |-> close_proof lit
+  in ((lits, p), mk_context repr rdata apps trace prf back) end
+
+(*
+  The explanation that t1 and t2 are distinct uses the negated equality u1 ~= u2 that
+  explains why the equivalence class containing t1 and u1 and the equivalence class
+  containing t2 and u2 are disjoint. The explanations for t1 = u1 and u2 = t2 are
+  constructed using the above algorithm. By transitivity, it follows that t1 ~= t2.  
+*)
+
+fun finish_proof (Flat ((lit, _), _)) lits p prf = close_proof lit lits (p, prf)
+  | finish_proof (Cong _) _ _ _ = raise Fail "bad equality"
+  | finish_proof (Symm eq) lits p prf = Argo_Proof.mk_symm p prf |-> finish_proof eq lits
+
+fun explain_neq eq eq' ({repr, rdata, apps, trace, prf, back}: context) =
+  let
+    val (t1, t2) = dest_eq eq
+    val (u1, u2) = dest_eq eq'
+
+    val (lits, (p, prf)) = proof_step trace eq' [] prf
+    val (lits, (p1, prf)) = mk_eq_proof trace t1 u1 lits prf
+    val (lits, (p2, prf)) = mk_eq_proof trace u2 t2 lits prf
+    val (lits, (p, prf)) = 
+      Argo_Proof.mk_trans p p2 prf |-> Argo_Proof.mk_trans p1 |-> finish_proof eq lits
+  in ((lits, p), mk_context repr rdata apps trace prf back) end
+
+
+(* propagating new equalities *)
+
+exception CONFLICT of Argo_Cls.clause * context
+
+(*
+  comment missing
+*)
+
+fun same_repr repr r (t, _) = Argo_Term.eq_term (r, repr_of repr t)
+
+fun has_atom rdata r eq =
+  (case #atoms (ritem_of rdata r) of
+    Eqs eqs => member (Argo_Term.eq_term o snd) eqs eq
+  | _ => false)
+
+fun add_implied mk_lit repr rdata r neqs (atom as (t, eq)) (eqs, ls) =
+  let val r' = repr_of repr t
+  in
+    if Argo_Term.eq_term (r, r') then (eqs, insert Argo_Lit.eq_lit (mk_lit eq) ls)
+    else if exists (same_repr repr r') neqs andalso has_atom rdata r' eq then
+      (eqs, Argo_Lit.Neg eq :: ls)
+    else (atom :: eqs, ls)
+  end
+
+(*
+  comment missing
+*)
+
+fun copy_occ repr app (eqs, occs, apps) =
+  let val rp = apply2 (repr_of repr) (dest_app app)
+  in
+    (case lookup_app apps rp of
+      SOME app' => (Cong (app, app') :: eqs, occs, apps)
+    | NONE => (eqs, app :: occs, put_app rp app apps))
+  end
+
+(*
+  comment missing
+*)
+
+fun add_lits (Argo_Lit.Pos _, _) = fold (cons o Argo_Lit.Pos)
+  | add_lits (Argo_Lit.Neg _, _) = fold (cons o Argo_Lit.Neg)
+
+fun join_atoms f (Eqs eqs1) (Eqs eqs2) ls = f eqs1 eqs2 ls
+  | join_atoms _ (Preds ts1) (Preds ts2) ls = (Preds (union Argo_Term.eq_term ts1 ts2), ls)
+  | join_atoms _ (Preds ts) (Cert lp) ls = (Cert lp, add_lits lp ts ls)
+  | join_atoms _ (Cert lp) (Preds ts) ls = (Cert lp, add_lits lp ts ls)
+  | join_atoms _ (Cert lp) (Cert _) ls = (Cert lp, ls)
+  | join_atoms _ _ _ _ = raise Fail "bad atoms"
+
+(*
+  comment missing
+*)
+
+fun join r1 ri1 r2 ri2 eq (eqs, ls, {repr, rdata, apps, trace, prf, back}: context) =
+  let
+    val {size=size1, class=class1, occs=occs1, neqs=neqs1, atoms=atoms1}: ritem = ri1
+    val {size=size2, class=class2, occs=occs2, neqs=neqs2, atoms=atoms2}: ritem = ri2
+
+    val repr = fold (fn t => put_repr t r1) class2 repr
+    val class = fold cons class2 class1
+    val (eqs, occs, apps) = fold (copy_occ repr) occs2 (eqs, occs1, apps)
+    val trace = add_edge eq trace
+    val neqs = AList.merge Argo_Term.eq_term (K true) (neqs1, neqs2)
+    fun add r neqs = fold (add_implied Argo_Lit.Pos repr rdata r neqs)
+    fun adds eqs1 eqs2 ls = ([], ls) |> add r2 neqs2 eqs1 |> add r1 neqs1 eqs2 |>> Eqs
+    val (atoms, ls) = join_atoms adds atoms1 atoms2 ls
+    (* TODO: make sure that all implied literals are propagated *)
+    val rdata = put_ritem r1 (mk_ritem (size1 + size2) class occs neqs atoms) rdata
+  in (eqs, ls, mk_context repr rdata apps trace prf back) end
+
+(*
+  comment missing
+*)
+
+fun find_neq ({repr, ...}: context) ({neqs, ...}: ritem) r = find_first (same_repr repr r) neqs
+
+fun check_join (r1, r2) (ri1, ri2) eq (ecx as (_, _, cx)) =
+  (case find_neq cx ri2 r1 of
+    SOME (_, eq') => raise CONFLICT (explain_neq (negate (symm eq)) eq' cx)
+  | NONE =>
+      (case find_neq cx ri1 r2 of
+        SOME (_, eq') => raise CONFLICT (explain_neq (negate eq) eq' cx)
+      | NONE => join r1 ri1 r2 ri2 eq ecx))
+
+(*
+  comment missing
+*)
+
+fun with_max_class f (rp as (r1, r2)) (rip as (ri1: ritem, ri2: ritem)) eq =
+  if #size ri1 >= #size ri2 then f rp rip eq else f (r2, r1) (ri2, ri1) (symm eq)
+
+(*
+  comment missing
+*)
+
+fun propagate ([], ls, cx) = (rev ls, cx)
+  | propagate (eq :: eqs, ls, cx) =
+      let val rp = apply2 (repr_of' cx) (dest_eq eq)
+      in 
+        if Argo_Term.eq_term rp then propagate (eqs, ls, cx)
+        else propagate (with_max_class check_join rp (apply2 (ritem_of' cx) rp) eq (eqs, ls, cx))
+      end
+
+fun without lit (lits, cx) = (Argo_Common.Implied (remove Argo_Lit.eq_lit lit lits), cx)
+
+fun flat_merge (lp as (lit, _)) eq cx = without lit (propagate ([Flat (lp, eq)], [], cx))
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+(*
+  comment missing
+*)
+
+fun app_merge app tp (cx as {repr, rdata, apps, trace, prf, back}: context) =
+  let val rp as (r1, r2) = apply2 (repr_of repr) tp
+  in
+    (case lookup_app apps rp of
+      SOME app' =>
+        (case propagate ([Cong (app, app')], [], cx) of
+          ([], cx) => cx
+        | _ => raise Fail "bad application merge")
+    | NONE =>
+        let val rdata = add_occ r1 app (add_occ r2 app rdata)
+        in mk_context repr rdata (put_app rp app apps) trace prf back end)
+  end
+
+(*
+  A negated equality between t1 and t2 is only recorded if t1 and t2 are not already known
+  to belong to the same class. In that case, a conflict is raised with an explanation
+  why t1 and t2 are equal. Otherwise, the classes of t1 and t2 are marked as disjoint by
+  storing the negated equality in the ritems of t1's and t2's representative. All equalities
+  between terms of t1's and t2's class are implied as negated equalities. Those equalities
+  are found in the ritems of t1's and t2's representative.
+*)
+
+fun note_neq eq (r1, r2) (t1, t2) ({repr, rdata, apps, trace, prf, back}: context) =
+  let
+    val {size=size1, class=class1, occs=occs1, neqs=neqs1, atoms=atoms1}: ritem = ritem_of rdata r1
+    val {size=size2, class=class2, occs=occs2, neqs=neqs2, atoms=atoms2}: ritem = ritem_of rdata r2
+
+    fun add r (Eqs eqs) ls = fold (add_implied Argo_Lit.Neg repr rdata r []) eqs ([], ls) |>> Eqs
+      | add _ _ _ = raise Fail "bad negated equality between predicates"
+    val ((atoms1, atoms2), ls) = [] |> add r2 atoms1 ||>> add r1 atoms2
+    val ri1 = mk_ritem size1 class1 occs1 ((t2, eq) :: neqs1) atoms1
+    val ri2 = mk_ritem size2 class2 occs2 ((t1, symm eq) :: neqs2) atoms2
+  in (ls, mk_context repr (put_ritem r1 ri1 (put_ritem r2 ri2 rdata)) apps trace prf back) end
+
+fun flat_neq (lp as (lit, _)) (tp as (t1, t2)) cx =
+  let val rp = apply2 (repr_of' cx) tp
+  in
+    if Argo_Term.eq_term rp then
+      let val (cls, cx) = explain_eq (Argo_Lit.negate lit) t1 t2 cx
+      in (Argo_Common.Conflict cls, cx) end
+    else without lit (note_neq (Flat (lp, tp)) rp tp cx)
+  end
+
+
+(* simplification *)
+
+(*
+  Only equalities are subject to normalizations. An equality between two expressions e1 and e2
+  is normalized, if e1 is less than e2 based on the expression ordering. If e1 and e2 are
+  syntactically equal, the equality between these two expressions is normalized to the true
+  expression.
+*)
+
+fun norm_eq env =
+  let val e1 = Argo_Rewr.get env 1 and e2 = Argo_Rewr.get env 2
+  in
+    (case Argo_Expr.expr_ord (e1, e2) of
+      EQUAL => SOME (Argo_Proof.Rewr_Eq_Refl, Argo_Rewr.E Argo_Expr.true_expr)
+    | LESS => NONE
+    | GREATER => SOME (Argo_Proof.Rewr_Eq_Symm, Argo_Rewr.E (Argo_Expr.mk_eq e2 e1)))
+  end
+
+val simplify = Argo_Rewr.func "(eq (? 1) (? 2))" norm_eq
+
+
+(* declaring atoms *)
+
+(*
+  Only a genuinely new equality term t for the equality "t1 = t2" is added. If t1 and t2 belong
+  to the same equality class or if the classes of t1 and t2 are known to be disjoint, the
+  respective literal is returned together with an unmodified context.
+*)
+
+fun add_eq_term t t1 t2 (rp as (r1, r2)) (cx as {repr, rdata, apps, trace, prf, back}: context) =
+  if Argo_Term.eq_term rp then (SOME (Argo_Lit.Pos t), cx)
+  else if is_some (find_neq cx (ritem_of rdata r1) r2) then (SOME (Argo_Lit.Neg t), cx)
+  else
+    let val rdata = add_eq_atom r1 (t2, t) (add_eq_atom r2 (t1, t) rdata)
+    in (NONE, mk_context repr rdata apps trace prf back) end
+
+(*
+  Only a genuinely new predicate t, which is an application "t1 t2", is added.
+  If there is a predicate that is known to be congruent to the representatives of t1 and t2,
+  and that predicate or its negation has already been assummed, the respective literal of t
+  is returned together with an unmodified context.
+*)
+
+fun add_pred_term t rp (cx as {repr, rdata, apps, trace, prf, back}: context) =
+  (case lookup_app apps rp of
+    NONE => (NONE, mk_context repr (put_ritem t (as_pred_ritem t) rdata) apps trace prf back)
+  | SOME app =>
+      (case `(ritem_of_pred rdata) (repr_of repr app) of
+        ({atoms=Cert (Argo_Lit.Pos _, _), ...}: ritem, _) => (SOME (Argo_Lit.Pos t), cx)
+      | ({atoms=Cert (Argo_Lit.Neg _, _), ...}: ritem, _) => (SOME (Argo_Lit.Neg t), cx)
+      | (ri as {atoms=Preds ts, ...}: ritem, r) =>
+          let val rdata = put_ritem r (put_atoms (Preds (t :: ts)) ri) rdata
+          in (NONE, mk_context repr rdata apps trace prf back) end
+      | ({atoms=Eqs _, ...}: ritem, _) => raise Fail "bad predicate"))
+
+(*
+  For each term t that is an application "t1 t2", the reflexive equality t = t1 t2 is added
+  to the context. This is required for propagations of congruences.
+*)
+
+fun flatten (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])) cx =
+      flatten t1 (flatten t2 (app_merge t (t1, t2) cx))
+  | flatten _ cx = cx
+
+(*
+  Atoms to be added to the context must either be an equality "t1 = t2" or
+  an application "t1 t2" (a predicate). Besides adding the equality or the application,
+  reflexive equalities for for all applications in the terms t1 and t2 are added.
+*)
+
+fun add_atom (t as Argo_Term.T (_, Argo_Expr.Eq, [t1, t2])) cx =
+      add_eq_term t t1 t2 (apply2 (repr_of' cx) (t1, t2)) (flatten t1 (flatten t2 cx))
+  | add_atom (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])) cx =
+      let val cx = flatten t1 (flatten t2 (app_merge t (t1, t2) cx))
+      in add_pred_term t (apply2 (repr_of' cx) (t1, t2)) cx end
+  | add_atom _ cx = (NONE, cx)
+
+
+(* assuming external knowledge *)
+
+(*
+  Assuming a predicate r replaces all predicate atoms of r's ritem with the assumed certificate.
+  The predicate atoms are implied, either with positive or with negative polarity based on
+  the assumption.
+
+  There must not be a certificate for r since otherwise r would have been assumed before already.
+*)
+
+fun assume_pred lit mk_lit cert r ({repr, rdata, apps, trace, prf, back}: context) =
+  (case ritem_of_pred rdata r of
+    {size, class, occs, neqs, atoms=Preds ts}: ritem =>
+      let val rdata = put_ritem r (mk_ritem size class occs neqs cert) rdata
+      in without lit (map mk_lit ts, mk_context repr rdata apps trace prf back) end
+  | _ => raise Fail "bad predicate assumption")
+
+(*
+  Assumed equalities "t1 = t2" are treated as flat equalities between terms t1 and t2.
+  If t1 and t2 are applications, congruences are propagated as part of the merge between t1 and t2.
+  Negated equalities are handled likewise.
+
+  Assumed predicates do not trigger congruences. Only predicates of the same class are implied.
+*)
+
+fun assume (lp as (Argo_Lit.Pos (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2])), _)) cx =
+      flat_merge lp (t1, t2) cx
+  | assume (lp as (Argo_Lit.Neg (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2])), _)) cx =
+      flat_neq lp (t1, t2) cx
+  | assume (lp as (lit as Argo_Lit.Pos (t as Argo_Term.T (_, Argo_Expr.App, [_, _])), _)) cx =
+      assume_pred lit Argo_Lit.Pos (Cert lp) (repr_of' cx t) cx
+  | assume (lp as (lit as Argo_Lit.Neg (t as Argo_Term.T (_, Argo_Expr.App, [_, _])), _)) cx =
+      assume_pred lit Argo_Lit.Neg (Cert lp) (repr_of' cx t) cx
+  | assume _ cx = (Argo_Common.Implied [], cx)
+
+
+(* checking for consistency and pending implications *)
+
+(*
+  The internal model is always kept consistent. All implications are propagated as soon as
+  new information is assumed. Hence, there is nothing to be done here.
+*)
+
+fun check cx = (Argo_Common.Implied [], cx)
+
+
+(* explanations *)
+
+(*
+  The explanation for the predicate t, which is an application of t1 and t2, is constructed
+  from the explanation of the predicate application "u1 u2" as well as the equalities "u1 = t1"
+  and "u2 = t2" which both are constructed from the proof forest. The substitution rule is
+  the proof step that concludes "t1 t2" from "u1 u2" and the two equalities "u1 = t1"
+  and "u2 = t2".
+
+  The atoms part of the ritem of t's representative must be a certificate of an already
+  assumed predicate for otherwise there would be no explanation for t.
+*)
+
+fun explain_pred lit t t1 t2 ({repr, rdata, apps, trace, prf, back}: context) =
+  (case ritem_of_pred rdata (repr_of repr t) of
+    {atoms=Cert (cert as (lit', _)), ...}: ritem =>
+      let
+        val (u1, u2) = dest_app (Argo_Lit.term_of lit')
+        val (lits, (p, prf)) = proof_of cert [] prf
+        val (lits, (p1, prf)) = mk_eq_proof trace u1 t1 lits prf
+        val (lits, (p2, prf)) = mk_eq_proof trace u2 t2 lits prf
+        val (lits, (p, prf)) = Argo_Proof.mk_subst p p1 p2 prf |> close_proof lit lits
+      in ((lits, p), mk_context repr rdata apps trace prf back) end
+  | _ => raise Fail "no explanation for bad predicate")
+
+(*
+  Explanations are produced based on the proof forest that is constructed while assuming new
+  information and propagating this among the internal data structures.
+  
+  For predicates, no distinction between both polarities needs to be done here. The atoms
+  part of the relevant ritem knows the assumed polarity.
+*)
+
+fun explain (lit as Argo_Lit.Pos (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2]))) cx =
+      SOME (explain_eq lit t1 t2 cx)
+  | explain (lit as Argo_Lit.Neg (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2]))) cx =
+      let val (_, eq) = the (find_neq cx (ritem_of' cx (repr_of' cx t1)) (repr_of' cx t2))
+      in SOME (explain_neq (Flat ((lit, NONE), (t1, t2))) eq cx) end
+  | explain (lit as (Argo_Lit.Pos (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])))) cx =
+      SOME (explain_pred lit t t1 t2 cx)
+  | explain (lit as (Argo_Lit.Neg (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])))) cx =
+      SOME (explain_pred lit t t1 t2 cx)
+  | explain _ _ = NONE
+
+
+(* backtracking *)
+
+(*
+  All information that needs to be reconstructed on backtracking is stored on the backtracking
+  stack. On backtracking any current information is replaced by what was stored before. No copying
+  nor subtle updates are required thanks to immutable data structures.
+*)
+
+fun add_level ({repr, rdata, apps, trace, prf, back}: context) =
+  mk_context repr rdata apps trace prf ((repr, rdata, apps, trace) :: back)
+
+fun backtrack ({back=[], ...}: context) = raise Empty
+  | backtrack ({prf, back=(repr, rdata, apps, trace) :: back, ...}: context) =
+      mk_context repr rdata apps trace prf back
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_cdcl.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,477 @@
+(*  Title:      Tools/Argo/argo_cdcl.ML
+    Author:     Sascha Boehme
+
+Propositional satisfiability solver in the style of conflict-driven
+clause-learning (CDCL). It features:
+
+ * conflict analysis and clause learning based on the first unique implication point
+ * nonchronological backtracking
+ * dynamic variable ordering (VSIDS)
+ * restarting
+ * polarity caching
+ * propagation via two watched literals
+ * special propagation of binary clauses 
+ * minimizing learned clauses
+ * support for external knowledge
+
+These features might be added:
+
+ * pruning of unnecessary learned clauses
+ * rebuilding the variable heap
+ * aligning the restart level with the decision heuristics: keep decisions that would
+   be recovered instead of backjumping to level 0
+
+The implementation is inspired by:
+
+  Niklas E'en and Niklas S"orensson. An Extensible SAT-solver. In Enrico
+  Giunchiglia and Armando Tacchella, editors, Theory and Applications of
+  Satisfiability Testing. Volume 2919 of Lecture Notes in Computer
+  Science, pages 502-518. Springer, 2003.
+
+  Niklas S"orensson and Armin Biere. Minimizing Learned Clauses. In
+  Oliver Kullmann, editor, Theory and Applications of Satisfiability
+  Testing. Volume 5584 of Lecture Notes in Computer Science,
+  pages 237-243. Springer, 2009.
+*)
+
+signature ARGO_CDCL =
+sig
+  (* types *)
+  type 'a explain = Argo_Lit.literal -> 'a -> Argo_Cls.clause * 'a
+
+  (* context *)
+  type context
+  val context: context
+  val assignment_of: context -> Argo_Lit.literal -> bool option
+
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> context
+  val add_axiom: Argo_Cls.clause -> context -> int * context
+
+  (* main operations *)
+  val assume: 'a explain -> Argo_Lit.literal -> context -> 'a ->
+    Argo_Cls.clause option * context * 'a
+  val propagate: context -> Argo_Common.literal Argo_Common.implied * context
+  val decide: context -> context option
+  val analyze: 'a explain -> Argo_Cls.clause -> context -> 'a -> int * context * 'a
+  val restart: context -> int * context
+end
+
+structure Argo_Cdcl: ARGO_CDCL =
+struct
+
+(* basic types and operations *)
+
+type 'a explain = Argo_Lit.literal -> 'a -> Argo_Cls.clause * 'a
+
+datatype reason =
+  Level0 of Argo_Proof.proof |
+  Decided of int * int * (bool * reason) Argo_Termtab.table |
+  Implied of int * int * (Argo_Lit.literal * reason) list * Argo_Proof.proof |
+  External of int
+
+fun level_of (Level0 _) = 0
+  | level_of (Decided (l, _, _)) = l
+  | level_of (Implied (l, _, _, _)) = l
+  | level_of (External l) = l
+
+type justified = Argo_Lit.literal * reason
+
+type watches = Argo_Cls.clause list * Argo_Cls.clause list
+
+fun get_watches wts t = Argo_Termtab.lookup wts t
+fun map_watches f t wts = Argo_Termtab.map_default (t, ([], [])) f wts
+
+fun map_lit_watches f (Argo_Lit.Pos t) = map_watches (apsnd f) t
+  | map_lit_watches f (Argo_Lit.Neg t) = map_watches (apfst f) t
+
+fun watches_of wts (Argo_Lit.Pos t) = (case get_watches wts t of SOME (ws, _) => ws | NONE => [])
+  | watches_of wts (Argo_Lit.Neg t) = (case get_watches wts t of SOME (_, ws) => ws | NONE => [])
+
+fun attach cls lit = map_lit_watches (cons cls) lit
+fun detach cls lit = map_lit_watches (remove Argo_Cls.eq_clause cls) lit
+
+
+(* literal values *)
+
+fun raw_val_of vals lit = Argo_Termtab.lookup vals (Argo_Lit.term_of lit)
+
+fun val_of vals (Argo_Lit.Pos t) = Argo_Termtab.lookup vals t
+  | val_of vals (Argo_Lit.Neg t) = Option.map (apfst not) (Argo_Termtab.lookup vals t)
+
+fun value_of vals (Argo_Lit.Pos t) = Option.map fst (Argo_Termtab.lookup vals t)
+  | value_of vals (Argo_Lit.Neg t) = Option.map (not o fst) (Argo_Termtab.lookup vals t)
+
+fun justified vals lit = Option.map (pair lit o snd) (raw_val_of vals lit)
+fun the_reason_of vals lit = snd (the (raw_val_of vals lit))
+
+fun assign (Argo_Lit.Pos t) r = Argo_Termtab.update (t, (true, r))
+  | assign (Argo_Lit.Neg t) r = Argo_Termtab.update (t, (false, r))
+
+
+(* context *)
+
+type trail = int * justified list (* the trail height and the sequence of assigned literals *)
+
+type context = {
+  units: Argo_Common.literal list, (* the literals that await propagation *)
+  level: int, (* the decision level *)
+  trail: int * justified list, (* the trail height and the sequence of assigned literals *)
+  vals: (bool * reason) Argo_Termtab.table, (* mapping of terms to polarity and reason *)
+  wts: watches Argo_Termtab.table, (* clauses watched by terms *)
+  heap: Argo_Heap.heap, (* max-priority heap for decision heuristics *)
+  clss: Argo_Cls.table, (* information about clauses *)
+  prf: Argo_Proof.context} (* the proof context *)
+
+fun mk_context units level trail vals wts heap clss prf: context =
+  {units=units, level=level, trail=trail, vals=vals, wts=wts, heap=heap, clss=clss, prf=prf}
+
+val context =
+  mk_context [] 0 (0, []) Argo_Termtab.empty Argo_Termtab.empty Argo_Heap.heap
+    Argo_Cls.table Argo_Proof.cdcl_context
+
+fun drop_levels n (Decided (l, h, vals)) trail heap =
+      if l = n + 1 then ((h, trail), vals, heap) else drop_literal n trail heap
+  | drop_levels n _ tr heap = drop_literal n tr heap
+
+and drop_literal n ((lit, r) :: trail) heap = drop_levels n r trail (Argo_Heap.insert lit heap)
+  | drop_literal _ [] _ = raise Fail "bad trail"
+
+fun backjump_to new_level (cx as {level, trail=(_, tr), wts, heap, clss, prf, ...}: context) =
+  if new_level >= level then (0, cx)
+  else
+    let val (trail, vals, heap) = drop_literal (Integer.max 0 new_level) tr heap
+    in (level - new_level, mk_context [] new_level trail vals wts heap clss prf) end
+
+
+(* proofs *)
+
+fun tag_clause (lits, p) prf = Argo_Proof.mk_clause lits p prf |>> pair lits
+
+fun level0_unit_proof (lit, Level0 p') (p, prf) = Argo_Proof.mk_unit_res lit p p' prf
+  | level0_unit_proof _ _ = raise Fail "bad reason"
+
+fun level0_unit_proofs lrs p prf = fold level0_unit_proof lrs (p, prf)
+
+fun unsat ({vals, prf, ...}: context) (lits, p) =
+  let val lrs = map (fn lit => (lit, the_reason_of vals lit)) lits
+  in Argo_Proof.unsat (fst (level0_unit_proofs lrs p prf)) end
+
+
+(* literal operations *)
+
+fun push lit p reason prf ({units, level, trail=(h, tr), vals, wts, heap, clss, ...}: context) =
+  let val vals = assign lit reason vals
+  in mk_context ((lit, p) :: units) level (h + 1, (lit, reason) :: tr) vals wts heap clss prf end
+
+fun push_level0 lit p lrs (cx as {prf, ...}: context) =
+  let val (p, prf) = level0_unit_proofs lrs p prf
+  in push lit (SOME p) (Level0 p) prf cx end
+
+fun push_implied lit p lrs (cx as {level, trail=(h, _), prf, ...}: context) =
+  if level > 0 then push lit NONE (Implied (level, h, lrs, p)) prf cx
+  else push_level0 lit p lrs cx
+
+fun push_decided lit (cx as {level, trail=(h, _), vals, prf, ...}: context) =
+  push lit NONE (Decided (level, h, vals)) prf cx
+
+fun assignment_of ({vals, ...}: context) = value_of vals
+
+fun replace_watches old new cls ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  mk_context units level trail vals (attach cls new (detach cls old wts)) heap clss prf
+
+
+(* clause operations *)
+
+fun as_clause cls ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  let val (cls, prf) = tag_clause cls prf
+  in (cls, mk_context units level trail vals wts heap clss prf) end
+
+fun note_watches ([_, _], _) _ clss = clss
+  | note_watches cls lp clss = Argo_Cls.put_watches cls lp clss
+
+fun attach_clause lit1 lit2 (cls as (lits, _)) cx =
+  let
+    val {units, level, trail, vals, wts, heap, clss, prf}: context = cx
+    val wts = attach cls lit1 (attach cls lit2 wts)
+    val clss = note_watches cls (lit1, lit2) clss
+  in mk_context units level trail vals wts (fold Argo_Heap.count lits heap) clss prf end
+
+fun change_watches _ (false, _, _) cx = cx
+  | change_watches cls (true, l1, l2) ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+      mk_context units level trail vals wts heap (Argo_Cls.put_watches cls (l1, l2) clss) prf
+
+fun add_asserting lit lit' (cls as (_, p)) lrs cx =
+  attach_clause lit lit' cls (push_implied lit p lrs cx)
+
+(*
+  When learning a non-unit clause, the context is backtracked to the highest decision level
+  of the assigned literals.
+*)
+
+fun learn_clause _ ([lit], p) cx = backjump_to 0 cx ||> push_level0 lit p []
+  | learn_clause lrs (cls as (lits, _)) cx =
+      let
+        fun max_level (l, r) (ll as (_, lvl)) = if level_of r > lvl then (l, level_of r) else ll
+        val (lit, lvl) = fold max_level lrs (hd lits, 0)
+      in backjump_to lvl cx ||> add_asserting (hd lits) lit cls lrs end
+
+(*
+  An axiom with one unassigned literal and all remaining literals being assigned to
+  false is asserting. An axiom with all literals assigned to false on level 0 makes the
+  context unsatisfiable. An axiom with all literals assigned to false on higher levels
+  causes backjumping before the highest level, and then the axiom might be asserting if
+  only one literal is unassigned on that level.
+*)
+
+fun min lit i NONE = SOME (lit, i)
+  | min lit i (SOME (lj as (_, j))) = SOME (if i < j then (lit, i) else lj)
+
+fun level_ord ((_, r1), (_, r2)) = int_ord (level_of r2, level_of r1)
+fun add_max lr lrs = Ord_List.insert level_ord lr lrs
+
+fun part [] [] t us fs = (t, us, fs)
+  | part (NONE :: vs) (l :: ls) t us fs = part vs ls t (l :: us) fs
+  | part (SOME (true, r) :: vs) (l :: ls) t us fs = part vs ls (min l (level_of r) t) us fs
+  | part (SOME (false, r) :: vs) (l :: ls) t us fs = part vs ls t us (add_max (l, r) fs)
+  | part _ _ _ _ _ = raise Fail "mismatch between values and literals"
+
+fun backjump_add (lit, r) (lit', r') cls lrs cx =
+  let
+    val add =
+      if level_of r = level_of r' then attach_clause lit lit' cls
+      else add_asserting lit lit' cls lrs
+  in backjump_to (level_of r - 1) cx ||> add end
+
+fun analyze_axiom vs (cls as (lits, p), cx) =
+  (case part vs lits NONE [] [] of
+    (SOME (lit, lvl), [], []) =>
+      if lvl > 0 then backjump_to 0 cx ||> push_implied lit p [] else (0, cx)
+  | (SOME (lit, lvl), [], (lit', _) :: _) => (0, cx |> (lvl > 0) ? attach_clause lit lit' cls)
+  | (SOME (lit, lvl), lit' :: _, _) => (0, cx |> (lvl > 0) ? attach_clause lit lit' cls)
+  | (NONE, [], (_, Level0 _) :: _) => unsat cx cls
+  | (NONE, [], [(lit, _)]) => backjump_to 0 cx ||> push_implied lit p []
+  | (NONE, [], lrs as (lr :: lr' :: _)) => backjump_add lr lr' cls lrs cx
+  | (NONE, [lit], []) => backjump_to 0 cx ||> push_implied lit p []
+  | (NONE, [lit], lrs as (lit', _) :: _) => (0, add_asserting lit lit' cls lrs cx)
+  | (NONE, lit1 :: lit2 :: _, _) => (0, attach_clause lit1 lit2 cls cx)
+  | _ => raise Fail "bad clause")
+
+
+(* enriching the context *)
+
+fun add_atom t ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  let val heap = Argo_Heap.insert (Argo_Lit.Pos t) heap
+  in mk_context units level trail vals wts heap clss prf end
+
+fun add_axiom ([], p) _ = Argo_Proof.unsat p
+  | add_axiom (cls as (lits, _)) (cx as {vals, ...}: context) =
+      if has_duplicates Argo_Lit.eq_lit lits then raise Fail "clause with duplicate literals"
+      else if has_duplicates Argo_Lit.dual_lit lits then (0, cx)
+      else analyze_axiom (map (val_of vals) lits) (as_clause cls cx)
+
+
+(* external knowledge *)
+
+fun assume explain lit (cx as {level, vals, prf, ...}: context) x =
+  (case value_of vals lit of
+    SOME true => (NONE, cx, x)
+  | SOME false => 
+      let val (cls, x) = explain lit x
+      in if level = 0 then unsat cx cls else (SOME cls, cx, x) end
+  | NONE =>
+      if level = 0 then
+        let val ((lits, p), x) = explain lit x
+        in (NONE, push_level0 lit p (map_filter (justified vals) lits) cx, x) end
+      else (NONE, push lit NONE (External level) prf cx, x))
+
+
+(* propagation *)
+
+exception CONFLICT of Argo_Cls.clause * context
+
+fun order_lits_by lit (l1, l2) =
+  if Argo_Lit.eq_id (l1, lit) then (true, l2, l1) else (false, l1, l2)
+
+fun prop_binary (_, implied_lit, other_lit) (cls as (_, p)) (cx as {level, vals, ...}: context) =
+  (case value_of vals implied_lit of
+    NONE => push_implied implied_lit p [(other_lit, the_reason_of vals other_lit)] cx
+  | SOME true => cx
+  | SOME false => if level = 0 then unsat cx cls else raise CONFLICT (cls, cx))
+
+datatype next = Lit of Argo_Lit.literal | None of justified list
+
+fun with_non_false f l (SOME (false, r)) lrs = f ((l, r) :: lrs)
+  | with_non_false _ l _ _ = Lit l
+
+fun first_non_false _ _ [] lrs = None lrs
+  | first_non_false vals lit (l :: ls) lrs =
+      if Argo_Lit.eq_lit (l, lit) then first_non_false vals lit ls lrs
+      else with_non_false (first_non_false vals lit ls) l (val_of vals l) lrs
+
+fun prop_nary (lp as (_, lit1, lit2)) (cls as (lits, p)) (cx as {level, vals, ...}: context) =
+  let val v = value_of vals lit1
+  in
+    if v = SOME true then change_watches cls lp cx
+    else
+      (case first_non_false vals lit1 lits [] of
+        Lit lit2' => change_watches cls (true, lit1, lit2') (replace_watches lit2 lit2' cls cx)
+      | None lrs =>
+          if v = NONE then push_implied lit1 p lrs (change_watches cls lp cx)
+          else if level = 0 then unsat cx cls
+          else raise CONFLICT (cls, change_watches cls lp cx))
+  end
+
+fun prop_cls lit (cls as ([l1, l2], _)) cx = prop_binary (order_lits_by lit (l1, l2)) cls cx
+  | prop_cls lit cls (cx as {clss, ...}: context) =
+      prop_nary (order_lits_by lit (Argo_Cls.get_watches clss cls)) cls cx
+
+fun prop_lit (lp as (lit, _)) (lps, cx as {wts, ...}: context) =
+  (lp :: lps, fold (prop_cls lit) (watches_of wts lit) cx)
+
+fun prop lps (cx as {units=[], ...}: context) = (Argo_Common.Implied (rev lps), cx)
+  | prop lps ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+      fold_rev prop_lit units (lps, mk_context [] level trail vals wts heap clss prf) |-> prop
+
+fun propagate cx = prop [] cx
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+
+(* decisions *)
+
+(*
+  Decisions are based on an activity heuristics. The most active variable that is
+  still unassigned is chosen.
+*)
+
+fun decide ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  let
+    fun check NONE = NONE
+      | check (SOME (lit, heap)) =
+          if Argo_Termtab.defined vals (Argo_Lit.term_of lit) then check (Argo_Heap.extract heap)
+          else SOME (push_decided lit (mk_context units (level + 1) trail vals wts heap clss prf))
+  in check (Argo_Heap.extract heap) end
+
+
+(* conflict analysis and clause learning *)
+
+(*
+  Learned clauses often contain literals that are redundant, because they are
+  subsumed by other literals of the clause. By analyzing the implication graph beyond
+  the unique implication point, such redundant literals can be identified and hence
+  removed from the learned clause. Only literals occurring in the learned clause and
+  their reasons need to be analyzed.
+*)
+
+exception ESSENTIAL of unit
+
+fun history_ord ((h1, lit1, _), (h2, lit2, _)) =
+  if h1 < 0 andalso h2 < 0 then int_ord (apply2 Argo_Lit.signed_id_of (lit1, lit2))
+  else int_ord (h2, h1)
+
+fun rec_redundant stop (lit, Implied (lvl, h, lrs, p)) lps =
+      if stop lit lvl then lps
+      else fold (rec_redundant stop) lrs ((h, lit, p) :: lps)
+  | rec_redundant stop (lit, Decided (lvl, _, _)) lps =
+      if stop lit lvl then lps
+      else raise ESSENTIAL ()
+  | rec_redundant _ (lit, Level0 p) lps = ((~1, lit, p) :: lps)
+  | rec_redundant _ _ _ = raise ESSENTIAL ()
+
+fun redundant stop (lr as (lit, Implied (_, h, lrs, p))) (lps, essential_lrs) = (
+      (fold (rec_redundant stop) lrs ((h, lit, p) :: lps), essential_lrs)
+      handle ESSENTIAL () => (lps, lr :: essential_lrs))
+  | redundant _ lr (lps, essential_lrs) = (lps, lr :: essential_lrs)
+
+fun resolve_step (_, l, p') (p, prf) = Argo_Proof.mk_unit_res l p p' prf
+
+fun reduce lrs p prf =
+  let
+    val lits = map fst lrs
+    val levels = fold (insert (op =) o level_of o snd) lrs []
+    fun stop lit level =
+      if member Argo_Lit.eq_lit lits lit then true
+      else if member (op =) levels level then false
+      else raise ESSENTIAL ()
+
+    val (lps, lrs) = fold (redundant stop) lrs ([], [])
+  in (lrs, fold resolve_step (sort_distinct history_ord lps) (p, prf)) end
+
+(*
+  Literals that are candidates for the learned lemma are marked and unmarked while
+  traversing backwards through the trail. The last remaining marked literal is the first
+  unique implication point.
+*)
+
+fun unmark lit ms = remove Argo_Lit.eq_id lit ms
+fun marked ms lit = member Argo_Lit.eq_id ms lit
+
+(*
+  Whenever an implication is recorded, the reason for the false literals of the
+  asserting clause are known. It is reasonable to store this justification list as part
+  of the implication reason. Consequently, the implementation of conflict analysis can
+  benefit from this information, which does not need to be re-computed here.
+*)
+
+fun justification_for _ _ _ (Implied (_, _, lrs, p)) x = (lrs, p, x)
+  | justification_for explain vals lit (External _) x =
+      let val ((lits, p), x) = explain lit x
+      in (map_filter (justified vals) lits, p, x) end
+  | justification_for _ _ _ _ _ = raise Fail "bad reason"
+
+fun first_lit pred ((lr as (lit, _)) :: lrs) = if pred lit then (lr, lrs) else first_lit pred lrs
+  | first_lit _ _ = raise Empty
+
+(*
+  Beginning from the conflicting clause, the implication graph is traversed to the first
+  unique implication point. This breadth-first search is controlled by the topological order of
+  the trail, which is traversed backwards. While traversing through the trail, the conflict
+  literals of lower levels are collected to form the conflict lemma together with the unique
+  implication point. Conflict literals assigned on level 0 are excluded from the conflict lemma.
+  Conflict literals assigned on the current level are candidates for the first unique
+  implication point.
+*)
+
+fun analyze explain cls (cx as {level, trail, vals, wts, heap, clss, prf, ...}: context) x =
+  let
+    fun from_clause [] trail ms lrs h p prf x =
+          from_trail (first_lit (marked ms) trail) ms lrs h p prf x
+      | from_clause ((lit, r) :: clause_lrs) trail ms lrs h p prf x =
+          from_reason r lit clause_lrs trail ms lrs h p prf x
+ 
+    and from_reason (Level0 p') lit clause_lrs trail ms lrs h p prf x =
+          let val (p, prf) = Argo_Proof.mk_unit_res lit p p' prf
+          in from_clause clause_lrs trail ms lrs h p prf x end
+      | from_reason r lit clause_lrs trail ms lrs h p prf x =
+          if level_of r = level then
+            if marked ms lit then from_clause clause_lrs trail ms lrs h p prf x
+            else from_clause clause_lrs trail (lit :: ms) lrs (Argo_Heap.increase lit h) p prf x
+          else
+            let
+              val (lrs, h) =
+                if AList.defined Argo_Lit.eq_id lrs lit then (lrs, h)
+                else ((lit, r) :: lrs, Argo_Heap.increase lit h)
+            in from_clause clause_lrs trail ms lrs h p prf x end
+
+    and from_trail ((lit, _), _) [_] lrs h p prf x =
+          let val (lrs, (p, prf)) = reduce lrs p prf
+          in (Argo_Lit.negate lit :: map fst lrs, lrs, h, p, prf, x) end
+      | from_trail ((lit, r), trail) ms lrs h p prf x =
+          let
+            val (clause_lrs, p', x) = justification_for explain vals lit r x
+            val (p, prf) = Argo_Proof.mk_unit_res lit p' p prf
+          in from_clause clause_lrs trail (unmark lit ms) lrs h p prf x end
+
+    val (ls, p) = cls
+    val lrs = if level = 0 then unsat cx cls else map (fn l => (l, the_reason_of vals l)) ls
+    val (lits, lrs, heap, p, prf, x) = from_clause lrs (snd trail) [] [] heap p prf x
+    val heap = Argo_Heap.decay heap
+    val (levels, cx) = learn_clause lrs (lits, p) (mk_context [] level trail vals wts heap clss prf)
+  in (levels, cx, x) end
+
+
+(* restarting *)
+
+fun restart cx = backjump_to 0 cx
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_clausify.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,168 @@
+(*  Title:      Tools/Argo/argo_clausify.ML
+    Author:     Sascha Boehme
+
+Conversion of propositional formulas to definitional CNF.
+
+The clausification implementation is based on:
+
+  G. S. Tseitin. On the complexity of derivation in propositional
+  calculus.  In A. O. Slisenko (editor) Structures in Constructive
+  Mathematics and Mathematical Logic, Part II, Seminars in Mathematics,
+  pages 115-125. Steklov Mathematic Institute, 1968.
+
+  D. A. Plaisted and S. Greenbaum. A Structure-Preserving Clause Form
+  Translation. Journal of Symbolic Computation, 1986.
+
+  L. de Moura and N. Bj\orner. Proofs and Refutations, and Z3. In
+  P. Rudnicki and G. Sutcliffe and B. Konev and R. A. Schmidt and
+  S. Schulz (editors) International Workshop on the Implementation of
+  Logics. CEUR Workshop Proceedings, 2008.
+*)
+
+signature ARGO_CLAUSIFY =
+sig
+  val clausify: Argo_Rewr.context -> Argo_Expr.expr * Argo_Proof.proof ->
+    Argo_Proof.context * Argo_Core.context -> Argo_Proof.context * Argo_Core.context
+end
+
+structure Argo_Clausify: ARGO_CLAUSIFY =
+struct
+
+(* lifting of if-then-else expressions *)
+
+(*
+  It is assumed that expressions are free of if-then-else expressions whose then- and else-branch
+  have boolean type. Such if-then-else expressions can be rewritten to expressions using only
+  negation, conjunction and disjunction.
+
+  All other modules treat if-then-else expressions as constant expressions. They do not analyze or
+  decend into sub-expressions of an if-then-else expression.
+
+  Lifting an if-then-else expression (ite P t u) introduces two new clauses
+    (or (not P) (= (ite P t u) t)) and
+    (or P (= (ite P t u) u))
+*)
+
+fun ite_clause simp k es (eps, (prf, core)) =
+  let
+    val e = Argo_Expr.mk_or es
+    val (p, prf) = Argo_Proof.mk_taut k e prf 
+    val (ep, prf) = Argo_Rewr.with_proof (Argo_Rewr.args (Argo_Rewr.rewrite_top simp)) (e, p) prf
+  in (ep :: eps, (prf, core)) end
+
+fun check_ite simp t (e as Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) (eps, (prf, core)) =
+      (case Argo_Core.identify (Argo_Term.Term t) core of
+        (Argo_Term.Known _, core) => (eps, (prf, core))
+      | (Argo_Term.New _, core) =>
+          (eps, (prf, core))
+          |> ite_clause simp Argo_Proof.Taut_Ite_Then [Argo_Expr.mk_not e1, Argo_Expr.mk_eq e e2]
+          |> ite_clause simp Argo_Proof.Taut_Ite_Else [e1, Argo_Expr.mk_eq e e3])
+  | check_ite _ _ _ cx = cx
+
+fun lift_ites simp (t as Argo_Term.T (_, _, ts)) =
+  check_ite simp t (Argo_Term.expr_of t) #>
+  fold (lift_ites simp) ts
+
+
+(* tagged expressions and terms *)
+
+fun pos x = (true, x)
+fun neg x = (false, x)
+
+fun mk_lit true t = Argo_Lit.Pos t
+  | mk_lit false t = Argo_Lit.Neg t
+
+fun expr_of (true, t) = Argo_Term.expr_of t
+  | expr_of (false, t) = Argo_Expr.mk_not (Argo_Term.expr_of t)
+
+
+(* adding literals *)
+
+fun lit_for (polarity, x) (new_atoms, core) =
+  (case Argo_Core.add_atom x core of
+    (Argo_Term.Known t, core) => (mk_lit polarity t, (new_atoms, core))
+  | (Argo_Term.New t, core) => (mk_lit polarity t, (t :: new_atoms, core)))
+
+fun lit_of (Argo_Expr.E (Argo_Expr.Not, [e])) = lit_for (neg (Argo_Term.Expr e))
+  | lit_of e = lit_for (pos (Argo_Term.Expr e))
+
+fun lit_of' (pol, Argo_Term.T (_, Argo_Expr.Not, [t])) = lit_for (not pol, Argo_Term.Term t)
+  | lit_of' (pol, t) = lit_for (pol, Argo_Term.Term t)
+
+
+(* adding clauses *)
+
+fun add_clause f xs p (new_atoms, (prf, core)) =
+  let val (lits, (new_atoms, core)) = fold_map f xs (new_atoms, core)
+  in (new_atoms, (prf, Argo_Core.add_axiom (lits, p) core)) end
+
+fun simp_lit (e as Argo_Expr.E (Argo_Expr.Not, [Argo_Expr.E (Argo_Expr.Not, [e'])])) =
+      Argo_Rewr.rewr Argo_Proof.Rewr_Not_Not e' e
+  | simp_lit e = Argo_Rewr.keep e
+
+fun simp_clause (e as Argo_Expr.E (Argo_Expr.Or, _)) = Argo_Rewr.args simp_lit e
+  | simp_clause e = Argo_Rewr.keep e
+
+fun new_clause k ls (new_atoms, (prf, core)) =
+  let
+    val e = Argo_Expr.mk_or (map expr_of ls)
+    val (p, prf) = Argo_Proof.mk_taut k e prf
+    val ((_, p), prf) = Argo_Rewr.with_proof simp_clause (e, p) prf
+  in add_clause lit_of' ls p (new_atoms, (prf, core)) end
+
+
+(* clausifying propositions *)
+
+fun clausify_and t ts cx =
+  let
+    val n = length ts
+    val k1 = Argo_Proof.Taut_And_1 n and k2 = Argo_Proof.Taut_And_2 o rpair n
+  in
+    cx
+    |> new_clause k1 (pos t :: map neg ts)
+    |> fold_index (fn (i, t') => new_clause (k2 i) [neg t, pos t']) ts
+  end
+
+fun clausify_or t ts cx =
+  let
+    val n = length ts
+    val k1 = Argo_Proof.Taut_Or_1 o rpair n and k2 = Argo_Proof.Taut_Or_2 n
+  in
+    cx
+    |> fold_index (fn (i, t') => new_clause (k1 i) [pos t, neg t']) ts
+    |> new_clause k2 (neg t :: map pos ts)
+  end
+
+fun clausify_iff t t1 t2 cx =
+  cx
+  |> new_clause Argo_Proof.Taut_Iff_1 [pos t, pos t1, pos t2]
+  |> new_clause Argo_Proof.Taut_Iff_2 [pos t, neg t1, neg t2]
+  |> new_clause Argo_Proof.Taut_Iff_3 [neg t, neg t1, pos t2]
+  |> new_clause Argo_Proof.Taut_Iff_4 [neg t, pos t1, neg t2]
+
+fun clausify_lit (t as Argo_Term.T (_, Argo_Expr.And, ts)) = clausify_and t ts
+  | clausify_lit (t as Argo_Term.T (_, Argo_Expr.Or, ts)) = clausify_or t ts
+  | clausify_lit (t as Argo_Term.T (_, Argo_Expr.Iff, [t1, t2])) = clausify_iff t t1 t2
+  | clausify_lit _ = I
+
+fun exhaust_new_atoms ([], cx) = cx
+  | exhaust_new_atoms (t :: new_atoms, cx) = exhaust_new_atoms (clausify_lit t (new_atoms, cx))
+
+fun clausify_expr _ (Argo_Expr.E (Argo_Expr.True, _), _) cx = cx
+  | clausify_expr _ (Argo_Expr.E (Argo_Expr.False, _), p) _ = Argo_Proof.unsat p
+  | clausify_expr f (Argo_Expr.E (Argo_Expr.And, es), p) cx =
+      fold_index (clausify_conj f (length es) p) es cx
+  | clausify_expr f (Argo_Expr.E (Argo_Expr.Or, es), p) cx = add_clausify f es p cx
+  | clausify_expr f (e, p) cx = add_clausify f [e] p cx
+
+and clausify_conj f n p (i, e) (prf, core) =
+  let val (p, prf) = Argo_Proof.mk_conj i n p prf
+  in clausify_expr f (e, p) (prf, core) end
+
+and add_clausify f es p cx =
+  let val ecx as (new_atoms, _) = add_clause lit_of es p ([], cx)
+  in fold f new_atoms ([], exhaust_new_atoms ecx) |-> fold (clausify_expr (K I)) end
+
+fun clausify simp ep cx = clausify_expr (lift_ites simp) ep cx
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_cls.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,45 @@
+(*  Title:      Tools/Argo/argo_cls.ML
+    Author:     Sascha Boehme
+
+Representation of clauses. Clauses are disjunctions of literals with a proof that explains
+why the disjunction holds.
+*)
+
+signature ARGO_CLS =
+sig
+  type clause = Argo_Lit.literal list * Argo_Proof.proof
+  val eq_clause: clause * clause -> bool
+
+  (* two-literal watches for clauses *)
+  type table
+  val table: table
+  val put_watches: clause -> Argo_Lit.literal * Argo_Lit.literal -> table -> table
+  val get_watches: table -> clause -> Argo_Lit.literal * Argo_Lit.literal
+end
+
+structure Argo_Cls: ARGO_CLS =
+struct
+
+type clause = Argo_Lit.literal list * Argo_Proof.proof
+
+fun eq_clause ((_, p1), (_, p2)) = Argo_Proof.eq_proof_id (apply2 Argo_Proof.id_of (p1, p2))
+fun clause_ord ((_, p1), (_, p2)) = Argo_Proof.proof_id_ord (apply2 Argo_Proof.id_of (p1, p2))
+
+
+(* two-literal watches for clauses *)
+
+(*
+  The CDCL solver keeps a mapping of some literals to clauses. Exactly two literals
+  of a clause are used to index the clause.
+*)
+
+structure Clstab = Table(type key = clause val ord = clause_ord)
+
+type table = (Argo_Lit.literal * Argo_Lit.literal) Clstab.table
+
+val table: table = Clstab.empty
+
+fun put_watches cls lp table = Clstab.update (cls, lp) table
+fun get_watches table cls = the (Clstab.lookup table cls)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_common.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,25 @@
+(*  Title:      Tools/Argo/argo_common.ML
+    Author:     Sascha Boehme
+
+Common infrastructure for the decision procedures of Argo.
+*)
+
+signature ARGO_COMMON =
+sig
+  type literal = Argo_Lit.literal * Argo_Proof.proof option
+  datatype 'a implied = Implied of 'a list | Conflict of Argo_Cls.clause
+end
+
+structure Argo_Common: ARGO_COMMON =
+struct
+
+type literal = Argo_Lit.literal * Argo_Proof.proof option
+  (* Implied new knowledge accompanied with an optional proof. If there is no proof,
+     the literal is to be treated hypothetically. If there is a proof, the literal is
+     to be treated as uni clause. *)
+
+datatype 'a implied = Implied of 'a list | Conflict of Argo_Cls.clause
+  (* A result of a step of a decision procedure, either an implication of new knowledge
+     or clause whose literals are known to be false. *)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_core.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,159 @@
+(*  Title:      Tools/Argo/argo_core.ML
+    Author:     Sascha Boehme
+
+Core of the Argo theorem prover implementing the DPLL(T) loop.
+
+The implementation is based on:
+
+  Harald Ganzinger, George Hagen, Robert Nieuwenhuis, Albert Oliveras,
+  Cesare Tinelli. DPLL(T): Fast decision procedures. In Lecture Notes in
+  Computer Science, volume 3114, pages 175-188. Springer, 2004.
+
+  Robert Nieuwenhuis, Albert Oliveras, Cesare Tinelli. Solving SAT and
+  SAT modulo theories: From an abstract Davis-Putnam-Logemann-Loveland
+  procedure to DPLL(T). In Journal of the ACM, volume 53(6), pages
+  937-977.  ACM, 2006.
+*)
+
+signature ARGO_CORE =
+sig
+  (* context *)
+  type context
+  val context: context
+
+  (* enriching the context *)
+  val identify: Argo_Term.item -> context -> Argo_Term.identified * context
+  val add_atom: Argo_Term.item -> context -> Argo_Term.identified * context
+  val add_axiom: Argo_Cls.clause -> context -> context
+
+  (* DPLL(T) loop *)
+  val run: context -> context (* raises Argo_Proof.UNSAT *)
+
+  (* model *)
+  val model_of: context -> string * Argo_Expr.typ -> bool option
+end
+
+structure Argo_Core: ARGO_CORE =
+struct
+
+(* context *)
+
+type context = {
+  terms: Argo_Term.context, (* the term context to identify equal expressions *)
+  iter: int, (* the current iteration of the search *)
+  cdcl: Argo_Cdcl.context, (* the context of the propositional solver *)
+  thy: Argo_Thy.context} (* the context of the theory solver *)
+
+fun mk_context terms iter cdcl thy: context = {terms=terms, iter=iter, cdcl=cdcl, thy=thy}
+
+val context = mk_context Argo_Term.context 1 Argo_Cdcl.context Argo_Thy.context
+
+fun backjump levels = funpow levels Argo_Thy.backtrack
+
+
+(* enriching the context *)
+
+fun identify i ({terms, iter, cdcl, thy}: context) =
+  let val (identified, terms) = Argo_Term.identify_item i terms
+  in (identified, mk_context terms iter cdcl thy) end
+
+fun add_atom i cx =
+  (case identify i cx of
+    known as (Argo_Term.Known _, _) => known
+  | (atom as Argo_Term.New t, {terms, iter, cdcl, thy}: context) =>
+      (case (Argo_Cdcl.add_atom t cdcl, Argo_Thy.add_atom t thy) of
+        (cdcl, (NONE, thy)) => (atom, mk_context terms iter cdcl thy)
+      | (cdcl, (SOME lit, thy)) =>
+          (case Argo_Cdcl.assume Argo_Thy.explain lit cdcl thy of
+            (NONE, cdcl, thy) => (atom, mk_context terms iter cdcl thy)
+          | (SOME _, _, _) => raise Fail "bad conflict with new atom")))
+
+fun add_axiom cls ({terms, iter, cdcl, thy}: context) =
+  let val (levels, cdcl) = Argo_Cdcl.add_axiom cls cdcl
+  in mk_context terms iter cdcl (backjump levels thy) end
+
+
+(* DPLL(T) loop: CDCL with theories *)
+
+datatype implications = None | Implications | Conflict of Argo_Cls.clause
+
+fun cdcl_assume [] cdcl thy = (NONE, cdcl, thy)
+  | cdcl_assume (lit :: lits) cdcl thy =
+      (* assume an assignment deduced by the theory solver *)
+      (case Argo_Cdcl.assume Argo_Thy.explain lit cdcl thy of
+        (NONE, cdcl, thy) => cdcl_assume lits cdcl thy
+      | (SOME cls, cdcl, thy) => (SOME cls, cdcl, thy))
+
+fun theory_deduce _ (conflict as (Conflict _, _, _)) = conflict
+  | theory_deduce f (result, cdcl, thy) =
+      (case f thy of
+        (Argo_Common.Implied [], thy) => (result, cdcl, thy)
+      | (Argo_Common.Implied lits, thy) => 
+          (* turn all implications of the theory solver into propositional assignments *)
+          (case cdcl_assume lits cdcl thy of
+            (NONE, cdcl, thy) => (Implications, cdcl, thy)
+          | (SOME cls, cdcl, thy) => (Conflict cls, cdcl, thy))
+      | (Argo_Common.Conflict cls, thy) => (Conflict cls, cdcl, thy))
+
+fun theory_assume [] cdcl thy = (None, cdcl, thy)
+  | theory_assume lps cdcl thy =
+      (None, cdcl, thy)
+      (* propagate all propositional implications to the theory solver *)
+      |> fold (theory_deduce o Argo_Thy.assume) lps
+      (* check the consistency of the theory model *)
+      |> theory_deduce Argo_Thy.check
+
+fun search limit cdcl thy =
+  (* collect all propositional implications of the last assignments *)
+  (case Argo_Cdcl.propagate cdcl of
+    (Argo_Common.Implied lps, cdcl) =>
+      (* propagate all propositional implications to the theory solver *)
+      (case theory_assume lps cdcl thy of
+        (None, cdcl, thy) =>
+          (* stop searching if the conflict limit has been exceeded *)
+          if limit <= 0 then (false, cdcl, thy)
+          else
+            (* no further propositional assignments, choose a value for the next unassigned atom *)
+            (case Argo_Cdcl.decide cdcl of
+              NONE => (true, cdcl, thy) (* the context is satisfiable *)
+            | SOME cdcl => search limit cdcl (Argo_Thy.add_level thy))
+      | (Implications, cdcl, thy) => search limit cdcl thy
+      | (Conflict ([], p), _, _) => Argo_Proof.unsat p
+      | (Conflict cls, cdcl, thy) => analyze cls limit cdcl thy)
+  | (Argo_Common.Conflict cls, cdcl) => analyze cls limit cdcl thy)
+
+and analyze cls limit cdcl thy =
+  (* analyze the conflict, probably using lazy explanations from the theory solver *)
+  let val (levels, cdcl, thy) = Argo_Cdcl.analyze Argo_Thy.explain cls cdcl thy
+  in search (limit - 1) cdcl (backjump levels thy) end
+
+fun luby_number i =
+  let
+    fun mult p = if p < i + 1 then mult (2 * p) else p
+    val p = mult 2
+  in if i = p - 1 then p div 2 else luby_number (i - (p div 2) + 1) end
+
+fun next_restart_limit iter = 100 * luby_number iter
+
+fun loop iter cdcl thy =
+  (* perform a limited search that is stopped after a certain number of conflicts *)
+  (case search (next_restart_limit iter) cdcl thy of
+    (true, cdcl, thy) => (iter + 1, cdcl, thy)
+  | (false, cdcl, thy) =>
+      (* restart the solvers to avoid that they get stuck in a fruitless search *)
+      let val (levels, cdcl) = Argo_Cdcl.restart cdcl
+      in loop (iter + 1) cdcl (backjump levels thy) end)
+
+fun run ({terms, iter, cdcl, thy}: context) =
+  let val (iter, cdcl, thy) = loop iter cdcl (Argo_Thy.prepare thy)
+  in mk_context terms iter cdcl thy end
+
+
+(* model *)
+
+fun model_of ({terms, cdcl, ...}: context) c =
+  (case Argo_Term.identify_item (Argo_Term.Expr (Argo_Expr.E (Argo_Expr.Con c, []))) terms of
+    (Argo_Term.Known t, _) => Argo_Cdcl.assignment_of cdcl (Argo_Lit.Pos t)
+  | (Argo_Term.New _, _) => NONE)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_expr.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,244 @@
+(*  Title:      Tools/Argo/sid_expr.ML
+    Author:     Sascha Boehme
+
+The input language of the Argo solver.
+*)
+
+signature ARGO_EXPR =
+sig
+  (* data types *)
+  datatype typ = Bool | Real | Func of typ * typ | Type of string
+  datatype kind =
+    True | False | Not | And | Or | Imp | Iff | Ite | Eq | App | Con of string * typ |
+    Le | Lt | Num of Rat.rat | Neg | Add | Sub | Mul | Div | Min | Max | Abs
+  datatype expr = E of kind * expr list
+
+  (* indices, equalities, orders *)
+  val int_of_kind: kind -> int
+  val con_ord: (string * typ) * (string * typ) -> order
+  val eq_kind: kind * kind -> bool
+  val kind_ord: kind * kind -> order
+  val eq_expr: expr * expr -> bool
+  val expr_ord: expr * expr -> order
+  val dual_expr: expr -> expr -> bool
+
+  (* constructors *)
+  val kind_of_string: string -> kind
+  val true_expr: expr
+  val false_expr: expr
+  val mk_not: expr -> expr
+  val mk_and: expr list -> expr
+  val mk_and2: expr -> expr -> expr
+  val mk_or: expr list -> expr
+  val mk_or2: expr -> expr -> expr
+  val mk_imp: expr -> expr -> expr
+  val mk_iff: expr -> expr -> expr
+  val mk_ite: expr -> expr -> expr -> expr
+  val mk_eq: expr -> expr -> expr
+  val mk_app: expr -> expr -> expr
+  val mk_con: string * typ -> expr
+  val mk_le: expr -> expr -> expr
+  val mk_lt: expr -> expr -> expr
+  val mk_num: Rat.rat -> expr
+  val mk_neg: expr -> expr
+  val mk_add: expr list -> expr
+  val mk_add2: expr -> expr -> expr
+  val mk_sub: expr -> expr -> expr
+  val mk_mul: expr -> expr -> expr
+  val mk_div: expr -> expr -> expr
+  val mk_min: expr -> expr -> expr
+  val mk_max: expr -> expr -> expr
+  val mk_abs: expr -> expr
+
+  (* type checking *)
+  exception TYPE of expr
+  exception EXPR of expr
+  val type_of: expr -> typ (* raises EXPR *)
+  val check: expr -> bool (* raises TYPE and EXPR *)
+end
+
+structure Argo_Expr: ARGO_EXPR =
+struct
+
+(* data types *)
+
+datatype typ = Bool | Real | Func of typ * typ | Type of string
+
+datatype kind =
+  True | False | Not | And | Or | Imp | Iff | Ite | Eq | App | Con of string * typ |
+  Le | Lt | Num of Rat.rat | Neg | Add | Sub | Mul | Div | Min | Max | Abs
+
+datatype expr = E of kind * expr list
+
+
+(* indices, equalities, orders *)
+
+fun int_of_type Bool = 0
+  | int_of_type Real = 1
+  | int_of_type (Func _) = 2
+  | int_of_type (Type _) = 3
+
+fun int_of_kind True = 0
+  | int_of_kind False = 1
+  | int_of_kind Not = 2
+  | int_of_kind And = 3
+  | int_of_kind Or = 4
+  | int_of_kind Imp = 5
+  | int_of_kind Iff = 6
+  | int_of_kind Ite = 7
+  | int_of_kind Eq = 8
+  | int_of_kind App = 9
+  | int_of_kind (Con _) = 10
+  | int_of_kind Le = 11
+  | int_of_kind Lt = 12
+  | int_of_kind (Num _) = 13
+  | int_of_kind Neg = 14
+  | int_of_kind Add = 15
+  | int_of_kind Sub = 16
+  | int_of_kind Mul = 17
+  | int_of_kind Div = 18
+  | int_of_kind Min = 19
+  | int_of_kind Max = 20
+  | int_of_kind Abs = 21
+
+fun eq_type (Bool, Bool) = true
+  | eq_type (Real, Real) = true
+  | eq_type (Func tys1, Func tys2) = eq_pair eq_type eq_type (tys1, tys2)
+  | eq_type (Type n1, Type n2) = (n1 = n2)
+  | eq_type _ = false
+
+fun type_ord (Bool, Bool) = EQUAL
+  | type_ord (Real, Real) = EQUAL
+  | type_ord (Type n1, Type n2) = fast_string_ord (n1, n2)
+  | type_ord (Func tys1, Func tys2) = prod_ord type_ord type_ord (tys1, tys2)
+  | type_ord (ty1, ty2) = int_ord (int_of_type ty1, int_of_type ty2)
+
+fun eq_con cp = eq_pair (op =) eq_type cp
+fun con_ord cp = prod_ord fast_string_ord type_ord cp
+
+fun eq_kind (Con c1, Con c2) = eq_con (c1, c2)
+  | eq_kind (Num n1, Num n2) = n1 = n2
+  | eq_kind (k1, k2) = (k1 = k2)
+
+fun kind_ord (Con c1, Con c2) = con_ord (c1, c2)
+  | kind_ord (Num n1, Num n2) = Rat.ord (n1, n2)
+  | kind_ord (k1, k2) = int_ord (int_of_kind k1, int_of_kind k2)
+
+fun eq_expr (E e1, E e2) = eq_pair eq_kind (eq_list eq_expr) (e1, e2)
+fun expr_ord (E e1, E e2) = prod_ord kind_ord (list_ord expr_ord) (e1, e2)
+
+fun dual_expr (E (Not, [e1])) e2 = eq_expr (e1, e2)
+  | dual_expr e1 (E (Not, [e2])) = eq_expr (e1, e2)
+  | dual_expr _ _ = false
+
+
+(* constructors *)
+
+val kind_of_string = the o Symtab.lookup (Symtab.make [
+  ("true", True),("false", False), ("not", Not), ("and", And), ("or", Or), ("imp", Imp),
+  ("iff", Iff), ("ite", Ite), ("eq", Eq), ("app", App), ("le", Le), ("lt", Lt), ("neg", Neg),
+  ("add", Add), ("sub", Sub), ("mul", Mul), ("div", Div), ("min", Min), ("max", Max), ("abs", Abs)])
+
+val true_expr = E (True, [])
+val false_expr = E (False, [])
+fun mk_not e = E (Not, [e])
+fun mk_and es = E (And, es)
+fun mk_and2 e1 e2 = mk_and [e1, e2]
+fun mk_or es = E (Or, es)
+fun mk_or2 e1 e2 = mk_or [e1, e2]
+fun mk_imp e1 e2 = E (Imp, [e1, e2])
+fun mk_iff e1 e2 = E (Iff, [e1, e2])
+fun mk_ite e1 e2 e3 = E (Ite, [e1, e2, e3])
+fun mk_eq e1 e2 = E (Eq, [e1, e2])
+fun mk_app e1 e2 = E (App, [e1, e2])
+fun mk_con n = E (Con n, [])
+fun mk_le e1 e2 = E (Le, [e1, e2])
+fun mk_lt e1 e2 = E (Lt, [e1, e2])
+fun mk_num r = E (Num r, [])
+fun mk_neg e = E (Neg, [e])
+fun mk_add es = E (Add, es)
+fun mk_add2 e1 e2 = mk_add [e1, e2]
+fun mk_sub e1 e2 = E (Sub, [e1, e2])
+fun mk_mul e1 e2 = E (Mul, [e1, e2])
+fun mk_div e1 e2 = E (Div, [e1, e2])
+fun mk_min e1 e2 = E (Min, [e1, e2])
+fun mk_max e1 e2 = E (Max, [e1, e2])
+fun mk_abs e = E (Abs, [e])
+
+
+(* type checking *)
+
+exception TYPE of expr
+exception EXPR of expr
+
+fun dest_func_type _ (Func tys) = tys
+  | dest_func_type e _ = raise TYPE e
+
+fun type_of (E (True, _)) = Bool
+  | type_of (E (False, _)) = Bool
+  | type_of (E (Not, _)) = Bool
+  | type_of (E (And, _)) = Bool
+  | type_of (E (Or, _)) = Bool
+  | type_of (E (Imp, _)) = Bool
+  | type_of (E (Iff, _)) = Bool
+  | type_of (E (Ite, [_, e, _])) = type_of e
+  | type_of (E (Eq, _)) = Bool
+  | type_of (E (App, [e, _])) = snd (dest_func_type e (type_of e))
+  | type_of (E (Con (_, ty), _)) = ty
+  | type_of (E (Le, _)) = Bool
+  | type_of (E (Lt, _)) = Bool
+  | type_of (E (Num _, _)) = Real
+  | type_of (E (Neg, _)) = Real
+  | type_of (E (Add, _)) = Real
+  | type_of (E (Sub, _)) = Real
+  | type_of (E (Mul, _)) = Real
+  | type_of (E (Div, _)) = Real
+  | type_of (E (Min, _)) = Real
+  | type_of (E (Max, _)) = Real
+  | type_of (E (Abs, _)) = Real
+  | type_of e = raise EXPR e
+
+fun all_type ty (E (_, es)) = forall (curry eq_type ty o type_of) es
+val all_bool = all_type Bool
+val all_real = all_type Real
+
+(*
+  Types as well as proper arities are checked.
+  Exception TYPE is raised for invalid types.
+  Exception EXPR is raised for invalid expressions and invalid arities.
+*)
+
+fun check (e as E (_, es)) = (forall check es andalso raw_check e) orelse raise TYPE e
+
+and raw_check (E (True, [])) = true
+  | raw_check (E (False, [])) = true
+  | raw_check (e as E (Not, [_])) = all_bool e
+  | raw_check (e as E (And, _ :: _)) = all_bool e
+  | raw_check (e as E (Or, _ :: _)) = all_bool e
+  | raw_check (e as E (Imp, [_, _])) = all_bool e
+  | raw_check (e as E (Iff, [_, _])) = all_bool e
+  | raw_check (E (Ite, [e1, e2, e3])) =
+      let val ty1 = type_of e1 and ty2 = type_of e2 and ty3 = type_of e3
+      in eq_type (ty1, Bool) andalso eq_type (ty2, ty3) end
+  | raw_check (E (Eq, [e1, e2])) =
+      let val ty1 = type_of e1 and ty2 = type_of e2
+      in eq_type (ty1, ty2) andalso not (eq_type (ty1, Bool)) end
+  | raw_check (E (App, [e1, e2])) =
+      eq_type (fst (dest_func_type e1 (type_of e1)), type_of e2)
+  | raw_check (E (Con _, [])) = true
+  | raw_check (E (Num _, [])) = true
+  | raw_check (e as E (Le, [_, _])) = all_real e
+  | raw_check (e as E (Lt, [_, _])) = all_real e
+  | raw_check (e as E (Neg, [_])) = all_real e
+  | raw_check (e as E (Add, _)) = all_real e
+  | raw_check (e as E (Sub, [_, _])) = all_real e
+  | raw_check (e as E (Mul, [_, _])) = all_real e
+  | raw_check (e as E (Div, [_, _])) = all_real e
+  | raw_check (e as E (Min, [_, _])) = all_real e
+  | raw_check (e as E (Max, [_, _])) = all_real e
+  | raw_check (e as E (Abs, [_])) = all_real e
+  | raw_check e = raise EXPR e
+
+end
+
+structure Argo_Exprtab = Table(type key = Argo_Expr.expr val ord = Argo_Expr.expr_ord)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_heap.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,186 @@
+(*  Title:      Tools/Argo/argo_heap.ML
+    Author:     Sascha Boehme
+
+A maximum-priority heap for literals with integer priorities and with inverse indices.
+The heap is intended to be used as VSIDS-like decision heuristics. This implementation
+is based on pairing heaps described in:
+
+  Chris Okasaki. Purely Functional Data Structures. Chapter 5.
+  Cambridge University Press, 1998.
+*)
+
+signature ARGO_HEAP =
+sig
+  type heap
+  val heap: heap
+  val insert: Argo_Lit.literal -> heap -> heap
+  val extract: heap -> (Argo_Lit.literal * heap) option
+  val increase: Argo_Lit.literal -> heap -> heap
+  val count: Argo_Lit.literal -> heap -> heap
+  val decay: heap -> heap
+  val rebuild: (Argo_Term.term -> bool) -> heap -> heap
+end
+
+structure Argo_Heap: ARGO_HEAP =
+struct
+
+(* heuristic activity constants *)
+
+val min_incr = 128
+fun decay_incr i = (i * 11) div 10
+val max_activity = Integer.pow 24 2
+val activity_rescale = Integer.pow 14 2
+
+
+(* data structures and basic operations *)
+
+datatype tree = E | T of Argo_Term.term * bool * tree list
+
+datatype parent = None | Root | Some of Argo_Term.term
+
+type heap = {
+  incr: int, (* the increment to apply in an increase operation *)
+  vals: ((int * int) * parent) Argo_Termtab.table, (* weights and parents of the stored terms *)
+  tree: tree} (* the pairing heap of literals; note: the tree caches literal polarities *)
+
+fun mk_heap incr vals tree: heap = {incr=incr, vals=vals, tree=tree}
+fun mk_heap' incr (tree, vals) = mk_heap incr vals tree
+
+val heap = mk_heap min_incr Argo_Termtab.empty E
+
+val empty_value = ((0, 0), None)
+fun value_of vals t = the_default empty_value (Argo_Termtab.lookup vals t)
+fun map_value t = Argo_Termtab.map_default (t, empty_value)
+
+
+(* weight operations *)
+
+(*
+  The weight of a term is a pair of activity and count. The activity describes how
+  often a term participated in conflicts. The count describes how often a term occurs
+  in clauses.
+*)
+
+val weight_ord = prod_ord int_ord int_ord
+
+fun weight_of vals t = fst (value_of vals t)
+
+fun less_than vals t1 t2 = weight_ord (weight_of vals t1, weight_of vals t2) = LESS
+
+fun rescale activity = activity div activity_rescale
+
+fun incr_activity incr t = map_value t (apfst (apfst (Integer.add incr)))
+fun incr_count t = map_value t (apfst (apsnd (Integer.add 1)))
+
+fun rescale_activities a incr vals =
+  if a <= max_activity then (incr, vals)
+  else (rescale incr, Argo_Termtab.map (fn _ => apfst (apfst rescale)) vals)
+
+
+(* reverse index operations *)
+
+(*
+  The reverse index is required to retrieve elements when increasing their priorities.
+*)
+
+fun contains vals t =
+  (case value_of vals t of
+    (_, None) => false
+  | _ => true)
+
+fun path_to vals t parents =
+  (case value_of vals t of
+    (_, Root) => parents
+  | (_, Some parent) => path_to vals parent (t :: parents)
+  | _ => raise Fail "bad heap")
+
+fun put_parent t parent = map_value t (apsnd (K parent))
+fun delete t = put_parent t None
+fun root t = put_parent t Root
+
+fun as_root (tree as T (t, _, _), vals) = (tree, root t vals)
+  | as_root x = x
+
+
+(* pairing heap operations *)
+
+fun merge E tree vals = (tree, vals)
+  | merge tree E vals = (tree, vals)
+  | merge (tree1 as T (t1, p1, trees1)) (tree2 as T (t2, p2, trees2)) vals =
+      if less_than vals t1 t2 then (T (t2, p2, tree1 :: trees2), put_parent t1 (Some t2) vals)
+      else (T (t1, p1, tree2 :: trees1), put_parent t2 (Some t1) vals)
+
+fun merge_pairs [] vals = (E, vals)
+  | merge_pairs [tree] vals = (tree, vals)
+  | merge_pairs (tree1 :: tree2 :: trees) vals =
+      vals |> merge tree1 tree2 ||>> merge_pairs trees |-> uncurry merge
+
+
+(* cutting subtrees specified by a path *)
+
+(*
+  The extractions are performed in such a way that the heap is changed in as few positions
+  as possible.
+*)
+
+fun with_index f u ((tree as T (t, _, _)) :: trees) =
+      if Argo_Term.eq_term (t, u) then f tree ||> (fn E => trees | tree => tree :: trees)
+      else with_index f u trees ||> cons tree
+  | with_index _ _ _ = raise Fail "bad heap"
+
+fun lift_index f u (T (t, p, trees)) = with_index f u trees ||> (fn trees => T (t, p, trees))
+  | lift_index _ _ E = raise Fail "bad heap"
+
+fun cut t [] tree = lift_index (fn tree => (tree, E)) t tree
+  | cut t (parent :: ts) tree = lift_index (cut t ts) parent tree
+
+
+(* filtering the heap *)
+
+val proper_trees = filter (fn E => false | T _ => true)
+
+fun filter_tree _ E vals = (E, vals)
+  | filter_tree pred (T (t, p, ts)) vals =
+      let val (ts, vals) = fold_map (filter_tree pred) ts vals |>> proper_trees
+      in if pred t then (T (t, p, ts), vals) else merge_pairs ts (delete t vals) end
+
+
+(* exported heap operations *)
+
+fun insert lit (h as {incr, vals, tree}: heap) = 
+  let val (t, p) = Argo_Lit.dest lit
+  in if contains vals t then h else mk_heap' incr (merge tree (T (t, p, [])) (root t vals)) end
+
+fun extract ({tree=E, ...}: heap) = NONE
+  | extract ({incr, vals, tree=T (t, p, ts)}: heap) =
+      SOME (Argo_Lit.literal t p, mk_heap' incr (as_root (merge_pairs ts (delete t vals))))
+
+fun with_term lit f = f (Argo_Lit.term_of lit)
+
+(*
+  If the changed weight violates the heap property, the corresponding tree
+  is extracted and merged with the root.
+*)
+
+fun fix t (w, Some parent) (incr, vals) tree =
+      if weight_ord (weight_of vals parent, w) = LESS then
+        let val (tree1, tree2) = cut t (path_to vals parent []) tree
+        in mk_heap' incr (merge tree1 tree2 (root t vals)) end
+      else mk_heap incr vals tree
+  | fix _ _ (incr, vals) tree = mk_heap incr vals tree
+
+fun increase lit ({incr, vals, tree}: heap) = with_term lit (fn t =>
+  let
+    val vals = incr_activity incr t vals
+    val value as ((a, _), _) = value_of vals t
+  in fix t value (rescale_activities a incr vals) tree end)
+
+fun count lit ({incr, vals, tree}: heap) = with_term lit (fn t =>
+  let val vals = incr_count t vals
+  in fix t (value_of vals t) (incr, vals) tree end)
+
+fun decay ({incr, vals, tree}: heap) = mk_heap (decay_incr incr) vals tree
+
+fun rebuild pred ({incr, vals, tree}: heap) = mk_heap' incr (filter_tree pred tree vals)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_lit.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,64 @@
+(*  Title:      Tools/Argo/argo_lit.ML
+    Author:     Sascha Boehme
+
+Representation of literals. Literals are terms with a polarity, either positive or negative.
+A literal for term t with positive polarity is equivalent to t.
+A literal for term t with negative polarity is equivalent to the propositional negation of t.
+*)
+
+signature ARGO_LIT =
+sig
+  datatype literal = Pos of Argo_Term.term | Neg of Argo_Term.term
+  val literal: Argo_Term.term -> bool -> literal
+  val dest: literal -> Argo_Term.term * bool
+  val term_of: literal -> Argo_Term.term
+  val signed_id_of: literal -> int
+  val signed_expr_of: literal -> Argo_Expr.expr
+  val negate: literal -> literal
+  val eq_id: literal * literal -> bool
+  val eq_lit: literal * literal -> bool
+  val dual_lit: literal * literal -> bool
+end
+
+structure Argo_Lit: ARGO_LIT =
+struct
+
+(* data type *)
+
+datatype literal = Pos of Argo_Term.term | Neg of Argo_Term.term
+
+
+(* operations *)
+
+fun literal t true = Pos t
+  | literal t false = Neg t
+
+fun dest (Pos t) = (t, true)
+  | dest (Neg t) = (t, false)
+
+fun term_of (Pos t) = t
+  | term_of (Neg t) = t
+
+fun signed_id_of (Pos t) = Argo_Term.id_of t
+  | signed_id_of (Neg t) = ~(Argo_Term.id_of t)
+
+fun signed_expr_of (Pos t) = Argo_Term.expr_of t
+  | signed_expr_of (Neg t) = Argo_Expr.mk_not (Argo_Term.expr_of t)
+
+fun id_of (Pos t) = Argo_Term.id_of t
+  | id_of (Neg t) = Argo_Term.id_of t
+
+fun negate (Pos t) = Neg t
+  | negate (Neg t) = Pos t
+
+fun eq_id (lit1, lit2) = (id_of lit1 = id_of lit2)
+
+fun eq_lit (Pos t1, Pos t2) = Argo_Term.eq_term (t1, t2)
+  | eq_lit (Neg t1, Neg t2) = Argo_Term.eq_term (t1, t2)
+  | eq_lit _ = false
+
+fun dual_lit (Pos t1, Neg t2) = Argo_Term.eq_term (t1, t2)
+  | dual_lit (Neg t1, Pos t2) = Argo_Term.eq_term (t1, t2)
+  | dual_lit _ = false
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_proof.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,387 @@
+(*  Title:      Tools/Argo/argo_proof.ML
+    Author:     Sascha Boehme
+
+The proof language of the Argo solver.
+
+Proofs trace the inferences of the solver. They can be used to check unsatisfiability results.
+
+The proof language is inspired by:
+
+  Leonardo  de  Moura  and  Nikolaj  Bj/orner. Proofs and Refutations, and Z3. In
+  Proceedings of the LPAR 2008 Workshops, Knowledge Exchange: Automated Provers and Proof
+  Assistants, and the 7th International Workshop on the Implementation of Logics,
+  volume 418 of CEUR Workshop Proceedings. CEUR-WS.org, 2008.
+*)
+
+signature ARGO_PROOF =
+sig
+  (* types *)
+  type proof_id
+  datatype tautology =
+    Taut_And_1 of int | Taut_And_2 of int * int | Taut_Or_1 of int * int | Taut_Or_2 of int |
+    Taut_Iff_1 | Taut_Iff_2 | Taut_Iff_3 | Taut_Iff_4 | Taut_Ite_Then | Taut_Ite_Else
+  datatype inequality = Le | Lt
+  datatype rewrite =
+    Rewr_Not_True | Rewr_Not_False | Rewr_Not_Not | Rewr_Not_And of int | Rewr_Not_Or of int |
+    Rewr_Not_Iff |
+    Rewr_And_False of int | Rewr_And_Dual of int * int | Rewr_And_Sort of int * int list list |
+    Rewr_Or_True of int | Rewr_Or_Dual of int * int | Rewr_Or_Sort of int * int list list |
+    Rewr_Iff_True | Rewr_Iff_False | Rewr_Iff_Not_Not | Rewr_Iff_Refl | Rewr_Iff_Symm |
+    Rewr_Iff_Dual |
+    Rewr_Imp | Rewr_Ite_Prop | Rewr_Ite_True | Rewr_Ite_False | Rewr_Ite_Eq |
+    Rewr_Eq_Refl | Rewr_Eq_Symm |
+    Rewr_Neg | Rewr_Add of (Rat.rat * int option) list * (Rat.rat * int option) list | Rewr_Sub |
+    Rewr_Mul_Nums of Rat.rat * Rat.rat | Rewr_Mul_Zero | Rewr_Mul_One | Rewr_Mul_Comm |
+    Rewr_Mul_Assoc | Rewr_Mul_Sum | Rewr_Div_Nums of Rat.rat * Rat.rat | Rewr_Div_Zero |
+    Rewr_Div_One | Rewr_Div_Mul | Rewr_Div_Inv | Rewr_Div_Left | Rewr_Div_Right | Rewr_Min |
+    Rewr_Max | Rewr_Abs | Rewr_Eq_Le | Rewr_Ineq_Nums of inequality * bool |
+    Rewr_Ineq_Add of inequality * Rat.rat | Rewr_Ineq_Sub of inequality |
+    Rewr_Ineq_Mul of inequality * Rat.rat | Rewr_Not_Ineq of inequality
+  datatype conv =
+    Keep_Conv | Then_Conv of conv * conv | Args_Conv of conv list | Rewr_Conv of rewrite
+  datatype rule =
+    Axiom of int | Taut of tautology * Argo_Expr.expr | Conjunct of int * int | Rewrite of conv |
+    Hyp of int * Argo_Expr.expr | Clause of int list | Lemma of int list | Unit_Res of int |
+    Refl of Argo_Expr.expr | Symm | Trans | Cong | Subst | Linear_Comb
+  type proof
+
+  (* equalities and orders *)
+  val eq_proof_id: proof_id * proof_id -> bool
+  val proof_id_ord: proof_id * proof_id -> order
+
+  (* conversion constructors *)
+  val keep_conv: conv
+  val mk_then_conv: conv -> conv -> conv
+  val mk_args_conv: conv list -> conv
+  val mk_rewr_conv: rewrite -> conv
+
+  (* context *)
+  type context
+  val cdcl_context: context
+  val cc_context: context
+  val simplex_context: context
+  val solver_context: context
+
+  (* proof constructors *)
+  val mk_axiom: int -> context -> proof * context
+  val mk_taut: tautology -> Argo_Expr.expr -> context -> proof * context
+  val mk_conj: int -> int -> proof -> context -> proof * context
+  val mk_rewrite: conv -> proof -> context -> proof * context
+  val mk_hyp: Argo_Lit.literal -> context -> proof * context
+  val mk_clause: Argo_Lit.literal list -> proof -> context -> proof * context
+  val mk_lemma: Argo_Lit.literal list -> proof -> context -> proof * context
+  val mk_unit_res: Argo_Lit.literal -> proof -> proof -> context -> proof * context
+  val mk_refl: Argo_Term.term -> context -> proof * context
+  val mk_symm: proof -> context -> proof * context
+  val mk_trans: proof -> proof -> context -> proof * context
+  val mk_cong: proof -> proof -> context -> proof * context
+  val mk_subst: proof -> proof -> proof -> context -> proof * context
+  val mk_linear_comb: proof list -> context -> proof * context
+
+  (* proof destructors *)
+  val id_of: proof -> proof_id
+  val dest: proof -> proof_id * rule * proof list
+
+  (* string representations *)
+  val string_of_proof_id: proof_id -> string
+  val string_of_taut: tautology -> string
+  val string_of_rule: rule -> string
+
+  (* unsatisfiability *)
+  exception UNSAT of proof
+  val unsat: proof -> 'a (* raises UNSAT *)
+end
+
+structure Argo_Proof: ARGO_PROOF =
+struct
+
+(* types *)
+
+datatype tautology =
+  Taut_And_1 of int | Taut_And_2 of int * int | Taut_Or_1 of int * int | Taut_Or_2 of int |
+  Taut_Iff_1 | Taut_Iff_2 | Taut_Iff_3 | Taut_Iff_4 | Taut_Ite_Then | Taut_Ite_Else
+
+datatype inequality = Le | Lt
+
+datatype rewrite =
+  Rewr_Not_True | Rewr_Not_False | Rewr_Not_Not | Rewr_Not_And of int | Rewr_Not_Or of int |
+  Rewr_Not_Iff |
+  Rewr_And_False of int | Rewr_And_Dual of int * int | Rewr_And_Sort of int * int list list |
+  Rewr_Or_True of int | Rewr_Or_Dual of int * int | Rewr_Or_Sort of int * int list list |
+  Rewr_Iff_True | Rewr_Iff_False | Rewr_Iff_Not_Not | Rewr_Iff_Refl | Rewr_Iff_Symm |
+  Rewr_Iff_Dual |
+  Rewr_Imp | Rewr_Ite_Prop | Rewr_Ite_True | Rewr_Ite_False | Rewr_Ite_Eq |
+  Rewr_Eq_Refl | Rewr_Eq_Symm |
+  Rewr_Neg | Rewr_Add of (Rat.rat * int option) list * (Rat.rat * int option) list | Rewr_Sub |
+  Rewr_Mul_Nums of Rat.rat * Rat.rat | Rewr_Mul_Zero | Rewr_Mul_One | Rewr_Mul_Comm |
+  Rewr_Mul_Assoc | Rewr_Mul_Sum | Rewr_Div_Nums of Rat.rat * Rat.rat | Rewr_Div_Zero |
+  Rewr_Div_One | Rewr_Div_Mul | Rewr_Div_Inv | Rewr_Div_Left | Rewr_Div_Right | Rewr_Min |
+  Rewr_Max | Rewr_Abs | Rewr_Eq_Le | Rewr_Ineq_Nums of inequality * bool |
+  Rewr_Ineq_Add of inequality * Rat.rat | Rewr_Ineq_Sub of inequality |
+  Rewr_Ineq_Mul of inequality * Rat.rat | Rewr_Not_Ineq of inequality
+
+datatype conv =
+  Keep_Conv | Then_Conv of conv * conv | Args_Conv of conv list | Rewr_Conv of rewrite
+
+datatype rule =
+  Axiom of int | Taut of tautology * Argo_Expr.expr | Conjunct of int * int | Rewrite of conv |
+  Hyp of int * Argo_Expr.expr | Clause of int list | Lemma of int list | Unit_Res of int |
+  Refl of Argo_Expr.expr | Symm | Trans | Cong | Subst | Linear_Comb
+
+(*
+  Proof identifiers are intentially hidden to prevent that functions outside of this structure
+  are able to build proofs. Proof can hence only be built by the functions provided by
+  this structure.
+*)
+
+datatype proof_id = Cdcl of int | Cc of int | Simplex of int | Solver of int
+
+datatype proof = Proof of proof_id * rule * proof list
+
+
+(* internal functions *)
+
+val proof_id_card = 4
+
+fun raw_proof_id (Cdcl i) = i
+  | raw_proof_id (Cc i) = i
+  | raw_proof_id (Simplex i) = i
+  | raw_proof_id (Solver i) = i
+
+
+(* equalities and orders *)
+
+fun int_of_proof_id (Cdcl _) = 0
+  | int_of_proof_id (Cc _) = 1
+  | int_of_proof_id (Simplex _) = 2
+  | int_of_proof_id (Solver _) = 3
+
+fun eq_proof_id (Cdcl i1, Cdcl i2) = (i1 = i2)
+  | eq_proof_id (Cc i1, Cc i2) = (i1 = i2)
+  | eq_proof_id (Simplex i1, Simplex i2) = (i1 = i2)
+  | eq_proof_id (Solver i1, Solver i2) = (i1 = i2)
+  | eq_proof_id _ = false
+
+fun proof_id_ord (Cdcl i1, Cdcl i2) = int_ord (i1, i2)
+  | proof_id_ord (Cc i1, Cc i2) = int_ord (i1, i2)
+  | proof_id_ord (Simplex i1, Simplex i2) = int_ord (i1, i2)
+  | proof_id_ord (Solver i1, Solver i2) = int_ord (i1, i2)
+  | proof_id_ord (id1, id2) = int_ord (int_of_proof_id id1, int_of_proof_id id2)
+
+
+(* conversion constructors *)
+
+val keep_conv = Keep_Conv
+
+fun mk_then_conv Keep_Conv c = c
+  | mk_then_conv c Keep_Conv = c
+  | mk_then_conv c1 c2 = Then_Conv (c1, c2)
+
+fun mk_args_conv cs =
+  if forall (fn Keep_Conv => true | _ => false) cs then Keep_Conv
+  else Args_Conv cs
+
+fun mk_rewr_conv r = Rewr_Conv r
+
+
+(* context *)
+
+(*
+  The proof context stores the next unused identifier. Incidentally, the same type as
+  for the proof identifier can be used as context. Every proof-producing module of the
+  solver has its own proof identifier domain to ensure globally unique identifiers
+  without sharing a single proof context.
+*)
+
+type context = proof_id
+
+val cdcl_context = Cdcl 0
+val cc_context = Cc 0
+val simplex_context = Simplex 0
+val solver_context = Solver 0
+
+fun next_id (id as Cdcl i) = (id, Cdcl (i + 1))
+  | next_id (id as Cc i) = (id, Cc (i + 1))
+  | next_id (id as Simplex i) = (id, Simplex (i + 1))
+  | next_id (id as Solver i) = (id, Solver (i + 1))
+
+
+(* proof destructors *)
+
+fun id_of (Proof (id, _, _)) = id
+
+fun dest (Proof p) = p
+
+
+(* proof constructors *)
+
+fun mk_proof r ps cx =
+  let val (id, cx) = next_id cx
+  in (Proof (id, r, ps), cx) end
+
+fun mk_axiom i = mk_proof (Axiom i) []
+fun mk_taut t e = mk_proof (Taut (t, e)) []
+fun mk_conj i n p = mk_proof (Conjunct (i, n)) [p]
+
+fun mk_rewrite Keep_Conv p cx = (p, cx)
+  | mk_rewrite c p cx = mk_proof (Rewrite c) [p] cx
+
+fun mk_hyp lit = mk_proof (Hyp (Argo_Lit.signed_id_of lit, Argo_Lit.signed_expr_of lit)) []
+fun mk_clause lits p cx = mk_proof (Clause (map Argo_Lit.signed_id_of lits)) [p] cx
+fun mk_lemma lits p = mk_proof (Lemma (map Argo_Lit.signed_id_of lits)) [p]
+
+(*
+  Replay of unit-resolution steps can be optimized if all premises follow a specific form.
+  Therefore, each premise is checked if it is in clausal form.
+*)
+
+fun check_clause (p as Proof (_, Clause _, _)) = p
+  | check_clause (p as Proof (_, Lemma _, _)) = p
+  | check_clause (p as Proof (_, Unit_Res _, _)) = p
+  | check_clause _ = raise Fail "bad clause proof"
+
+fun mk_unit t p1 p2 = mk_proof (Unit_Res (Argo_Term.id_of t)) (map check_clause [p1, p2])
+
+fun mk_unit_res (Argo_Lit.Pos t) p1 p2 = mk_unit t p1 p2
+  | mk_unit_res (Argo_Lit.Neg t) p1 p2 = mk_unit t p2 p1
+
+fun mk_refl t = mk_proof (Refl (Argo_Term.expr_of t)) []
+fun mk_symm p = mk_proof Symm [p]
+
+fun mk_trans (Proof (_, Refl _, _)) p2 = pair p2
+  | mk_trans p1 (Proof (_, Refl _, _)) = pair p1
+  | mk_trans p1 p2 = mk_proof Trans [p1, p2]
+
+fun mk_cong p1 p2 = mk_proof Cong [p1, p2]
+
+fun mk_subst p1 (Proof (_, Refl _, _)) (Proof (_, Refl _, _)) = pair p1
+  | mk_subst p1 p2 p3 = mk_proof Subst [p1, p2, p3]
+
+fun mk_linear_comb ps = mk_proof Linear_Comb ps
+
+
+(* string representations *)
+
+fun string_of_proof_id id = string_of_int (proof_id_card * raw_proof_id id + int_of_proof_id id)
+
+fun string_of_list l r f xs = enclose l r (space_implode ", " (map f xs))
+fun parens f xs = string_of_list "(" ")" f xs
+fun brackets f xs = string_of_list "[" "]" f xs
+
+fun string_of_taut (Taut_And_1 n) = "and " ^ string_of_int n
+  | string_of_taut (Taut_And_2 (i, n)) = "and " ^ parens string_of_int [i, n]
+  | string_of_taut (Taut_Or_1 (i, n)) = "or " ^ parens string_of_int [i, n]
+  | string_of_taut (Taut_Or_2 n) = "or " ^ string_of_int n
+  | string_of_taut Taut_Iff_1 = "(p1 == p2) | p1 | p2"
+  | string_of_taut Taut_Iff_2 = "(p1 == p2) | ~p1 | ~p2"
+  | string_of_taut Taut_Iff_3 = "~(p1 == p2) | ~p1 | p2"
+  | string_of_taut Taut_Iff_4 = "~(p1 == p2) | p1 | ~p2"
+  | string_of_taut Taut_Ite_Then = "~p | (ite p t1 t2) = t1"
+  | string_of_taut Taut_Ite_Else = "p | (ite p t1 t2) = t2"
+
+fun string_of_rewr Rewr_Not_True = "~T = F"
+  | string_of_rewr Rewr_Not_False = "~F = T"
+  | string_of_rewr Rewr_Not_Not = "~~p = p"
+  | string_of_rewr (Rewr_Not_And n) =
+      "~(and [" ^ string_of_int n ^ "]) = (or [" ^ string_of_int n ^ "])" 
+  | string_of_rewr (Rewr_Not_Or n) =
+      "~(or [" ^ string_of_int n ^ "]) = (and [" ^ string_of_int n ^ "])"
+  | string_of_rewr Rewr_Not_Iff = "~(p1 == p2) = (~p1 == ~p2)"
+  | string_of_rewr (Rewr_And_False i) = "(and ... F(" ^ string_of_int i ^ ") ...) = F"
+  | string_of_rewr (Rewr_And_Dual (i1, i2)) =
+      "(and ... p(" ^ string_of_int i1 ^ ") ... ~p(" ^ string_of_int i2 ^ ") ...) = F"
+  | string_of_rewr (Rewr_And_Sort (n, iss)) =
+      "(and [" ^ string_of_int n ^ "]) = " ^
+      "(and " ^ brackets (brackets string_of_int) iss ^ ")" 
+  | string_of_rewr (Rewr_Or_True i) = "(or ... T(" ^ string_of_int i ^ ") ...) = T"
+  | string_of_rewr (Rewr_Or_Dual (i1, i2)) =
+      "(or ... p(" ^ string_of_int i1 ^ ") ... ~p(" ^ string_of_int i2 ^ ") ...) = T"
+  | string_of_rewr (Rewr_Or_Sort (n, iss)) =
+      "(or [" ^ string_of_int n ^ "]) = " ^
+      "(or " ^ brackets (brackets string_of_int) iss ^ ")" 
+  | string_of_rewr Rewr_Iff_True = "(p == T) = p"
+  | string_of_rewr Rewr_Iff_False = "(p == F) = ~p"
+  | string_of_rewr Rewr_Iff_Not_Not = "(~p1 == ~p2) = (p1 == p2)"
+  | string_of_rewr Rewr_Iff_Refl = "(p == p) = T"
+  | string_of_rewr Rewr_Iff_Symm = "(p1 == p2) = (p2 == p1)"
+  | string_of_rewr Rewr_Iff_Dual = "(p == ~p) = F"
+  | string_of_rewr Rewr_Imp = "(p1 --> p2) = (~p1 | p2)"
+  | string_of_rewr Rewr_Ite_Prop = "(if p1 p2 p2) = ((~p1 | p2) & (p1 | p3) & (p2 | p3))"
+  | string_of_rewr Rewr_Ite_True = "(if T t1 t2) = t1"
+  | string_of_rewr Rewr_Ite_False = "(if F t1 t2) = t2"
+  | string_of_rewr Rewr_Ite_Eq = "(if p t t) = t"
+  | string_of_rewr Rewr_Eq_Refl = "(e = e) = T"
+  | string_of_rewr Rewr_Eq_Symm = "(e1 = e2) = (e2 = e1)"
+  | string_of_rewr Rewr_Neg = "-e = -1 * e"
+  | string_of_rewr (Rewr_Add (p1, p2)) =
+      let
+        fun string_of_monom (n, NONE) = Rat.string_of_rat n
+          | string_of_monom (n, SOME i) =
+              (if n = @1 then "" else Rat.string_of_rat n ^ " * ") ^ "e" ^ string_of_int i
+        fun string_of_polynom ms = space_implode " + " (map string_of_monom ms)
+      in string_of_polynom p1 ^ " = " ^ string_of_polynom p2 end
+  | string_of_rewr Rewr_Sub = "e1 - e2 = e1 + -1 * e2"
+  | string_of_rewr (Rewr_Mul_Nums (n1, n2)) =
+      Rat.string_of_rat n1 ^ " * " ^ Rat.string_of_rat n2 ^ " = " ^ Rat.string_of_rat (n1 * n2)
+  | string_of_rewr Rewr_Mul_Zero = "0 * e = 0"
+  | string_of_rewr Rewr_Mul_One = "1 * e = e"
+  | string_of_rewr Rewr_Mul_Comm = "e1 * e2 = e2 * e1"
+  | string_of_rewr Rewr_Mul_Assoc = "n1 * (n2 * e) = (n1 * n2) * e"
+  | string_of_rewr Rewr_Mul_Sum = "n * (e1 + ... + em) = n * e1 + ... n * em"
+  | string_of_rewr (Rewr_Div_Nums (n1, n2)) =
+      Rat.string_of_rat n1 ^ " / " ^ Rat.string_of_rat n2 ^ " = " ^ Rat.string_of_rat (n1 / n2)
+  | string_of_rewr Rewr_Div_Zero = "0 / e = 0"
+  | string_of_rewr Rewr_Div_One = "e / 1 = e"
+  | string_of_rewr Rewr_Div_Mul = "n / e = n * (1 / e)"
+  | string_of_rewr Rewr_Div_Inv = "e / n = 1/n * e"
+  | string_of_rewr Rewr_Div_Left = "(n * e1) / e2 = n * (e1 / e2)"
+  | string_of_rewr Rewr_Div_Right = "e1 / (n * e2) = 1/n * (e1 / e2)"
+  | string_of_rewr Rewr_Min = "min e1 e2 = (if e1 <= e2 then e1 else e2)"
+  | string_of_rewr Rewr_Max = "max e1 e2 = (if e1 < e2 then e2 else e1)"
+  | string_of_rewr Rewr_Abs = "abs e = (if 0 <= e then e else -e)"
+  | string_of_rewr Rewr_Eq_Le = "(e1 = e2) = (and (e1 <= e2) (e2 <= e1))"
+  | string_of_rewr (Rewr_Ineq_Nums (Le, true)) = "(n1 <= n2) = true"
+  | string_of_rewr (Rewr_Ineq_Nums (Le, false)) = "(n1 <= n2) = false"
+  | string_of_rewr (Rewr_Ineq_Nums (Lt, true)) = "(n1 < n2) = true"
+  | string_of_rewr (Rewr_Ineq_Nums (Lt, false)) = "(n1 < n2) = false"
+  | string_of_rewr (Rewr_Ineq_Add (Le, _)) = "(e1 <= e2) = (e1 + n <= e2 + n)"
+  | string_of_rewr (Rewr_Ineq_Add (Lt, _)) = "(e1 < e2) = (e1 + n < e2 + n)"
+  | string_of_rewr (Rewr_Ineq_Sub Le) = "(e1 <= e2) = (e1 - e2 <= 0)"
+  | string_of_rewr (Rewr_Ineq_Sub Lt) = "(e1 < e2) = (e1 - e2 < 0)"
+  | string_of_rewr (Rewr_Ineq_Mul (Le, _)) = "(e1 <= e2) = (n * e1 <= n * e2)"
+  | string_of_rewr (Rewr_Ineq_Mul (Lt, _)) = "(e1 < e2) = (n * e1 < n * e2)"
+  | string_of_rewr (Rewr_Not_Ineq Le) = "~(e1 <= e2) = (e2 < e1)"
+  | string_of_rewr (Rewr_Not_Ineq Lt) = "~(e1 < e2) = (e2 <= e1)"
+
+fun flatten_then_conv (Then_Conv (c1, c2)) = flatten_then_conv c1 @ flatten_then_conv c2
+  | flatten_then_conv c = [c]
+
+fun string_of_conv Keep_Conv = "_"
+  | string_of_conv (c as Then_Conv _) =
+      space_implode " then " (map (enclose "(" ")" o string_of_conv) (flatten_then_conv c))
+  | string_of_conv (Args_Conv cs) = "args " ^ brackets string_of_conv cs
+  | string_of_conv (Rewr_Conv r) = string_of_rewr r
+
+fun string_of_rule (Axiom i) = "axiom " ^ string_of_int i
+  | string_of_rule (Taut (t, _)) = "tautology: " ^ string_of_taut t
+  | string_of_rule (Conjunct (i, n)) = "conjunct " ^ string_of_int i ^ " of " ^ string_of_int n
+  | string_of_rule (Rewrite c) = "rewrite: " ^ string_of_conv c
+  | string_of_rule (Hyp (i, _)) = "hypothesis " ^ string_of_int i
+  | string_of_rule (Clause is) = "clause " ^ brackets string_of_int is
+  | string_of_rule (Lemma is) = "lemma " ^ brackets string_of_int is
+  | string_of_rule (Unit_Res i) = "unit-resolution " ^ string_of_int i
+  | string_of_rule (Refl _) = "reflexivity"
+  | string_of_rule Symm = "symmetry"
+  | string_of_rule Trans = "transitivity"
+  | string_of_rule Cong = "congruence"
+  | string_of_rule Subst = "substitution"
+  | string_of_rule Linear_Comb = "linear-combination"
+
+
+(* unsatisfiability *)
+
+exception UNSAT of proof
+
+fun unsat p = raise UNSAT p
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_rewr.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,266 @@
+(*  Title:      Tools/Argo/argo_rewr.ML
+    Author:     Sascha Boehme
+
+Bottom-up rewriting of expressions based on rewrite rules and rewrite functions.
+*)
+
+signature ARGO_REWR =
+sig
+  (* patterns *)
+  datatype pattern =
+    V of int |
+    E of Argo_Expr.expr |
+    A of Argo_Expr.kind |
+    P of Argo_Expr.kind * pattern list |
+    M of pattern |
+    X
+
+  (* scanning patterns from strings *)
+  val scan: string -> pattern
+
+  (* pattern matching *)
+  type env
+  val get_all: env -> Argo_Expr.expr list
+  val get: env -> int -> Argo_Expr.expr
+
+  (* conversions *)
+  type conv = Argo_Expr.expr -> Argo_Expr.expr * Argo_Proof.conv
+  val keep: conv
+  val seq: conv list -> conv
+  val args: conv -> conv
+  val rewr: Argo_Proof.rewrite -> Argo_Expr.expr -> conv
+
+  (* context *)
+  type context
+  val context: context
+  val flat: string ->
+    (int -> Argo_Expr.expr list -> (Argo_Proof.rewrite * Argo_Expr.expr) option) ->
+    context -> context
+  val func: string -> (env -> (Argo_Proof.rewrite * pattern) option) -> context -> context
+  val rule: Argo_Proof.rewrite -> string -> string -> context -> context
+
+  (* rewriting *)
+  val rewrite: context -> conv
+  val rewrite_top: context -> conv
+  val with_proof: conv -> Argo_Expr.expr * Argo_Proof.proof -> Argo_Proof.context ->
+    (Argo_Expr.expr * Argo_Proof.proof) * Argo_Proof.context
+end
+
+structure Argo_Rewr: ARGO_REWR =
+struct
+
+(* patterns *)
+
+(*
+  Patterns are used for matching against expressions and as a schema to construct
+  expressions. For matching, only V, E, A and P must be used. For the schema,
+  additionally M and X can be used.
+*)
+
+datatype pattern =
+  V of int | (* indexed placeholder where the index must be greater than 0 *)
+  E of Argo_Expr.expr | (* expression without placeholders *)
+  A of Argo_Expr.kind | (* placeholder for the arguments of an n-ary expression *)
+  P of Argo_Expr.kind * pattern list | (* expression with optional placeholders *)
+  M of pattern | (* mapping operator to iterate over an argument list of an n-ary expression *)
+  X (* closure argument under a mapping operator representing an expression *)
+
+fun int_of_pattern (E _) = 0
+  | int_of_pattern (P _) = 1
+  | int_of_pattern (A _) = 2
+  | int_of_pattern (V _) = 3
+  | int_of_pattern _ = raise Fail "bad pattern"
+
+(*
+  Specific patterns are ordered before generic patterns, since pattern matching
+  performs a linear search for the most suitable pattern.
+*)
+
+fun pattern_ord (E e1, E e2) = Argo_Expr.expr_ord (e1, e2)
+  | pattern_ord (P (k1, ps1), P (k2, ps2)) =
+      (case Argo_Expr.kind_ord (k1, k2) of EQUAL => list_ord pattern_ord (ps1, ps2) | x => x)
+  | pattern_ord (A k1, A k2) = Argo_Expr.kind_ord (k1, k2)
+  | pattern_ord (V i1, V i2) = int_ord (i1, i2)
+  | pattern_ord (p1, p2) = int_ord (int_of_pattern p1, int_of_pattern p2)
+
+
+(* scanning patterns from strings *)
+
+(*
+  The pattern language is cumbersome to use in other structures. Strings are a more
+  lightweight representation. Scanning patterns from strings should be performed at
+  compile time to avoid the parsing overhead at runtime.
+*)
+
+val kind = Scan.many1 Symbol.is_ascii_letter >> (Argo_Expr.kind_of_string o implode)
+val num = Scan.many1 Symbol.is_ascii_digit >> (the o Int.fromString o implode)
+val integer = $$ "-" |-- num >> ~ || num
+val blank = Scan.many1 Symbol.is_ascii_blank >> K ()
+
+fun pattern xs = (
+  kind >> (P o rpair []) ||
+  $$ "!" >> K X ||
+  $$ "(" -- $$ "#" -- blank |-- pattern --| $$ ")" >> M ||
+  $$ "(" -- $$ "?" -- blank |-- num --| $$ ")" >> V ||
+  $$ "(" -- Scan.this_string "num" -- blank |-- integer --| $$ ")" >>
+    (E o Argo_Expr.mk_num o Rat.of_int) ||
+  $$ "(" |-- kind --| blank --| $$ "_" --| $$ ")" >> A ||
+  $$ "(" |-- kind -- Scan.repeat1 (blank |-- pattern) --| $$ ")" >> P) xs
+
+fun scan s = fst (pattern (map str (String.explode s) @ [""]))
+
+
+(* pattern matching *)
+
+exception PATTERN of unit
+
+(*
+  The environment stores the matched expressions for the pattern placeholders.
+  The expression for an indexed placeholder (V i) can be retrieved by "get env i".
+  The argument expressions for an n-ary placeholder (A k) can be retrieved by "get_all env".
+*)
+
+type env = Argo_Expr.expr list Inttab.table
+
+val empty_env: env = Inttab.empty
+fun get_all env = Inttab.lookup_list env 0
+fun get env i = hd (Inttab.lookup_list env i)
+
+fun depth_of (V _) = 0
+  | depth_of (E _) = 0
+  | depth_of (A _) = 1
+  | depth_of (P (_, ps)) = 1 + fold (Integer.max o depth_of) ps 0
+  | depth_of (M p) = depth_of p
+  | depth_of X = 0
+
+fun match_list f k k' env = if k = k' then f env else raise PATTERN ()
+
+fun match (V i) e env = Inttab.update_new (i, [e]) env
+  | match (A k) (Argo_Expr.E (k', es)) env = match_list (Inttab.update_new (0, es)) k k' env
+  | match (P (k, ps)) (Argo_Expr.E (k', es)) env = match_list (fold2 match ps es) k k' env
+  | match _ _ _ = raise Fail "bad pattern"
+
+fun unfold_index env (V i) _ = get env i
+  | unfold_index _ (E e) _ = e
+  | unfold_index env (P (k, ps)) e = Argo_Expr.E (k, map (fn p => unfold_index env p e) ps)
+  | unfold_index _ X e = e
+  | unfold_index _ _ _ = raise Fail "bad pattern"
+
+fun unfold_pattern env (V i) = get env i
+  | unfold_pattern _ (E e) = e
+  | unfold_pattern env (A k) = Argo_Expr.E (k, get_all env)
+  | unfold_pattern env (P (k, [M p])) = Argo_Expr.E (k, map (unfold_index env p) (get_all env))
+  | unfold_pattern env (P (k, ps)) = Argo_Expr.E (k, map (unfold_pattern env) ps)
+  | unfold_pattern _ _ = raise Fail "bad pattern"
+
+
+(* conversions *)
+
+(*
+  Conversions are atomic rewrite steps. For every conversion there is a corresponding
+  inference step.
+*)
+
+type conv = Argo_Expr.expr -> Argo_Expr.expr * Argo_Proof.conv
+
+fun keep e = (e, Argo_Proof.keep_conv)
+
+fun seq [] e = keep e
+  | seq [cv] e = cv e
+  | seq (cv :: cvs) e =
+      let val ((e, c2), c1) = cv e |>> seq cvs
+      in (e, Argo_Proof.mk_then_conv c1 c2) end
+
+fun args cv (Argo_Expr.E (k, es)) =
+  let val (es, cs) = split_list (map cv es)
+  in (Argo_Expr.E (k, es), Argo_Proof.mk_args_conv cs) end
+
+fun rewr r e _ = (e, Argo_Proof.mk_rewr_conv r)
+
+
+(* context *)
+
+(*
+  The context stores functions to flatten expressions and functions to rewrite expressions.
+  Flattening an n-ary expression of kind k produces an expression whose arguments are not
+  of kind k. For instance, flattening (and p (and q r)) produces (and p q r) where p, q and r
+  are not conjunctions.
+*)
+
+structure Pattab = Table(type key = pattern val ord = pattern_ord)
+
+type context = {
+  flats:
+    (Argo_Expr.kind * (int -> Argo_Expr.expr list -> (Argo_Proof.rewrite * Argo_Expr.expr) option))
+      list, (* expressions that should be flattened before rewriting *)
+  rewrs: (env -> (Argo_Proof.rewrite * pattern) option) list Pattab.table}
+    (* Looking up matching rules is O(n). This could be optimized. *)
+
+fun mk_context flats rewrs: context = {flats=flats, rewrs=rewrs}
+val context = mk_context [] Pattab.empty
+
+fun map_context map_flats map_rewrs ({flats, rewrs}: context) =
+  mk_context (map_flats flats) (map_rewrs rewrs)
+
+fun flat lhs f =
+  (case scan lhs of
+    A k => map_context (cons (k, f)) I
+  | _ => raise Fail "bad pattern")
+
+fun func lhs f = map_context I (Pattab.map_default (scan lhs, []) (fn fs => fs @ [f]))
+fun rule r lhs rhs = func lhs (K (SOME (r, scan rhs)))
+
+
+(* rewriting *)
+
+(*
+  Rewriting proceeds bottom-up. Whenever a rewrite rule with placeholders is used,
+  the arguments are again rewritten, but only up to depth of the placeholders within the
+  matched pattern.
+*)
+
+fun rewr_rule r env p = rewr r (unfold_pattern env p)
+
+fun try_apply p e f =
+  let val env = match p e empty_env
+  in (case f env of NONE => NONE | SOME (r, p) => SOME (r, env, p)) end
+  handle PATTERN () => NONE
+
+fun all_args cv k (e as Argo_Expr.E (k', _)) = if k = k' then args (all_args cv k) e else cv e
+fun all_args_of k (e as Argo_Expr.E (k', es)) = if k = k' then maps (all_args_of k) es else [e]
+fun kind_depth_of k (Argo_Expr.E (k', es)) =
+  if k = k' then 1 + fold (Integer.max o kind_depth_of k) es 0 else 0
+
+fun descend cv flats (e as Argo_Expr.E (k, _)) =
+  if AList.defined Argo_Expr.eq_kind flats k then all_args cv k e
+  else args cv e
+
+fun flatten flats (e as Argo_Expr.E (k, _)) =
+  (case AList.lookup Argo_Expr.eq_kind flats k of
+    NONE => keep e
+  | SOME f =>
+      (case f (kind_depth_of k e) (all_args_of k e) of
+        NONE => keep e
+      | SOME (r, e') => rewr r e' e))
+
+fun exhaust cv rewrs e =
+  (case Pattab.get_first (fn (p, fs) => get_first (try_apply p e) fs) rewrs of
+    NONE => keep e
+  | SOME (r, env, p) => seq [rewr_rule r env p, cv (depth_of p)] e)
+
+fun norm (cx as {flats, rewrs}: context) d e =
+  seq [
+    if d = 0 then keep else descend (norm cx (d - 1)) flats,
+    flatten flats,
+    exhaust (norm cx) rewrs] e
+
+fun rewrite cx = norm cx ~1   (* bottom-up rewriting *)
+fun rewrite_top cx = norm cx 0   (* top-down rewriting *)
+
+fun with_proof cv (e, p) prf =
+  let
+    val (e, c) = cv e
+    val (p, prf) = Argo_Proof.mk_rewrite c p prf
+  in ((e, p), prf) end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_simplex.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,672 @@
+(*  Title:      Tools/Argo/argo_simplex.ML
+    Author:     Sascha Boehme
+
+Linear arithmetic reasoning based on the simplex algorithm. It features:
+
+ * simplification and normalization of arithmetic expressions
+ * decision procedure for reals
+
+These features might be added:
+
+ * propagating implied inequality literals while assuming external knowledge
+ * propagating equalities for fixed variables to all other theory solvers
+ * pruning the tableau after new atoms have been added: eliminate unnecessary
+   variables
+
+The implementation is inspired by:
+
+  Bruno Dutertre and Leonardo de Moura. A fast linear-arithmetic solver
+  for DPLL(T). In Computer Aided Verification, pages 81-94. Springer, 2006.
+*)
+
+signature ARGO_SIMPLEX =
+sig
+  (* simplification *)
+  val simplify: Argo_Rewr.context -> Argo_Rewr.context
+  
+  (* context *)
+  type context
+  val context: context
+
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> Argo_Lit.literal option * context
+
+  (* main operations *)
+  val prepare: context -> context
+  val assume: Argo_Common.literal -> context -> Argo_Lit.literal Argo_Common.implied * context
+  val check: context -> Argo_Lit.literal Argo_Common.implied * context
+  val explain: Argo_Lit.literal -> context -> (Argo_Cls.clause * context) option
+  val add_level: context -> context
+  val backtrack: context -> context
+end
+
+structure Argo_Simplex: ARGO_SIMPLEX =
+struct
+
+(* expression functions *)
+
+fun mk_mul n e =
+  if n = @0 then Argo_Expr.mk_num n
+  else if n = @1 then e
+  else Argo_Expr.mk_mul (Argo_Expr.mk_num n) e
+
+fun dest_mul (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) = (n, e)
+  | dest_mul e = (@1, e)
+
+fun dest_factor e = fst (dest_mul e)
+
+fun mk_add [] = raise Fail "bad addition"
+  | mk_add [e] = e
+  | mk_add es = Argo_Expr.mk_add es
+
+fun dest_ineq (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (false, e1, e2)
+  | dest_ineq (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (true, e1, e2)
+  | dest_ineq _ = NONE
+
+
+(* simplification *)
+
+(*
+  Products are normalized either to a number or to the monomial form
+    a * x
+  where a is a non-zero number and x is a variable. If x is a product,
+  it contains no number factors. The coefficient a is dropped if it is equal
+  to one; instead of 1 * x the expression x is used.
+
+  No further normalization to non-linear products is applied. Hence, the
+  products x * y and y * x will be treated as two different variables by the
+  arithmetic decision procedure.
+*)
+
+val mul_comm_rhs = Argo_Rewr.scan "(mul (? 2) (? 1))"
+
+fun norm_mul env =
+  (case apply2 (Argo_Rewr.get env) (1, 2) of
+    (Argo_Expr.E (Argo_Expr.Num n1, _), Argo_Expr.E (Argo_Expr.Num n2, _)) =>
+      SOME (Argo_Proof.Rewr_Mul_Nums (n1, n2), Argo_Rewr.E (Argo_Expr.mk_num (n1 * n2)))
+  | (Argo_Expr.E (Argo_Expr.Num n, _), _) =>
+      if n = @0 then SOME (Argo_Proof.Rewr_Mul_Zero, Argo_Rewr.V 1)
+      else if n = @1 then SOME (Argo_Proof.Rewr_Mul_One, Argo_Rewr.V 2)
+      else NONE
+  | (_, Argo_Expr.E (Argo_Expr.Num _, _)) => SOME (Argo_Proof.Rewr_Mul_Comm, mul_comm_rhs)
+  | _ => NONE)
+
+(*
+  Quotients are normalized either to a number or to the monomial form
+    a * x
+  where a is a non-zero number and x is a variable. If x is a quotient,
+  both divident and divisor are not a number. The coefficient a is dropped
+  if it is equal to one; instead of 1 * x the expression x is used.
+
+  No further normalization to quotients is applied. Hence, the
+  expressions (x * z) / y and x * (z / y) will be treated as two different
+  variables by the arithmetic decision procedure.
+*)
+
+fun div_mul_rhs n e =
+  let val pat = Argo_Rewr.P (Argo_Expr.Div, [Argo_Rewr.E e, Argo_Rewr.V 2])
+  in Argo_Rewr.P (Argo_Expr.Mul, [Argo_Rewr.E (Argo_Expr.mk_num n), pat])  end
+
+fun div_inv_rhs n pat =
+  let val inv_pat = Argo_Rewr.P (Argo_Expr.Div, map (Argo_Rewr.E o Argo_Expr.mk_num) [@1, n]) 
+  in Argo_Rewr.P (Argo_Expr.Mul, [inv_pat, pat]) end
+
+fun norm_div env =
+  (case apply2 (Argo_Rewr.get env) (1, 2) of
+    (Argo_Expr.E (Argo_Expr.Num n1, _), Argo_Expr.E (Argo_Expr.Num n2, _)) =>
+      if n2 = @0 then NONE
+      else SOME (Argo_Proof.Rewr_Div_Nums (n1, n2), Argo_Rewr.E (Argo_Expr.mk_num (n1 / n2)))
+  | (Argo_Expr.E (Argo_Expr.Num n, _), _) =>
+      if n = @0 then SOME (Argo_Proof.Rewr_Div_Zero, Argo_Rewr.E (Argo_Expr.mk_num @0))
+      else if n = @1 then NONE
+      else SOME (Argo_Proof.Rewr_Div_Mul, div_mul_rhs n (Argo_Expr.mk_num @1))
+  | (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e]), _) =>
+      if n = @1 then NONE
+      else SOME (Argo_Proof.Rewr_Div_Left, div_mul_rhs n e)
+  | (_, Argo_Expr.E (Argo_Expr.Num n, _)) =>
+      if n = @0 then NONE
+      else if n = @1 then SOME (Argo_Proof.Rewr_Div_One, Argo_Rewr.V 1)
+      else SOME (Argo_Proof.Rewr_Div_Inv, div_inv_rhs n (Argo_Rewr.V 1))
+  | (_, Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) =>
+      let val pat = Argo_Rewr.P (Argo_Expr.Div, [Argo_Rewr.V 1, Argo_Rewr.E e])
+      in SOME (Argo_Proof.Rewr_Div_Right, div_inv_rhs n pat) end
+  | _ => NONE)
+
+(*
+  Sums are flatten and normalized to the polynomial form
+    a_0 + a_1 * x_1 + ... + a_n * x_n
+  where all variables x_i are disjoint and where all coefficients a_i are
+  non-zero numbers. Coefficients equal to one are dropped; instead of 1 * x_i
+  the reduced monom x_i is used. The variables x_i are ordered based on the
+  expression order to reduce the number of problem literals by sharing equal
+  expressions.
+*)
+
+fun add_monom_expr i n e (p, s, etab) =
+  let val etab = Argo_Exprtab.map_default (e, (i, @0)) (apsnd (Rat.add n)) etab
+  in ((n, Option.map fst (Argo_Exprtab.lookup etab e)) :: p, s, etab) end
+
+fun add_monom (_, Argo_Expr.E (Argo_Expr.Num n, _)) (p, s, etab) = ((n, NONE) :: p, s + n, etab)
+  | add_monom (i, Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) x =
+      add_monom_expr i n e x
+  | add_monom (i, e) x = add_monom_expr i @1 e x
+
+fun norm_add d es =
+  let
+    val (p1, s, etab) = fold_index add_monom es ([], @0, Argo_Exprtab.empty)
+    val (p2, es) =
+      []
+      |> Argo_Exprtab.fold_rev (fn (e, (i, n)) => n <> @0 ? cons ((n, SOME i), mk_mul n e)) etab
+      |> s <> @0 ? cons ((s, NONE), Argo_Expr.mk_num s)
+      |> (fn [] => [((@0, NONE), Argo_Expr.mk_num @0)] | xs => xs)
+      |> split_list
+    val ps = (rev p1, p2)
+  in
+    if d = 1 andalso eq_list (op =) ps then NONE
+    else SOME (Argo_Proof.Rewr_Add ps, mk_add es)
+  end
+
+(*
+  An equation between two arithmetic expressions is rewritten to a conjunction of two
+  non-strict inequalities.
+*)
+
+val eq_rhs = Argo_Rewr.scan "(and (le (? 1) (? 2)) (le (? 2) (? 1)))"
+fun is_arith e = member (op =) [Argo_Expr.Real] (Argo_Expr.type_of e)
+
+fun norm_eq env =
+  if is_arith (Argo_Rewr.get env 1) then SOME (Argo_Proof.Rewr_Eq_Le, eq_rhs)
+  else NONE
+
+(*
+  Arithmetic inequalities (less and less-than) are normalized to the normal form
+    a_0 + a_1 * x_1 + ... + a_n * x_n ~ b
+  or
+    b ~ a_0 + a_1 * x_1 + ... + a_n * x_n
+  such that most of the coefficients a_i are positive.
+
+  Arithmetic inequalities of the form
+    a * x ~ b
+  or
+    b ~ a * x
+  are normalized to the form
+    x ~ c
+  or
+    c ~ x
+  where c is a number.
+*)
+
+fun mk_num_pat n = Argo_Rewr.E (Argo_Expr.mk_num n)
+fun mk_cmp_pat k ps = Argo_Rewr.P (k, ps)
+
+fun norm_cmp_mul k r n =
+  let
+    fun mult i = Argo_Rewr.P (Argo_Expr.Mul, [mk_num_pat n, Argo_Rewr.V i])
+    val ps = if n > @0 then [mult 1, mult 2] else [mult 2, mult 1]
+  in SOME (Argo_Proof.Rewr_Ineq_Mul (r, n), mk_cmp_pat k ps) end
+
+fun count_factors pred es = fold (fn e => if pred (dest_factor e) then Integer.add 1 else I) es 0
+
+fun norm_cmp_swap k r es =
+  let
+    val pos = count_factors (fn n => n > @0) es
+    val neg = count_factors (fn n => n < @0) es
+    val keep = pos > neg orelse (pos = neg andalso dest_factor (hd es) > @0)
+  in if keep then NONE else norm_cmp_mul k r @~1 end
+
+fun norm_cmp1 k r (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), _])) =
+      norm_cmp_mul k r (Rat.inv n)
+  | norm_cmp1 k r (Argo_Expr.E (Argo_Expr.Add, Argo_Expr.E (Argo_Expr.Num n, _) :: _)) =
+      let fun mk i = Argo_Rewr.P (Argo_Expr.Add, [Argo_Rewr.V i, mk_num_pat (~ n)])
+      in SOME (Argo_Proof.Rewr_Ineq_Add (r, ~ n), mk_cmp_pat k [mk 1, mk 2]) end
+  | norm_cmp1 k r (Argo_Expr.E (Argo_Expr.Add, es)) = norm_cmp_swap k r es
+  | norm_cmp1 _ _ _ = NONE
+
+val nums_true_rhs = Argo_Rewr.scan "true"
+val nums_false_rhs = Argo_Rewr.scan "false"
+
+val cmp_sub_rhs = map Argo_Rewr.scan ["(sub (? 1) (? 2))", "(num 0)"]
+
+fun norm_cmp k r pred env =
+  (case apply2 (Argo_Rewr.get env) (1, 2) of
+    (Argo_Expr.E (Argo_Expr.Num n1, _), Argo_Expr.E (Argo_Expr.Num n2, _)) =>
+      let val b = pred n1 n2
+      in SOME (Argo_Proof.Rewr_Ineq_Nums (r, b), if b then nums_true_rhs else nums_false_rhs) end
+  | (Argo_Expr.E (Argo_Expr.Num _, _), e) => norm_cmp1 k r e
+  | (e, Argo_Expr.E (Argo_Expr.Num _, _)) => norm_cmp1 k r e
+  | _ => SOME (Argo_Proof.Rewr_Ineq_Sub r, mk_cmp_pat k cmp_sub_rhs))
+
+(*
+  Arithmetic expressions are normalized in order to reduce the number of
+  problem literals. Arithmetically equal expressions are normalized into
+  syntactically equal expressions as much as possible.
+*)
+
+val simplify =
+  Argo_Rewr.rule Argo_Proof.Rewr_Neg
+    "(neg (? 1))"
+    "(mul (num -1) (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Sub
+    "(sub (? 1) (? 2))"
+    "(add (? 1) (mul (num -1) (? 2)))" #>
+  Argo_Rewr.func "(mul (? 1) (? 2))" norm_mul #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Mul_Assoc
+    "(mul (? 1) (mul (? 2) (? 3)))"
+    "(mul (mul (? 1) (? 2)) (? 3))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Mul_Sum
+    "(mul (? 1) (add _))"
+    "(add (# (mul (? 1) !)))" #>
+  Argo_Rewr.func "(div (? 1) (? 2))" norm_div #>
+  Argo_Rewr.flat "(add _)" norm_add #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Min
+    "(min (? 1) (? 2))"
+    "(ite (le (? 1) (? 2)) (? 1) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Max
+    "(max (? 1) (? 2))"
+    "(ite (le (? 1) (? 2)) (? 2) (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Abs
+    "(abs (? 1))"
+    "(ite (le (num 0) (? 1)) (? 1) (neg (? 1)))" #>
+  Argo_Rewr.func "(eq (? 1) (? 2))" norm_eq #>
+  Argo_Rewr.func "(le (? 1) (? 2))" (norm_cmp Argo_Expr.Le Argo_Proof.Le Rat.le) #>
+  Argo_Rewr.func "(lt (? 1) (? 2))" (norm_cmp Argo_Expr.Lt Argo_Proof.Lt Rat.lt)
+
+
+(* extended rationals *)
+
+(*
+  Extended rationals (c, k) are reals (c + k * e) where e is some small positive real number.
+  Extended rationals are used to represent a strict inequality by a non-strict inequality:
+    c < x  ~~  c + k * e <= e
+    x < c  ~~  x <= c - k * e
+*)
+
+type erat = Rat.rat * Rat.rat
+
+val erat_zero = (@0, @0)
+
+fun add (c1, k1) (c2, k2) = (c1 + c2, k1 + k2)
+fun sub (c1, k1) (c2, k2) = (c1 - c2, k1 - k2)
+fun mul n (c, k) = (n * c, n * k)
+
+val erat_ord = prod_ord Rat.ord Rat.ord
+
+fun less_eq n1 n2 = (erat_ord (n1, n2) <> GREATER)
+fun less n1 n2 = (erat_ord (n1, n2) = LESS)
+
+
+(* term functions *)
+
+fun dest_monom (Argo_Term.T (_, Argo_Expr.Mul, [Argo_Term.T (_, Argo_Expr.Num n, _), t])) = (t, n)
+  | dest_monom t = (t, @1)
+
+datatype node = Var of Argo_Term.term | Num of Rat.rat
+datatype ineq = Lower of Argo_Term.term * erat | Upper of Argo_Term.term * erat
+
+fun dest_node (Argo_Term.T (_, Argo_Expr.Num n, _)) = Num n
+  | dest_node t = Var t
+
+fun dest_atom true (k as Argo_Expr.Le) t1 t2 = SOME (k, dest_node t1, dest_node t2)
+  | dest_atom true (k as Argo_Expr.Lt) t1 t2 = SOME (k, dest_node t1, dest_node t2)
+  | dest_atom false Argo_Expr.Le t1 t2 = SOME (Argo_Expr.Lt, dest_node t2, dest_node t1)
+  | dest_atom false Argo_Expr.Lt t1 t2 = SOME (Argo_Expr.Le, dest_node t2, dest_node t1)
+  | dest_atom _ _ _ _ = NONE
+
+fun ineq_of pol (Argo_Term.T (_, k, [t1, t2])) =
+      (case dest_atom pol k t1 t2 of
+        SOME (Argo_Expr.Le, Var x, Num n) => SOME (Upper (x, (n, @0)))
+      | SOME (Argo_Expr.Le, Num n, Var x) => SOME (Lower (x, (n, @0)))
+      | SOME (Argo_Expr.Lt, Var x, Num n) => SOME (Upper (x, (n, @~1)))
+      | SOME (Argo_Expr.Lt, Num n, Var x) => SOME (Lower (x, (n, @1)))
+      | _ => NONE)
+  | ineq_of _ _ = NONE
+
+
+(* proofs *)
+
+(*
+  comment missing
+*)
+
+fun mk_ineq is_lt = if is_lt then Argo_Expr.mk_lt else Argo_Expr.mk_le
+fun ineq_rule_of is_lt = if is_lt then Argo_Proof.Lt else Argo_Proof.Le
+
+fun rewrite_top f = Argo_Rewr.rewrite_top (f Argo_Rewr.context)
+
+fun unnegate_conv (e as Argo_Expr.E (Argo_Expr.Not, [Argo_Expr.E (Argo_Expr.Le, [e1, e2])])) =
+      Argo_Rewr.rewr (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) (Argo_Expr.mk_lt e2 e1) e
+  | unnegate_conv (e as Argo_Expr.E (Argo_Expr.Not, [Argo_Expr.E (Argo_Expr.Lt, [e1, e2])])) =
+      Argo_Rewr.rewr (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) (Argo_Expr.mk_le e2 e1) e
+  | unnegate_conv e = Argo_Rewr.keep e
+
+val norm_scale_conv = rewrite_top (
+  Argo_Rewr.rule Argo_Proof.Rewr_Mul_Sum
+    "(mul (? 1) (add _))"
+    "(add (# (mul (? 1) !)))" #>
+  Argo_Rewr.func "(mul (? 1) (? 2))" norm_mul)
+
+fun scale_conv r mk n e1 e2 =
+  let
+    fun scale e = Argo_Expr.mk_mul (Argo_Expr.mk_num n) e
+    val (e1, e2) = if n > @0 then (scale e1, scale e2) else (scale e2, scale e1)
+    val conv = Argo_Rewr.rewr (Argo_Proof.Rewr_Ineq_Mul (r, n)) (mk e1 e2)
+  in Argo_Rewr.seq [conv, Argo_Rewr.args norm_scale_conv] end
+
+fun scale_ineq_conv n e =
+  if n = @1 then Argo_Rewr.keep e
+  else
+    (case dest_ineq e of
+      NONE => raise Fail "bad inequality"
+    | SOME (is_lt, e1, e2) => scale_conv (ineq_rule_of is_lt) (mk_ineq is_lt) n e1 e2 e)
+
+fun simp_lit (n, (lit, p)) =
+  let val conv = Argo_Rewr.seq [unnegate_conv, scale_ineq_conv n]
+  in Argo_Rewr.with_proof conv (Argo_Lit.signed_expr_of lit, p) end
+
+val combine_conv = rewrite_top (Argo_Rewr.flat "(add _)" norm_add)
+fun reduce_conv r = Argo_Rewr.rewr (Argo_Proof.Rewr_Ineq_Nums (r, false)) Argo_Expr.false_expr
+
+fun simp_combine es p prf =
+  let
+    fun dest e (is_lt, (es1, es2)) =
+      let val (is_lt', e1, e2) = the (dest_ineq e)
+      in (is_lt orelse is_lt', (e1 :: es1, e2 :: es2)) end
+    val (is_lt, (es1, es2)) = fold_rev dest es (false, ([], []))
+    val e = uncurry (mk_ineq is_lt) (apply2 Argo_Expr.mk_add (es1, es2))
+    val conv = Argo_Rewr.seq [Argo_Rewr.args combine_conv, reduce_conv (ineq_rule_of is_lt)]
+  in prf |> Argo_Rewr.with_proof conv (e, p) |>> snd end
+
+fun linear_combination nlps prf =
+  let val ((es, ps), prf) = fold_map simp_lit nlps prf |>> split_list
+  in prf |> Argo_Proof.mk_linear_comb ps |-> simp_combine es |-> Argo_Proof.mk_lemma [] end
+
+fun proof_of (lit, SOME p) (ls, prf) = ((lit, p), (ls, prf))
+  | proof_of (lit, NONE) (ls, prf) =
+      let val (p, prf) = Argo_Proof.mk_hyp lit prf
+      in ((lit, p), (Argo_Lit.negate lit :: ls, prf)) end
+
+
+(* tableau *)
+
+(*
+  The tableau consists of equations x_i = a_i1 * x_1 + ... a_ik * x_k where
+  the variable on the left-hand side is called a basic variable and
+  the variables on the right-hand side are called non-basic variables.
+
+  For each basic variable, the polynom on the right-hand side is stored as a map
+  from variables to coefficients. Only variables with non-zero coefficients are stored.
+  The map is sorted by the term order of the variables for a deterministic order when
+  analyzing a polynom.
+
+  Additionally, for each basic variable a boolean flag is kept that, when false,
+  indicates that the current value of the basic variable might be outside its bounds.
+  The value of a non-basic variable is always within its bounds.
+
+  The tableau is stored as a table indexed by variables. For each variable,
+  both basic and non-basic, its current value is stored as extended rational
+  along with either the equations or the occurrences.
+*)
+
+type basic = bool * (Argo_Term.term * Rat.rat) Ord_List.T
+type entry = erat * basic option
+type tableau = entry Argo_Termtab.table
+
+fun dirty ms = SOME (false, ms)
+fun checked ms = SOME (true, ms)
+
+fun basic_entry ms = (erat_zero, dirty ms)
+val non_basic_entry: entry = (erat_zero, NONE)
+
+fun value_of tableau x =
+  (case Argo_Termtab.lookup tableau x of
+    NONE => erat_zero
+  | SOME (v, _) => v)
+
+fun first_unchecked_basic tableau =
+  Argo_Termtab.get_first (fn (y, (v, SOME (false, ms))) => SOME (y, v, ms) | _ => NONE) tableau
+
+local
+
+fun coeff_of ms x = the (AList.lookup Argo_Term.eq_term ms x)
+
+val eq_var = Argo_Term.eq_term
+fun monom_ord sp = prod_ord Argo_Term.term_ord (K EQUAL) sp
+
+fun add_monom m ms = Ord_List.insert monom_ord m ms
+fun update_monom (m as (x, a)) = if a = @0 then AList.delete eq_var x else AList.update eq_var m
+
+fun add_scaled_monom n (x, a) ms =
+  (case AList.lookup eq_var ms x of
+    NONE => add_monom (x, n * a) ms
+  | SOME b => update_monom (x, n * a + b) ms)
+
+fun replace_polynom x n ms' ms = fold (add_scaled_monom n) ms' (AList.delete eq_var x ms)
+
+fun map_basic f (v, SOME (_, ms)) = f v ms
+  | map_basic _ e = e
+
+fun map_basic_entries x f =
+  let
+    fun apply (e as (v, SOME (_, ms))) = if AList.defined eq_var ms x then f v ms else e
+      | apply ve = ve
+  in Argo_Termtab.map (K apply) end
+
+fun put_entry x e = Argo_Termtab.update (x, e)
+
+fun add_new_entry (y as Argo_Term.T (_, Argo_Expr.Add, ts)) tableau =
+      let val ms = Ord_List.make monom_ord (map dest_monom ts)
+      in fold (fn (x, _) => put_entry x non_basic_entry) ms (put_entry y (basic_entry ms) tableau) end
+  | add_new_entry x tableau = put_entry x non_basic_entry tableau
+
+fun with_non_basic update_basic x f tableau =
+  (case Argo_Termtab.lookup tableau x of
+    NONE => tableau
+  | SOME (v, NONE) => f v tableau
+  | SOME (v, SOME (_, ms)) => if update_basic then put_entry x (v, dirty ms) tableau else tableau)
+
+in
+
+fun add_entry x tableau =
+  if Argo_Termtab.defined tableau x then tableau
+  else add_new_entry x tableau
+
+fun basic_within_bounds y = Argo_Termtab.map_entry y (map_basic (fn v => fn ms => (v, checked ms)))
+
+fun eliminate _ tableau = tableau
+
+fun update_non_basic pred x v' = with_non_basic true x (fn v =>
+  let fun update_basic n v ms = (add v (mul (coeff_of ms x) n), dirty ms)
+  in pred v ? put_entry x (v', NONE) o map_basic_entries x (update_basic (sub v' v)) end)
+
+fun update_pivot y vy ms x c v = with_non_basic false x (fn vx =>
+  let
+    val a = Rat.inv c
+    val v' = mul a (sub v vy)
+
+    fun scale_or_drop (x', b) = if Argo_Term.eq_term (x', x) then NONE else SOME (x', ~ a * b)
+    val ms = add_monom (y, a) (map_filter scale_or_drop ms)
+
+    fun update_basic v ms' =
+      let val n = coeff_of ms' x
+      in (add v (mul n v'), dirty (replace_polynom x n ms ms')) end
+  in
+    put_entry x (add vx v', dirty ms) #>
+    put_entry y (v, NONE) #>
+    map_basic_entries x update_basic
+  end)
+
+end
+
+
+(* bounds *)
+
+(*
+  comment missing
+*)
+
+type bound = (erat * Argo_Common.literal) option
+type atoms = (erat * Argo_Term.term) list
+type bounds_atoms = ((bound * bound) * (atoms * atoms))
+type bounds = bounds_atoms Argo_Termtab.table
+
+val empty_bounds_atoms: bounds_atoms = ((NONE, NONE), ([], []))
+
+fun on_some pred (SOME (n, _)) = pred n
+  | on_some _ NONE = false
+
+fun none_or_some pred (SOME (n, _)) = pred n
+  | none_or_some _ NONE = true
+
+fun bound_of (SOME (n, _)) = n
+  | bound_of NONE = raise Fail "bad bound"
+
+fun reason_of (SOME (_, r)) = r
+  | reason_of NONE = raise Fail "bad reason"
+
+fun bounds_atoms_of bounds x = the_default empty_bounds_atoms (Argo_Termtab.lookup bounds x)
+fun bounds_of bounds x = fst (bounds_atoms_of bounds x)
+
+fun put_bounds x bs bounds = Argo_Termtab.map_default (x, empty_bounds_atoms) (apfst (K bs)) bounds
+
+fun has_bound_atoms bounds x =
+  (case Argo_Termtab.lookup bounds x of
+    NONE => false
+  | SOME (_, ([], [])) => false
+  | _ => true)
+
+fun add_new_atom f x n t =
+  let val ins = f (insert (eq_snd Argo_Term.eq_term) (n, t))
+  in Argo_Termtab.map_default (x, empty_bounds_atoms) (apsnd ins) end
+
+fun del_atom x t =
+  let fun eq_atom (t1, (_, t2)) = Argo_Term.eq_term (t1, t2)
+  in Argo_Termtab.map_entry x (apsnd (apply2 (remove eq_atom t))) end
+
+
+(* context *)
+
+type context = {
+  tableau: tableau, (* values of variables and tableau entries for each variable *)
+  bounds: bounds, (* bounds and unassigned atoms for each variable *)
+  prf: Argo_Proof.context, (* proof context *)
+  back: bounds list} (* stack storing previous bounds and unassigned atoms *)
+
+fun mk_context tableau bounds prf back: context =
+  {tableau=tableau, bounds=bounds, prf=prf, back=back}
+
+val context = mk_context Argo_Termtab.empty Argo_Termtab.empty Argo_Proof.simplex_context []
+
+
+(* declaring atoms *)
+
+fun add_ineq_atom f t x n ({tableau, bounds, prf, back}: context) =
+  (* TODO: check whether the atom is already known to hold *)
+  (NONE, mk_context (add_entry x tableau) (add_new_atom f x n t bounds) prf back)
+
+fun add_atom t cx =
+  (case ineq_of true t of
+    SOME (Lower (x, n)) => add_ineq_atom apfst t x n cx
+  | SOME (Upper (x, n)) => add_ineq_atom apsnd t x n cx
+  | NONE => (NONE, cx))
+
+
+(* preparing the solver after new atoms have been added *)
+
+(*
+  Variables that do not directly occur in atoms can be eliminated from the tableau
+  since no bounds will ever limit their value. This can reduce the tableau size
+  substantially.
+*)
+
+fun prepare ({tableau, bounds, prf, back}: context) =
+  let fun drop (xe as (x, _)) = not (has_bound_atoms bounds x) ? eliminate xe
+  in mk_context (Argo_Termtab.fold drop tableau tableau) bounds prf back end
+
+
+(* assuming external knowledge *)
+
+fun bounds_conflict r1 r2 ({tableau, bounds, prf, back}: context) =
+  let
+    val ((lp2, lp1), (lits, prf)) = ([], prf) |> proof_of r2 ||>> proof_of r1
+    val (p, prf) = linear_combination [(@~1, lp1), (@1, lp2)] prf
+  in (Argo_Common.Conflict (lits, p), mk_context tableau bounds prf back) end
+
+fun assume_bounds order x c bs ({tableau, bounds, prf, back}: context) =
+  let
+    val lits = []
+    val bounds = put_bounds x bs bounds
+    val tableau = update_non_basic (fn v => erat_ord (v, c) = order) x c tableau
+  in (Argo_Common.Implied lits, mk_context tableau bounds prf back) end
+
+fun assume_lower r x c (low, upp) cx =
+  if on_some (fn l => less_eq c l) low then (Argo_Common.Implied [], cx)
+  else if on_some (fn u => less u c) upp then bounds_conflict r (reason_of upp) cx
+  else assume_bounds LESS x c (SOME (c, r), upp) cx
+
+fun assume_upper r x c (low, upp) cx =
+  if on_some (fn u => less_eq u c) upp then (Argo_Common.Implied [], cx)
+  else if on_some (fn l => less c l) low then bounds_conflict (reason_of low) r cx
+  else assume_bounds GREATER x c (low, SOME (c, r)) cx
+
+fun with_bounds r t f x n ({tableau, bounds, prf, back}: context) =
+  f r x n (bounds_of bounds x) (mk_context tableau (del_atom x t bounds) prf back)
+
+fun choose f (SOME (Lower (x, n))) cx = f assume_lower x n cx
+  | choose f (SOME (Upper (x, n))) cx = f assume_upper x n cx
+  | choose _ NONE cx = (Argo_Common.Implied [], cx)
+
+fun assume (r as (lit, _)) cx =
+  let val (t, pol) = Argo_Lit.dest lit
+  in choose (with_bounds r t) (ineq_of pol t) cx end
+
+
+(* checking for consistency and pending implications *)
+
+fun basic_bounds_conflict lower y ms ({tableau, bounds, prf, back}: context) =
+  let
+    val (a, low, upp) = if lower then (@1, fst, snd) else (@~1, snd, fst)
+    fun coeff_proof f a x = apfst (pair a) o proof_of (reason_of (f (bounds_of bounds x)))
+    fun monom_proof (x, a) = coeff_proof (if a < @0 then low else upp) a x
+    val ((alp, alps), (lits, prf)) = ([], prf) |> coeff_proof low a y ||>> fold_map monom_proof ms
+    val (p, prf) = linear_combination (alp :: alps) prf
+  in (Argo_Common.Conflict (lits, p), mk_context tableau bounds prf back) end
+
+fun can_compensate ord tableau bounds (x, a) =
+  let val (low, upp) = bounds_of bounds x
+  in
+    if Rat.ord (a, @0) = ord then none_or_some (fn u => less (value_of tableau x) u) upp
+    else none_or_some (fn l => less l (value_of tableau x)) low
+  end
+
+fun check (cx as {tableau, bounds, prf, back}: context) =
+  (case first_unchecked_basic tableau of
+    NONE => (Argo_Common.Implied [], cx)
+  | SOME (y, v, ms) =>
+      let val (low, upp) = bounds_of bounds y
+      in
+        if on_some (fn l => less v l) low then adjust GREATER true y v ms (bound_of low) cx
+        else if on_some (fn u => less u v) upp then adjust LESS false y v ms (bound_of upp) cx
+        else check (mk_context (basic_within_bounds y tableau) bounds prf back)
+      end)
+
+and adjust ord lower y vy ms v (cx as {tableau, bounds, prf, back}: context) =
+  (case find_first (can_compensate ord tableau bounds) ms of
+    NONE => basic_bounds_conflict lower y ms cx
+  | SOME (x, a) => check (mk_context (update_pivot y vy ms x a v tableau) bounds prf back))
+
+
+(* explanations *)
+
+fun explain _ _ = NONE
+
+
+(* backtracking *)
+
+fun add_level ({tableau, bounds, prf, back}: context) =
+  mk_context tableau bounds prf (bounds :: back)
+
+fun backtrack ({back=[], ...}: context) = raise Empty
+  | backtrack ({tableau, prf, back=bounds :: back, ...}: context) =
+      mk_context tableau bounds prf back
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_solver.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,168 @@
+(*  Title:      Tools/Argo/argo_solver.ML
+    Author:     Sascha Boehme
+
+The main interface to the Argo solver.
+
+The solver performs satisfiability checking for a given set of assertions. If these assertions
+are unsatisfiable, a proof trace is returned. If these assertions are satisfiable, the computed
+model can be queried or further assertions may be added.
+*)
+
+signature ARGO_SOLVER =
+sig
+  type context
+  val context: context
+  val assert: Argo_Expr.expr list -> context -> context (* raises Argo_Expr.TYPE, Argo_Expr.EXPR
+    and Argo_Proof.UNSAT *)
+  val model_of: context -> string * Argo_Expr.typ -> bool option
+end
+
+structure Argo_Solver: ARGO_SOLVER =
+struct
+
+(* context *)
+
+type context = {
+  next_axiom: int,
+  prf: Argo_Proof.context,
+  core: Argo_Core.context}
+
+fun mk_context next_axiom prf core: context = {next_axiom=next_axiom, prf=prf, core=core}
+
+val context = mk_context 0 Argo_Proof.solver_context Argo_Core.context
+
+
+(* negation normal form *)
+
+fun nnf_nary r rhs env = SOME (r (length (Argo_Rewr.get_all env)), rhs)
+
+val not_and_rhs = Argo_Rewr.scan "(or (# (not !)))"
+val not_or_rhs = Argo_Rewr.scan "(and (# (not !)))"
+
+val nnf =
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_True "(not true)" "false" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_False "(not false)" "true" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Not "(not (not (? 1)))" "(? 1)" #>
+  Argo_Rewr.func "(not (and _))" (nnf_nary Argo_Proof.Rewr_Not_And not_and_rhs) #>
+  Argo_Rewr.func "(not (or _))" (nnf_nary Argo_Proof.Rewr_Not_Or not_or_rhs) #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Iff "(not (iff (not (? 1)) (? 2)))" "(iff (? 1) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Iff "(not (iff (? 1) (not (? 2))))" "(iff (? 1) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Iff "(not (iff (? 1) (? 2)))" "(iff (not (? 1)) (? 2))"
+
+
+(* propositional normalization *)
+
+(*
+  Propositional expressions are transformed into literals in the clausifier. Having
+  fewer literals results in faster solver execution. Normalizing propositional expressions
+  turns similar expressions into equal expressions, for which the same literal can be used.
+  The clausifier expects that only negation, disjunction, conjunction and equivalence form
+  propositional expressions. Expressions may be simplified to truth or falsity, but both
+  truth and falsity eventually occur nowhere inside expressions.
+*)
+
+val e_true = Argo_Expr.true_expr
+val e_false = Argo_Expr.false_expr
+
+fun first_index pred xs =
+  let val i = find_index pred xs
+  in if i >= 0 then SOME i else NONE end
+
+fun find_zero r_zero zero es =
+  Option.map (fn i => (r_zero i, zero)) (first_index (curry Argo_Expr.eq_expr zero) es)
+
+fun find_duals _ _ _ [] = NONE
+  | find_duals _ _ _ [_] = NONE
+  | find_duals r_dual zero i (e :: es) =
+      (case first_index (Argo_Expr.dual_expr e) es of
+        SOME i' => SOME (r_dual (i, i + i' + 1), zero)
+      | NONE => find_duals r_dual zero (i + 1) es)
+
+fun sort_nary r_sort one mk l es =
+  let
+    val n = length es
+    fun add (i, e) = if Argo_Expr.eq_expr (e, one) then I else Argo_Exprtab.cons_list (e, i)
+    fun dest (e, i) (es, is) = (e :: es, i :: is)
+    val (es, iss) = Argo_Exprtab.fold_rev dest (fold_index add es Argo_Exprtab.empty) ([], [])
+    fun identity is = length is = n andalso forall (op =) (map_index I is)
+  in if l = 1 andalso identity (map hd iss) then NONE else (SOME (r_sort (n, iss), mk es)) end
+
+fun apply_first fs es = get_first (fn f => f es) fs
+
+fun norm_nary r_zero r_dual r_sort zero one mk l =
+  apply_first [find_zero r_zero zero, find_duals r_dual zero 0, sort_nary r_sort one mk l]
+
+val norm_and = norm_nary Argo_Proof.Rewr_And_False Argo_Proof.Rewr_And_Dual Argo_Proof.Rewr_And_Sort
+  e_false e_true Argo_Expr.mk_and
+val norm_or = norm_nary Argo_Proof.Rewr_Or_True Argo_Proof.Rewr_Or_Dual Argo_Proof.Rewr_Or_Sort
+  e_true e_false Argo_Expr.mk_or
+
+fun norm_iff env =
+  let val e1 = Argo_Rewr.get env 1 and e2 = Argo_Rewr.get env 2
+  in
+    if Argo_Expr.dual_expr e1 e2 then SOME (Argo_Proof.Rewr_Iff_Dual, Argo_Rewr.E e_false)
+    else
+      (case Argo_Expr.expr_ord (e1, e2) of
+        EQUAL => SOME (Argo_Proof.Rewr_Iff_Refl, Argo_Rewr.E e_true)
+      | LESS => NONE
+      | GREATER => SOME (Argo_Proof.Rewr_Iff_Symm, Argo_Rewr.E (Argo_Expr.mk_iff e2 e1)))
+  end
+
+val norm_prop =
+  Argo_Rewr.flat "(and _)" norm_and #>
+  Argo_Rewr.flat "(or _)" norm_or #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Imp "(imp (? 1) (? 2))" "(or (not (? 1)) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_True "(iff true (? 1))" "(? 1)" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_False "(iff false (? 1))" "(not (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_True "(iff (? 1) true)" "(? 1)" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_False "(iff (? 1) false)" "(not (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_Not_Not "(iff (not (? 1)) (not (? 2)))" "(iff (? 1) (? 2))" #>
+  Argo_Rewr.func "(iff (? 1) (? 2))" norm_iff
+
+
+(* normalization of if-then-else expressions *)
+
+val simp_prop_ite_result =
+  Argo_Rewr.scan "(and (or (not (? 1)) (? 2)) (or (? 1) (? 3)) (or (? 2) (? 3)))"
+
+val simp_ite_eq_result = Argo_Rewr.scan "(? 2)"
+
+fun simp_ite env =
+  if Argo_Expr.type_of (Argo_Rewr.get env 2) = Argo_Expr.Bool then
+    SOME (Argo_Proof.Rewr_Ite_Prop, simp_prop_ite_result)
+  else if Argo_Expr.eq_expr (Argo_Rewr.get env 2, Argo_Rewr.get env 3) then
+    SOME (Argo_Proof.Rewr_Ite_Eq, simp_ite_eq_result)
+  else NONE
+
+val norm_ite =
+  Argo_Rewr.rule Argo_Proof.Rewr_Ite_True "(ite true (? 1) (? 2))" "(? 1)" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Ite_False "(ite false (? 1) (? 2))" "(? 2)" #>
+  Argo_Rewr.func "(ite (? 1) (? 2) (? 3))" simp_ite
+
+
+(* rewriting and normalizing axioms *)
+
+val simp_context = Argo_Rewr.context |> nnf |> norm_prop |> norm_ite |> Argo_Thy.simplify
+val simp_axiom = Argo_Rewr.with_proof (Argo_Rewr.rewrite simp_context)
+
+
+(* asserting axioms *)
+
+fun add_axiom e ({next_axiom, prf, core}: context) =
+  let
+    val _ = Argo_Expr.check e
+    val (p, prf) = Argo_Proof.mk_axiom next_axiom prf
+    val (ep, prf) = simp_axiom (e, p) prf 
+    val (prf, core) = Argo_Clausify.clausify simp_context ep (prf, core)
+  in mk_context (next_axiom + 1) prf core end
+
+fun assert es cx =
+  let val {next_axiom, prf, core}: context = fold add_axiom es cx
+  in mk_context next_axiom prf (Argo_Core.run core) end
+
+
+(* models *)
+
+fun model_of ({core, ...}: context) = Argo_Core.model_of core
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_term.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,152 @@
+(*  Title:      Tools/Argo/argo_term.ML
+    Author:     Sascha Boehme
+
+Internal language of the Argo solver.
+
+Terms are fully-shared via hash-consing. Alpha-equivalent terms have the same identifier.
+*)
+
+signature ARGO_TERM =
+sig
+  (* data types *)
+  type meta
+  datatype term = T of meta * Argo_Expr.kind * term list
+
+  (* term operations *)
+  val id_of: term -> int
+  val expr_of: term -> Argo_Expr.expr
+  val type_of: term -> Argo_Expr.typ
+  val eq_term: term * term -> bool
+  val term_ord: term * term -> order
+
+  (* context *)
+  type context
+  val context: context
+
+  (* identifying expressions *)
+  datatype item = Expr of Argo_Expr.expr | Term of term
+  datatype identified = New of term | Known of term
+  val identify_item: item -> context -> identified * context
+end
+
+structure Argo_Term: ARGO_TERM =
+struct
+
+(* data types *)
+
+(*
+  The type meta is intentionally hidden to prevent that functions outside of this structure
+  are able to build terms. Meta stores the identifier of the term as well as the complete
+  expression underlying the term.
+*)
+
+datatype meta = M of int * Argo_Expr.expr
+datatype term = T of meta * Argo_Expr.kind * term list
+
+
+(* term operations *)
+
+fun id_of (T (M (id, _), _, _)) = id
+fun expr_of (T (M (_, e), _, _)) = e
+fun type_of t = Argo_Expr.type_of (expr_of t)
+
+(*
+  Comparing terms is fast as only the identifiers are compared. No expressions need to
+  be taken into account.
+*)
+
+fun eq_term (t1, t2) = (id_of t1 = id_of t2)
+fun term_ord (t1, t2) = int_ord (id_of t1, id_of t2)
+
+
+(* sharing of terms *)
+
+(*
+  Kinds are short representation of expressions. Constants and numbers carry additional
+  information and have no arguments. Their kind is hence similar to them. All other expressions
+  are stored in a flat way with identifiers of shared terms as arguments instead of expression
+  as arguments.
+*)
+
+datatype kind =
+  Con of string * Argo_Expr.typ |
+  Num of Rat.rat |
+  Exp of int list
+
+fun kind_of (Argo_Expr.E (Argo_Expr.Con c, _)) _ = Con c
+  | kind_of (Argo_Expr.E (Argo_Expr.Num n, _)) _ = Num n
+  | kind_of (Argo_Expr.E (k, _)) is = Exp (Argo_Expr.int_of_kind k :: is)
+
+fun int_of_kind (Con _) = 1
+  | int_of_kind (Num _) = 2
+  | int_of_kind (Exp _) = 3
+
+fun kind_ord (Con c1, Con c2) = Argo_Expr.con_ord (c1, c2)
+  | kind_ord (Num n1, Num n2) = Rat.ord (n1, n2)
+  | kind_ord (Exp is1, Exp is2) = dict_ord int_ord (is1, is2)
+  | kind_ord (k1, k2) = int_ord (int_of_kind k1, int_of_kind k2)
+
+structure Kindtab = Table(type key = kind val ord = kind_ord)
+
+(*
+  The context keeps track of the next unused identifier as well as all shared terms,
+  which are indexed by their unique kind. For each term, a boolean marker flag is stored.
+  When set to true on an atom, the atom is already asserted to the solver core. When set to
+  true on an if-then-else term, the term has already been lifted.
+
+  Zero is intentionally avoided as identifier, since literals use term identifiers
+  with a sign as literal identifiers.
+*)
+
+type context = {
+  next_id: int,
+  terms: (term * bool) Kindtab.table}
+
+fun mk_context next_id terms: context = {next_id=next_id, terms=terms}
+
+val context = mk_context 1 Kindtab.empty
+
+fun note_atom true kind (t, false) ({next_id, terms}: context) =
+      mk_context next_id (Kindtab.update (kind, (t, true)) terms)
+  | note_atom _ _ _ cx = cx
+
+fun with_unique_id kind is_atom (e as Argo_Expr.E (k, _)) ts ({next_id, terms}: context) =
+  let val t = T (M (next_id, e), k, ts)
+  in ((t, false), mk_context (next_id + 1) (Kindtab.update (kind, (t, is_atom)) terms)) end
+
+fun unique kind is_atom e ts (cx as {terms, ...}: context) =
+  (case Kindtab.lookup terms kind of
+    SOME tp => (tp, note_atom is_atom kind tp cx)
+  | NONE => with_unique_id kind is_atom e ts cx)
+
+
+(* identifying expressions *)
+
+(*
+  Only atoms, i.e., boolean propositons, and if-then-else expressions need to be identified.
+  Other terms are identified implicitly. The identification process works bottom-up.
+
+  The solver core needs to know whether an atom has already been added. Likewise, the clausifier
+  needs to know whether an if-then-else expression has already been lifted. Therefore,
+  the identified term is marked as either "new" when identified for the first time or
+  "known" when it has already been identified before.
+*)
+
+datatype item = Expr of Argo_Expr.expr | Term of term
+datatype identified = New of term | Known of term
+
+fun identify_head is_atom e (ts, cx) = unique (kind_of e (map id_of ts)) is_atom e ts cx
+
+fun identify is_atom (e as Argo_Expr.E (_, es)) cx =
+  identify_head is_atom e (fold_map (apfst fst oo identify false) es cx)
+
+fun identified (t, true) = Known t
+  | identified (t, false) = New t
+
+fun identify_item (Expr e) cx = identify true e cx |>> identified
+  | identify_item (Term (t as T (_, _, ts))) cx =
+      identify_head true (expr_of t) (ts, cx) |>> identified
+
+end
+
+structure Argo_Termtab = Table(type key = Argo_Term.term val ord = Argo_Term.term_ord)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_thy.ML	Fri Sep 30 10:00:49 2016 +0200
@@ -0,0 +1,125 @@
+(*  Title:      Tools/Argo/argo_theory.ML
+    Author:     Sascha Boehme
+
+Combination of all theory solvers.
+
+Currently, it is assumed that theories have distinct domains. Theory solvers do not
+exchange knowledge among each other. This should be changed in the future. Relevant work is:
+
+  Greg Nelson and Derek C. Oppen. Simplification by cooperating decision procedures. In ACM
+  Transactions on Programming Languages and Systems, 1(2):245-257, 1979.
+
+  Leonardo de Moura and Nikolaj Bj/orner. Model-based Theory Combination. In Electronic Notes
+  in Theoretical Computer Science, volume 198(2), pages 37-49, 2008.
+*)
+
+signature ARGO_THY =
+sig
+  (* context *)
+  type context
+  val context: context
+
+  (* simplification *)
+  val simplify: Argo_Rewr.context -> Argo_Rewr.context
+  
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> Argo_Lit.literal option * context
+
+  (* main operations *)
+  val prepare: context -> context
+  val assume: Argo_Common.literal -> context -> Argo_Lit.literal Argo_Common.implied * context
+  val check: context -> Argo_Lit.literal Argo_Common.implied * context
+  val explain: Argo_Lit.literal -> context -> Argo_Cls.clause * context
+  val add_level: context -> context
+  val backtrack: context -> context
+end
+
+structure Argo_Thy: ARGO_THY =
+struct
+
+(* context *)
+
+type context = Argo_Cc.context * Argo_Simplex.context
+
+val context = (Argo_Cc.context, Argo_Simplex.context)
+
+fun map_cc f (cc, simplex) =
+  let val (x, cc) = f cc
+  in (x, (cc, simplex)) end
+
+fun map_simplex f (cc, simplex) =
+  let val (x, simplex) = f simplex
+  in (x, (cc, simplex)) end
+
+
+(* simplification *)
+
+val simplify = Argo_Cc.simplify #> Argo_Simplex.simplify
+
+
+(* enriching the context *)
+
+fun add_atom t (cc, simplex) =
+  let
+    val (lit1, cc) = Argo_Cc.add_atom t cc
+    val (lit2, simplex) = Argo_Simplex.add_atom t simplex
+  in
+    (case fold (union Argo_Lit.eq_lit o the_list) [lit1, lit2] [] of
+      [] => (NONE, (cc, simplex))
+    | [lit] => (SOME lit, (cc, simplex))
+    | _ => raise Fail "unsynchronized theory solvers")
+  end
+    
+
+
+(* main operations *)
+
+fun prepare (cc, simplex) = (cc, Argo_Simplex.prepare simplex)
+
+local
+
+exception CONFLICT of Argo_Cls.clause * context
+
+datatype tag = All | Cc | Simplex
+
+fun apply f cx =
+  (case f cx of
+    (Argo_Common.Conflict cls, cx) => raise CONFLICT (cls, cx)
+  | (Argo_Common.Implied lits, cx) => (lits, cx))
+
+fun with_lits tag f (txs, lits, cx) =
+  let val (lits', cx) = f cx
+  in (fold (fn l => cons (tag, (l, NONE))) lits' txs, union Argo_Lit.eq_lit lits' lits, cx) end
+
+fun apply0 (tag, f) = with_lits tag (apply f)
+fun apply1 (tag, f) (tag', x) = if tag <> tag' then with_lits tag (apply (f x)) else I
+
+val assumes = [(Cc, map_cc o Argo_Cc.assume), (Simplex, map_simplex o Argo_Simplex.assume)]
+val checks = [(Cc, map_cc Argo_Cc.check), (Simplex, map_simplex Argo_Simplex.check)]
+
+fun propagate ([], lits, cx) = (Argo_Common.Implied lits, cx)
+  | propagate (txs, lits, cx) = propagate (fold_product apply1 assumes txs ([], lits, cx))
+
+in
+
+fun assume lp cx = propagate ([(All, lp)], [], cx)
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+fun check cx = propagate (fold apply0 checks ([], [], cx))
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+end
+
+fun explain lit (cc, simplex) =
+  (case Argo_Cc.explain lit cc of
+    SOME (cls, cc) => (cls, (cc, simplex))
+  | NONE =>
+      (case Argo_Simplex.explain lit simplex of
+        SOME (cls, simplex) => (cls, (cc, simplex))
+      | NONE => raise Fail "bad literal without explanation"))
+
+fun add_level (cc, simplex) = (Argo_Cc.add_level cc, Argo_Simplex.add_level simplex)
+
+fun backtrack (cc, simplex) = (Argo_Cc.backtrack cc, Argo_Simplex.backtrack simplex)
+
+end