merged
authorhoelzl
Tue, 23 Feb 2010 17:55:00 +0100
changeset 35329 cac5a37fb638
parent 35327 c76b7dcd77ce (diff)
parent 35328 e8888458dce3 (current diff)
child 35330 e7eb254db165
child 35337 48e23510a3d8
child 35349 f9801fdeb789
merged
src/HOL/IsaMakefile
--- a/NEWS	Tue Feb 23 17:33:03 2010 +0100
+++ b/NEWS	Tue Feb 23 17:55:00 2010 +0100
@@ -122,9 +122,6 @@
 
 INCOMPATIBILITY.
 
-* New theory Algebras contains generic algebraic structures and
-generic algebraic operations.
-
 * HOLogic.strip_psplit: types are returned in syntactic order, similar
 to other strip and tuple operations.  INCOMPATIBILITY.
 
--- a/doc-src/Nitpick/nitpick.tex	Tue Feb 23 17:33:03 2010 +0100
+++ b/doc-src/Nitpick/nitpick.tex	Tue Feb 23 17:55:00 2010 +0100
@@ -150,11 +150,11 @@
 \postw
 
 The results presented here were obtained using the JNI version of MiniSat and
-with multithreading disabled to reduce nondeterminism. This was done by adding
-the line
+with multithreading disabled to reduce nondeterminism and a time limit of
+15~seconds (instead of 30~seconds). This was done by adding the line
 
 \prew
-\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSat\_JNI}, \,\textit{max\_threads}~= 1]
+\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSat\_JNI}, \,\textit{max\_threads}~= 1, \,\textit{timeout} = 15$\,s$]
 \postw
 
 after the \textbf{begin} keyword. The JNI version of MiniSat is bundled with
@@ -499,7 +499,7 @@
 
 \prew
 \textbf{lemma} ``$P~\textit{Suc}$'' \\
-\textbf{nitpick} [\textit{card} = 1--6] \\[2\smallskipamount]
+\textbf{nitpick} \\[2\smallskipamount]
 \slshape
 Nitpick found no counterexample.
 \postw
@@ -1617,7 +1617,7 @@
 ``$w \in A \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = b] + 1$'' \\
 ``$w \in B \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = b] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] + 1$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
-\slshape Nitpick found no counterexample.
+\slshape Nitpick ran out of time after checking 7 of 8 scopes.
 \postw
 
 \subsection{AA Trees}
@@ -1726,7 +1726,7 @@
 ``$\textit{dataset}~(\textit{skew}~t) = \textit{dataset}~t$'' \\
 ``$\textit{dataset}~(\textit{split}~t) = \textit{dataset}~t$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
-{\slshape Nitpick found no counterexample.}
+{\slshape Nitpick ran out of time after checking 7 of 8 scopes.}
 \postw
 
 Furthermore, applying \textit{skew} or \textit{split} to a well-formed tree
--- a/src/HOL/Algebra/abstract/Ideal2.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Algebra/abstract/Ideal2.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,6 +1,5 @@
 (*
     Ideals for commutative rings
-    $Id$
     Author: Clemens Ballarin, started 24 September 1999
 *)
 
@@ -23,9 +22,8 @@
 
 text {* Principle ideal domains *}
 
-axclass pid < "domain"
-  pid_ax: "is_ideal I ==> is_pideal I"
-
+class pid =
+  assumes pid_ax: "is_ideal (I :: 'a::domain \<Rightarrow> _) \<Longrightarrow> is_pideal I"
 
 (* is_ideal *)
 
--- a/src/HOL/Algebras.thy	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,55 +0,0 @@
-(*  Title:      HOL/Algebras.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Generic algebraic structures and operations *}
-
-theory Algebras
-imports HOL
-begin
-
-text {*
-  These locales provide basic structures for interpretation into
-  bigger structures;  extensions require careful thinking, otherwise
-  undesired effects may occur due to interpretation.
-*}
-
-ML {*
-structure Ac_Simps = Named_Thms(
-  val name = "ac_simps"
-  val description = "associativity and commutativity simplification rules"
-)
-*}
-
-setup Ac_Simps.setup
-
-locale semigroup =
-  fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)
-  assumes assoc [ac_simps]: "a * b * c = a * (b * c)"
-
-locale abel_semigroup = semigroup +
-  assumes commute [ac_simps]: "a * b = b * a"
-begin
-
-lemma left_commute [ac_simps]:
-  "b * (a * c) = a * (b * c)"
-proof -
-  have "(b * a) * c = (a * b) * c"
-    by (simp only: commute)
-  then show ?thesis
-    by (simp only: assoc)
-qed
-
-end
-
-locale semilattice = abel_semigroup +
-  assumes idem [simp]: "a * a = a"
-begin
-
-lemma left_idem [simp]:
-  "a * (a * b) = a * b"
-  by (simp add: assoc [symmetric])
-
-end
-
-end
\ No newline at end of file
--- a/src/HOL/Bali/Decl.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Bali/Decl.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -230,18 +230,22 @@
 
 datatype memberid = fid vname | mid sig
 
-axclass has_memberid < "type"
-consts
- memberid :: "'a::has_memberid \<Rightarrow> memberid"
+class has_memberid =
+  fixes memberid :: "'a \<Rightarrow> memberid"
 
-instance memberdecl::has_memberid ..
+instantiation memberdecl :: has_memberid
+begin
 
-defs (overloaded)
+definition
 memberdecl_memberid_def:
   "memberid m \<equiv> (case m of
                     fdecl (vn,f)  \<Rightarrow> fid vn
                   | mdecl (sig,m) \<Rightarrow> mid sig)"
 
+instance ..
+
+end
+
 lemma memberid_fdecl_simp[simp]: "memberid (fdecl (vn,f)) = fid vn"
 by (simp add: memberdecl_memberid_def)
 
@@ -254,12 +258,17 @@
 lemma memberid_mdecl_simp1: "memberid (mdecl m) = mid (fst m)"
 by (cases m) (simp add: memberdecl_memberid_def)
 
-instance * :: (type, has_memberid) has_memberid ..
+instantiation * :: (type, has_memberid) has_memberid
+begin
 
-defs (overloaded)
+definition
 pair_memberid_def:
   "memberid p \<equiv> memberid (snd p)"
 
+instance ..
+
+end
+
 lemma memberid_pair_simp[simp]: "memberid (c,m) = memberid m"
 by (simp add: pair_memberid_def)
 
--- a/src/HOL/Bali/DeclConcepts.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Bali/DeclConcepts.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -79,41 +79,60 @@
 text {* overloaded selector @{text accmodi} to select the access modifier 
 out of various HOL types *}
 
-axclass has_accmodi < "type"
-consts accmodi:: "'a::has_accmodi \<Rightarrow> acc_modi"
+class has_accmodi =
+  fixes accmodi:: "'a \<Rightarrow> acc_modi"
+
+instantiation acc_modi :: has_accmodi
+begin
 
-instance acc_modi::has_accmodi ..
+definition
+  acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
 
-defs (overloaded)
-acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
+instance ..
+
+end
 
 lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a"
 by (simp add: acc_modi_accmodi_def)
 
-instance decl_ext_type:: ("type") has_accmodi ..
+instantiation decl_ext_type:: (type) has_accmodi
+begin
 
-defs (overloaded)
-decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
+definition
+  decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
 
+instance ..
+
+end
 
 lemma decl_acc_modi_simp[simp]: "accmodi (d::('a::type) decl_scheme) = access d"
 by (simp add: decl_acc_modi_def)
 
-instance * :: ("type",has_accmodi) has_accmodi ..
+instantiation * :: (type, has_accmodi) has_accmodi
+begin
 
-defs (overloaded)
-pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
+definition
+  pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
+
+instance ..
+
+end
 
 lemma pair_acc_modi_simp[simp]: "accmodi (x,a) = (accmodi a)"
 by (simp add: pair_acc_modi_def)
 
-instance memberdecl :: has_accmodi ..
+instantiation memberdecl :: has_accmodi
+begin
 
-defs (overloaded)
-memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
+definition
+  memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
                                           fdecl f \<Rightarrow> accmodi f
                                         | mdecl m \<Rightarrow> accmodi m)"
 
+instance ..
+
+end
+
 lemma memberdecl_fdecl_acc_modi_simp[simp]:
  "accmodi (fdecl m) = accmodi m"
 by (simp add: memberdecl_acc_modi_def)
@@ -125,21 +144,35 @@
 text {* overloaded selector @{text declclass} to select the declaring class 
 out of various HOL types *}
 
-axclass has_declclass < "type"
-consts declclass:: "'a::has_declclass \<Rightarrow> qtname"
+class has_declclass =
+  fixes declclass:: "'a \<Rightarrow> qtname"
+
+instantiation qtname_ext_type :: (type) has_declclass
+begin
 
-instance qtname_ext_type::("type") has_declclass ..
+definition
+  "declclass q \<equiv> \<lparr> pid = pid q, tid = tid q \<rparr>"
+
+instance ..
 
-defs (overloaded)
-qtname_declclass_def: "declclass (q::qtname) \<equiv> q"
+end
+
+lemma qtname_declclass_def:
+  "declclass q \<equiv> (q::qtname)"
+  by (induct q) (simp add: declclass_qtname_ext_type_def)
 
 lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q"
 by (simp add: qtname_declclass_def)
 
-instance * :: ("has_declclass","type") has_declclass ..
+instantiation * :: (has_declclass, type) has_declclass
+begin
 
-defs (overloaded)
-pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
+definition
+  pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
+
+instance ..
+
+end
 
 lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c" 
 by (simp add: pair_declclass_def)
@@ -147,25 +180,38 @@
 text {* overloaded selector @{text is_static} to select the static modifier 
 out of various HOL types *}
 
+class has_static =
+  fixes is_static :: "'a \<Rightarrow> bool"
 
-axclass has_static < "type"
-consts is_static :: "'a::has_static \<Rightarrow> bool"
+instantiation decl_ext_type :: (has_static) has_static
+begin
 
-instance decl_ext_type :: ("has_static") has_static ..
+instance ..
+
+end
 
-instance member_ext_type :: ("type") has_static ..
+instantiation member_ext_type :: (type) has_static
+begin
+
+instance ..
 
-defs (overloaded)
-static_field_type_is_static_def: 
- "is_static (m::('b member_scheme)) \<equiv> static m"
+end
+
+axiomatization where
+  static_field_type_is_static_def: "is_static (m::('a member_scheme)) \<equiv> static m"
 
 lemma member_is_static_simp: "is_static (m::'a member_scheme) = static m"
 by (simp add: static_field_type_is_static_def)
 
-instance * :: ("type","has_static") has_static ..
+instantiation * :: (type, has_static) has_static
+begin
 
-defs (overloaded)
-pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
+definition
+  pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
+
+instance ..
+
+end
 
 lemma pair_is_static_simp [simp]: "is_static (x,s) = is_static s"
 by (simp add: pair_is_static_def)
@@ -173,14 +219,19 @@
 lemma pair_is_static_simp1: "is_static p = is_static (snd p)"
 by (simp add: pair_is_static_def)
 
-instance memberdecl:: has_static ..
+instantiation memberdecl :: has_static
+begin
 
-defs (overloaded)
+definition
 memberdecl_is_static_def: 
  "is_static m \<equiv> (case m of
                     fdecl f \<Rightarrow> is_static f
                   | mdecl m \<Rightarrow> is_static m)"
 
+instance ..
+
+end
+
 lemma memberdecl_is_static_fdecl_simp[simp]:
  "is_static (fdecl f) = is_static f"
 by (simp add: memberdecl_is_static_def)
@@ -389,18 +440,32 @@
 text {* overloaded selector @{text resTy} to select the result type 
 out of various HOL types *}
 
-axclass has_resTy < "type"
-consts resTy:: "'a::has_resTy \<Rightarrow> ty"
+class has_resTy =
+  fixes resTy:: "'a \<Rightarrow> ty"
+
+instantiation decl_ext_type :: (has_resTy) has_resTy
+begin
 
-instance decl_ext_type :: ("has_resTy") has_resTy ..
+instance ..
+
+end
+
+instantiation member_ext_type :: (has_resTy) has_resTy
+begin
 
-instance member_ext_type :: ("has_resTy") has_resTy ..
+instance ..
 
-instance mhead_ext_type :: ("type") has_resTy ..
+end
+
+instantiation mhead_ext_type :: (type) has_resTy
+begin
 
-defs (overloaded)
-mhead_ext_type_resTy_def: 
- "resTy (m::('b mhead_scheme)) \<equiv> resT m"
+instance ..
+
+end
+
+axiomatization where
+  mhead_ext_type_resTy_def: "resTy (m::('b mhead_scheme)) \<equiv> resT m"
 
 lemma mhead_resTy_simp: "resTy (m::'a mhead_scheme) = resT m"
 by (simp add: mhead_ext_type_resTy_def)
@@ -408,10 +473,15 @@
 lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m"
 by (simp add: mhead_def mhead_resTy_simp)
 
-instance * :: ("type","has_resTy") has_resTy ..
+instantiation * :: ("type","has_resTy") has_resTy
+begin
 
-defs (overloaded)
-pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
+definition
+  pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
+
+instance ..
+
+end
 
 lemma pair_resTy_simp[simp]: "resTy (x,m) = resTy m"
 by (simp add: pair_resTy_def)
--- a/src/HOL/Bali/Name.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Bali/Name.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -48,29 +48,34 @@
           pid :: pname  
           tid :: tname
 
-axclass has_pname < "type"
-consts pname::"'a::has_pname \<Rightarrow> pname"
+class has_pname =
+  fixes pname :: "'a \<Rightarrow> pname"
 
-instance pname::has_pname ..
+instantiation pname :: has_pname
+begin
 
-defs (overloaded)
-pname_pname_def: "pname (p::pname) \<equiv> p"
+definition
+  pname_pname_def: "pname (p::pname) \<equiv> p"
 
-axclass has_tname < "type"
-consts tname::"'a::has_tname \<Rightarrow> tname"
+instance ..
+
+end
 
-instance tname::has_tname ..
+class has_tname =
+  fixes tname :: "'a \<Rightarrow> tname"
 
-defs (overloaded)
-tname_tname_def: "tname (t::tname) \<equiv> t"
+instantiation tname :: has_tname
+begin
 
-axclass has_qtname < type
-consts qtname:: "'a::has_qtname \<Rightarrow> qtname"
+definition
+  tname_tname_def: "tname (t::tname) \<equiv> t"
+
+instance ..
 
-instance qtname_ext_type :: (type) has_qtname ..
+end
 
-defs (overloaded)
-qtname_qtname_def: "qtname (q::qtname) \<equiv> q"
+definition
+  qtname_qtname_def: "qtname (q::'a qtname_ext_type) \<equiv> q"
 
 translations
   "mname"  <= "Name.mname"
--- a/src/HOL/Bali/Term.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Bali/Term.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -295,8 +295,8 @@
 following.
 *}
 
-axclass inj_term < "type"
-consts inj_term:: "'a::inj_term \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
+class inj_term =
+  fixes inj_term:: "'a \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
 
 text {* How this overloaded injections work can be seen in the theory 
 @{text DefiniteAssignment}. Other big inductive relations on
@@ -308,10 +308,15 @@
 as bridge between the different conventions.  
 *}
 
-instance stmt::inj_term ..
+instantiation stmt :: inj_term
+begin
 
-defs (overloaded)
-stmt_inj_term_def: "\<langle>c::stmt\<rangle> \<equiv> In1r c"
+definition
+  stmt_inj_term_def: "\<langle>c::stmt\<rangle> \<equiv> In1r c"
+
+instance ..
+
+end
 
 lemma stmt_inj_term_simp: "\<langle>c::stmt\<rangle> = In1r c"
 by (simp add: stmt_inj_term_def)
@@ -319,10 +324,15 @@
 lemma  stmt_inj_term [iff]: "\<langle>x::stmt\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
   by (simp add: stmt_inj_term_simp)
 
-instance expr::inj_term ..
+instantiation expr :: inj_term
+begin
 
-defs (overloaded)
-expr_inj_term_def: "\<langle>e::expr\<rangle> \<equiv> In1l e"
+definition
+  expr_inj_term_def: "\<langle>e::expr\<rangle> \<equiv> In1l e"
+
+instance ..
+
+end
 
 lemma expr_inj_term_simp: "\<langle>e::expr\<rangle> = In1l e"
 by (simp add: expr_inj_term_def)
@@ -330,10 +340,15 @@
 lemma expr_inj_term [iff]: "\<langle>x::expr\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
   by (simp add: expr_inj_term_simp)
 
-instance var::inj_term ..
+instantiation var :: inj_term
+begin
 
-defs (overloaded)
-var_inj_term_def: "\<langle>v::var\<rangle> \<equiv> In2 v"
+definition
+  var_inj_term_def: "\<langle>v::var\<rangle> \<equiv> In2 v"
+
+instance ..
+
+end
 
 lemma var_inj_term_simp: "\<langle>v::var\<rangle> = In2 v"
 by (simp add: var_inj_term_def)
@@ -341,10 +356,32 @@
 lemma var_inj_term [iff]: "\<langle>x::var\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
   by (simp add: var_inj_term_simp)
 
-instance "list":: (type) inj_term ..
+class expr_of =
+  fixes expr_of :: "'a \<Rightarrow> expr"
+
+instantiation expr :: expr_of
+begin
+
+definition
+  "expr_of = (\<lambda>(e::expr). e)"
+
+instance ..
+
+end 
 
-defs (overloaded)
-expr_list_inj_term_def: "\<langle>es::expr list\<rangle> \<equiv> In3 es"
+instantiation list :: (expr_of) inj_term
+begin
+
+definition
+  "\<langle>es::'a list\<rangle> \<equiv> In3 (map expr_of es)"
+
+instance ..
+
+end
+
+lemma expr_list_inj_term_def:
+  "\<langle>es::expr list\<rangle> \<equiv> In3 es"
+  by (simp add: inj_term_list_def expr_of_expr_def)
 
 lemma expr_list_inj_term_simp: "\<langle>es::expr list\<rangle> = In3 es"
 by (simp add: expr_list_inj_term_def)
--- a/src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -56,10 +56,12 @@
 
 text {* Let's find out which assertions of @{text max} are hard to prove: *}
 
-boogie_status (narrow timeout: 4) max
+boogie_status (narrow step_timeout: 4 final_timeout: 15) max
   (simp add: labels, (fast | metis)?)
-  -- {* The argument @{text timeout} is optional, restricting the runtime of
-        each narrowing step by the given number of seconds. *}
+  -- {* The argument @{text step_timeout} is optional, restricting the runtime
+        of each intermediate narrowing step by the given number of seconds. *}
+  -- {* The argument @{text final_timeout} is optional, restricting the
+        runtime of the final narrowing steps by the given number of seconds. *}
   -- {* The last argument should be a method to be applied at each narrowing
         step. Note that different methods may be composed to one method by
         a language similar to regular expressions. *}
--- a/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -156,26 +156,27 @@
     |> Seq.hd
     |> Proof.global_done_proof
 in
-fun boogie_narrow_vc timeout vc_name meth thy =
+fun boogie_narrow_vc (quick, timeout) vc_name meth thy =
   let
-    val tp = TimeLimit.timeLimit (Time.fromSeconds timeout) (prove thy meth)
+    fun tp t = TimeLimit.timeLimit (Time.fromSeconds t) (prove thy meth)
 
-    fun try_vc (tag, split_tag) split vc = (trying tag;
-      (case try tp vc of
+    fun try_vc t (tag, split_tag) split vc = (trying tag;
+      (case try (tp t) vc of
         SOME _ => (success_on tag; [])
       | NONE => (failure_on tag split_tag; split vc)))
 
     fun some_asserts vc =
-      try_vc (string_of_asserts vc,
-        if Boogie_VCs.size_of vc = 1 then "." else ", further splitting ...")
-        (divide some_asserts) vc
+      let
+        val (t, sep) = if Boogie_VCs.size_of vc = 1 then (timeout, ".")
+          else (quick, ", further splitting ...")
+      in try_vc t (string_of_asserts vc, sep) (divide some_asserts) vc end
 
     fun single_path p =
-      try_vc (string_of_path p, ", splitting into assertions ...")
+      try_vc quick (string_of_path p, ", splitting into assertions ...")
         (divide some_asserts)
 
     val complete_vc =
-      try_vc ("full goal", ", splitting into paths ...")
+      try_vc quick ("full goal", ", splitting into paths ...")
         (par_map (uncurry single_path) o itemize_paths o Boogie_VCs.paths_of)
 
     val unsolved = complete_vc (get_vc thy vc_name)
@@ -262,15 +263,18 @@
       (Toplevel.print o Toplevel.local_theory_to_proof NONE (boogie_vc args))))
 
 
-val default_timeout = 5
+val quick_timeout = 5
+val default_timeout = 20
+
+fun timeout name = Scan.optional (scan_val name OuterParse.nat)
 
 val status_test =
   scan_arg (
-    (Args.$$$ "scan" >> K boogie_scan_vc ||
-     Args.$$$ "narrow" >> K boogie_narrow_vc) --
-    Scan.optional (scan_val "timeout" OuterParse.nat) default_timeout) --
+    Args.$$$ "scan" |-- timeout "timeout" quick_timeout >> boogie_scan_vc ||
+    Args.$$$ "narrow" |-- timeout "step_timeout" quick_timeout --
+      timeout "final_timeout" default_timeout >> boogie_narrow_vc) --
   vc_name -- Method.parse >>
-  (fn (((f, timeout), vc_name), meth) => f timeout vc_name meth)
+  (fn ((f, vc_name), meth) => f vc_name meth)
 
 val status_vc =
   (scan_arg
--- a/src/HOL/Code_Evaluation.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Code_Evaluation.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -76,7 +76,8 @@
         andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
     in if need_inst then add_term_of tyco raw_vs thy else thy end;
 in
-  Code.type_interpretation ensure_term_of
+  Code.datatype_interpretation ensure_term_of
+  #> Code.abstype_interpretation ensure_term_of
 end
 *}
 
@@ -114,7 +115,7 @@
       val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
     in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
 in
-  Code.type_interpretation ensure_term_of_code
+  Code.datatype_interpretation ensure_term_of_code
 end
 *}
 
--- a/src/HOL/Complex_Main.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Complex_Main.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -9,7 +9,7 @@
   Log
   Ln
   Taylor
-  Integration
+  Deriv
 begin
 
 end
--- a/src/HOL/Groups.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Groups.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -9,6 +9,65 @@
 uses ("~~/src/Provers/Arith/abel_cancel.ML")
 begin
 
+subsection {* Fact collections *}
+
+ML {*
+structure Algebra_Simps = Named_Thms(
+  val name = "algebra_simps"
+  val description = "algebra simplification rules"
+)
+*}
+
+setup Algebra_Simps.setup
+
+text{* The rewrites accumulated in @{text algebra_simps} deal with the
+classical algebraic structures of groups, rings and family. They simplify
+terms by multiplying everything out (in case of a ring) and bringing sums and
+products into a canonical form (by ordered rewriting). As a result it decides
+group and ring equalities but also helps with inequalities.
+
+Of course it also works for fields, but it knows nothing about multiplicative
+inverses or division. This is catered for by @{text field_simps}. *}
+
+
+ML {*
+structure Ac_Simps = Named_Thms(
+  val name = "ac_simps"
+  val description = "associativity and commutativity simplification rules"
+)
+*}
+
+setup Ac_Simps.setup
+
+
+subsection {* Abstract structures *}
+
+text {*
+  These locales provide basic structures for interpretation into
+  bigger structures;  extensions require careful thinking, otherwise
+  undesired effects may occur due to interpretation.
+*}
+
+locale semigroup =
+  fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)
+  assumes assoc [ac_simps]: "a * b * c = a * (b * c)"
+
+locale abel_semigroup = semigroup +
+  assumes commute [ac_simps]: "a * b = b * a"
+begin
+
+lemma left_commute [ac_simps]:
+  "b * (a * c) = a * (b * c)"
+proof -
+  have "(b * a) * c = (a * b) * c"
+    by (simp only: commute)
+  then show ?thesis
+    by (simp only: assoc)
+qed
+
+end
+
+
 subsection {* Generic operations *}
 
 class zero = 
@@ -64,37 +123,6 @@
 
 use "~~/src/Provers/Arith/abel_cancel.ML"
 
-text {*
-  The theory of partially ordered groups is taken from the books:
-  \begin{itemize}
-  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
-  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
-  \end{itemize}
-  Most of the used notions can also be looked up in 
-  \begin{itemize}
-  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
-  \item \emph{Algebra I} by van der Waerden, Springer.
-  \end{itemize}
-*}
-
-ML {*
-structure Algebra_Simps = Named_Thms(
-  val name = "algebra_simps"
-  val description = "algebra simplification rules"
-)
-*}
-
-setup Algebra_Simps.setup
-
-text{* The rewrites accumulated in @{text algebra_simps} deal with the
-classical algebraic structures of groups, rings and family. They simplify
-terms by multiplying everything out (in case of a ring) and bringing sums and
-products into a canonical form (by ordered rewriting). As a result it decides
-group and ring equalities but also helps with inequalities.
-
-Of course it also works for fields, but it knows nothing about multiplicative
-inverses or division. This is catered for by @{text field_simps}. *}
-
 
 subsection {* Semigroups and Monoids *}
 
@@ -144,19 +172,6 @@
 
 theorems mult_ac = mult_assoc mult_commute mult_left_commute
 
-class ab_semigroup_idem_mult = ab_semigroup_mult +
-  assumes mult_idem: "x * x = x"
-
-sublocale ab_semigroup_idem_mult < times!: semilattice times proof
-qed (fact mult_idem)
-
-context ab_semigroup_idem_mult
-begin
-
-lemmas mult_left_idem = times.left_idem
-
-end
-
 class monoid_add = zero + semigroup_add +
   assumes add_0_left [simp]: "0 + a = a"
     and add_0_right [simp]: "a + 0 = a"
@@ -411,6 +426,19 @@
 
 subsection {* (Partially) Ordered Groups *} 
 
+text {*
+  The theory of partially ordered groups is taken from the books:
+  \begin{itemize}
+  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
+  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
+  \end{itemize}
+  Most of the used notions can also be looked up in 
+  \begin{itemize}
+  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
+  \item \emph{Algebra I} by van der Waerden, Springer.
+  \end{itemize}
+*}
+
 class ordered_ab_semigroup_add = order + ab_semigroup_add +
   assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"
 begin
--- a/src/HOL/Hoare/Examples.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/Examples.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Hoare/Examples.thy
-    ID:         $Id$
     Author:     Norbert Galm
     Copyright   1998 TUM
 
 Various examples.
 *)
 
-theory Examples imports Hoare Arith2 begin
+theory Examples imports Hoare_Logic Arith2 begin
 
 (*** ARITHMETIC ***)
 
--- a/src/HOL/Hoare/ExamplesAbort.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/ExamplesAbort.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Hoare/ExamplesAbort.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   1998 TUM
 
 Some small examples for programs that may abort.
 *)
 
-theory ExamplesAbort imports HoareAbort begin
+theory ExamplesAbort imports Hoare_Logic_Abort begin
 
 lemma "VARS x y z::nat
  {y = z & z \<noteq> 0} z \<noteq> 0 \<rightarrow> x := y div z {x = 1}"
--- a/src/HOL/Hoare/HeapSyntax.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/HeapSyntax.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -3,7 +3,7 @@
     Copyright   2002 TUM
 *)
 
-theory HeapSyntax imports Hoare Heap begin
+theory HeapSyntax imports Hoare_Logic Heap begin
 
 subsection "Field access and update"
 
--- a/src/HOL/Hoare/HeapSyntaxAbort.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/HeapSyntaxAbort.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -3,7 +3,7 @@
     Copyright   2002 TUM
 *)
 
-theory HeapSyntaxAbort imports HoareAbort Heap begin
+theory HeapSyntaxAbort imports Hoare_Logic_Abort Heap begin
 
 subsection "Field access and update"
 
--- a/src/HOL/Hoare/Hoare.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/Hoare.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,245 +1,9 @@
-(*  Title:      HOL/Hoare/Hoare.thy
-    Author:     Leonor Prensa Nieto & Tobias Nipkow
-    Copyright   1998 TUM
-
-Sugared semantic embedding of Hoare logic.
-Strictly speaking a shallow embedding (as implemented by Norbert Galm
-following Mike Gordon) would suffice. Maybe the datatype com comes in useful
-later.
+(*  Author:     Tobias Nipkow
+    Copyright   1998-2003 TUM
 *)
 
 theory Hoare
-imports Main
-uses ("hoare_tac.ML")
+imports Examples ExamplesAbort Pointers0 Pointer_Examples Pointer_ExamplesAbort SchorrWaite Separation
 begin
 
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
-
-datatype
- 'a com = Basic "'a \<Rightarrow> 'a"
-   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
-   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
-   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
-
-abbreviation annskip ("SKIP") where "SKIP == Basic id"
-
-types 'a sem = "'a => 'a => bool"
-
-consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
-primrec
-"iter 0 b S = (%s s'. s ~: b & (s=s'))"
-"iter (Suc n) b S = (%s s'. s : b & (? s''. S s s'' & iter n b S s'' s'))"
-
-consts Sem :: "'a com => 'a sem"
-primrec
-"Sem(Basic f) s s' = (s' = f s)"
-"Sem(c1;c2) s s' = (? s''. Sem c1 s s'' & Sem c2 s'' s')"
-"Sem(IF b THEN c1 ELSE c2 FI) s s' = ((s  : b --> Sem c1 s s') &
-                                      (s ~: b --> Sem c2 s s'))"
-"Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')"
-
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
-  "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
-
-
-
-(** parse translations **)
-
-syntax
-  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
-
-syntax
- "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
-                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
-syntax ("" output)
- "_hoare"      :: "['a assn,'a com,'a assn] => bool"
-                 ("{_} // _ // {_}" [0,55,0] 50)
-ML {*
-
-local
-
-fun abs((a,T),body) =
-  let val a = absfree(a, dummyT, body)
-  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
-in
-
-fun mk_abstuple [x] body = abs (x, body)
-  | mk_abstuple (x::xs) body =
-      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
-
-fun mk_fbody a e [x as (b,_)] = if a=b then e else Syntax.free b
-  | mk_fbody a e ((b,_)::xs) =
-      Syntax.const @{const_syntax Pair} $ (if a=b then e else Syntax.free b) $ mk_fbody a e xs;
-
-fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
 end
-*}
-
-(* bexp_tr & assn_tr *)
-(*all meta-variables for bexp except for TRUE are translated as if they
-  were boolean expressions*)
-ML{*
-fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
-  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
-
-fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
-*}
-(* com_tr *)
-ML{*
-fun com_tr (Const(@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
-      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
-  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
-  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
-      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
-  | com_tr t _ = t (* if t is just a Free/Var *)
-*}
-
-(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
-ML{*
-local
-
-fun var_tr(Free(a,_)) = (a,Bound 0) (* Bound 0 = dummy term *)
-  | var_tr(Const (@{syntax_const "_constrain"}, _) $ (Free (a,_)) $ T) = (a,T);
-
-fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
-  | vars_tr t = [var_tr t]
-
-in
-fun hoare_vars_tr [vars, pre, prg, post] =
-      let val xs = vars_tr vars
-      in Syntax.const @{const_syntax Valid} $
-         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
-      end
-  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
-end
-*}
-
-parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
-
-
-(*****************************************************************************)
-
-(*** print translations ***)
-ML{*
-fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
-                            subst_bound (Syntax.free v, dest_abstuple body)
-  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
-  | dest_abstuple trm = trm;
-
-fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
-  | abs2list (Abs(x,T,t)) = [Free (x, T)]
-  | abs2list _ = [];
-
-fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
-  | mk_ts (Abs(x,_,t)) = mk_ts t
-  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
-  | mk_ts t = [t];
-
-fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
-           ((Syntax.free x)::(abs2list t), mk_ts t)
-  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
-  | mk_vts t = raise Match;
-
-fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
-  | find_ch ((v,t)::vts) i xs =
-      if t = Bound i then find_ch vts (i-1) xs
-      else (true, (v, subst_bounds (xs, t)));
-
-fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
-  | is_f (Abs(x,_,t)) = true
-  | is_f t = false;
-*}
-
-(* assn_tr' & bexp_tr'*)
-ML{*
-fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | assn_tr' (Const (@{const_syntax inter}, _) $
-        (Const (@{const_syntax Collect},_) $ T1) $ (Const (@{const_syntax Collect},_) $ T2)) =
-      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
-  | assn_tr' t = t;
-
-fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | bexp_tr' t = t;
-*}
-
-(*com_tr' *)
-ML{*
-fun mk_assign f =
-  let val (vs, ts) = mk_vts f;
-      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
-  in
-    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
-    else Syntax.const @{const_syntax annskip}
-  end;
-
-fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
-      if is_f f then mk_assign f
-      else Syntax.const @{const_syntax Basic} $ f
-  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
-      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
-      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
-      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
-  | com_tr' t = t;
-
-fun spec_tr' [p, c, q] =
-  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
-*}
-
-print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
-
-lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
-by (auto simp:Valid_def)
-
-lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
-by (auto simp:Valid_def)
-
-lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
-by (auto simp:Valid_def)
-
-lemma CondRule:
- "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
-  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
-by (auto simp:Valid_def)
-
-lemma iter_aux: "! s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
-       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)";
-apply(induct n)
- apply clarsimp
-apply(simp (no_asm_use))
-apply blast
-done
-
-lemma WhileRule:
- "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
-apply (clarsimp simp:Valid_def)
-apply(drule iter_aux)
-  prefer 2 apply assumption
- apply blast
-apply blast
-done
-
-
-lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
-  by blast
-
-lemmas AbortRule = SkipRule  -- "dummy version"
-use "hoare_tac.ML"
-
-method_setup vcg = {*
-  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
-  "verification condition generator"
-
-method_setup vcg_simp = {*
-  Scan.succeed (fn ctxt =>
-    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
-  "verification condition generator plus simplification"
-
-end
--- a/src/HOL/Hoare/HoareAbort.thy	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,269 +0,0 @@
-(*  Title:      HOL/Hoare/HoareAbort.thy
-    Author:     Leonor Prensa Nieto & Tobias Nipkow
-    Copyright   2003 TUM
-
-Like Hoare.thy, but with an Abort statement for modelling run time errors.
-*)
-
-theory HoareAbort
-imports Main
-uses ("hoare_tac.ML")
-begin
-
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
-
-datatype
- 'a com = Basic "'a \<Rightarrow> 'a"
-   | Abort
-   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
-   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
-   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
-
-abbreviation annskip ("SKIP") where "SKIP == Basic id"
-
-types 'a sem = "'a option => 'a option => bool"
-
-consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
-primrec
-"iter 0 b S = (\<lambda>s s'. s \<notin> Some ` b \<and> s=s')"
-"iter (Suc n) b S =
-  (\<lambda>s s'. s \<in> Some ` b \<and> (\<exists>s''. S s s'' \<and> iter n b S s'' s'))"
-
-consts Sem :: "'a com => 'a sem"
-primrec
-"Sem(Basic f) s s' = (case s of None \<Rightarrow> s' = None | Some t \<Rightarrow> s' = Some(f t))"
-"Sem Abort s s' = (s' = None)"
-"Sem(c1;c2) s s' = (\<exists>s''. Sem c1 s s'' \<and> Sem c2 s'' s')"
-"Sem(IF b THEN c1 ELSE c2 FI) s s' =
- (case s of None \<Rightarrow> s' = None
-  | Some t \<Rightarrow> ((t \<in> b \<longrightarrow> Sem c1 s s') \<and> (t \<notin> b \<longrightarrow> Sem c2 s s')))"
-"Sem(While b x c) s s' =
- (if s = None then s' = None else \<exists>n. iter n b (Sem c) s s')"
-
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
-  "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
-
-
-
-(** parse translations **)
-
-syntax
-  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
-
-syntax
-  "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
-                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
-syntax ("" output)
-  "_hoare"      :: "['a assn,'a com,'a assn] => bool"
-                 ("{_} // _ // {_}" [0,55,0] 50)
-ML {*
-
-local
-fun free a = Free(a,dummyT)
-fun abs((a,T),body) =
-  let val a = absfree(a, dummyT, body)
-  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
-in
-
-fun mk_abstuple [x] body = abs (x, body)
-  | mk_abstuple (x::xs) body =
-      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
-
-fun mk_fbody a e [x as (b,_)] = if a=b then e else free b
-  | mk_fbody a e ((b,_)::xs) =
-      Syntax.const @{const_syntax Pair} $ (if a=b then e else free b) $ mk_fbody a e xs;
-
-fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
-end
-*}
-
-(* bexp_tr & assn_tr *)
-(*all meta-variables for bexp except for TRUE are translated as if they
-  were boolean expressions*)
-ML{*
-fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
-  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
-
-fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
-*}
-(* com_tr *)
-ML{*
-fun com_tr (Const (@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
-      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
-  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
-  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
-      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
-  | com_tr t _ = t (* if t is just a Free/Var *)
-*}
-
-(* triple_tr *)  (* FIXME does not handle "_idtdummy" *)
-ML{*
-local
-
-fun var_tr (Free (a, _)) = (a, Bound 0) (* Bound 0 = dummy term *)
-  | var_tr (Const (@{syntax_const "_constrain"}, _) $ Free (a, _) $ T) = (a, T);
-
-fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
-  | vars_tr t = [var_tr t]
-
-in
-fun hoare_vars_tr [vars, pre, prg, post] =
-      let val xs = vars_tr vars
-      in Syntax.const @{const_syntax Valid} $
-         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
-      end
-  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
-end
-*}
-
-parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
-
-
-(*****************************************************************************)
-
-(*** print translations ***)
-ML{*
-fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
-      subst_bound (Syntax.free v, dest_abstuple body)
-  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
-  | dest_abstuple trm = trm;
-
-fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
-  | abs2list (Abs(x,T,t)) = [Free (x, T)]
-  | abs2list _ = [];
-
-fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
-  | mk_ts (Abs(x,_,t)) = mk_ts t
-  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
-  | mk_ts t = [t];
-
-fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
-           ((Syntax.free x)::(abs2list t), mk_ts t)
-  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
-  | mk_vts t = raise Match;
-
-fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
-  | find_ch ((v,t)::vts) i xs =
-      if t = Bound i then find_ch vts (i-1) xs
-      else (true, (v, subst_bounds (xs,t)));
-
-fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
-  | is_f (Abs(x,_,t)) = true
-  | is_f t = false;
-*}
-
-(* assn_tr' & bexp_tr'*)
-ML{*
-fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | assn_tr' (Const (@{const_syntax inter},_) $ (Const (@{const_syntax Collect},_) $ T1) $
-        (Const (@{const_syntax Collect},_) $ T2)) =
-      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
-  | assn_tr' t = t;
-
-fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | bexp_tr' t = t;
-*}
-
-(*com_tr' *)
-ML{*
-fun mk_assign f =
-  let val (vs, ts) = mk_vts f;
-      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
-  in
-    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
-    else Syntax.const @{const_syntax annskip}
-  end;
-
-fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
-      if is_f f then mk_assign f else Syntax.const @{const_syntax Basic} $ f
-  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
-      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
-      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
-      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
-  | com_tr' t = t;
-
-fun spec_tr' [p, c, q] =
-  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
-*}
-
-print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
-
-(*** The proof rules ***)
-
-lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
-by (auto simp:Valid_def)
-
-lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
-by (auto simp:Valid_def)
-
-lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
-by (auto simp:Valid_def)
-
-lemma CondRule:
- "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
-  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
-by (fastsimp simp:Valid_def image_def)
-
-lemma iter_aux:
- "! s s'. Sem c s s' \<longrightarrow> s \<in> Some ` (I \<inter> b) \<longrightarrow> s' \<in> Some ` I \<Longrightarrow>
-  (\<And>s s'. s \<in> Some ` I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' \<in> Some ` (I \<inter> -b))";
-apply(unfold image_def)
-apply(induct n)
- apply clarsimp
-apply(simp (no_asm_use))
-apply blast
-done
-
-lemma WhileRule:
- "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
-apply(simp add:Valid_def)
-apply(simp (no_asm) add:image_def)
-apply clarify
-apply(drule iter_aux)
-  prefer 2 apply assumption
- apply blast
-apply blast
-done
-
-lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
-by(auto simp:Valid_def)
-
-
-subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
-
-lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
-  by blast
-
-use "hoare_tac.ML"
-
-method_setup vcg = {*
-  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
-  "verification condition generator"
-
-method_setup vcg_simp = {*
-  Scan.succeed (fn ctxt =>
-    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
-  "verification condition generator plus simplification"
-
-(* Special syntax for guarded statements and guarded array updates: *)
-
-syntax
-  guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
-  array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
-translations
-  "P \<rightarrow> c" == "IF P THEN c ELSE CONST Abort FI"
-  "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
-  (* reverse translation not possible because of duplicate "a" *)
-
-text{* Note: there is no special syntax for guarded array access. Thus
-you must write @{text"j < length a \<rightarrow> a[i] := a!j"}. *}
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare/Hoare_Logic.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,245 @@
+(*  Title:      HOL/Hoare/Hoare.thy
+    Author:     Leonor Prensa Nieto & Tobias Nipkow
+    Copyright   1998 TUM
+
+Sugared semantic embedding of Hoare logic.
+Strictly speaking a shallow embedding (as implemented by Norbert Galm
+following Mike Gordon) would suffice. Maybe the datatype com comes in useful
+later.
+*)
+
+theory Hoare_Logic
+imports Main
+uses ("hoare_tac.ML")
+begin
+
+types
+    'a bexp = "'a set"
+    'a assn = "'a set"
+
+datatype
+ 'a com = Basic "'a \<Rightarrow> 'a"
+   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
+   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
+   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
+
+abbreviation annskip ("SKIP") where "SKIP == Basic id"
+
+types 'a sem = "'a => 'a => bool"
+
+consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
+primrec
+"iter 0 b S = (%s s'. s ~: b & (s=s'))"
+"iter (Suc n) b S = (%s s'. s : b & (? s''. S s s'' & iter n b S s'' s'))"
+
+consts Sem :: "'a com => 'a sem"
+primrec
+"Sem(Basic f) s s' = (s' = f s)"
+"Sem(c1;c2) s s' = (? s''. Sem c1 s s'' & Sem c2 s'' s')"
+"Sem(IF b THEN c1 ELSE c2 FI) s s' = ((s  : b --> Sem c1 s s') &
+                                      (s ~: b --> Sem c2 s s'))"
+"Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')"
+
+constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+  "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
+
+
+
+(** parse translations **)
+
+syntax
+  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
+
+syntax
+ "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
+                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
+syntax ("" output)
+ "_hoare"      :: "['a assn,'a com,'a assn] => bool"
+                 ("{_} // _ // {_}" [0,55,0] 50)
+ML {*
+
+local
+
+fun abs((a,T),body) =
+  let val a = absfree(a, dummyT, body)
+  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
+in
+
+fun mk_abstuple [x] body = abs (x, body)
+  | mk_abstuple (x::xs) body =
+      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
+
+fun mk_fbody a e [x as (b,_)] = if a=b then e else Syntax.free b
+  | mk_fbody a e ((b,_)::xs) =
+      Syntax.const @{const_syntax Pair} $ (if a=b then e else Syntax.free b) $ mk_fbody a e xs;
+
+fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
+end
+*}
+
+(* bexp_tr & assn_tr *)
+(*all meta-variables for bexp except for TRUE are translated as if they
+  were boolean expressions*)
+ML{*
+fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
+  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
+
+fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
+*}
+(* com_tr *)
+ML{*
+fun com_tr (Const(@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
+      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
+  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
+  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
+      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
+  | com_tr t _ = t (* if t is just a Free/Var *)
+*}
+
+(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
+ML{*
+local
+
+fun var_tr(Free(a,_)) = (a,Bound 0) (* Bound 0 = dummy term *)
+  | var_tr(Const (@{syntax_const "_constrain"}, _) $ (Free (a,_)) $ T) = (a,T);
+
+fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
+  | vars_tr t = [var_tr t]
+
+in
+fun hoare_vars_tr [vars, pre, prg, post] =
+      let val xs = vars_tr vars
+      in Syntax.const @{const_syntax Valid} $
+         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
+      end
+  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
+end
+*}
+
+parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
+
+
+(*****************************************************************************)
+
+(*** print translations ***)
+ML{*
+fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
+                            subst_bound (Syntax.free v, dest_abstuple body)
+  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
+  | dest_abstuple trm = trm;
+
+fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
+  | abs2list (Abs(x,T,t)) = [Free (x, T)]
+  | abs2list _ = [];
+
+fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
+  | mk_ts (Abs(x,_,t)) = mk_ts t
+  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
+  | mk_ts t = [t];
+
+fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
+           ((Syntax.free x)::(abs2list t), mk_ts t)
+  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
+  | mk_vts t = raise Match;
+
+fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
+  | find_ch ((v,t)::vts) i xs =
+      if t = Bound i then find_ch vts (i-1) xs
+      else (true, (v, subst_bounds (xs, t)));
+
+fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
+  | is_f (Abs(x,_,t)) = true
+  | is_f t = false;
+*}
+
+(* assn_tr' & bexp_tr'*)
+ML{*
+fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | assn_tr' (Const (@{const_syntax inter}, _) $
+        (Const (@{const_syntax Collect},_) $ T1) $ (Const (@{const_syntax Collect},_) $ T2)) =
+      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
+  | assn_tr' t = t;
+
+fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | bexp_tr' t = t;
+*}
+
+(*com_tr' *)
+ML{*
+fun mk_assign f =
+  let val (vs, ts) = mk_vts f;
+      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
+  in
+    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
+    else Syntax.const @{const_syntax annskip}
+  end;
+
+fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
+      if is_f f then mk_assign f
+      else Syntax.const @{const_syntax Basic} $ f
+  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
+      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
+      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
+      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
+  | com_tr' t = t;
+
+fun spec_tr' [p, c, q] =
+  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
+*}
+
+print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
+
+lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
+by (auto simp:Valid_def)
+
+lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
+by (auto simp:Valid_def)
+
+lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
+by (auto simp:Valid_def)
+
+lemma CondRule:
+ "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
+  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
+by (auto simp:Valid_def)
+
+lemma iter_aux: "! s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
+       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)";
+apply(induct n)
+ apply clarsimp
+apply(simp (no_asm_use))
+apply blast
+done
+
+lemma WhileRule:
+ "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
+apply (clarsimp simp:Valid_def)
+apply(drule iter_aux)
+  prefer 2 apply assumption
+ apply blast
+apply blast
+done
+
+
+lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
+  by blast
+
+lemmas AbortRule = SkipRule  -- "dummy version"
+use "hoare_tac.ML"
+
+method_setup vcg = {*
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
+  "verification condition generator"
+
+method_setup vcg_simp = {*
+  Scan.succeed (fn ctxt =>
+    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
+  "verification condition generator plus simplification"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,269 @@
+(*  Title:      HOL/Hoare/HoareAbort.thy
+    Author:     Leonor Prensa Nieto & Tobias Nipkow
+    Copyright   2003 TUM
+
+Like Hoare.thy, but with an Abort statement for modelling run time errors.
+*)
+
+theory Hoare_Logic_Abort
+imports Main
+uses ("hoare_tac.ML")
+begin
+
+types
+    'a bexp = "'a set"
+    'a assn = "'a set"
+
+datatype
+ 'a com = Basic "'a \<Rightarrow> 'a"
+   | Abort
+   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
+   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
+   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
+
+abbreviation annskip ("SKIP") where "SKIP == Basic id"
+
+types 'a sem = "'a option => 'a option => bool"
+
+consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
+primrec
+"iter 0 b S = (\<lambda>s s'. s \<notin> Some ` b \<and> s=s')"
+"iter (Suc n) b S =
+  (\<lambda>s s'. s \<in> Some ` b \<and> (\<exists>s''. S s s'' \<and> iter n b S s'' s'))"
+
+consts Sem :: "'a com => 'a sem"
+primrec
+"Sem(Basic f) s s' = (case s of None \<Rightarrow> s' = None | Some t \<Rightarrow> s' = Some(f t))"
+"Sem Abort s s' = (s' = None)"
+"Sem(c1;c2) s s' = (\<exists>s''. Sem c1 s s'' \<and> Sem c2 s'' s')"
+"Sem(IF b THEN c1 ELSE c2 FI) s s' =
+ (case s of None \<Rightarrow> s' = None
+  | Some t \<Rightarrow> ((t \<in> b \<longrightarrow> Sem c1 s s') \<and> (t \<notin> b \<longrightarrow> Sem c2 s s')))"
+"Sem(While b x c) s s' =
+ (if s = None then s' = None else \<exists>n. iter n b (Sem c) s s')"
+
+constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+  "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
+
+
+
+(** parse translations **)
+
+syntax
+  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
+
+syntax
+  "_hoare_abort_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
+                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
+syntax ("" output)
+  "_hoare_abort"      :: "['a assn,'a com,'a assn] => bool"
+                 ("{_} // _ // {_}" [0,55,0] 50)
+ML {*
+
+local
+fun free a = Free(a,dummyT)
+fun abs((a,T),body) =
+  let val a = absfree(a, dummyT, body)
+  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
+in
+
+fun mk_abstuple [x] body = abs (x, body)
+  | mk_abstuple (x::xs) body =
+      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
+
+fun mk_fbody a e [x as (b,_)] = if a=b then e else free b
+  | mk_fbody a e ((b,_)::xs) =
+      Syntax.const @{const_syntax Pair} $ (if a=b then e else free b) $ mk_fbody a e xs;
+
+fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
+end
+*}
+
+(* bexp_tr & assn_tr *)
+(*all meta-variables for bexp except for TRUE are translated as if they
+  were boolean expressions*)
+ML{*
+fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
+  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
+
+fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
+*}
+(* com_tr *)
+ML{*
+fun com_tr (Const (@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
+      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
+  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
+  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
+      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
+  | com_tr t _ = t (* if t is just a Free/Var *)
+*}
+
+(* triple_tr *)  (* FIXME does not handle "_idtdummy" *)
+ML{*
+local
+
+fun var_tr (Free (a, _)) = (a, Bound 0) (* Bound 0 = dummy term *)
+  | var_tr (Const (@{syntax_const "_constrain"}, _) $ Free (a, _) $ T) = (a, T);
+
+fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
+  | vars_tr t = [var_tr t]
+
+in
+fun hoare_vars_tr [vars, pre, prg, post] =
+      let val xs = vars_tr vars
+      in Syntax.const @{const_syntax Valid} $
+         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
+      end
+  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
+end
+*}
+
+parse_translation {* [(@{syntax_const "_hoare_abort_vars"}, hoare_vars_tr)] *}
+
+
+(*****************************************************************************)
+
+(*** print translations ***)
+ML{*
+fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
+      subst_bound (Syntax.free v, dest_abstuple body)
+  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
+  | dest_abstuple trm = trm;
+
+fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
+  | abs2list (Abs(x,T,t)) = [Free (x, T)]
+  | abs2list _ = [];
+
+fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
+  | mk_ts (Abs(x,_,t)) = mk_ts t
+  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
+  | mk_ts t = [t];
+
+fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
+           ((Syntax.free x)::(abs2list t), mk_ts t)
+  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
+  | mk_vts t = raise Match;
+
+fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
+  | find_ch ((v,t)::vts) i xs =
+      if t = Bound i then find_ch vts (i-1) xs
+      else (true, (v, subst_bounds (xs,t)));
+
+fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
+  | is_f (Abs(x,_,t)) = true
+  | is_f t = false;
+*}
+
+(* assn_tr' & bexp_tr'*)
+ML{*
+fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | assn_tr' (Const (@{const_syntax inter},_) $ (Const (@{const_syntax Collect},_) $ T1) $
+        (Const (@{const_syntax Collect},_) $ T2)) =
+      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
+  | assn_tr' t = t;
+
+fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | bexp_tr' t = t;
+*}
+
+(*com_tr' *)
+ML{*
+fun mk_assign f =
+  let val (vs, ts) = mk_vts f;
+      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
+  in
+    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
+    else Syntax.const @{const_syntax annskip}
+  end;
+
+fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
+      if is_f f then mk_assign f else Syntax.const @{const_syntax Basic} $ f
+  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
+      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
+      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
+      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
+  | com_tr' t = t;
+
+fun spec_tr' [p, c, q] =
+  Syntax.const @{syntax_const "_hoare_abort"} $ assn_tr' p $ com_tr' c $ assn_tr' q
+*}
+
+print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
+
+(*** The proof rules ***)
+
+lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
+by (auto simp:Valid_def)
+
+lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
+by (auto simp:Valid_def)
+
+lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
+by (auto simp:Valid_def)
+
+lemma CondRule:
+ "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
+  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
+by (fastsimp simp:Valid_def image_def)
+
+lemma iter_aux:
+ "! s s'. Sem c s s' \<longrightarrow> s \<in> Some ` (I \<inter> b) \<longrightarrow> s' \<in> Some ` I \<Longrightarrow>
+  (\<And>s s'. s \<in> Some ` I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' \<in> Some ` (I \<inter> -b))";
+apply(unfold image_def)
+apply(induct n)
+ apply clarsimp
+apply(simp (no_asm_use))
+apply blast
+done
+
+lemma WhileRule:
+ "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
+apply(simp add:Valid_def)
+apply(simp (no_asm) add:image_def)
+apply clarify
+apply(drule iter_aux)
+  prefer 2 apply assumption
+ apply blast
+apply blast
+done
+
+lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
+by(auto simp:Valid_def)
+
+
+subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
+
+lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
+  by blast
+
+use "hoare_tac.ML"
+
+method_setup vcg = {*
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
+  "verification condition generator"
+
+method_setup vcg_simp = {*
+  Scan.succeed (fn ctxt =>
+    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
+  "verification condition generator plus simplification"
+
+(* Special syntax for guarded statements and guarded array updates: *)
+
+syntax
+  guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
+  array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
+translations
+  "P \<rightarrow> c" == "IF P THEN c ELSE CONST Abort FI"
+  "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
+  (* reverse translation not possible because of duplicate "a" *)
+
+text{* Note: there is no special syntax for guarded array access. Thus
+you must write @{text"j < length a \<rightarrow> a[i] := a!j"}. *}
+
+end
--- a/src/HOL/Hoare/Pointer_Examples.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/Pointer_Examples.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Hoare/Pointers.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   2002 TUM
 
--- a/src/HOL/Hoare/Pointers0.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/Pointers0.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -9,12 +9,12 @@
 - in fact in some case they appear to get (a bit) more complicated.
 *)
 
-theory Pointers0 imports Hoare begin
+theory Pointers0 imports Hoare_Logic begin
 
 subsection "References"
 
-axclass ref < type
-consts Null :: "'a::ref"
+class ref =
+  fixes Null :: 'a
 
 subsection "Field access and update"
 
--- a/src/HOL/Hoare/ROOT.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/ROOT.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -1,8 +1,2 @@
-(*  Title:      HOL/Hoare/ROOT.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1998-2003 TUM
-*)
 
-use_thys ["Examples", "ExamplesAbort", "Pointers0", "Pointer_Examples",
-  "Pointer_ExamplesAbort", "SchorrWaite", "Separation"];
+use_thy "Hoare";
--- a/src/HOL/Hoare/Separation.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Hoare/Separation.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -12,7 +12,7 @@
 
 *)
 
-theory Separation imports HoareAbort SepLogHeap begin
+theory Separation imports Hoare_Logic_Abort SepLogHeap begin
 
 text{* The semantic definition of a few connectives: *}
 
--- a/src/HOL/IsaMakefile	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/IsaMakefile	Tue Feb 23 17:55:00 2010 +0100
@@ -66,7 +66,6 @@
       TLA-Memory \
   HOL-UNITY \
   HOL-Unix \
-  HOL-W0 \
   HOL-Word-Examples \
   HOL-ZF
     # ^ this is the sort position
@@ -142,7 +141,6 @@
 	@$(ISABELLE_TOOL) usedir -b -f base.ML -d false -g false $(OUT)/Pure HOL-Base
 
 PLAIN_DEPENDENCIES = $(BASE_DEPENDENCIES)\
-  Algebras.thy \
   Complete_Lattice.thy \
   Datatype.thy \
   Extraction.thy \
@@ -344,7 +342,6 @@
   Deriv.thy \
   Fact.thy \
   GCD.thy \
-  Integration.thy \
   Lim.thy \
   Limits.thy \
   Ln.thy \
@@ -386,7 +383,7 @@
 $(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy		\
   Library/Abstract_Rat.thy Library/BigO.thy Library/ContNotDenum.thy	\
   Library/Efficient_Nat.thy Library/Sum_Of_Squares.thy			\
-  Library/Sum_Of_Squares/sos_wrapper.ML					\
+  Library/Dlist.thy Library/Sum_Of_Squares/sos_wrapper.ML		\
   Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy		\
   Library/Glbs.thy Library/normarith.ML Library/Executable_Set.thy	\
   Library/Infinite_Set.thy Library/FuncSet.thy				\
@@ -592,9 +589,10 @@
 
 $(LOG)/HOL-Hoare.gz: $(OUT)/HOL Hoare/Arith2.thy Hoare/Examples.thy	\
   Hoare/Hoare.thy Hoare/hoare_tac.ML Hoare/Heap.thy			\
+  Hoare/Hoare_Logic.thy	Hoare/Hoare_Logic_Abort.thy			\
   Hoare/HeapSyntax.thy Hoare/Pointer_Examples.thy Hoare/ROOT.ML		\
   Hoare/ExamplesAbort.thy Hoare/HeapSyntaxAbort.thy			\
-  Hoare/HoareAbort.thy Hoare/SchorrWaite.thy Hoare/Separation.thy	\
+  Hoare/SchorrWaite.thy Hoare/Separation.thy				\
   Hoare/SepLogHeap.thy Hoare/document/root.tex Hoare/document/root.bib
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Hoare
 
@@ -848,14 +846,6 @@
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Prolog
 
 
-## HOL-W0
-
-HOL-W0: HOL $(LOG)/HOL-W0.gz
-
-$(LOG)/HOL-W0.gz: $(OUT)/HOL W0/ROOT.ML W0/W0.thy W0/document/root.tex
-	@$(ISABELLE_TOOL) usedir $(OUT)/HOL W0
-
-
 ## HOL-MicroJava
 
 HOL-MicroJava: HOL $(LOG)/HOL-MicroJava.gz
@@ -965,12 +955,12 @@
   ex/Codegenerator_Pretty.thy ex/Codegenerator_Pretty_Test.thy		\
   ex/Codegenerator_Test.thy ex/Coherent.thy				\
   ex/Efficient_Nat_examples.thy ex/Eval_Examples.thy ex/Fundefs.thy	\
-  ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy        \
-  ex/HarmonicSeries.thy	ex/Hebrew.thy ex/Hex_Bin_Examples.thy           \
-  ex/Higher_Order_Logic.thy ex/Hilbert_Classical.thy                    \
-  ex/Induction_Schema.thy ex/InductiveInvariant.thy                     \
-  ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy              \
-  ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy                          \
+  ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy	\
+  ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy		\
+  ex/Higher_Order_Logic.thy ex/Hilbert_Classical.thy			\
+  ex/Induction_Schema.thy ex/InductiveInvariant.thy			\
+  ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy		\
+  ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy				\
   ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy			\
   ex/Multiquote.thy ex/NatSum.thy ex/Numeral.thy ex/PER.thy		\
   ex/Predicate_Compile_ex.thy ex/Predicate_Compile_Quickcheck.thy	\
@@ -1322,7 +1312,7 @@
 		$(LOG)/HOL-SMT-Examples.gz $(LOG)/HOL-SMT.gz		\
 		$(LOG)/HOL-Statespace.gz $(LOG)/HOL-Subst.gz		\
 		$(LOG)/HOL-UNITY.gz $(LOG)/HOL-Unix.gz			\
-		$(LOG)/HOL-W0.gz $(LOG)/HOL-Word-Examples.gz		\
+		$(LOG)/HOL-Word-Examples.gz		\
 		$(LOG)/HOL-Word.gz $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz	\
 		$(LOG)/HOL.gz $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz	\
 		$(LOG)/TLA-Inc.gz $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz	\
--- a/src/HOL/Isar_Examples/Group.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Isar_Examples/Group.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -17,15 +17,12 @@
  $\idt{times}$ is provided by the basic HOL theory.
 *}
 
-consts
-  one :: "'a"
-  inverse :: "'a => 'a"
+notation Groups.one ("one")
 
-axclass
-  group < times
-  group_assoc:         "(x * y) * z = x * (y * z)"
-  group_left_one:      "one * x = x"
-  group_left_inverse:  "inverse x * x = one"
+class group = times + one + inverse +
+  assumes group_assoc:         "(x * y) * z = x * (y * z)"
+  assumes group_left_one:      "one * x = x"
+  assumes group_left_inverse:  "inverse x * x = one"
 
 text {*
  The group axioms only state the properties of left one and inverse,
@@ -144,10 +141,10 @@
  \idt{one} :: \alpha)$ are defined like this.
 *}
 
-axclass monoid < times
-  monoid_assoc:       "(x * y) * z = x * (y * z)"
-  monoid_left_one:   "one * x = x"
-  monoid_right_one:  "x * one = x"
+class monoid = times + one +
+  assumes monoid_assoc:       "(x * y) * z = x * (y * z)"
+  assumes monoid_left_one:   "one * x = x"
+  assumes monoid_right_one:  "x * one = x"
 
 text {*
  Groups are \emph{not} yet monoids directly from the definition.  For
--- a/src/HOL/Lattice/Bounds.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Lattice/Bounds.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/Bounds.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
--- a/src/HOL/Lattice/CompleteLattice.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Lattice/CompleteLattice.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/CompleteLattice.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -16,8 +15,8 @@
   bounds (see \S\ref{sec:connect-bounds}).
 *}
 
-axclass complete_lattice \<subseteq> partial_order
-  ex_Inf: "\<exists>inf. is_Inf A inf"
+class complete_lattice =
+  assumes ex_Inf: "\<exists>inf. is_Inf A inf"
 
 theorem ex_Sup: "\<exists>sup::'a::complete_lattice. is_Sup A sup"
 proof -
--- a/src/HOL/Lattice/Lattice.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Lattice/Lattice.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/Lattice.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -15,9 +14,9 @@
   as well).
 *}
 
-axclass lattice \<subseteq> partial_order
-  ex_inf: "\<exists>inf. is_inf x y inf"
-  ex_sup: "\<exists>sup. is_sup x y sup"
+class lattice =
+  assumes ex_inf: "\<exists>inf. is_inf x y inf"
+  assumes ex_sup: "\<exists>sup. is_sup x y sup"
 
 text {*
   The @{text \<sqinter>} (meet) and @{text \<squnion>} (join) operations select such
--- a/src/HOL/Lattice/Orders.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Lattice/Orders.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/Orders.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -18,21 +17,21 @@
   required to be related (in either direction).
 *}
 
-axclass leq < type
-consts
-  leq :: "'a::leq \<Rightarrow> 'a \<Rightarrow> bool"  (infixl "[=" 50)
+class leq =
+  fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infixl "[=" 50)
+
 notation (xsymbols)
   leq  (infixl "\<sqsubseteq>" 50)
 
-axclass quasi_order < leq
-  leq_refl [intro?]: "x \<sqsubseteq> x"
-  leq_trans [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
+class quasi_order = leq +
+  assumes leq_refl [intro?]: "x \<sqsubseteq> x"
+  assumes leq_trans [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
 
-axclass partial_order < quasi_order
-  leq_antisym [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
+class partial_order = quasi_order +
+  assumes leq_antisym [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
 
-axclass linear_order < partial_order
-  leq_linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
+class linear_order = partial_order +
+  assumes leq_linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
 
 lemma linear_order_cases:
     "((x::'a::linear_order) \<sqsubseteq> y \<Longrightarrow> C) \<Longrightarrow> (y \<sqsubseteq> x \<Longrightarrow> C) \<Longrightarrow> C"
@@ -54,11 +53,16 @@
 primrec
   undual_dual: "undual (dual x) = x"
 
-instance dual :: (leq) leq ..
+instantiation dual :: (leq) leq
+begin
 
-defs (overloaded)
+definition
   leq_dual_def: "x' \<sqsubseteq> y' \<equiv> undual y' \<sqsubseteq> undual x'"
 
+instance ..
+
+end
+
 lemma undual_leq [iff?]: "(undual x' \<sqsubseteq> undual y') = (y' \<sqsubseteq> x')"
   by (simp add: leq_dual_def)
 
@@ -192,11 +196,16 @@
   \emph{not} be linear in general.
 *}
 
-instance * :: (leq, leq) leq ..
+instantiation * :: (leq, leq) leq
+begin
 
-defs (overloaded)
+definition
   leq_prod_def: "p \<sqsubseteq> q \<equiv> fst p \<sqsubseteq> fst q \<and> snd p \<sqsubseteq> snd q"
 
+instance ..
+
+end
+
 lemma leq_prodI [intro?]:
     "fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q"
   by (unfold leq_prod_def) blast
@@ -249,11 +258,16 @@
   orders need \emph{not} be linear in general.
 *}
 
-instance "fun" :: (type, leq) leq ..
+instantiation "fun" :: (type, leq) leq
+begin
 
-defs (overloaded)
+definition
   leq_fun_def: "f \<sqsubseteq> g \<equiv> \<forall>x. f x \<sqsubseteq> g x"
 
+instance ..
+
+end
+
 lemma leq_funI [intro?]: "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> f \<sqsubseteq> g"
   by (unfold leq_fun_def) blast
 
--- a/src/HOL/Lattices.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Lattices.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -8,7 +8,42 @@
 imports Orderings Groups
 begin
 
-subsection {* Lattices *}
+subsection {* Abstract semilattice *}
+
+text {*
+  This locales provide a basic structure for interpretation into
+  bigger structures;  extensions require careful thinking, otherwise
+  undesired effects may occur due to interpretation.
+*}
+
+locale semilattice = abel_semigroup +
+  assumes idem [simp]: "f a a = a"
+begin
+
+lemma left_idem [simp]:
+  "f a (f a b) = f a b"
+  by (simp add: assoc [symmetric])
+
+end
+
+
+subsection {* Idempotent semigroup *}
+
+class ab_semigroup_idem_mult = ab_semigroup_mult +
+  assumes mult_idem: "x * x = x"
+
+sublocale ab_semigroup_idem_mult < times!: semilattice times proof
+qed (fact mult_idem)
+
+context ab_semigroup_idem_mult
+begin
+
+lemmas mult_left_idem = times.left_idem
+
+end
+
+
+subsection {* Conrete lattices *}
 
 notation
   less_eq  (infix "\<sqsubseteq>" 50) and
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Dlist.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,256 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Lists with elements distinct as canonical example for datatype invariants *}
+
+theory Dlist
+imports Main Fset
+begin
+
+section {* Prelude *}
+
+text {* Without canonical argument order, higher-order things tend to get confusing quite fast: *}
+
+setup {* Sign.map_naming (Name_Space.add_path "List") *}
+
+primrec member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
+    "member [] y \<longleftrightarrow> False"
+  | "member (x#xs) y \<longleftrightarrow> x = y \<or> member xs y"
+
+lemma member_set:
+  "member = set"
+proof (rule ext)+
+  fix xs :: "'a list" and x :: 'a
+  have "member xs x \<longleftrightarrow> x \<in> set xs" by (induct xs) auto
+  then show "member xs x = set xs x" by (simp add: mem_def)
+qed
+
+lemma not_set_compl:
+  "Not \<circ> set xs = - set xs"
+  by (simp add: fun_Compl_def bool_Compl_def comp_def expand_fun_eq)
+
+primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
+    "fold f [] s = s"
+  | "fold f (x#xs) s = fold f xs (f x s)"
+
+lemma foldl_fold:
+  "foldl f s xs = List.fold (\<lambda>x s. f s x) xs s"
+  by (induct xs arbitrary: s) simp_all
+
+setup {* Sign.map_naming Name_Space.parent_path *}
+
+
+section {* The type of distinct lists *}
+
+typedef (open) 'a dlist = "{xs::'a list. distinct xs}"
+  morphisms list_of_dlist Abs_dlist
+proof
+  show "[] \<in> ?dlist" by simp
+qed
+
+text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
+
+definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
+  [code del]: "Dlist xs = Abs_dlist (remdups xs)"
+
+lemma distinct_list_of_dlist [simp]:
+  "distinct (list_of_dlist dxs)"
+  using list_of_dlist [of dxs] by simp
+
+lemma list_of_dlist_Dlist [simp]:
+  "list_of_dlist (Dlist xs) = remdups xs"
+  by (simp add: Dlist_def Abs_dlist_inverse)
+
+lemma Dlist_list_of_dlist [simp]:
+  "Dlist (list_of_dlist dxs) = dxs"
+  by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
+
+
+text {* Fundamental operations: *}
+
+definition empty :: "'a dlist" where
+  "empty = Dlist []"
+
+definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
+  "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"
+
+definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
+  "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"
+
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
+  "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"
+
+definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
+  "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
+
+
+text {* Derived operations: *}
+
+definition null :: "'a dlist \<Rightarrow> bool" where
+  "null dxs = List.null (list_of_dlist dxs)"
+
+definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where
+  "member dxs = List.member (list_of_dlist dxs)"
+
+definition length :: "'a dlist \<Rightarrow> nat" where
+  "length dxs = List.length (list_of_dlist dxs)"
+
+definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
+  "fold f dxs = List.fold f (list_of_dlist dxs)"
+
+
+section {* Executable version obeying invariant *}
+
+code_abstype Dlist list_of_dlist
+  by simp
+
+lemma list_of_dlist_empty [simp, code abstract]:
+  "list_of_dlist empty = []"
+  by (simp add: empty_def)
+
+lemma list_of_dlist_insert [simp, code abstract]:
+  "list_of_dlist (insert x dxs) = List.insert x (list_of_dlist dxs)"
+  by (simp add: insert_def)
+
+lemma list_of_dlist_remove [simp, code abstract]:
+  "list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)"
+  by (simp add: remove_def)
+
+lemma list_of_dlist_map [simp, code abstract]:
+  "list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))"
+  by (simp add: map_def)
+
+lemma list_of_dlist_filter [simp, code abstract]:
+  "list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"
+  by (simp add: filter_def)
+
+declare null_def [code] member_def [code] length_def [code] fold_def [code] -- {* explicit is better than implicit *}
+
+
+section {* Implementation of sets by distinct lists -- canonical! *}
+
+definition Set :: "'a dlist \<Rightarrow> 'a fset" where
+  "Set dxs = Fset.Set (list_of_dlist dxs)"
+
+definition Coset :: "'a dlist \<Rightarrow> 'a fset" where
+  "Coset dxs = Fset.Coset (list_of_dlist dxs)"
+
+code_datatype Set Coset
+
+declare member_code [code del]
+declare is_empty_Set [code del]
+declare empty_Set [code del]
+declare UNIV_Set [code del]
+declare insert_Set [code del]
+declare remove_Set [code del]
+declare map_Set [code del]
+declare filter_Set [code del]
+declare forall_Set [code del]
+declare exists_Set [code del]
+declare card_Set [code del]
+declare subfset_eq_forall [code del]
+declare subfset_subfset_eq [code del]
+declare eq_fset_subfset_eq [code del]
+declare inter_project [code del]
+declare subtract_remove [code del]
+declare union_insert [code del]
+declare Infimum_inf [code del]
+declare Supremum_sup [code del]
+
+lemma Set_Dlist [simp]:
+  "Set (Dlist xs) = Fset (set xs)"
+  by (simp add: Set_def Fset.Set_def)
+
+lemma Coset_Dlist [simp]:
+  "Coset (Dlist xs) = Fset (- set xs)"
+  by (simp add: Coset_def Fset.Coset_def)
+
+lemma member_Set [simp]:
+  "Fset.member (Set dxs) = List.member (list_of_dlist dxs)"
+  by (simp add: Set_def member_set)
+
+lemma member_Coset [simp]:
+  "Fset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"
+  by (simp add: Coset_def member_set not_set_compl)
+
+lemma is_empty_Set [code]:
+  "Fset.is_empty (Set dxs) \<longleftrightarrow> null dxs"
+  by (simp add: null_def null_empty member_set)
+
+lemma bot_code [code]:
+  "bot = Set empty"
+  by (simp add: empty_def)
+
+lemma top_code [code]:
+  "top = Coset empty"
+  by (simp add: empty_def)
+
+lemma insert_code [code]:
+  "Fset.insert x (Set dxs) = Set (insert x dxs)"
+  "Fset.insert x (Coset dxs) = Coset (remove x dxs)"
+  by (simp_all add: insert_def remove_def member_set not_set_compl)
+
+lemma remove_code [code]:
+  "Fset.remove x (Set dxs) = Set (remove x dxs)"
+  "Fset.remove x (Coset dxs) = Coset (insert x dxs)"
+  by (auto simp add: insert_def remove_def member_set not_set_compl)
+
+lemma member_code [code]:
+  "Fset.member (Set dxs) = member dxs"
+  "Fset.member (Coset dxs) = Not \<circ> member dxs"
+  by (simp_all add: member_def)
+
+lemma map_code [code]:
+  "Fset.map f (Set dxs) = Set (map f dxs)"
+  by (simp add: member_set)
+  
+lemma filter_code [code]:
+  "Fset.filter f (Set dxs) = Set (filter f dxs)"
+  by (simp add: member_set)
+
+lemma forall_Set [code]:
+  "Fset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"
+  by (simp add: member_set list_all_iff)
+
+lemma exists_Set [code]:
+  "Fset.exists P (Set xs) \<longleftrightarrow> list_ex P (list_of_dlist xs)"
+  by (simp add: member_set list_ex_iff)
+
+lemma card_code [code]:
+  "Fset.card (Set dxs) = length dxs"
+  by (simp add: length_def member_set distinct_card)
+
+lemma foldl_list_of_dlist:
+  "foldl f s (list_of_dlist dxs) = fold (\<lambda>x s. f s x) dxs s"
+  by (simp add: foldl_fold fold_def)
+
+lemma inter_code [code]:
+  "inf A (Set xs) = Set (filter (Fset.member A) xs)"
+  "inf A (Coset xs) = fold Fset.remove xs A"
+  by (simp_all only: Set_def Coset_def foldl_list_of_dlist inter_project list_of_dlist_filter)
+
+lemma subtract_code [code]:
+  "A - Set xs = fold Fset.remove xs A"
+  "A - Coset xs = Set (filter (Fset.member A) xs)"
+  by (simp_all only: Set_def Coset_def foldl_list_of_dlist subtract_remove list_of_dlist_filter)
+
+lemma union_code [code]:
+  "sup (Set xs) A = fold Fset.insert xs A"
+  "sup (Coset xs) A = Coset (filter (Not \<circ> Fset.member A) xs)"
+  by (simp_all only: Set_def Coset_def foldl_list_of_dlist union_insert list_of_dlist_filter)
+
+context complete_lattice
+begin
+
+lemma Infimum_code [code]:
+  "Infimum (Set As) = fold inf As top"
+  by (simp only: Set_def Infimum_inf foldl_list_of_dlist inf.commute)
+
+lemma Supremum_code [code]:
+  "Supremum (Set As) = fold sup As bot"
+  by (simp only: Set_def Supremum_sup foldl_list_of_dlist sup.commute)
+
+end
+
+hide (open) const member fold empty insert remove map filter null member length fold
+
+end
--- a/src/HOL/Library/Library.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Library/Library.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -15,6 +15,7 @@
   ContNotDenum
   Countable
   Diagonalize
+  Dlist
   Efficient_Nat
   Enum
   Eval_Witness
--- a/src/HOL/Library/Multiset.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Library/Multiset.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1212,8 +1212,8 @@
 definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where
   "M' <=# M \<longleftrightarrow> M' <# M \<or> M' = M"
 
-notation (xsymbol) less_multiset (infix "\<subset>#" 50)
-notation (xsymbol) le_multiset (infix "\<subseteq>#" 50)
+notation (xsymbols) less_multiset (infix "\<subset>#" 50)
+notation (xsymbols) le_multiset (infix "\<subseteq>#" 50)
 
 interpretation multiset_order: order le_multiset less_multiset
 proof -
--- a/src/HOL/List.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/List.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -250,7 +250,7 @@
 @{lemma "distinct [2,0,1::nat]" by simp}\\
 @{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
 @{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
-@{lemma "List.insert 3 [0::nat,1,2] = [3, 0,1,2]" by (simp add: List.insert_def)}\\
+@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
 @{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
 @{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
 @{lemma "nth [a,b,c,d] 2 = c" by simp}\\
@@ -2900,10 +2900,14 @@
   "List.insert x [] = [x]"
   by simp
 
-lemma set_insert:
+lemma set_insert [simp]:
   "set (List.insert x xs) = insert x (set xs)"
   by (auto simp add: List.insert_def)
 
+lemma distinct_insert [simp]:
+  "distinct xs \<Longrightarrow> distinct (List.insert x xs)"
+  by (simp add: List.insert_def)
+
 
 subsubsection {* @{text remove1} *}
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Integration.cert	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,3296 @@
+tB2Atlor9W4pSnrAz5nHpw 907 0
+#2 := false
+#299 := 0::real
+decl uf_1 :: (-> T3 T2 real)
+decl uf_10 :: (-> T4 T2)
+decl uf_7 :: T4
+#15 := uf_7
+#22 := (uf_10 uf_7)
+decl uf_2 :: (-> T1 T3)
+decl uf_4 :: T1
+#11 := uf_4
+#91 := (uf_2 uf_4)
+#902 := (uf_1 #91 #22)
+#297 := -1::real
+#1084 := (* -1::real #902)
+decl uf_16 :: T1
+#50 := uf_16
+#78 := (uf_2 uf_16)
+#799 := (uf_1 #78 #22)
+#1267 := (+ #799 #1084)
+#1272 := (>= #1267 0::real)
+#1266 := (= #799 #902)
+decl uf_9 :: T3
+#21 := uf_9
+#23 := (uf_1 uf_9 #22)
+#905 := (= #23 #902)
+decl uf_11 :: T3
+#24 := uf_11
+#850 := (uf_1 uf_11 #22)
+#904 := (= #850 #902)
+decl uf_6 :: (-> T2 T4)
+#74 := (uf_6 #22)
+#281 := (= uf_7 #74)
+#922 := (ite #281 #905 #904)
+decl uf_8 :: T3
+#18 := uf_8
+#848 := (uf_1 uf_8 #22)
+#903 := (= #848 #902)
+#60 := 0::int
+decl uf_5 :: (-> T4 int)
+#803 := (uf_5 #74)
+#117 := -1::int
+#813 := (* -1::int #803)
+#16 := (uf_5 uf_7)
+#916 := (+ #16 #813)
+#917 := (<= #916 0::int)
+#925 := (ite #917 #922 #903)
+#6 := (:var 0 T2)
+#19 := (uf_1 uf_8 #6)
+#544 := (pattern #19)
+#25 := (uf_1 uf_11 #6)
+#543 := (pattern #25)
+#92 := (uf_1 #91 #6)
+#542 := (pattern #92)
+#13 := (uf_6 #6)
+#541 := (pattern #13)
+#447 := (= #19 #92)
+#445 := (= #25 #92)
+#444 := (= #23 #92)
+#20 := (= #13 uf_7)
+#446 := (ite #20 #444 #445)
+#120 := (* -1::int #16)
+#14 := (uf_5 #13)
+#121 := (+ #14 #120)
+#119 := (>= #121 0::int)
+#448 := (ite #119 #446 #447)
+#545 := (forall (vars (?x3 T2)) (:pat #541 #542 #543 #544) #448)
+#451 := (forall (vars (?x3 T2)) #448)
+#548 := (iff #451 #545)
+#546 := (iff #448 #448)
+#547 := [refl]: #546
+#549 := [quant-intro #547]: #548
+#26 := (ite #20 #23 #25)
+#127 := (ite #119 #26 #19)
+#368 := (= #92 #127)
+#369 := (forall (vars (?x3 T2)) #368)
+#452 := (iff #369 #451)
+#449 := (iff #368 #448)
+#450 := [rewrite]: #449
+#453 := [quant-intro #450]: #452
+#392 := (~ #369 #369)
+#390 := (~ #368 #368)
+#391 := [refl]: #390
+#366 := [nnf-pos #391]: #392
+decl uf_3 :: (-> T1 T2 real)
+#12 := (uf_3 uf_4 #6)
+#132 := (= #12 #127)
+#135 := (forall (vars (?x3 T2)) #132)
+#370 := (iff #135 #369)
+#4 := (:var 1 T1)
+#8 := (uf_3 #4 #6)
+#5 := (uf_2 #4)
+#7 := (uf_1 #5 #6)
+#9 := (= #7 #8)
+#10 := (forall (vars (?x1 T1) (?x2 T2)) #9)
+#113 := [asserted]: #10
+#371 := [rewrite* #113]: #370
+#17 := (< #14 #16)
+#27 := (ite #17 #19 #26)
+#28 := (= #12 #27)
+#29 := (forall (vars (?x3 T2)) #28)
+#136 := (iff #29 #135)
+#133 := (iff #28 #132)
+#130 := (= #27 #127)
+#118 := (not #119)
+#124 := (ite #118 #19 #26)
+#128 := (= #124 #127)
+#129 := [rewrite]: #128
+#125 := (= #27 #124)
+#122 := (iff #17 #118)
+#123 := [rewrite]: #122
+#126 := [monotonicity #123]: #125
+#131 := [trans #126 #129]: #130
+#134 := [monotonicity #131]: #133
+#137 := [quant-intro #134]: #136
+#114 := [asserted]: #29
+#138 := [mp #114 #137]: #135
+#372 := [mp #138 #371]: #369
+#367 := [mp~ #372 #366]: #369
+#454 := [mp #367 #453]: #451
+#550 := [mp #454 #549]: #545
+#738 := (not #545)
+#928 := (or #738 #925)
+#75 := (= #74 uf_7)
+#906 := (ite #75 #905 #904)
+#907 := (+ #803 #120)
+#908 := (>= #907 0::int)
+#909 := (ite #908 #906 #903)
+#929 := (or #738 #909)
+#931 := (iff #929 #928)
+#933 := (iff #928 #928)
+#934 := [rewrite]: #933
+#926 := (iff #909 #925)
+#923 := (iff #906 #922)
+#283 := (iff #75 #281)
+#284 := [rewrite]: #283
+#924 := [monotonicity #284]: #923
+#920 := (iff #908 #917)
+#910 := (+ #120 #803)
+#913 := (>= #910 0::int)
+#918 := (iff #913 #917)
+#919 := [rewrite]: #918
+#914 := (iff #908 #913)
+#911 := (= #907 #910)
+#912 := [rewrite]: #911
+#915 := [monotonicity #912]: #914
+#921 := [trans #915 #919]: #920
+#927 := [monotonicity #921 #924]: #926
+#932 := [monotonicity #927]: #931
+#935 := [trans #932 #934]: #931
+#930 := [quant-inst]: #929
+#936 := [mp #930 #935]: #928
+#1300 := [unit-resolution #936 #550]: #925
+#989 := (= #16 #803)
+#1277 := (= #803 #16)
+#280 := [asserted]: #75
+#287 := [mp #280 #284]: #281
+#1276 := [symm #287]: #75
+#1278 := [monotonicity #1276]: #1277
+#1301 := [symm #1278]: #989
+#1302 := (not #989)
+#1303 := (or #1302 #917)
+#1304 := [th-lemma]: #1303
+#1305 := [unit-resolution #1304 #1301]: #917
+#950 := (not #917)
+#949 := (not #925)
+#951 := (or #949 #950 #922)
+#952 := [def-axiom]: #951
+#1306 := [unit-resolution #952 #1305 #1300]: #922
+#937 := (not #922)
+#1307 := (or #937 #905)
+#938 := (not #281)
+#939 := (or #937 #938 #905)
+#940 := [def-axiom]: #939
+#1308 := [unit-resolution #940 #287]: #1307
+#1309 := [unit-resolution #1308 #1306]: #905
+#1356 := (= #799 #23)
+#800 := (= #23 #799)
+decl uf_15 :: T4
+#40 := uf_15
+#41 := (uf_5 uf_15)
+#814 := (+ #41 #813)
+#815 := (<= #814 0::int)
+#836 := (not #815)
+#158 := (* -1::int #41)
+#1270 := (+ #16 #158)
+#1265 := (>= #1270 0::int)
+#1339 := (not #1265)
+#1269 := (= #16 #41)
+#1298 := (not #1269)
+#286 := (= uf_7 uf_15)
+#44 := (uf_10 uf_15)
+#72 := (uf_6 #44)
+#73 := (= #72 uf_15)
+#277 := (= uf_15 #72)
+#278 := (iff #73 #277)
+#279 := [rewrite]: #278
+#276 := [asserted]: #73
+#282 := [mp #276 #279]: #277
+#1274 := [symm #282]: #73
+#729 := (= uf_7 #72)
+decl uf_17 :: (-> int T4)
+#611 := (uf_5 #72)
+#991 := (uf_17 #611)
+#1289 := (= #991 #72)
+#992 := (= #72 #991)
+#55 := (:var 0 T4)
+#56 := (uf_5 #55)
+#574 := (pattern #56)
+#57 := (uf_17 #56)
+#177 := (= #55 #57)
+#575 := (forall (vars (?x7 T4)) (:pat #574) #177)
+#195 := (forall (vars (?x7 T4)) #177)
+#578 := (iff #195 #575)
+#576 := (iff #177 #177)
+#577 := [refl]: #576
+#579 := [quant-intro #577]: #578
+#405 := (~ #195 #195)
+#403 := (~ #177 #177)
+#404 := [refl]: #403
+#406 := [nnf-pos #404]: #405
+#58 := (= #57 #55)
+#59 := (forall (vars (?x7 T4)) #58)
+#196 := (iff #59 #195)
+#193 := (iff #58 #177)
+#194 := [rewrite]: #193
+#197 := [quant-intro #194]: #196
+#155 := [asserted]: #59
+#200 := [mp #155 #197]: #195
+#407 := [mp~ #200 #406]: #195
+#580 := [mp #407 #579]: #575
+#995 := (not #575)
+#996 := (or #995 #992)
+#997 := [quant-inst]: #996
+#1273 := [unit-resolution #997 #580]: #992
+#1290 := [symm #1273]: #1289
+#1293 := (= uf_7 #991)
+#993 := (uf_17 #803)
+#1287 := (= #993 #991)
+#1284 := (= #803 #611)
+#987 := (= #41 #611)
+#1279 := (= #611 #41)
+#1280 := [monotonicity #1274]: #1279
+#1281 := [symm #1280]: #987
+#1282 := (= #803 #41)
+#1275 := [hypothesis]: #1269
+#1283 := [trans #1278 #1275]: #1282
+#1285 := [trans #1283 #1281]: #1284
+#1288 := [monotonicity #1285]: #1287
+#1291 := (= uf_7 #993)
+#994 := (= #74 #993)
+#1000 := (or #995 #994)
+#1001 := [quant-inst]: #1000
+#1286 := [unit-resolution #1001 #580]: #994
+#1292 := [trans #287 #1286]: #1291
+#1294 := [trans #1292 #1288]: #1293
+#1295 := [trans #1294 #1290]: #729
+#1296 := [trans #1295 #1274]: #286
+#290 := (not #286)
+#76 := (= uf_15 uf_7)
+#77 := (not #76)
+#291 := (iff #77 #290)
+#288 := (iff #76 #286)
+#289 := [rewrite]: #288
+#292 := [monotonicity #289]: #291
+#285 := [asserted]: #77
+#295 := [mp #285 #292]: #290
+#1297 := [unit-resolution #295 #1296]: false
+#1299 := [lemma #1297]: #1298
+#1342 := (or #1269 #1339)
+#1271 := (<= #1270 0::int)
+#621 := (* -1::int #611)
+#723 := (+ #16 #621)
+#724 := (<= #723 0::int)
+decl uf_12 :: T1
+#30 := uf_12
+#88 := (uf_2 uf_12)
+#771 := (uf_1 #88 #44)
+#45 := (uf_1 uf_9 #44)
+#772 := (= #45 #771)
+#796 := (not #772)
+decl uf_14 :: T1
+#38 := uf_14
+#83 := (uf_2 uf_14)
+#656 := (uf_1 #83 #44)
+#1239 := (= #656 #771)
+#1252 := (not #1239)
+#1324 := (iff #1252 #796)
+#1322 := (iff #1239 #772)
+#1320 := (= #656 #45)
+#661 := (= #45 #656)
+#659 := (uf_1 uf_11 #44)
+#664 := (= #656 #659)
+#667 := (ite #277 #661 #664)
+#657 := (uf_1 uf_8 #44)
+#670 := (= #656 #657)
+#622 := (+ #41 #621)
+#623 := (<= #622 0::int)
+#673 := (ite #623 #667 #670)
+#84 := (uf_1 #83 #6)
+#560 := (pattern #84)
+#467 := (= #19 #84)
+#465 := (= #25 #84)
+#464 := (= #45 #84)
+#43 := (= #13 uf_15)
+#466 := (ite #43 #464 #465)
+#159 := (+ #14 #158)
+#157 := (>= #159 0::int)
+#468 := (ite #157 #466 #467)
+#561 := (forall (vars (?x5 T2)) (:pat #541 #560 #543 #544) #468)
+#471 := (forall (vars (?x5 T2)) #468)
+#564 := (iff #471 #561)
+#562 := (iff #468 #468)
+#563 := [refl]: #562
+#565 := [quant-intro #563]: #564
+#46 := (ite #43 #45 #25)
+#165 := (ite #157 #46 #19)
+#378 := (= #84 #165)
+#379 := (forall (vars (?x5 T2)) #378)
+#472 := (iff #379 #471)
+#469 := (iff #378 #468)
+#470 := [rewrite]: #469
+#473 := [quant-intro #470]: #472
+#359 := (~ #379 #379)
+#361 := (~ #378 #378)
+#358 := [refl]: #361
+#356 := [nnf-pos #358]: #359
+#39 := (uf_3 uf_14 #6)
+#170 := (= #39 #165)
+#173 := (forall (vars (?x5 T2)) #170)
+#380 := (iff #173 #379)
+#381 := [rewrite* #113]: #380
+#42 := (< #14 #41)
+#47 := (ite #42 #19 #46)
+#48 := (= #39 #47)
+#49 := (forall (vars (?x5 T2)) #48)
+#174 := (iff #49 #173)
+#171 := (iff #48 #170)
+#168 := (= #47 #165)
+#156 := (not #157)
+#162 := (ite #156 #19 #46)
+#166 := (= #162 #165)
+#167 := [rewrite]: #166
+#163 := (= #47 #162)
+#160 := (iff #42 #156)
+#161 := [rewrite]: #160
+#164 := [monotonicity #161]: #163
+#169 := [trans #164 #167]: #168
+#172 := [monotonicity #169]: #171
+#175 := [quant-intro #172]: #174
+#116 := [asserted]: #49
+#176 := [mp #116 #175]: #173
+#382 := [mp #176 #381]: #379
+#357 := [mp~ #382 #356]: #379
+#474 := [mp #357 #473]: #471
+#566 := [mp #474 #565]: #561
+#676 := (not #561)
+#677 := (or #676 #673)
+#658 := (= #657 #656)
+#660 := (= #659 #656)
+#662 := (ite #73 #661 #660)
+#612 := (+ #611 #158)
+#613 := (>= #612 0::int)
+#663 := (ite #613 #662 #658)
+#678 := (or #676 #663)
+#680 := (iff #678 #677)
+#682 := (iff #677 #677)
+#683 := [rewrite]: #682
+#674 := (iff #663 #673)
+#671 := (iff #658 #670)
+#672 := [rewrite]: #671
+#668 := (iff #662 #667)
+#665 := (iff #660 #664)
+#666 := [rewrite]: #665
+#669 := [monotonicity #279 #666]: #668
+#626 := (iff #613 #623)
+#615 := (+ #158 #611)
+#618 := (>= #615 0::int)
+#624 := (iff #618 #623)
+#625 := [rewrite]: #624
+#619 := (iff #613 #618)
+#616 := (= #612 #615)
+#617 := [rewrite]: #616
+#620 := [monotonicity #617]: #619
+#627 := [trans #620 #625]: #626
+#675 := [monotonicity #627 #669 #672]: #674
+#681 := [monotonicity #675]: #680
+#684 := [trans #681 #683]: #680
+#679 := [quant-inst]: #678
+#685 := [mp #679 #684]: #677
+#1311 := [unit-resolution #685 #566]: #673
+#1312 := (not #987)
+#1313 := (or #1312 #623)
+#1314 := [th-lemma]: #1313
+#1315 := [unit-resolution #1314 #1281]: #623
+#645 := (not #623)
+#698 := (not #673)
+#699 := (or #698 #645 #667)
+#700 := [def-axiom]: #699
+#1316 := [unit-resolution #700 #1315 #1311]: #667
+#686 := (not #667)
+#1317 := (or #686 #661)
+#687 := (not #277)
+#688 := (or #686 #687 #661)
+#689 := [def-axiom]: #688
+#1318 := [unit-resolution #689 #282]: #1317
+#1319 := [unit-resolution #1318 #1316]: #661
+#1321 := [symm #1319]: #1320
+#1323 := [monotonicity #1321]: #1322
+#1325 := [monotonicity #1323]: #1324
+#1145 := (* -1::real #771)
+#1240 := (+ #656 #1145)
+#1241 := (<= #1240 0::real)
+#1249 := (not #1241)
+#1243 := [hypothesis]: #1241
+decl uf_18 :: T3
+#80 := uf_18
+#1040 := (uf_1 uf_18 #44)
+#1043 := (* -1::real #1040)
+#1156 := (+ #771 #1043)
+#1157 := (>= #1156 0::real)
+#1189 := (not #1157)
+#708 := (uf_1 #91 #44)
+#1168 := (+ #708 #1043)
+#1169 := (<= #1168 0::real)
+#1174 := (or #1157 #1169)
+#1177 := (not #1174)
+#89 := (uf_1 #88 #6)
+#552 := (pattern #89)
+#81 := (uf_1 uf_18 #6)
+#594 := (pattern #81)
+#324 := (* -1::real #92)
+#325 := (+ #81 #324)
+#323 := (>= #325 0::real)
+#317 := (* -1::real #89)
+#318 := (+ #81 #317)
+#319 := (<= #318 0::real)
+#436 := (or #319 #323)
+#437 := (not #436)
+#601 := (forall (vars (?x11 T2)) (:pat #594 #552 #542) #437)
+#440 := (forall (vars (?x11 T2)) #437)
+#604 := (iff #440 #601)
+#602 := (iff #437 #437)
+#603 := [refl]: #602
+#605 := [quant-intro #603]: #604
+#326 := (not #323)
+#320 := (not #319)
+#329 := (and #320 #326)
+#332 := (forall (vars (?x11 T2)) #329)
+#441 := (iff #332 #440)
+#438 := (iff #329 #437)
+#439 := [rewrite]: #438
+#442 := [quant-intro #439]: #441
+#425 := (~ #332 #332)
+#423 := (~ #329 #329)
+#424 := [refl]: #423
+#426 := [nnf-pos #424]: #425
+#306 := (* -1::real #84)
+#307 := (+ #81 #306)
+#305 := (>= #307 0::real)
+#308 := (not #305)
+#301 := (* -1::real #81)
+#79 := (uf_1 #78 #6)
+#302 := (+ #79 #301)
+#300 := (>= #302 0::real)
+#298 := (not #300)
+#311 := (and #298 #308)
+#314 := (forall (vars (?x10 T2)) #311)
+#335 := (and #314 #332)
+#93 := (< #81 #92)
+#90 := (< #89 #81)
+#94 := (and #90 #93)
+#95 := (forall (vars (?x11 T2)) #94)
+#85 := (< #81 #84)
+#82 := (< #79 #81)
+#86 := (and #82 #85)
+#87 := (forall (vars (?x10 T2)) #86)
+#96 := (and #87 #95)
+#336 := (iff #96 #335)
+#333 := (iff #95 #332)
+#330 := (iff #94 #329)
+#327 := (iff #93 #326)
+#328 := [rewrite]: #327
+#321 := (iff #90 #320)
+#322 := [rewrite]: #321
+#331 := [monotonicity #322 #328]: #330
+#334 := [quant-intro #331]: #333
+#315 := (iff #87 #314)
+#312 := (iff #86 #311)
+#309 := (iff #85 #308)
+#310 := [rewrite]: #309
+#303 := (iff #82 #298)
+#304 := [rewrite]: #303
+#313 := [monotonicity #304 #310]: #312
+#316 := [quant-intro #313]: #315
+#337 := [monotonicity #316 #334]: #336
+#293 := [asserted]: #96
+#338 := [mp #293 #337]: #335
+#340 := [and-elim #338]: #332
+#427 := [mp~ #340 #426]: #332
+#443 := [mp #427 #442]: #440
+#606 := [mp #443 #605]: #601
+#1124 := (not #601)
+#1180 := (or #1124 #1177)
+#1142 := (* -1::real #708)
+#1143 := (+ #1040 #1142)
+#1144 := (>= #1143 0::real)
+#1146 := (+ #1040 #1145)
+#1147 := (<= #1146 0::real)
+#1148 := (or #1147 #1144)
+#1149 := (not #1148)
+#1181 := (or #1124 #1149)
+#1183 := (iff #1181 #1180)
+#1185 := (iff #1180 #1180)
+#1186 := [rewrite]: #1185
+#1178 := (iff #1149 #1177)
+#1175 := (iff #1148 #1174)
+#1172 := (iff #1144 #1169)
+#1162 := (+ #1142 #1040)
+#1165 := (>= #1162 0::real)
+#1170 := (iff #1165 #1169)
+#1171 := [rewrite]: #1170
+#1166 := (iff #1144 #1165)
+#1163 := (= #1143 #1162)
+#1164 := [rewrite]: #1163
+#1167 := [monotonicity #1164]: #1166
+#1173 := [trans #1167 #1171]: #1172
+#1160 := (iff #1147 #1157)
+#1150 := (+ #1145 #1040)
+#1153 := (<= #1150 0::real)
+#1158 := (iff #1153 #1157)
+#1159 := [rewrite]: #1158
+#1154 := (iff #1147 #1153)
+#1151 := (= #1146 #1150)
+#1152 := [rewrite]: #1151
+#1155 := [monotonicity #1152]: #1154
+#1161 := [trans #1155 #1159]: #1160
+#1176 := [monotonicity #1161 #1173]: #1175
+#1179 := [monotonicity #1176]: #1178
+#1184 := [monotonicity #1179]: #1183
+#1187 := [trans #1184 #1186]: #1183
+#1182 := [quant-inst]: #1181
+#1188 := [mp #1182 #1187]: #1180
+#1244 := [unit-resolution #1188 #606]: #1177
+#1190 := (or #1174 #1189)
+#1191 := [def-axiom]: #1190
+#1245 := [unit-resolution #1191 #1244]: #1189
+#1054 := (+ #656 #1043)
+#1055 := (<= #1054 0::real)
+#1079 := (not #1055)
+#607 := (uf_1 #78 #44)
+#1044 := (+ #607 #1043)
+#1045 := (>= #1044 0::real)
+#1060 := (or #1045 #1055)
+#1063 := (not #1060)
+#567 := (pattern #79)
+#428 := (or #300 #305)
+#429 := (not #428)
+#595 := (forall (vars (?x10 T2)) (:pat #567 #594 #560) #429)
+#432 := (forall (vars (?x10 T2)) #429)
+#598 := (iff #432 #595)
+#596 := (iff #429 #429)
+#597 := [refl]: #596
+#599 := [quant-intro #597]: #598
+#433 := (iff #314 #432)
+#430 := (iff #311 #429)
+#431 := [rewrite]: #430
+#434 := [quant-intro #431]: #433
+#420 := (~ #314 #314)
+#418 := (~ #311 #311)
+#419 := [refl]: #418
+#421 := [nnf-pos #419]: #420
+#339 := [and-elim #338]: #314
+#422 := [mp~ #339 #421]: #314
+#435 := [mp #422 #434]: #432
+#600 := [mp #435 #599]: #595
+#1066 := (not #595)
+#1067 := (or #1066 #1063)
+#1039 := (* -1::real #656)
+#1041 := (+ #1040 #1039)
+#1042 := (>= #1041 0::real)
+#1046 := (or #1045 #1042)
+#1047 := (not #1046)
+#1068 := (or #1066 #1047)
+#1070 := (iff #1068 #1067)
+#1072 := (iff #1067 #1067)
+#1073 := [rewrite]: #1072
+#1064 := (iff #1047 #1063)
+#1061 := (iff #1046 #1060)
+#1058 := (iff #1042 #1055)
+#1048 := (+ #1039 #1040)
+#1051 := (>= #1048 0::real)
+#1056 := (iff #1051 #1055)
+#1057 := [rewrite]: #1056
+#1052 := (iff #1042 #1051)
+#1049 := (= #1041 #1048)
+#1050 := [rewrite]: #1049
+#1053 := [monotonicity #1050]: #1052
+#1059 := [trans #1053 #1057]: #1058
+#1062 := [monotonicity #1059]: #1061
+#1065 := [monotonicity #1062]: #1064
+#1071 := [monotonicity #1065]: #1070
+#1074 := [trans #1071 #1073]: #1070
+#1069 := [quant-inst]: #1068
+#1075 := [mp #1069 #1074]: #1067
+#1246 := [unit-resolution #1075 #600]: #1063
+#1080 := (or #1060 #1079)
+#1081 := [def-axiom]: #1080
+#1247 := [unit-resolution #1081 #1246]: #1079
+#1248 := [th-lemma #1247 #1245 #1243]: false
+#1250 := [lemma #1248]: #1249
+#1253 := (or #1252 #1241)
+#1254 := [th-lemma]: #1253
+#1310 := [unit-resolution #1254 #1250]: #1252
+#1326 := [mp #1310 #1325]: #796
+#1328 := (or #724 #772)
+decl uf_13 :: T3
+#33 := uf_13
+#609 := (uf_1 uf_13 #44)
+#773 := (= #609 #771)
+#775 := (ite #724 #773 #772)
+#32 := (uf_1 uf_9 #6)
+#553 := (pattern #32)
+#34 := (uf_1 uf_13 #6)
+#551 := (pattern #34)
+#456 := (= #32 #89)
+#455 := (= #34 #89)
+#457 := (ite #119 #455 #456)
+#554 := (forall (vars (?x4 T2)) (:pat #541 #551 #552 #553) #457)
+#460 := (forall (vars (?x4 T2)) #457)
+#557 := (iff #460 #554)
+#555 := (iff #457 #457)
+#556 := [refl]: #555
+#558 := [quant-intro #556]: #557
+#143 := (ite #119 #34 #32)
+#373 := (= #89 #143)
+#374 := (forall (vars (?x4 T2)) #373)
+#461 := (iff #374 #460)
+#458 := (iff #373 #457)
+#459 := [rewrite]: #458
+#462 := [quant-intro #459]: #461
+#362 := (~ #374 #374)
+#364 := (~ #373 #373)
+#365 := [refl]: #364
+#363 := [nnf-pos #365]: #362
+#31 := (uf_3 uf_12 #6)
+#148 := (= #31 #143)
+#151 := (forall (vars (?x4 T2)) #148)
+#375 := (iff #151 #374)
+#376 := [rewrite* #113]: #375
+#35 := (ite #17 #32 #34)
+#36 := (= #31 #35)
+#37 := (forall (vars (?x4 T2)) #36)
+#152 := (iff #37 #151)
+#149 := (iff #36 #148)
+#146 := (= #35 #143)
+#140 := (ite #118 #32 #34)
+#144 := (= #140 #143)
+#145 := [rewrite]: #144
+#141 := (= #35 #140)
+#142 := [monotonicity #123]: #141
+#147 := [trans #142 #145]: #146
+#150 := [monotonicity #147]: #149
+#153 := [quant-intro #150]: #152
+#115 := [asserted]: #37
+#154 := [mp #115 #153]: #151
+#377 := [mp #154 #376]: #374
+#360 := [mp~ #377 #363]: #374
+#463 := [mp #360 #462]: #460
+#559 := [mp #463 #558]: #554
+#778 := (not #554)
+#779 := (or #778 #775)
+#714 := (+ #611 #120)
+#715 := (>= #714 0::int)
+#774 := (ite #715 #773 #772)
+#780 := (or #778 #774)
+#782 := (iff #780 #779)
+#784 := (iff #779 #779)
+#785 := [rewrite]: #784
+#776 := (iff #774 #775)
+#727 := (iff #715 #724)
+#717 := (+ #120 #611)
+#720 := (>= #717 0::int)
+#725 := (iff #720 #724)
+#726 := [rewrite]: #725
+#721 := (iff #715 #720)
+#718 := (= #714 #717)
+#719 := [rewrite]: #718
+#722 := [monotonicity #719]: #721
+#728 := [trans #722 #726]: #727
+#777 := [monotonicity #728]: #776
+#783 := [monotonicity #777]: #782
+#786 := [trans #783 #785]: #782
+#781 := [quant-inst]: #780
+#787 := [mp #781 #786]: #779
+#1327 := [unit-resolution #787 #559]: #775
+#788 := (not #775)
+#791 := (or #788 #724 #772)
+#792 := [def-axiom]: #791
+#1329 := [unit-resolution #792 #1327]: #1328
+#1330 := [unit-resolution #1329 #1326]: #724
+#988 := (>= #622 0::int)
+#1331 := (or #1312 #988)
+#1332 := [th-lemma]: #1331
+#1333 := [unit-resolution #1332 #1281]: #988
+#761 := (not #724)
+#1334 := (not #988)
+#1335 := (or #1271 #1334 #761)
+#1336 := [th-lemma]: #1335
+#1337 := [unit-resolution #1336 #1333 #1330]: #1271
+#1338 := (not #1271)
+#1340 := (or #1269 #1338 #1339)
+#1341 := [th-lemma]: #1340
+#1343 := [unit-resolution #1341 #1337]: #1342
+#1344 := [unit-resolution #1343 #1299]: #1339
+#990 := (>= #916 0::int)
+#1345 := (or #1302 #990)
+#1346 := [th-lemma]: #1345
+#1347 := [unit-resolution #1346 #1301]: #990
+#1348 := (not #990)
+#1349 := (or #836 #1348 #1265)
+#1350 := [th-lemma]: #1349
+#1351 := [unit-resolution #1350 #1347 #1344]: #836
+#1353 := (or #815 #800)
+#801 := (uf_1 uf_13 #22)
+#820 := (= #799 #801)
+#823 := (ite #815 #820 #800)
+#476 := (= #32 #79)
+#475 := (= #34 #79)
+#477 := (ite #157 #475 #476)
+#568 := (forall (vars (?x6 T2)) (:pat #541 #551 #567 #553) #477)
+#480 := (forall (vars (?x6 T2)) #477)
+#571 := (iff #480 #568)
+#569 := (iff #477 #477)
+#570 := [refl]: #569
+#572 := [quant-intro #570]: #571
+#181 := (ite #157 #34 #32)
+#383 := (= #79 #181)
+#384 := (forall (vars (?x6 T2)) #383)
+#481 := (iff #384 #480)
+#478 := (iff #383 #477)
+#479 := [rewrite]: #478
+#482 := [quant-intro #479]: #481
+#352 := (~ #384 #384)
+#354 := (~ #383 #383)
+#355 := [refl]: #354
+#353 := [nnf-pos #355]: #352
+#51 := (uf_3 uf_16 #6)
+#186 := (= #51 #181)
+#189 := (forall (vars (?x6 T2)) #186)
+#385 := (iff #189 #384)
+#386 := [rewrite* #113]: #385
+#52 := (ite #42 #32 #34)
+#53 := (= #51 #52)
+#54 := (forall (vars (?x6 T2)) #53)
+#190 := (iff #54 #189)
+#187 := (iff #53 #186)
+#184 := (= #52 #181)
+#178 := (ite #156 #32 #34)
+#182 := (= #178 #181)
+#183 := [rewrite]: #182
+#179 := (= #52 #178)
+#180 := [monotonicity #161]: #179
+#185 := [trans #180 #183]: #184
+#188 := [monotonicity #185]: #187
+#191 := [quant-intro #188]: #190
+#139 := [asserted]: #54
+#192 := [mp #139 #191]: #189
+#387 := [mp #192 #386]: #384
+#402 := [mp~ #387 #353]: #384
+#483 := [mp #402 #482]: #480
+#573 := [mp #483 #572]: #568
+#634 := (not #568)
+#826 := (or #634 #823)
+#802 := (= #801 #799)
+#804 := (+ #803 #158)
+#805 := (>= #804 0::int)
+#806 := (ite #805 #802 #800)
+#827 := (or #634 #806)
+#829 := (iff #827 #826)
+#831 := (iff #826 #826)
+#832 := [rewrite]: #831
+#824 := (iff #806 #823)
+#821 := (iff #802 #820)
+#822 := [rewrite]: #821
+#818 := (iff #805 #815)
+#807 := (+ #158 #803)
+#810 := (>= #807 0::int)
+#816 := (iff #810 #815)
+#817 := [rewrite]: #816
+#811 := (iff #805 #810)
+#808 := (= #804 #807)
+#809 := [rewrite]: #808
+#812 := [monotonicity #809]: #811
+#819 := [trans #812 #817]: #818
+#825 := [monotonicity #819 #822]: #824
+#830 := [monotonicity #825]: #829
+#833 := [trans #830 #832]: #829
+#828 := [quant-inst]: #827
+#834 := [mp #828 #833]: #826
+#1352 := [unit-resolution #834 #573]: #823
+#835 := (not #823)
+#839 := (or #835 #815 #800)
+#840 := [def-axiom]: #839
+#1354 := [unit-resolution #840 #1352]: #1353
+#1355 := [unit-resolution #1354 #1351]: #800
+#1357 := [symm #1355]: #1356
+#1358 := [trans #1357 #1309]: #1266
+#1359 := (not #1266)
+#1360 := (or #1359 #1272)
+#1361 := [th-lemma]: #1360
+#1362 := [unit-resolution #1361 #1358]: #1272
+#1085 := (uf_1 uf_18 #22)
+#1099 := (* -1::real #1085)
+#1112 := (+ #902 #1099)
+#1113 := (<= #1112 0::real)
+#1137 := (not #1113)
+#960 := (uf_1 #88 #22)
+#1100 := (+ #960 #1099)
+#1101 := (>= #1100 0::real)
+#1118 := (or #1101 #1113)
+#1121 := (not #1118)
+#1125 := (or #1124 #1121)
+#1086 := (+ #1085 #1084)
+#1087 := (>= #1086 0::real)
+#1088 := (* -1::real #960)
+#1089 := (+ #1085 #1088)
+#1090 := (<= #1089 0::real)
+#1091 := (or #1090 #1087)
+#1092 := (not #1091)
+#1126 := (or #1124 #1092)
+#1128 := (iff #1126 #1125)
+#1130 := (iff #1125 #1125)
+#1131 := [rewrite]: #1130
+#1122 := (iff #1092 #1121)
+#1119 := (iff #1091 #1118)
+#1116 := (iff #1087 #1113)
+#1106 := (+ #1084 #1085)
+#1109 := (>= #1106 0::real)
+#1114 := (iff #1109 #1113)
+#1115 := [rewrite]: #1114
+#1110 := (iff #1087 #1109)
+#1107 := (= #1086 #1106)
+#1108 := [rewrite]: #1107
+#1111 := [monotonicity #1108]: #1110
+#1117 := [trans #1111 #1115]: #1116
+#1104 := (iff #1090 #1101)
+#1093 := (+ #1088 #1085)
+#1096 := (<= #1093 0::real)
+#1102 := (iff #1096 #1101)
+#1103 := [rewrite]: #1102
+#1097 := (iff #1090 #1096)
+#1094 := (= #1089 #1093)
+#1095 := [rewrite]: #1094
+#1098 := [monotonicity #1095]: #1097
+#1105 := [trans #1098 #1103]: #1104
+#1120 := [monotonicity #1105 #1117]: #1119
+#1123 := [monotonicity #1120]: #1122
+#1129 := [monotonicity #1123]: #1128
+#1132 := [trans #1129 #1131]: #1128
+#1127 := [quant-inst]: #1126
+#1133 := [mp #1127 #1132]: #1125
+#1363 := [unit-resolution #1133 #606]: #1121
+#1138 := (or #1118 #1137)
+#1139 := [def-axiom]: #1138
+#1364 := [unit-resolution #1139 #1363]: #1137
+#1200 := (+ #799 #1099)
+#1201 := (>= #1200 0::real)
+#1231 := (not #1201)
+#847 := (uf_1 #83 #22)
+#1210 := (+ #847 #1099)
+#1211 := (<= #1210 0::real)
+#1216 := (or #1201 #1211)
+#1219 := (not #1216)
+#1222 := (or #1066 #1219)
+#1197 := (* -1::real #847)
+#1198 := (+ #1085 #1197)
+#1199 := (>= #1198 0::real)
+#1202 := (or #1201 #1199)
+#1203 := (not #1202)
+#1223 := (or #1066 #1203)
+#1225 := (iff #1223 #1222)
+#1227 := (iff #1222 #1222)
+#1228 := [rewrite]: #1227
+#1220 := (iff #1203 #1219)
+#1217 := (iff #1202 #1216)
+#1214 := (iff #1199 #1211)
+#1204 := (+ #1197 #1085)
+#1207 := (>= #1204 0::real)
+#1212 := (iff #1207 #1211)
+#1213 := [rewrite]: #1212
+#1208 := (iff #1199 #1207)
+#1205 := (= #1198 #1204)
+#1206 := [rewrite]: #1205
+#1209 := [monotonicity #1206]: #1208
+#1215 := [trans #1209 #1213]: #1214
+#1218 := [monotonicity #1215]: #1217
+#1221 := [monotonicity #1218]: #1220
+#1226 := [monotonicity #1221]: #1225
+#1229 := [trans #1226 #1228]: #1225
+#1224 := [quant-inst]: #1223
+#1230 := [mp #1224 #1229]: #1222
+#1365 := [unit-resolution #1230 #600]: #1219
+#1232 := (or #1216 #1231)
+#1233 := [def-axiom]: #1232
+#1366 := [unit-resolution #1233 #1365]: #1231
+[th-lemma #1366 #1364 #1362]: false
+unsat
+NQHwTeL311Tq3wf2s5BReA 419 0
+#2 := false
+#194 := 0::real
+decl uf_4 :: (-> T2 T3 real)
+decl uf_6 :: (-> T1 T3)
+decl uf_3 :: T1
+#21 := uf_3
+#25 := (uf_6 uf_3)
+decl uf_5 :: T2
+#24 := uf_5
+#26 := (uf_4 uf_5 #25)
+decl uf_7 :: T2
+#27 := uf_7
+#28 := (uf_4 uf_7 #25)
+decl uf_10 :: T1
+#38 := uf_10
+#42 := (uf_6 uf_10)
+decl uf_9 :: T2
+#33 := uf_9
+#43 := (uf_4 uf_9 #42)
+#41 := (= uf_3 uf_10)
+#44 := (ite #41 #43 #28)
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+#39 := (uf_2 uf_10)
+#226 := -1::int
+#229 := (* -1::int #39)
+#22 := (uf_2 uf_3)
+#230 := (+ #22 #229)
+#228 := (>= #230 0::int)
+#236 := (ite #228 #44 #26)
+#192 := -1::real
+#244 := (* -1::real #236)
+#642 := (+ #26 #244)
+#643 := (<= #642 0::real)
+#567 := (= #26 #236)
+#227 := (not #228)
+decl uf_1 :: (-> int T1)
+#593 := (uf_1 #39)
+#660 := (= #593 uf_10)
+#594 := (= uf_10 #593)
+#4 := (:var 0 T1)
+#5 := (uf_2 #4)
+#546 := (pattern #5)
+#6 := (uf_1 #5)
+#93 := (= #4 #6)
+#547 := (forall (vars (?x1 T1)) (:pat #546) #93)
+#96 := (forall (vars (?x1 T1)) #93)
+#550 := (iff #96 #547)
+#548 := (iff #93 #93)
+#549 := [refl]: #548
+#551 := [quant-intro #549]: #550
+#448 := (~ #96 #96)
+#450 := (~ #93 #93)
+#451 := [refl]: #450
+#449 := [nnf-pos #451]: #448
+#7 := (= #6 #4)
+#8 := (forall (vars (?x1 T1)) #7)
+#97 := (iff #8 #96)
+#94 := (iff #7 #93)
+#95 := [rewrite]: #94
+#98 := [quant-intro #95]: #97
+#92 := [asserted]: #8
+#101 := [mp #92 #98]: #96
+#446 := [mp~ #101 #449]: #96
+#552 := [mp #446 #551]: #547
+#595 := (not #547)
+#600 := (or #595 #594)
+#601 := [quant-inst]: #600
+#654 := [unit-resolution #601 #552]: #594
+#680 := [symm #654]: #660
+#681 := (= uf_3 #593)
+#591 := (uf_1 #22)
+#658 := (= #591 #593)
+#656 := (= #593 #591)
+#652 := (= #39 #22)
+#647 := (= #22 #39)
+#290 := (<= #230 0::int)
+#70 := (<= #22 #39)
+#388 := (iff #70 #290)
+#389 := [rewrite]: #388
+#341 := [asserted]: #70
+#390 := [mp #341 #389]: #290
+#646 := [hypothesis]: #228
+#648 := [th-lemma #646 #390]: #647
+#653 := [symm #648]: #652
+#657 := [monotonicity #653]: #656
+#659 := [symm #657]: #658
+#592 := (= uf_3 #591)
+#596 := (or #595 #592)
+#597 := [quant-inst]: #596
+#655 := [unit-resolution #597 #552]: #592
+#682 := [trans #655 #659]: #681
+#683 := [trans #682 #680]: #41
+#570 := (not #41)
+decl uf_11 :: T2
+#47 := uf_11
+#59 := (uf_4 uf_11 #42)
+#278 := (ite #41 #26 #59)
+#459 := (* -1::real #278)
+#637 := (+ #26 #459)
+#639 := (>= #637 0::real)
+#585 := (= #26 #278)
+#661 := [hypothesis]: #41
+#587 := (or #570 #585)
+#588 := [def-axiom]: #587
+#662 := [unit-resolution #588 #661]: #585
+#663 := (not #585)
+#664 := (or #663 #639)
+#665 := [th-lemma]: #664
+#666 := [unit-resolution #665 #662]: #639
+decl uf_8 :: T2
+#30 := uf_8
+#56 := (uf_4 uf_8 #42)
+#357 := (* -1::real #56)
+#358 := (+ #43 #357)
+#356 := (>= #358 0::real)
+#355 := (not #356)
+#374 := (* -1::real #59)
+#375 := (+ #56 #374)
+#373 := (>= #375 0::real)
+#376 := (not #373)
+#381 := (and #355 #376)
+#64 := (< #39 #39)
+#67 := (ite #64 #43 #59)
+#68 := (< #56 #67)
+#53 := (uf_4 uf_5 #42)
+#65 := (ite #64 #53 #43)
+#66 := (< #65 #56)
+#69 := (and #66 #68)
+#382 := (iff #69 #381)
+#379 := (iff #68 #376)
+#370 := (< #56 #59)
+#377 := (iff #370 #376)
+#378 := [rewrite]: #377
+#371 := (iff #68 #370)
+#368 := (= #67 #59)
+#363 := (ite false #43 #59)
+#366 := (= #363 #59)
+#367 := [rewrite]: #366
+#364 := (= #67 #363)
+#343 := (iff #64 false)
+#344 := [rewrite]: #343
+#365 := [monotonicity #344]: #364
+#369 := [trans #365 #367]: #368
+#372 := [monotonicity #369]: #371
+#380 := [trans #372 #378]: #379
+#361 := (iff #66 #355)
+#352 := (< #43 #56)
+#359 := (iff #352 #355)
+#360 := [rewrite]: #359
+#353 := (iff #66 #352)
+#350 := (= #65 #43)
+#345 := (ite false #53 #43)
+#348 := (= #345 #43)
+#349 := [rewrite]: #348
+#346 := (= #65 #345)
+#347 := [monotonicity #344]: #346
+#351 := [trans #347 #349]: #350
+#354 := [monotonicity #351]: #353
+#362 := [trans #354 #360]: #361
+#383 := [monotonicity #362 #380]: #382
+#340 := [asserted]: #69
+#384 := [mp #340 #383]: #381
+#385 := [and-elim #384]: #355
+#394 := (* -1::real #53)
+#395 := (+ #43 #394)
+#393 := (>= #395 0::real)
+#54 := (uf_4 uf_7 #42)
+#402 := (* -1::real #54)
+#403 := (+ #53 #402)
+#401 := (>= #403 0::real)
+#397 := (+ #43 #374)
+#398 := (<= #397 0::real)
+#412 := (and #393 #398 #401)
+#73 := (<= #43 #59)
+#72 := (<= #53 #43)
+#74 := (and #72 #73)
+#71 := (<= #54 #53)
+#75 := (and #71 #74)
+#415 := (iff #75 #412)
+#406 := (and #393 #398)
+#409 := (and #401 #406)
+#413 := (iff #409 #412)
+#414 := [rewrite]: #413
+#410 := (iff #75 #409)
+#407 := (iff #74 #406)
+#399 := (iff #73 #398)
+#400 := [rewrite]: #399
+#392 := (iff #72 #393)
+#396 := [rewrite]: #392
+#408 := [monotonicity #396 #400]: #407
+#404 := (iff #71 #401)
+#405 := [rewrite]: #404
+#411 := [monotonicity #405 #408]: #410
+#416 := [trans #411 #414]: #415
+#342 := [asserted]: #75
+#417 := [mp #342 #416]: #412
+#418 := [and-elim #417]: #393
+#650 := (+ #26 #394)
+#651 := (<= #650 0::real)
+#649 := (= #26 #53)
+#671 := (= #53 #26)
+#669 := (= #42 #25)
+#667 := (= #25 #42)
+#668 := [monotonicity #661]: #667
+#670 := [symm #668]: #669
+#672 := [monotonicity #670]: #671
+#673 := [symm #672]: #649
+#674 := (not #649)
+#675 := (or #674 #651)
+#676 := [th-lemma]: #675
+#677 := [unit-resolution #676 #673]: #651
+#462 := (+ #56 #459)
+#465 := (>= #462 0::real)
+#438 := (not #465)
+#316 := (ite #290 #278 #43)
+#326 := (* -1::real #316)
+#327 := (+ #56 #326)
+#325 := (>= #327 0::real)
+#324 := (not #325)
+#439 := (iff #324 #438)
+#466 := (iff #325 #465)
+#463 := (= #327 #462)
+#460 := (= #326 #459)
+#457 := (= #316 #278)
+#1 := true
+#452 := (ite true #278 #43)
+#455 := (= #452 #278)
+#456 := [rewrite]: #455
+#453 := (= #316 #452)
+#444 := (iff #290 true)
+#445 := [iff-true #390]: #444
+#454 := [monotonicity #445]: #453
+#458 := [trans #454 #456]: #457
+#461 := [monotonicity #458]: #460
+#464 := [monotonicity #461]: #463
+#467 := [monotonicity #464]: #466
+#468 := [monotonicity #467]: #439
+#297 := (ite #290 #54 #53)
+#305 := (* -1::real #297)
+#306 := (+ #56 #305)
+#307 := (<= #306 0::real)
+#308 := (not #307)
+#332 := (and #308 #324)
+#58 := (= uf_10 uf_3)
+#60 := (ite #58 #26 #59)
+#52 := (< #39 #22)
+#61 := (ite #52 #43 #60)
+#62 := (< #56 #61)
+#55 := (ite #52 #53 #54)
+#57 := (< #55 #56)
+#63 := (and #57 #62)
+#335 := (iff #63 #332)
+#281 := (ite #52 #43 #278)
+#284 := (< #56 #281)
+#287 := (and #57 #284)
+#333 := (iff #287 #332)
+#330 := (iff #284 #324)
+#321 := (< #56 #316)
+#328 := (iff #321 #324)
+#329 := [rewrite]: #328
+#322 := (iff #284 #321)
+#319 := (= #281 #316)
+#291 := (not #290)
+#313 := (ite #291 #43 #278)
+#317 := (= #313 #316)
+#318 := [rewrite]: #317
+#314 := (= #281 #313)
+#292 := (iff #52 #291)
+#293 := [rewrite]: #292
+#315 := [monotonicity #293]: #314
+#320 := [trans #315 #318]: #319
+#323 := [monotonicity #320]: #322
+#331 := [trans #323 #329]: #330
+#311 := (iff #57 #308)
+#302 := (< #297 #56)
+#309 := (iff #302 #308)
+#310 := [rewrite]: #309
+#303 := (iff #57 #302)
+#300 := (= #55 #297)
+#294 := (ite #291 #53 #54)
+#298 := (= #294 #297)
+#299 := [rewrite]: #298
+#295 := (= #55 #294)
+#296 := [monotonicity #293]: #295
+#301 := [trans #296 #299]: #300
+#304 := [monotonicity #301]: #303
+#312 := [trans #304 #310]: #311
+#334 := [monotonicity #312 #331]: #333
+#288 := (iff #63 #287)
+#285 := (iff #62 #284)
+#282 := (= #61 #281)
+#279 := (= #60 #278)
+#225 := (iff #58 #41)
+#277 := [rewrite]: #225
+#280 := [monotonicity #277]: #279
+#283 := [monotonicity #280]: #282
+#286 := [monotonicity #283]: #285
+#289 := [monotonicity #286]: #288
+#336 := [trans #289 #334]: #335
+#179 := [asserted]: #63
+#337 := [mp #179 #336]: #332
+#339 := [and-elim #337]: #324
+#469 := [mp #339 #468]: #438
+#678 := [th-lemma #469 #677 #418 #385 #666]: false
+#679 := [lemma #678]: #570
+#684 := [unit-resolution #679 #683]: false
+#685 := [lemma #684]: #227
+#577 := (or #228 #567)
+#578 := [def-axiom]: #577
+#645 := [unit-resolution #578 #685]: #567
+#686 := (not #567)
+#687 := (or #686 #643)
+#688 := [th-lemma]: #687
+#689 := [unit-resolution #688 #645]: #643
+#31 := (uf_4 uf_8 #25)
+#245 := (+ #31 #244)
+#246 := (<= #245 0::real)
+#247 := (not #246)
+#34 := (uf_4 uf_9 #25)
+#48 := (uf_4 uf_11 #25)
+#255 := (ite #228 #48 #34)
+#264 := (* -1::real #255)
+#265 := (+ #31 #264)
+#263 := (>= #265 0::real)
+#266 := (not #263)
+#271 := (and #247 #266)
+#40 := (< #22 #39)
+#49 := (ite #40 #34 #48)
+#50 := (< #31 #49)
+#45 := (ite #40 #26 #44)
+#46 := (< #45 #31)
+#51 := (and #46 #50)
+#272 := (iff #51 #271)
+#269 := (iff #50 #266)
+#260 := (< #31 #255)
+#267 := (iff #260 #266)
+#268 := [rewrite]: #267
+#261 := (iff #50 #260)
+#258 := (= #49 #255)
+#252 := (ite #227 #34 #48)
+#256 := (= #252 #255)
+#257 := [rewrite]: #256
+#253 := (= #49 #252)
+#231 := (iff #40 #227)
+#232 := [rewrite]: #231
+#254 := [monotonicity #232]: #253
+#259 := [trans #254 #257]: #258
+#262 := [monotonicity #259]: #261
+#270 := [trans #262 #268]: #269
+#250 := (iff #46 #247)
+#241 := (< #236 #31)
+#248 := (iff #241 #247)
+#249 := [rewrite]: #248
+#242 := (iff #46 #241)
+#239 := (= #45 #236)
+#233 := (ite #227 #26 #44)
+#237 := (= #233 #236)
+#238 := [rewrite]: #237
+#234 := (= #45 #233)
+#235 := [monotonicity #232]: #234
+#240 := [trans #235 #238]: #239
+#243 := [monotonicity #240]: #242
+#251 := [trans #243 #249]: #250
+#273 := [monotonicity #251 #270]: #272
+#178 := [asserted]: #51
+#274 := [mp #178 #273]: #271
+#275 := [and-elim #274]: #247
+#196 := (* -1::real #31)
+#212 := (+ #26 #196)
+#213 := (<= #212 0::real)
+#214 := (not #213)
+#197 := (+ #28 #196)
+#195 := (>= #197 0::real)
+#193 := (not #195)
+#219 := (and #193 #214)
+#23 := (< #22 #22)
+#35 := (ite #23 #34 #26)
+#36 := (< #31 #35)
+#29 := (ite #23 #26 #28)
+#32 := (< #29 #31)
+#37 := (and #32 #36)
+#220 := (iff #37 #219)
+#217 := (iff #36 #214)
+#209 := (< #31 #26)
+#215 := (iff #209 #214)
+#216 := [rewrite]: #215
+#210 := (iff #36 #209)
+#207 := (= #35 #26)
+#202 := (ite false #34 #26)
+#205 := (= #202 #26)
+#206 := [rewrite]: #205
+#203 := (= #35 #202)
+#180 := (iff #23 false)
+#181 := [rewrite]: #180
+#204 := [monotonicity #181]: #203
+#208 := [trans #204 #206]: #207
+#211 := [monotonicity #208]: #210
+#218 := [trans #211 #216]: #217
+#200 := (iff #32 #193)
+#189 := (< #28 #31)
+#198 := (iff #189 #193)
+#199 := [rewrite]: #198
+#190 := (iff #32 #189)
+#187 := (= #29 #28)
+#182 := (ite false #26 #28)
+#185 := (= #182 #28)
+#186 := [rewrite]: #185
+#183 := (= #29 #182)
+#184 := [monotonicity #181]: #183
+#188 := [trans #184 #186]: #187
+#191 := [monotonicity #188]: #190
+#201 := [trans #191 #199]: #200
+#221 := [monotonicity #201 #218]: #220
+#177 := [asserted]: #37
+#222 := [mp #177 #221]: #219
+#224 := [and-elim #222]: #214
+[th-lemma #224 #275 #689]: false
+unsat
+NX/HT1QOfbspC2LtZNKpBA 428 0
+#2 := false
+decl uf_10 :: T1
+#38 := uf_10
+decl uf_3 :: T1
+#21 := uf_3
+#45 := (= uf_3 uf_10)
+decl uf_1 :: (-> int T1)
+decl uf_2 :: (-> T1 int)
+#39 := (uf_2 uf_10)
+#588 := (uf_1 #39)
+#686 := (= #588 uf_10)
+#589 := (= uf_10 #588)
+#4 := (:var 0 T1)
+#5 := (uf_2 #4)
+#541 := (pattern #5)
+#6 := (uf_1 #5)
+#93 := (= #4 #6)
+#542 := (forall (vars (?x1 T1)) (:pat #541) #93)
+#96 := (forall (vars (?x1 T1)) #93)
+#545 := (iff #96 #542)
+#543 := (iff #93 #93)
+#544 := [refl]: #543
+#546 := [quant-intro #544]: #545
+#454 := (~ #96 #96)
+#456 := (~ #93 #93)
+#457 := [refl]: #456
+#455 := [nnf-pos #457]: #454
+#7 := (= #6 #4)
+#8 := (forall (vars (?x1 T1)) #7)
+#97 := (iff #8 #96)
+#94 := (iff #7 #93)
+#95 := [rewrite]: #94
+#98 := [quant-intro #95]: #97
+#92 := [asserted]: #8
+#101 := [mp #92 #98]: #96
+#452 := [mp~ #101 #455]: #96
+#547 := [mp #452 #546]: #542
+#590 := (not #542)
+#595 := (or #590 #589)
+#596 := [quant-inst]: #595
+#680 := [unit-resolution #596 #547]: #589
+#687 := [symm #680]: #686
+#688 := (= uf_3 #588)
+#22 := (uf_2 uf_3)
+#586 := (uf_1 #22)
+#684 := (= #586 #588)
+#682 := (= #588 #586)
+#678 := (= #39 #22)
+#676 := (= #22 #39)
+#9 := 0::int
+#227 := -1::int
+#230 := (* -1::int #39)
+#231 := (+ #22 #230)
+#296 := (<= #231 0::int)
+#70 := (<= #22 #39)
+#393 := (iff #70 #296)
+#394 := [rewrite]: #393
+#347 := [asserted]: #70
+#395 := [mp #347 #394]: #296
+#229 := (>= #231 0::int)
+decl uf_4 :: (-> T2 T3 real)
+decl uf_6 :: (-> T1 T3)
+#25 := (uf_6 uf_3)
+decl uf_7 :: T2
+#27 := uf_7
+#28 := (uf_4 uf_7 #25)
+decl uf_9 :: T2
+#33 := uf_9
+#34 := (uf_4 uf_9 #25)
+#46 := (uf_6 uf_10)
+decl uf_5 :: T2
+#24 := uf_5
+#47 := (uf_4 uf_5 #46)
+#48 := (ite #45 #47 #34)
+#256 := (ite #229 #48 #28)
+#568 := (= #28 #256)
+#648 := (not #568)
+#194 := 0::real
+#192 := -1::real
+#265 := (* -1::real #256)
+#640 := (+ #28 #265)
+#642 := (>= #640 0::real)
+#645 := (not #642)
+#643 := [hypothesis]: #642
+decl uf_8 :: T2
+#30 := uf_8
+#31 := (uf_4 uf_8 #25)
+#266 := (+ #31 #265)
+#264 := (>= #266 0::real)
+#267 := (not #264)
+#26 := (uf_4 uf_5 #25)
+decl uf_11 :: T2
+#41 := uf_11
+#42 := (uf_4 uf_11 #25)
+#237 := (ite #229 #42 #26)
+#245 := (* -1::real #237)
+#246 := (+ #31 #245)
+#247 := (<= #246 0::real)
+#248 := (not #247)
+#272 := (and #248 #267)
+#40 := (< #22 #39)
+#49 := (ite #40 #28 #48)
+#50 := (< #31 #49)
+#43 := (ite #40 #26 #42)
+#44 := (< #43 #31)
+#51 := (and #44 #50)
+#273 := (iff #51 #272)
+#270 := (iff #50 #267)
+#261 := (< #31 #256)
+#268 := (iff #261 #267)
+#269 := [rewrite]: #268
+#262 := (iff #50 #261)
+#259 := (= #49 #256)
+#228 := (not #229)
+#253 := (ite #228 #28 #48)
+#257 := (= #253 #256)
+#258 := [rewrite]: #257
+#254 := (= #49 #253)
+#232 := (iff #40 #228)
+#233 := [rewrite]: #232
+#255 := [monotonicity #233]: #254
+#260 := [trans #255 #258]: #259
+#263 := [monotonicity #260]: #262
+#271 := [trans #263 #269]: #270
+#251 := (iff #44 #248)
+#242 := (< #237 #31)
+#249 := (iff #242 #248)
+#250 := [rewrite]: #249
+#243 := (iff #44 #242)
+#240 := (= #43 #237)
+#234 := (ite #228 #26 #42)
+#238 := (= #234 #237)
+#239 := [rewrite]: #238
+#235 := (= #43 #234)
+#236 := [monotonicity #233]: #235
+#241 := [trans #236 #239]: #240
+#244 := [monotonicity #241]: #243
+#252 := [trans #244 #250]: #251
+#274 := [monotonicity #252 #271]: #273
+#178 := [asserted]: #51
+#275 := [mp #178 #274]: #272
+#277 := [and-elim #275]: #267
+#196 := (* -1::real #31)
+#197 := (+ #28 #196)
+#195 := (>= #197 0::real)
+#193 := (not #195)
+#213 := (* -1::real #34)
+#214 := (+ #31 #213)
+#212 := (>= #214 0::real)
+#215 := (not #212)
+#220 := (and #193 #215)
+#23 := (< #22 #22)
+#35 := (ite #23 #28 #34)
+#36 := (< #31 #35)
+#29 := (ite #23 #26 #28)
+#32 := (< #29 #31)
+#37 := (and #32 #36)
+#221 := (iff #37 #220)
+#218 := (iff #36 #215)
+#209 := (< #31 #34)
+#216 := (iff #209 #215)
+#217 := [rewrite]: #216
+#210 := (iff #36 #209)
+#207 := (= #35 #34)
+#202 := (ite false #28 #34)
+#205 := (= #202 #34)
+#206 := [rewrite]: #205
+#203 := (= #35 #202)
+#180 := (iff #23 false)
+#181 := [rewrite]: #180
+#204 := [monotonicity #181]: #203
+#208 := [trans #204 #206]: #207
+#211 := [monotonicity #208]: #210
+#219 := [trans #211 #217]: #218
+#200 := (iff #32 #193)
+#189 := (< #28 #31)
+#198 := (iff #189 #193)
+#199 := [rewrite]: #198
+#190 := (iff #32 #189)
+#187 := (= #29 #28)
+#182 := (ite false #26 #28)
+#185 := (= #182 #28)
+#186 := [rewrite]: #185
+#183 := (= #29 #182)
+#184 := [monotonicity #181]: #183
+#188 := [trans #184 #186]: #187
+#191 := [monotonicity #188]: #190
+#201 := [trans #191 #199]: #200
+#222 := [monotonicity #201 #219]: #221
+#177 := [asserted]: #37
+#223 := [mp #177 #222]: #220
+#224 := [and-elim #223]: #193
+#644 := [th-lemma #224 #277 #643]: false
+#646 := [lemma #644]: #645
+#647 := [hypothesis]: #568
+#649 := (or #648 #642)
+#650 := [th-lemma]: #649
+#651 := [unit-resolution #650 #647 #646]: false
+#652 := [lemma #651]: #648
+#578 := (or #229 #568)
+#579 := [def-axiom]: #578
+#675 := [unit-resolution #579 #652]: #229
+#677 := [th-lemma #675 #395]: #676
+#679 := [symm #677]: #678
+#683 := [monotonicity #679]: #682
+#685 := [symm #683]: #684
+#587 := (= uf_3 #586)
+#591 := (or #590 #587)
+#592 := [quant-inst]: #591
+#681 := [unit-resolution #592 #547]: #587
+#689 := [trans #681 #685]: #688
+#690 := [trans #689 #687]: #45
+#571 := (not #45)
+#54 := (uf_4 uf_11 #46)
+#279 := (ite #45 #28 #54)
+#465 := (* -1::real #279)
+#632 := (+ #28 #465)
+#633 := (<= #632 0::real)
+#580 := (= #28 #279)
+#656 := [hypothesis]: #45
+#582 := (or #571 #580)
+#583 := [def-axiom]: #582
+#657 := [unit-resolution #583 #656]: #580
+#658 := (not #580)
+#659 := (or #658 #633)
+#660 := [th-lemma]: #659
+#661 := [unit-resolution #660 #657]: #633
+#57 := (uf_4 uf_8 #46)
+#363 := (* -1::real #57)
+#379 := (+ #47 #363)
+#380 := (<= #379 0::real)
+#381 := (not #380)
+#364 := (+ #54 #363)
+#362 := (>= #364 0::real)
+#361 := (not #362)
+#386 := (and #361 #381)
+#59 := (uf_4 uf_7 #46)
+#64 := (< #39 #39)
+#67 := (ite #64 #59 #47)
+#68 := (< #57 #67)
+#65 := (ite #64 #47 #54)
+#66 := (< #65 #57)
+#69 := (and #66 #68)
+#387 := (iff #69 #386)
+#384 := (iff #68 #381)
+#376 := (< #57 #47)
+#382 := (iff #376 #381)
+#383 := [rewrite]: #382
+#377 := (iff #68 #376)
+#374 := (= #67 #47)
+#369 := (ite false #59 #47)
+#372 := (= #369 #47)
+#373 := [rewrite]: #372
+#370 := (= #67 #369)
+#349 := (iff #64 false)
+#350 := [rewrite]: #349
+#371 := [monotonicity #350]: #370
+#375 := [trans #371 #373]: #374
+#378 := [monotonicity #375]: #377
+#385 := [trans #378 #383]: #384
+#367 := (iff #66 #361)
+#358 := (< #54 #57)
+#365 := (iff #358 #361)
+#366 := [rewrite]: #365
+#359 := (iff #66 #358)
+#356 := (= #65 #54)
+#351 := (ite false #47 #54)
+#354 := (= #351 #54)
+#355 := [rewrite]: #354
+#352 := (= #65 #351)
+#353 := [monotonicity #350]: #352
+#357 := [trans #353 #355]: #356
+#360 := [monotonicity #357]: #359
+#368 := [trans #360 #366]: #367
+#388 := [monotonicity #368 #385]: #387
+#346 := [asserted]: #69
+#389 := [mp #346 #388]: #386
+#391 := [and-elim #389]: #381
+#397 := (* -1::real #59)
+#398 := (+ #47 #397)
+#399 := (<= #398 0::real)
+#409 := (* -1::real #54)
+#410 := (+ #47 #409)
+#408 := (>= #410 0::real)
+#60 := (uf_4 uf_9 #46)
+#402 := (* -1::real #60)
+#403 := (+ #59 #402)
+#404 := (<= #403 0::real)
+#418 := (and #399 #404 #408)
+#73 := (<= #59 #60)
+#72 := (<= #47 #59)
+#74 := (and #72 #73)
+#71 := (<= #54 #47)
+#75 := (and #71 #74)
+#421 := (iff #75 #418)
+#412 := (and #399 #404)
+#415 := (and #408 #412)
+#419 := (iff #415 #418)
+#420 := [rewrite]: #419
+#416 := (iff #75 #415)
+#413 := (iff #74 #412)
+#405 := (iff #73 #404)
+#406 := [rewrite]: #405
+#400 := (iff #72 #399)
+#401 := [rewrite]: #400
+#414 := [monotonicity #401 #406]: #413
+#407 := (iff #71 #408)
+#411 := [rewrite]: #407
+#417 := [monotonicity #411 #414]: #416
+#422 := [trans #417 #420]: #421
+#348 := [asserted]: #75
+#423 := [mp #348 #422]: #418
+#424 := [and-elim #423]: #399
+#637 := (+ #28 #397)
+#639 := (>= #637 0::real)
+#636 := (= #28 #59)
+#666 := (= #59 #28)
+#664 := (= #46 #25)
+#662 := (= #25 #46)
+#663 := [monotonicity #656]: #662
+#665 := [symm #663]: #664
+#667 := [monotonicity #665]: #666
+#668 := [symm #667]: #636
+#669 := (not #636)
+#670 := (or #669 #639)
+#671 := [th-lemma]: #670
+#672 := [unit-resolution #671 #668]: #639
+#468 := (+ #57 #465)
+#471 := (<= #468 0::real)
+#444 := (not #471)
+#322 := (ite #296 #279 #47)
+#330 := (* -1::real #322)
+#331 := (+ #57 #330)
+#332 := (<= #331 0::real)
+#333 := (not #332)
+#445 := (iff #333 #444)
+#472 := (iff #332 #471)
+#469 := (= #331 #468)
+#466 := (= #330 #465)
+#463 := (= #322 #279)
+#1 := true
+#458 := (ite true #279 #47)
+#461 := (= #458 #279)
+#462 := [rewrite]: #461
+#459 := (= #322 #458)
+#450 := (iff #296 true)
+#451 := [iff-true #395]: #450
+#460 := [monotonicity #451]: #459
+#464 := [trans #460 #462]: #463
+#467 := [monotonicity #464]: #466
+#470 := [monotonicity #467]: #469
+#473 := [monotonicity #470]: #472
+#474 := [monotonicity #473]: #445
+#303 := (ite #296 #60 #59)
+#313 := (* -1::real #303)
+#314 := (+ #57 #313)
+#312 := (>= #314 0::real)
+#311 := (not #312)
+#338 := (and #311 #333)
+#52 := (< #39 #22)
+#61 := (ite #52 #59 #60)
+#62 := (< #57 #61)
+#53 := (= uf_10 uf_3)
+#55 := (ite #53 #28 #54)
+#56 := (ite #52 #47 #55)
+#58 := (< #56 #57)
+#63 := (and #58 #62)
+#341 := (iff #63 #338)
+#282 := (ite #52 #47 #279)
+#285 := (< #282 #57)
+#291 := (and #62 #285)
+#339 := (iff #291 #338)
+#336 := (iff #285 #333)
+#327 := (< #322 #57)
+#334 := (iff #327 #333)
+#335 := [rewrite]: #334
+#328 := (iff #285 #327)
+#325 := (= #282 #322)
+#297 := (not #296)
+#319 := (ite #297 #47 #279)
+#323 := (= #319 #322)
+#324 := [rewrite]: #323
+#320 := (= #282 #319)
+#298 := (iff #52 #297)
+#299 := [rewrite]: #298
+#321 := [monotonicity #299]: #320
+#326 := [trans #321 #324]: #325
+#329 := [monotonicity #326]: #328
+#337 := [trans #329 #335]: #336
+#317 := (iff #62 #311)
+#308 := (< #57 #303)
+#315 := (iff #308 #311)
+#316 := [rewrite]: #315
+#309 := (iff #62 #308)
+#306 := (= #61 #303)
+#300 := (ite #297 #59 #60)
+#304 := (= #300 #303)
+#305 := [rewrite]: #304
+#301 := (= #61 #300)
+#302 := [monotonicity #299]: #301
+#307 := [trans #302 #305]: #306
+#310 := [monotonicity #307]: #309
+#318 := [trans #310 #316]: #317
+#340 := [monotonicity #318 #337]: #339
+#294 := (iff #63 #291)
+#288 := (and #285 #62)
+#292 := (iff #288 #291)
+#293 := [rewrite]: #292
+#289 := (iff #63 #288)
+#286 := (iff #58 #285)
+#283 := (= #56 #282)
+#280 := (= #55 #279)
+#226 := (iff #53 #45)
+#278 := [rewrite]: #226
+#281 := [monotonicity #278]: #280
+#284 := [monotonicity #281]: #283
+#287 := [monotonicity #284]: #286
+#290 := [monotonicity #287]: #289
+#295 := [trans #290 #293]: #294
+#342 := [trans #295 #340]: #341
+#179 := [asserted]: #63
+#343 := [mp #179 #342]: #338
+#345 := [and-elim #343]: #333
+#475 := [mp #345 #474]: #444
+#673 := [th-lemma #475 #672 #424 #391 #661]: false
+#674 := [lemma #673]: #571
+[unit-resolution #674 #690]: false
+unsat
+IL2powemHjRpCJYwmXFxyw 211 0
+#2 := false
+#33 := 0::real
+decl uf_11 :: (-> T5 T6 real)
+decl uf_15 :: T6
+#28 := uf_15
+decl uf_16 :: T5
+#30 := uf_16
+#31 := (uf_11 uf_16 uf_15)
+decl uf_12 :: (-> T7 T8 T5)
+decl uf_14 :: T8
+#26 := uf_14
+decl uf_13 :: (-> T1 T7)
+decl uf_8 :: T1
+#16 := uf_8
+#25 := (uf_13 uf_8)
+#27 := (uf_12 #25 uf_14)
+#29 := (uf_11 #27 uf_15)
+#73 := -1::real
+#84 := (* -1::real #29)
+#85 := (+ #84 #31)
+#74 := (* -1::real #31)
+#75 := (+ #29 #74)
+#112 := (>= #75 0::real)
+#119 := (ite #112 #75 #85)
+#127 := (* -1::real #119)
+decl uf_17 :: T5
+#37 := uf_17
+#38 := (uf_11 uf_17 uf_15)
+#102 := -1/3::real
+#103 := (* -1/3::real #38)
+#128 := (+ #103 #127)
+#100 := 1/3::real
+#101 := (* 1/3::real #31)
+#129 := (+ #101 #128)
+#130 := (<= #129 0::real)
+#131 := (not #130)
+#40 := 3::real
+#39 := (- #31 #38)
+#41 := (/ #39 3::real)
+#32 := (- #29 #31)
+#35 := (- #32)
+#34 := (< #32 0::real)
+#36 := (ite #34 #35 #32)
+#42 := (< #36 #41)
+#136 := (iff #42 #131)
+#104 := (+ #101 #103)
+#78 := (< #75 0::real)
+#90 := (ite #78 #85 #75)
+#109 := (< #90 #104)
+#134 := (iff #109 #131)
+#124 := (< #119 #104)
+#132 := (iff #124 #131)
+#133 := [rewrite]: #132
+#125 := (iff #109 #124)
+#122 := (= #90 #119)
+#113 := (not #112)
+#116 := (ite #113 #85 #75)
+#120 := (= #116 #119)
+#121 := [rewrite]: #120
+#117 := (= #90 #116)
+#114 := (iff #78 #113)
+#115 := [rewrite]: #114
+#118 := [monotonicity #115]: #117
+#123 := [trans #118 #121]: #122
+#126 := [monotonicity #123]: #125
+#135 := [trans #126 #133]: #134
+#110 := (iff #42 #109)
+#107 := (= #41 #104)
+#93 := (* -1::real #38)
+#94 := (+ #31 #93)
+#97 := (/ #94 3::real)
+#105 := (= #97 #104)
+#106 := [rewrite]: #105
+#98 := (= #41 #97)
+#95 := (= #39 #94)
+#96 := [rewrite]: #95
+#99 := [monotonicity #96]: #98
+#108 := [trans #99 #106]: #107
+#91 := (= #36 #90)
+#76 := (= #32 #75)
+#77 := [rewrite]: #76
+#88 := (= #35 #85)
+#81 := (- #75)
+#86 := (= #81 #85)
+#87 := [rewrite]: #86
+#82 := (= #35 #81)
+#83 := [monotonicity #77]: #82
+#89 := [trans #83 #87]: #88
+#79 := (iff #34 #78)
+#80 := [monotonicity #77]: #79
+#92 := [monotonicity #80 #89 #77]: #91
+#111 := [monotonicity #92 #108]: #110
+#137 := [trans #111 #135]: #136
+#72 := [asserted]: #42
+#138 := [mp #72 #137]: #131
+decl uf_1 :: T1
+#4 := uf_1
+#43 := (uf_13 uf_1)
+#44 := (uf_12 #43 uf_14)
+#45 := (uf_11 #44 uf_15)
+#149 := (* -1::real #45)
+#150 := (+ #38 #149)
+#140 := (+ #93 #45)
+#161 := (<= #150 0::real)
+#168 := (ite #161 #140 #150)
+#176 := (* -1::real #168)
+#177 := (+ #103 #176)
+#178 := (+ #101 #177)
+#179 := (<= #178 0::real)
+#180 := (not #179)
+#46 := (- #45 #38)
+#48 := (- #46)
+#47 := (< #46 0::real)
+#49 := (ite #47 #48 #46)
+#50 := (< #49 #41)
+#185 := (iff #50 #180)
+#143 := (< #140 0::real)
+#155 := (ite #143 #150 #140)
+#158 := (< #155 #104)
+#183 := (iff #158 #180)
+#173 := (< #168 #104)
+#181 := (iff #173 #180)
+#182 := [rewrite]: #181
+#174 := (iff #158 #173)
+#171 := (= #155 #168)
+#162 := (not #161)
+#165 := (ite #162 #150 #140)
+#169 := (= #165 #168)
+#170 := [rewrite]: #169
+#166 := (= #155 #165)
+#163 := (iff #143 #162)
+#164 := [rewrite]: #163
+#167 := [monotonicity #164]: #166
+#172 := [trans #167 #170]: #171
+#175 := [monotonicity #172]: #174
+#184 := [trans #175 #182]: #183
+#159 := (iff #50 #158)
+#156 := (= #49 #155)
+#141 := (= #46 #140)
+#142 := [rewrite]: #141
+#153 := (= #48 #150)
+#146 := (- #140)
+#151 := (= #146 #150)
+#152 := [rewrite]: #151
+#147 := (= #48 #146)
+#148 := [monotonicity #142]: #147
+#154 := [trans #148 #152]: #153
+#144 := (iff #47 #143)
+#145 := [monotonicity #142]: #144
+#157 := [monotonicity #145 #154 #142]: #156
+#160 := [monotonicity #157 #108]: #159
+#186 := [trans #160 #184]: #185
+#139 := [asserted]: #50
+#187 := [mp #139 #186]: #180
+#299 := (+ #140 #176)
+#300 := (<= #299 0::real)
+#290 := (= #140 #168)
+#329 := [hypothesis]: #162
+#191 := (+ #29 #149)
+#192 := (<= #191 0::real)
+#51 := (<= #29 #45)
+#193 := (iff #51 #192)
+#194 := [rewrite]: #193
+#188 := [asserted]: #51
+#195 := [mp #188 #194]: #192
+#298 := (+ #75 #127)
+#301 := (<= #298 0::real)
+#284 := (= #75 #119)
+#302 := [hypothesis]: #113
+#296 := (+ #85 #127)
+#297 := (<= #296 0::real)
+#285 := (= #85 #119)
+#288 := (or #112 #285)
+#289 := [def-axiom]: #288
+#303 := [unit-resolution #289 #302]: #285
+#304 := (not #285)
+#305 := (or #304 #297)
+#306 := [th-lemma]: #305
+#307 := [unit-resolution #306 #303]: #297
+#315 := (not #290)
+#310 := (not #300)
+#311 := (or #310 #112)
+#308 := [hypothesis]: #300
+#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false
+#312 := [lemma #309]: #311
+#322 := [unit-resolution #312 #302]: #310
+#316 := (or #315 #300)
+#313 := [hypothesis]: #310
+#314 := [hypothesis]: #290
+#317 := [th-lemma]: #316
+#318 := [unit-resolution #317 #314 #313]: false
+#319 := [lemma #318]: #316
+#323 := [unit-resolution #319 #322]: #315
+#292 := (or #162 #290)
+#293 := [def-axiom]: #292
+#324 := [unit-resolution #293 #323]: #162
+#325 := [th-lemma #324 #307 #138 #302 #195]: false
+#326 := [lemma #325]: #112
+#286 := (or #113 #284)
+#287 := [def-axiom]: #286
+#330 := [unit-resolution #287 #326]: #284
+#331 := (not #284)
+#332 := (or #331 #301)
+#333 := [th-lemma]: #332
+#334 := [unit-resolution #333 #330]: #301
+#335 := [th-lemma #326 #334 #195 #329 #138]: false
+#336 := [lemma #335]: #161
+#327 := [unit-resolution #293 #336]: #290
+#328 := [unit-resolution #319 #327]: #300
+[th-lemma #326 #334 #195 #328 #187 #138]: false
+unsat
+GX51o3DUO/UBS3eNP2P9kA 285 0
+#2 := false
+#7 := 0::real
+decl uf_4 :: real
+#16 := uf_4
+#40 := -1::real
+#116 := (* -1::real uf_4)
+decl uf_3 :: real
+#11 := uf_3
+#117 := (+ uf_3 #116)
+#128 := (<= #117 0::real)
+#129 := (not #128)
+#220 := 2/3::real
+#221 := (* 2/3::real uf_3)
+#222 := (+ #221 #116)
+decl uf_2 :: real
+#5 := uf_2
+#67 := 1/3::real
+#68 := (* 1/3::real uf_2)
+#233 := (+ #68 #222)
+#243 := (<= #233 0::real)
+#268 := (not #243)
+#287 := [hypothesis]: #268
+#41 := (* -1::real uf_2)
+decl uf_1 :: real
+#4 := uf_1
+#42 := (+ uf_1 #41)
+#79 := (>= #42 0::real)
+#80 := (not #79)
+#297 := (or #80 #243)
+#158 := (+ uf_1 #116)
+#159 := (<= #158 0::real)
+#22 := (<= uf_1 uf_4)
+#160 := (iff #22 #159)
+#161 := [rewrite]: #160
+#155 := [asserted]: #22
+#162 := [mp #155 #161]: #159
+#200 := (* 1/3::real uf_3)
+#198 := -4/3::real
+#199 := (* -4/3::real uf_2)
+#201 := (+ #199 #200)
+#202 := (+ uf_1 #201)
+#203 := (>= #202 0::real)
+#258 := (not #203)
+#292 := [hypothesis]: #79
+#293 := (or #80 #258)
+#69 := -1/3::real
+#70 := (* -1/3::real uf_3)
+#186 := -2/3::real
+#187 := (* -2/3::real uf_2)
+#188 := (+ #187 #70)
+#189 := (+ uf_1 #188)
+#204 := (<= #189 0::real)
+#205 := (ite #79 #203 #204)
+#210 := (not #205)
+#51 := (* -1::real uf_1)
+#52 := (+ #51 uf_2)
+#86 := (ite #79 #42 #52)
+#94 := (* -1::real #86)
+#95 := (+ #70 #94)
+#96 := (+ #68 #95)
+#97 := (<= #96 0::real)
+#98 := (not #97)
+#211 := (iff #98 #210)
+#208 := (iff #97 #205)
+#182 := 4/3::real
+#183 := (* 4/3::real uf_2)
+#184 := (+ #183 #70)
+#185 := (+ #51 #184)
+#190 := (ite #79 #185 #189)
+#195 := (<= #190 0::real)
+#206 := (iff #195 #205)
+#207 := [rewrite]: #206
+#196 := (iff #97 #195)
+#193 := (= #96 #190)
+#172 := (+ #41 #70)
+#173 := (+ uf_1 #172)
+#170 := (+ uf_2 #70)
+#171 := (+ #51 #170)
+#174 := (ite #79 #171 #173)
+#179 := (+ #68 #174)
+#191 := (= #179 #190)
+#192 := [rewrite]: #191
+#180 := (= #96 #179)
+#177 := (= #95 #174)
+#164 := (ite #79 #52 #42)
+#167 := (+ #70 #164)
+#175 := (= #167 #174)
+#176 := [rewrite]: #175
+#168 := (= #95 #167)
+#156 := (= #94 #164)
+#165 := [rewrite]: #156
+#169 := [monotonicity #165]: #168
+#178 := [trans #169 #176]: #177
+#181 := [monotonicity #178]: #180
+#194 := [trans #181 #192]: #193
+#197 := [monotonicity #194]: #196
+#209 := [trans #197 #207]: #208
+#212 := [monotonicity #209]: #211
+#13 := 3::real
+#12 := (- uf_2 uf_3)
+#14 := (/ #12 3::real)
+#6 := (- uf_1 uf_2)
+#9 := (- #6)
+#8 := (< #6 0::real)
+#10 := (ite #8 #9 #6)
+#15 := (< #10 #14)
+#103 := (iff #15 #98)
+#71 := (+ #68 #70)
+#45 := (< #42 0::real)
+#57 := (ite #45 #52 #42)
+#76 := (< #57 #71)
+#101 := (iff #76 #98)
+#91 := (< #86 #71)
+#99 := (iff #91 #98)
+#100 := [rewrite]: #99
+#92 := (iff #76 #91)
+#89 := (= #57 #86)
+#83 := (ite #80 #52 #42)
+#87 := (= #83 #86)
+#88 := [rewrite]: #87
+#84 := (= #57 #83)
+#81 := (iff #45 #80)
+#82 := [rewrite]: #81
+#85 := [monotonicity #82]: #84
+#90 := [trans #85 #88]: #89
+#93 := [monotonicity #90]: #92
+#102 := [trans #93 #100]: #101
+#77 := (iff #15 #76)
+#74 := (= #14 #71)
+#60 := (* -1::real uf_3)
+#61 := (+ uf_2 #60)
+#64 := (/ #61 3::real)
+#72 := (= #64 #71)
+#73 := [rewrite]: #72
+#65 := (= #14 #64)
+#62 := (= #12 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#75 := [trans #66 #73]: #74
+#58 := (= #10 #57)
+#43 := (= #6 #42)
+#44 := [rewrite]: #43
+#55 := (= #9 #52)
+#48 := (- #42)
+#53 := (= #48 #52)
+#54 := [rewrite]: #53
+#49 := (= #9 #48)
+#50 := [monotonicity #44]: #49
+#56 := [trans #50 #54]: #55
+#46 := (iff #8 #45)
+#47 := [monotonicity #44]: #46
+#59 := [monotonicity #47 #56 #44]: #58
+#78 := [monotonicity #59 #75]: #77
+#104 := [trans #78 #102]: #103
+#39 := [asserted]: #15
+#105 := [mp #39 #104]: #98
+#213 := [mp #105 #212]: #210
+#259 := (or #205 #80 #258)
+#260 := [def-axiom]: #259
+#294 := [unit-resolution #260 #213]: #293
+#295 := [unit-resolution #294 #292]: #258
+#296 := [th-lemma #287 #292 #295 #162]: false
+#298 := [lemma #296]: #297
+#299 := [unit-resolution #298 #287]: #80
+#261 := (not #204)
+#281 := (or #79 #261)
+#262 := (or #205 #79 #261)
+#263 := [def-axiom]: #262
+#282 := [unit-resolution #263 #213]: #281
+#300 := [unit-resolution #282 #299]: #261
+#290 := (or #79 #204 #243)
+#276 := [hypothesis]: #261
+#288 := [hypothesis]: #80
+#289 := [th-lemma #288 #276 #162 #287]: false
+#291 := [lemma #289]: #290
+#301 := [unit-resolution #291 #300 #299 #287]: false
+#302 := [lemma #301]: #243
+#303 := (or #129 #268)
+#223 := (* -4/3::real uf_3)
+#224 := (+ #223 uf_4)
+#234 := (+ #68 #224)
+#244 := (<= #234 0::real)
+#245 := (ite #128 #243 #244)
+#250 := (not #245)
+#107 := (+ #60 uf_4)
+#135 := (ite #128 #107 #117)
+#143 := (* -1::real #135)
+#144 := (+ #70 #143)
+#145 := (+ #68 #144)
+#146 := (<= #145 0::real)
+#147 := (not #146)
+#251 := (iff #147 #250)
+#248 := (iff #146 #245)
+#235 := (ite #128 #233 #234)
+#240 := (<= #235 0::real)
+#246 := (iff #240 #245)
+#247 := [rewrite]: #246
+#241 := (iff #146 #240)
+#238 := (= #145 #235)
+#225 := (ite #128 #222 #224)
+#230 := (+ #68 #225)
+#236 := (= #230 #235)
+#237 := [rewrite]: #236
+#231 := (= #145 #230)
+#228 := (= #144 #225)
+#214 := (ite #128 #117 #107)
+#217 := (+ #70 #214)
+#226 := (= #217 #225)
+#227 := [rewrite]: #226
+#218 := (= #144 #217)
+#215 := (= #143 #214)
+#216 := [rewrite]: #215
+#219 := [monotonicity #216]: #218
+#229 := [trans #219 #227]: #228
+#232 := [monotonicity #229]: #231
+#239 := [trans #232 #237]: #238
+#242 := [monotonicity #239]: #241
+#249 := [trans #242 #247]: #248
+#252 := [monotonicity #249]: #251
+#17 := (- uf_4 uf_3)
+#19 := (- #17)
+#18 := (< #17 0::real)
+#20 := (ite #18 #19 #17)
+#21 := (< #20 #14)
+#152 := (iff #21 #147)
+#110 := (< #107 0::real)
+#122 := (ite #110 #117 #107)
+#125 := (< #122 #71)
+#150 := (iff #125 #147)
+#140 := (< #135 #71)
+#148 := (iff #140 #147)
+#149 := [rewrite]: #148
+#141 := (iff #125 #140)
+#138 := (= #122 #135)
+#132 := (ite #129 #117 #107)
+#136 := (= #132 #135)
+#137 := [rewrite]: #136
+#133 := (= #122 #132)
+#130 := (iff #110 #129)
+#131 := [rewrite]: #130
+#134 := [monotonicity #131]: #133
+#139 := [trans #134 #137]: #138
+#142 := [monotonicity #139]: #141
+#151 := [trans #142 #149]: #150
+#126 := (iff #21 #125)
+#123 := (= #20 #122)
+#108 := (= #17 #107)
+#109 := [rewrite]: #108
+#120 := (= #19 #117)
+#113 := (- #107)
+#118 := (= #113 #117)
+#119 := [rewrite]: #118
+#114 := (= #19 #113)
+#115 := [monotonicity #109]: #114
+#121 := [trans #115 #119]: #120
+#111 := (iff #18 #110)
+#112 := [monotonicity #109]: #111
+#124 := [monotonicity #112 #121 #109]: #123
+#127 := [monotonicity #124 #75]: #126
+#153 := [trans #127 #151]: #152
+#106 := [asserted]: #21
+#154 := [mp #106 #153]: #147
+#253 := [mp #154 #252]: #250
+#269 := (or #245 #129 #268)
+#270 := [def-axiom]: #269
+#304 := [unit-resolution #270 #253]: #303
+#305 := [unit-resolution #304 #302]: #129
+#271 := (not #244)
+#306 := (or #128 #271)
+#272 := (or #245 #128 #271)
+#273 := [def-axiom]: #272
+#307 := [unit-resolution #273 #253]: #306
+#308 := [unit-resolution #307 #305]: #271
+#285 := (or #128 #244)
+#274 := [hypothesis]: #271
+#275 := [hypothesis]: #129
+#278 := (or #204 #128 #244)
+#277 := [th-lemma #276 #275 #274 #162]: false
+#279 := [lemma #277]: #278
+#280 := [unit-resolution #279 #275 #274]: #204
+#283 := [unit-resolution #282 #280]: #79
+#284 := [th-lemma #275 #274 #283 #162]: false
+#286 := [lemma #284]: #285
+[unit-resolution #286 #308 #305]: false
+unsat
+cebG074uorSr8ODzgTmcKg 97 0
+#2 := false
+#18 := 0::real
+decl uf_1 :: (-> T2 T1 real)
+decl uf_5 :: T1
+#11 := uf_5
+decl uf_2 :: T2
+#4 := uf_2
+#20 := (uf_1 uf_2 uf_5)
+#42 := -1::real
+#53 := (* -1::real #20)
+decl uf_3 :: T2
+#7 := uf_3
+#19 := (uf_1 uf_3 uf_5)
+#54 := (+ #19 #53)
+#63 := (<= #54 0::real)
+#21 := (- #19 #20)
+#22 := (< 0::real #21)
+#23 := (not #22)
+#74 := (iff #23 #63)
+#57 := (< 0::real #54)
+#60 := (not #57)
+#72 := (iff #60 #63)
+#64 := (not #63)
+#67 := (not #64)
+#70 := (iff #67 #63)
+#71 := [rewrite]: #70
+#68 := (iff #60 #67)
+#65 := (iff #57 #64)
+#66 := [rewrite]: #65
+#69 := [monotonicity #66]: #68
+#73 := [trans #69 #71]: #72
+#61 := (iff #23 #60)
+#58 := (iff #22 #57)
+#55 := (= #21 #54)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#62 := [monotonicity #59]: #61
+#75 := [trans #62 #73]: #74
+#41 := [asserted]: #23
+#76 := [mp #41 #75]: #63
+#5 := (:var 0 T1)
+#8 := (uf_1 uf_3 #5)
+#141 := (pattern #8)
+#6 := (uf_1 uf_2 #5)
+#140 := (pattern #6)
+#45 := (* -1::real #8)
+#46 := (+ #6 #45)
+#44 := (>= #46 0::real)
+#43 := (not #44)
+#142 := (forall (vars (?x1 T1)) (:pat #140 #141) #43)
+#49 := (forall (vars (?x1 T1)) #43)
+#145 := (iff #49 #142)
+#143 := (iff #43 #43)
+#144 := [refl]: #143
+#146 := [quant-intro #144]: #145
+#80 := (~ #49 #49)
+#82 := (~ #43 #43)
+#83 := [refl]: #82
+#81 := [nnf-pos #83]: #80
+#9 := (< #6 #8)
+#10 := (forall (vars (?x1 T1)) #9)
+#50 := (iff #10 #49)
+#47 := (iff #9 #43)
+#48 := [rewrite]: #47
+#51 := [quant-intro #48]: #50
+#39 := [asserted]: #10
+#52 := [mp #39 #51]: #49
+#79 := [mp~ #52 #81]: #49
+#147 := [mp #79 #146]: #142
+#164 := (not #142)
+#165 := (or #164 #64)
+#148 := (* -1::real #19)
+#149 := (+ #20 #148)
+#150 := (>= #149 0::real)
+#151 := (not #150)
+#166 := (or #164 #151)
+#168 := (iff #166 #165)
+#170 := (iff #165 #165)
+#171 := [rewrite]: #170
+#162 := (iff #151 #64)
+#160 := (iff #150 #63)
+#152 := (+ #148 #20)
+#155 := (>= #152 0::real)
+#158 := (iff #155 #63)
+#159 := [rewrite]: #158
+#156 := (iff #150 #155)
+#153 := (= #149 #152)
+#154 := [rewrite]: #153
+#157 := [monotonicity #154]: #156
+#161 := [trans #157 #159]: #160
+#163 := [monotonicity #161]: #162
+#169 := [monotonicity #163]: #168
+#172 := [trans #169 #171]: #168
+#167 := [quant-inst]: #166
+#173 := [mp #167 #172]: #165
+[unit-resolution #173 #147 #76]: false
+unsat
+DKRtrJ2XceCkITuNwNViRw 57 0
+#2 := false
+#4 := 0::real
+decl uf_1 :: (-> T2 real)
+decl uf_2 :: (-> T1 T1 T2)
+decl uf_12 :: (-> T4 T1)
+decl uf_4 :: T4
+#11 := uf_4
+#39 := (uf_12 uf_4)
+decl uf_10 :: T4
+#27 := uf_10
+#38 := (uf_12 uf_10)
+#40 := (uf_2 #38 #39)
+#41 := (uf_1 #40)
+#264 := (>= #41 0::real)
+#266 := (not #264)
+#43 := (= #41 0::real)
+#44 := (not #43)
+#131 := [asserted]: #44
+#272 := (or #43 #266)
+#42 := (<= #41 0::real)
+#130 := [asserted]: #42
+#265 := (not #42)
+#270 := (or #43 #265 #266)
+#271 := [th-lemma]: #270
+#273 := [unit-resolution #271 #130]: #272
+#274 := [unit-resolution #273 #131]: #266
+#6 := (:var 0 T1)
+#5 := (:var 1 T1)
+#7 := (uf_2 #5 #6)
+#241 := (pattern #7)
+#8 := (uf_1 #7)
+#65 := (>= #8 0::real)
+#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65)
+#66 := (forall (vars (?x1 T1) (?x2 T1)) #65)
+#245 := (iff #66 #242)
+#243 := (iff #65 #65)
+#244 := [refl]: #243
+#246 := [quant-intro #244]: #245
+#149 := (~ #66 #66)
+#151 := (~ #65 #65)
+#152 := [refl]: #151
+#150 := [nnf-pos #152]: #149
+#9 := (<= 0::real #8)
+#10 := (forall (vars (?x1 T1) (?x2 T1)) #9)
+#67 := (iff #10 #66)
+#63 := (iff #9 #65)
+#64 := [rewrite]: #63
+#68 := [quant-intro #64]: #67
+#60 := [asserted]: #10
+#69 := [mp #60 #68]: #66
+#147 := [mp~ #69 #150]: #66
+#247 := [mp #147 #246]: #242
+#267 := (not #242)
+#268 := (or #267 #264)
+#269 := [quant-inst]: #268
+[unit-resolution #269 #247 #274]: false
+unsat
+97KJAJfUio+nGchEHWvgAw 91 0
+#2 := false
+#38 := 0::real
+decl uf_1 :: (-> T1 T2 real)
+decl uf_3 :: T2
+#5 := uf_3
+decl uf_4 :: T1
+#7 := uf_4
+#8 := (uf_1 uf_4 uf_3)
+#35 := -1::real
+#36 := (* -1::real #8)
+decl uf_2 :: T1
+#4 := uf_2
+#6 := (uf_1 uf_2 uf_3)
+#37 := (+ #6 #36)
+#130 := (>= #37 0::real)
+#155 := (not #130)
+#43 := (= #6 #8)
+#55 := (not #43)
+#15 := (= #8 #6)
+#16 := (not #15)
+#56 := (iff #16 #55)
+#53 := (iff #15 #43)
+#54 := [rewrite]: #53
+#57 := [monotonicity #54]: #56
+#34 := [asserted]: #16
+#60 := [mp #34 #57]: #55
+#158 := (or #43 #155)
+#39 := (<= #37 0::real)
+#9 := (<= #6 #8)
+#40 := (iff #9 #39)
+#41 := [rewrite]: #40
+#32 := [asserted]: #9
+#42 := [mp #32 #41]: #39
+#154 := (not #39)
+#156 := (or #43 #154 #155)
+#157 := [th-lemma]: #156
+#159 := [unit-resolution #157 #42]: #158
+#160 := [unit-resolution #159 #60]: #155
+#10 := (:var 0 T2)
+#12 := (uf_1 uf_2 #10)
+#123 := (pattern #12)
+#11 := (uf_1 uf_4 #10)
+#122 := (pattern #11)
+#44 := (* -1::real #12)
+#45 := (+ #11 #44)
+#46 := (<= #45 0::real)
+#124 := (forall (vars (?x1 T2)) (:pat #122 #123) #46)
+#49 := (forall (vars (?x1 T2)) #46)
+#127 := (iff #49 #124)
+#125 := (iff #46 #46)
+#126 := [refl]: #125
+#128 := [quant-intro #126]: #127
+#62 := (~ #49 #49)
+#64 := (~ #46 #46)
+#65 := [refl]: #64
+#63 := [nnf-pos #65]: #62
+#13 := (<= #11 #12)
+#14 := (forall (vars (?x1 T2)) #13)
+#50 := (iff #14 #49)
+#47 := (iff #13 #46)
+#48 := [rewrite]: #47
+#51 := [quant-intro #48]: #50
+#33 := [asserted]: #14
+#52 := [mp #33 #51]: #49
+#61 := [mp~ #52 #63]: #49
+#129 := [mp #61 #128]: #124
+#144 := (not #124)
+#145 := (or #144 #130)
+#131 := (* -1::real #6)
+#132 := (+ #8 #131)
+#133 := (<= #132 0::real)
+#146 := (or #144 #133)
+#148 := (iff #146 #145)
+#150 := (iff #145 #145)
+#151 := [rewrite]: #150
+#142 := (iff #133 #130)
+#134 := (+ #131 #8)
+#137 := (<= #134 0::real)
+#140 := (iff #137 #130)
+#141 := [rewrite]: #140
+#138 := (iff #133 #137)
+#135 := (= #132 #134)
+#136 := [rewrite]: #135
+#139 := [monotonicity #136]: #138
+#143 := [trans #139 #141]: #142
+#149 := [monotonicity #143]: #148
+#152 := [trans #149 #151]: #148
+#147 := [quant-inst]: #146
+#153 := [mp #147 #152]: #145
+[unit-resolution #153 #129 #160]: false
+unsat
+flJYbeWfe+t2l/zsRqdujA 149 0
+#2 := false
+#19 := 0::real
+decl uf_1 :: (-> T1 T2 real)
+decl uf_3 :: T2
+#5 := uf_3
+decl uf_4 :: T1
+#7 := uf_4
+#8 := (uf_1 uf_4 uf_3)
+#44 := -1::real
+#156 := (* -1::real #8)
+decl uf_2 :: T1
+#4 := uf_2
+#6 := (uf_1 uf_2 uf_3)
+#203 := (+ #6 #156)
+#205 := (>= #203 0::real)
+#9 := (= #6 #8)
+#40 := [asserted]: #9
+#208 := (not #9)
+#209 := (or #208 #205)
+#210 := [th-lemma]: #209
+#211 := [unit-resolution #210 #40]: #205
+decl uf_5 :: T1
+#12 := uf_5
+#22 := (uf_1 uf_5 uf_3)
+#160 := (* -1::real #22)
+#161 := (+ #6 #160)
+#207 := (>= #161 0::real)
+#222 := (not #207)
+#206 := (= #6 #22)
+#216 := (not #206)
+#62 := (= #8 #22)
+#70 := (not #62)
+#217 := (iff #70 #216)
+#214 := (iff #62 #206)
+#212 := (iff #206 #62)
+#213 := [monotonicity #40]: #212
+#215 := [symm #213]: #214
+#218 := [monotonicity #215]: #217
+#23 := (= #22 #8)
+#24 := (not #23)
+#71 := (iff #24 #70)
+#68 := (iff #23 #62)
+#69 := [rewrite]: #68
+#72 := [monotonicity #69]: #71
+#43 := [asserted]: #24
+#75 := [mp #43 #72]: #70
+#219 := [mp #75 #218]: #216
+#225 := (or #206 #222)
+#162 := (<= #161 0::real)
+#172 := (+ #8 #160)
+#173 := (>= #172 0::real)
+#178 := (not #173)
+#163 := (not #162)
+#181 := (or #163 #178)
+#184 := (not #181)
+#10 := (:var 0 T2)
+#15 := (uf_1 uf_4 #10)
+#149 := (pattern #15)
+#13 := (uf_1 uf_5 #10)
+#148 := (pattern #13)
+#11 := (uf_1 uf_2 #10)
+#147 := (pattern #11)
+#50 := (* -1::real #15)
+#51 := (+ #13 #50)
+#52 := (<= #51 0::real)
+#76 := (not #52)
+#45 := (* -1::real #13)
+#46 := (+ #11 #45)
+#47 := (<= #46 0::real)
+#78 := (not #47)
+#73 := (or #78 #76)
+#83 := (not #73)
+#150 := (forall (vars (?x1 T2)) (:pat #147 #148 #149) #83)
+#86 := (forall (vars (?x1 T2)) #83)
+#153 := (iff #86 #150)
+#151 := (iff #83 #83)
+#152 := [refl]: #151
+#154 := [quant-intro #152]: #153
+#55 := (and #47 #52)
+#58 := (forall (vars (?x1 T2)) #55)
+#87 := (iff #58 #86)
+#84 := (iff #55 #83)
+#85 := [rewrite]: #84
+#88 := [quant-intro #85]: #87
+#79 := (~ #58 #58)
+#81 := (~ #55 #55)
+#82 := [refl]: #81
+#80 := [nnf-pos #82]: #79
+#16 := (<= #13 #15)
+#14 := (<= #11 #13)
+#17 := (and #14 #16)
+#18 := (forall (vars (?x1 T2)) #17)
+#59 := (iff #18 #58)
+#56 := (iff #17 #55)
+#53 := (iff #16 #52)
+#54 := [rewrite]: #53
+#48 := (iff #14 #47)
+#49 := [rewrite]: #48
+#57 := [monotonicity #49 #54]: #56
+#60 := [quant-intro #57]: #59
+#41 := [asserted]: #18
+#61 := [mp #41 #60]: #58
+#77 := [mp~ #61 #80]: #58
+#89 := [mp #77 #88]: #86
+#155 := [mp #89 #154]: #150
+#187 := (not #150)
+#188 := (or #187 #184)
+#157 := (+ #22 #156)
+#158 := (<= #157 0::real)
+#159 := (not #158)
+#164 := (or #163 #159)
+#165 := (not #164)
+#189 := (or #187 #165)
+#191 := (iff #189 #188)
+#193 := (iff #188 #188)
+#194 := [rewrite]: #193
+#185 := (iff #165 #184)
+#182 := (iff #164 #181)
+#179 := (iff #159 #178)
+#176 := (iff #158 #173)
+#166 := (+ #156 #22)
+#169 := (<= #166 0::real)
+#174 := (iff #169 #173)
+#175 := [rewrite]: #174
+#170 := (iff #158 #169)
+#167 := (= #157 #166)
+#168 := [rewrite]: #167
+#171 := [monotonicity #168]: #170
+#177 := [trans #171 #175]: #176
+#180 := [monotonicity #177]: #179
+#183 := [monotonicity #180]: #182
+#186 := [monotonicity #183]: #185
+#192 := [monotonicity #186]: #191
+#195 := [trans #192 #194]: #191
+#190 := [quant-inst]: #189
+#196 := [mp #190 #195]: #188
+#220 := [unit-resolution #196 #155]: #184
+#197 := (or #181 #162)
+#198 := [def-axiom]: #197
+#221 := [unit-resolution #198 #220]: #162
+#223 := (or #206 #163 #222)
+#224 := [th-lemma]: #223
+#226 := [unit-resolution #224 #221]: #225
+#227 := [unit-resolution #226 #219]: #222
+#199 := (or #181 #173)
+#200 := [def-axiom]: #199
+#228 := [unit-resolution #200 #220]: #173
+[th-lemma #228 #227 #211]: false
+unsat
+rbrrQuQfaijtLkQizgEXnQ 222 0
+#2 := false
+#4 := 0::real
+decl uf_2 :: (-> T2 T1 real)
+decl uf_5 :: T1
+#15 := uf_5
+decl uf_3 :: T2
+#7 := uf_3
+#20 := (uf_2 uf_3 uf_5)
+decl uf_6 :: T2
+#17 := uf_6
+#18 := (uf_2 uf_6 uf_5)
+#59 := -1::real
+#73 := (* -1::real #18)
+#106 := (+ #73 #20)
+decl uf_1 :: real
+#5 := uf_1
+#78 := (* -1::real #20)
+#79 := (+ #18 #78)
+#144 := (+ uf_1 #79)
+#145 := (<= #144 0::real)
+#148 := (ite #145 uf_1 #106)
+#279 := (* -1::real #148)
+#280 := (+ uf_1 #279)
+#281 := (<= #280 0::real)
+#289 := (not #281)
+#72 := 1/2::real
+#151 := (* 1/2::real #148)
+#248 := (<= #151 0::real)
+#162 := (= #151 0::real)
+#24 := 2::real
+#27 := (- #20 #18)
+#28 := (<= uf_1 #27)
+#29 := (ite #28 uf_1 #27)
+#30 := (/ #29 2::real)
+#31 := (+ #18 #30)
+#32 := (= #31 #18)
+#33 := (not #32)
+#34 := (not #33)
+#165 := (iff #34 #162)
+#109 := (<= uf_1 #106)
+#112 := (ite #109 uf_1 #106)
+#118 := (* 1/2::real #112)
+#123 := (+ #18 #118)
+#129 := (= #18 #123)
+#163 := (iff #129 #162)
+#154 := (+ #18 #151)
+#157 := (= #18 #154)
+#160 := (iff #157 #162)
+#161 := [rewrite]: #160
+#158 := (iff #129 #157)
+#155 := (= #123 #154)
+#152 := (= #118 #151)
+#149 := (= #112 #148)
+#146 := (iff #109 #145)
+#147 := [rewrite]: #146
+#150 := [monotonicity #147]: #149
+#153 := [monotonicity #150]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#164 := [trans #159 #161]: #163
+#142 := (iff #34 #129)
+#134 := (not #129)
+#137 := (not #134)
+#140 := (iff #137 #129)
+#141 := [rewrite]: #140
+#138 := (iff #34 #137)
+#135 := (iff #33 #134)
+#132 := (iff #32 #129)
+#126 := (= #123 #18)
+#130 := (iff #126 #129)
+#131 := [rewrite]: #130
+#127 := (iff #32 #126)
+#124 := (= #31 #123)
+#121 := (= #30 #118)
+#115 := (/ #112 2::real)
+#119 := (= #115 #118)
+#120 := [rewrite]: #119
+#116 := (= #30 #115)
+#113 := (= #29 #112)
+#107 := (= #27 #106)
+#108 := [rewrite]: #107
+#110 := (iff #28 #109)
+#111 := [monotonicity #108]: #110
+#114 := [monotonicity #111 #108]: #113
+#117 := [monotonicity #114]: #116
+#122 := [trans #117 #120]: #121
+#125 := [monotonicity #122]: #124
+#128 := [monotonicity #125]: #127
+#133 := [trans #128 #131]: #132
+#136 := [monotonicity #133]: #135
+#139 := [monotonicity #136]: #138
+#143 := [trans #139 #141]: #142
+#166 := [trans #143 #164]: #165
+#105 := [asserted]: #34
+#167 := [mp #105 #166]: #162
+#283 := (not #162)
+#284 := (or #283 #248)
+#285 := [th-lemma]: #284
+#286 := [unit-resolution #285 #167]: #248
+#287 := [hypothesis]: #281
+#53 := (<= uf_1 0::real)
+#54 := (not #53)
+#6 := (< 0::real uf_1)
+#55 := (iff #6 #54)
+#56 := [rewrite]: #55
+#50 := [asserted]: #6
+#57 := [mp #50 #56]: #54
+#288 := [th-lemma #57 #287 #286]: false
+#290 := [lemma #288]: #289
+#241 := (= uf_1 #148)
+#242 := (= #106 #148)
+#299 := (not #242)
+#282 := (+ #106 #279)
+#291 := (<= #282 0::real)
+#296 := (not #291)
+decl uf_4 :: T2
+#10 := uf_4
+#16 := (uf_2 uf_4 uf_5)
+#260 := (+ #16 #78)
+#261 := (>= #260 0::real)
+#266 := (not #261)
+#8 := (:var 0 T1)
+#11 := (uf_2 uf_4 #8)
+#234 := (pattern #11)
+#9 := (uf_2 uf_3 #8)
+#233 := (pattern #9)
+#60 := (* -1::real #11)
+#61 := (+ #9 #60)
+#62 := (<= #61 0::real)
+#179 := (not #62)
+#235 := (forall (vars (?x1 T1)) (:pat #233 #234) #179)
+#178 := (forall (vars (?x1 T1)) #179)
+#238 := (iff #178 #235)
+#236 := (iff #179 #179)
+#237 := [refl]: #236
+#239 := [quant-intro #237]: #238
+#65 := (exists (vars (?x1 T1)) #62)
+#68 := (not #65)
+#175 := (~ #68 #178)
+#180 := (~ #179 #179)
+#177 := [refl]: #180
+#176 := [nnf-neg #177]: #175
+#12 := (<= #9 #11)
+#13 := (exists (vars (?x1 T1)) #12)
+#14 := (not #13)
+#69 := (iff #14 #68)
+#66 := (iff #13 #65)
+#63 := (iff #12 #62)
+#64 := [rewrite]: #63
+#67 := [quant-intro #64]: #66
+#70 := [monotonicity #67]: #69
+#51 := [asserted]: #14
+#71 := [mp #51 #70]: #68
+#173 := [mp~ #71 #176]: #178
+#240 := [mp #173 #239]: #235
+#269 := (not #235)
+#270 := (or #269 #266)
+#250 := (* -1::real #16)
+#251 := (+ #20 #250)
+#252 := (<= #251 0::real)
+#253 := (not #252)
+#271 := (or #269 #253)
+#273 := (iff #271 #270)
+#275 := (iff #270 #270)
+#276 := [rewrite]: #275
+#267 := (iff #253 #266)
+#264 := (iff #252 #261)
+#254 := (+ #250 #20)
+#257 := (<= #254 0::real)
+#262 := (iff #257 #261)
+#263 := [rewrite]: #262
+#258 := (iff #252 #257)
+#255 := (= #251 #254)
+#256 := [rewrite]: #255
+#259 := [monotonicity #256]: #258
+#265 := [trans #259 #263]: #264
+#268 := [monotonicity #265]: #267
+#274 := [monotonicity #268]: #273
+#277 := [trans #274 #276]: #273
+#272 := [quant-inst]: #271
+#278 := [mp #272 #277]: #270
+#293 := [unit-resolution #278 #240]: #266
+#90 := (* 1/2::real #20)
+#102 := (+ #73 #90)
+#89 := (* 1/2::real #16)
+#103 := (+ #89 #102)
+#100 := (>= #103 0::real)
+#23 := (+ #16 #20)
+#25 := (/ #23 2::real)
+#26 := (<= #18 #25)
+#98 := (iff #26 #100)
+#91 := (+ #89 #90)
+#94 := (<= #18 #91)
+#97 := (iff #94 #100)
+#99 := [rewrite]: #97
+#95 := (iff #26 #94)
+#92 := (= #25 #91)
+#93 := [rewrite]: #92
+#96 := [monotonicity #93]: #95
+#101 := [trans #96 #99]: #98
+#58 := [asserted]: #26
+#104 := [mp #58 #101]: #100
+#294 := [hypothesis]: #291
+#295 := [th-lemma #294 #104 #293 #286]: false
+#297 := [lemma #295]: #296
+#298 := [hypothesis]: #242
+#300 := (or #299 #291)
+#301 := [th-lemma]: #300
+#302 := [unit-resolution #301 #298 #297]: false
+#303 := [lemma #302]: #299
+#246 := (or #145 #242)
+#247 := [def-axiom]: #246
+#304 := [unit-resolution #247 #303]: #145
+#243 := (not #145)
+#244 := (or #243 #241)
+#245 := [def-axiom]: #244
+#305 := [unit-resolution #245 #304]: #241
+#306 := (not #241)
+#307 := (or #306 #281)
+#308 := [th-lemma]: #307
+[unit-resolution #308 #305 #290]: false
+unsat
+hwh3oeLAWt56hnKIa8Wuow 248 0
+#2 := false
+#4 := 0::real
+decl uf_2 :: (-> T2 T1 real)
+decl uf_5 :: T1
+#15 := uf_5
+decl uf_6 :: T2
+#17 := uf_6
+#18 := (uf_2 uf_6 uf_5)
+decl uf_4 :: T2
+#10 := uf_4
+#16 := (uf_2 uf_4 uf_5)
+#66 := -1::real
+#137 := (* -1::real #16)
+#138 := (+ #137 #18)
+decl uf_1 :: real
+#5 := uf_1
+#80 := (* -1::real #18)
+#81 := (+ #16 #80)
+#201 := (+ uf_1 #81)
+#202 := (<= #201 0::real)
+#205 := (ite #202 uf_1 #138)
+#352 := (* -1::real #205)
+#353 := (+ uf_1 #352)
+#354 := (<= #353 0::real)
+#362 := (not #354)
+#79 := 1/2::real
+#244 := (* 1/2::real #205)
+#322 := (<= #244 0::real)
+#245 := (= #244 0::real)
+#158 := -1/2::real
+#208 := (* -1/2::real #205)
+#211 := (+ #18 #208)
+decl uf_3 :: T2
+#7 := uf_3
+#20 := (uf_2 uf_3 uf_5)
+#117 := (+ #80 #20)
+#85 := (* -1::real #20)
+#86 := (+ #18 #85)
+#188 := (+ uf_1 #86)
+#189 := (<= #188 0::real)
+#192 := (ite #189 uf_1 #117)
+#195 := (* 1/2::real #192)
+#198 := (+ #18 #195)
+#97 := (* 1/2::real #20)
+#109 := (+ #80 #97)
+#96 := (* 1/2::real #16)
+#110 := (+ #96 #109)
+#107 := (>= #110 0::real)
+#214 := (ite #107 #198 #211)
+#217 := (= #18 #214)
+#248 := (iff #217 #245)
+#241 := (= #18 #211)
+#246 := (iff #241 #245)
+#247 := [rewrite]: #246
+#242 := (iff #217 #241)
+#239 := (= #214 #211)
+#234 := (ite false #198 #211)
+#237 := (= #234 #211)
+#238 := [rewrite]: #237
+#235 := (= #214 #234)
+#232 := (iff #107 false)
+#104 := (not #107)
+#24 := 2::real
+#23 := (+ #16 #20)
+#25 := (/ #23 2::real)
+#26 := (< #25 #18)
+#108 := (iff #26 #104)
+#98 := (+ #96 #97)
+#101 := (< #98 #18)
+#106 := (iff #101 #104)
+#105 := [rewrite]: #106
+#102 := (iff #26 #101)
+#99 := (= #25 #98)
+#100 := [rewrite]: #99
+#103 := [monotonicity #100]: #102
+#111 := [trans #103 #105]: #108
+#65 := [asserted]: #26
+#112 := [mp #65 #111]: #104
+#233 := [iff-false #112]: #232
+#236 := [monotonicity #233]: #235
+#240 := [trans #236 #238]: #239
+#243 := [monotonicity #240]: #242
+#249 := [trans #243 #247]: #248
+#33 := (- #18 #16)
+#34 := (<= uf_1 #33)
+#35 := (ite #34 uf_1 #33)
+#36 := (/ #35 2::real)
+#37 := (- #18 #36)
+#28 := (- #20 #18)
+#29 := (<= uf_1 #28)
+#30 := (ite #29 uf_1 #28)
+#31 := (/ #30 2::real)
+#32 := (+ #18 #31)
+#27 := (<= #18 #25)
+#38 := (ite #27 #32 #37)
+#39 := (= #38 #18)
+#40 := (not #39)
+#41 := (not #40)
+#220 := (iff #41 #217)
+#141 := (<= uf_1 #138)
+#144 := (ite #141 uf_1 #138)
+#159 := (* -1/2::real #144)
+#160 := (+ #18 #159)
+#120 := (<= uf_1 #117)
+#123 := (ite #120 uf_1 #117)
+#129 := (* 1/2::real #123)
+#134 := (+ #18 #129)
+#114 := (<= #18 #98)
+#165 := (ite #114 #134 #160)
+#171 := (= #18 #165)
+#218 := (iff #171 #217)
+#215 := (= #165 #214)
+#212 := (= #160 #211)
+#209 := (= #159 #208)
+#206 := (= #144 #205)
+#203 := (iff #141 #202)
+#204 := [rewrite]: #203
+#207 := [monotonicity #204]: #206
+#210 := [monotonicity #207]: #209
+#213 := [monotonicity #210]: #212
+#199 := (= #134 #198)
+#196 := (= #129 #195)
+#193 := (= #123 #192)
+#190 := (iff #120 #189)
+#191 := [rewrite]: #190
+#194 := [monotonicity #191]: #193
+#197 := [monotonicity #194]: #196
+#200 := [monotonicity #197]: #199
+#187 := (iff #114 #107)
+#186 := [rewrite]: #187
+#216 := [monotonicity #186 #200 #213]: #215
+#219 := [monotonicity #216]: #218
+#184 := (iff #41 #171)
+#176 := (not #171)
+#179 := (not #176)
+#182 := (iff #179 #171)
+#183 := [rewrite]: #182
+#180 := (iff #41 #179)
+#177 := (iff #40 #176)
+#174 := (iff #39 #171)
+#168 := (= #165 #18)
+#172 := (iff #168 #171)
+#173 := [rewrite]: #172
+#169 := (iff #39 #168)
+#166 := (= #38 #165)
+#163 := (= #37 #160)
+#150 := (* 1/2::real #144)
+#155 := (- #18 #150)
+#161 := (= #155 #160)
+#162 := [rewrite]: #161
+#156 := (= #37 #155)
+#153 := (= #36 #150)
+#147 := (/ #144 2::real)
+#151 := (= #147 #150)
+#152 := [rewrite]: #151
+#148 := (= #36 #147)
+#145 := (= #35 #144)
+#139 := (= #33 #138)
+#140 := [rewrite]: #139
+#142 := (iff #34 #141)
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143 #140]: #145
+#149 := [monotonicity #146]: #148
+#154 := [trans #149 #152]: #153
+#157 := [monotonicity #154]: #156
+#164 := [trans #157 #162]: #163
+#135 := (= #32 #134)
+#132 := (= #31 #129)
+#126 := (/ #123 2::real)
+#130 := (= #126 #129)
+#131 := [rewrite]: #130
+#127 := (= #31 #126)
+#124 := (= #30 #123)
+#118 := (= #28 #117)
+#119 := [rewrite]: #118
+#121 := (iff #29 #120)
+#122 := [monotonicity #119]: #121
+#125 := [monotonicity #122 #119]: #124
+#128 := [monotonicity #125]: #127
+#133 := [trans #128 #131]: #132
+#136 := [monotonicity #133]: #135
+#115 := (iff #27 #114)
+#116 := [monotonicity #100]: #115
+#167 := [monotonicity #116 #136 #164]: #166
+#170 := [monotonicity #167]: #169
+#175 := [trans #170 #173]: #174
+#178 := [monotonicity #175]: #177
+#181 := [monotonicity #178]: #180
+#185 := [trans #181 #183]: #184
+#221 := [trans #185 #219]: #220
+#113 := [asserted]: #41
+#222 := [mp #113 #221]: #217
+#250 := [mp #222 #249]: #245
+#356 := (not #245)
+#357 := (or #356 #322)
+#358 := [th-lemma]: #357
+#359 := [unit-resolution #358 #250]: #322
+#360 := [hypothesis]: #354
+#60 := (<= uf_1 0::real)
+#61 := (not #60)
+#6 := (< 0::real uf_1)
+#62 := (iff #6 #61)
+#63 := [rewrite]: #62
+#57 := [asserted]: #6
+#64 := [mp #57 #63]: #61
+#361 := [th-lemma #64 #360 #359]: false
+#363 := [lemma #361]: #362
+#315 := (= uf_1 #205)
+#316 := (= #138 #205)
+#371 := (not #316)
+#355 := (+ #138 #352)
+#364 := (<= #355 0::real)
+#368 := (not #364)
+#87 := (<= #86 0::real)
+#82 := (<= #81 0::real)
+#90 := (and #82 #87)
+#21 := (<= #18 #20)
+#19 := (<= #16 #18)
+#22 := (and #19 #21)
+#91 := (iff #22 #90)
+#88 := (iff #21 #87)
+#89 := [rewrite]: #88
+#83 := (iff #19 #82)
+#84 := [rewrite]: #83
+#92 := [monotonicity #84 #89]: #91
+#59 := [asserted]: #22
+#93 := [mp #59 #92]: #90
+#95 := [and-elim #93]: #87
+#366 := [hypothesis]: #364
+#367 := [th-lemma #366 #95 #112 #359]: false
+#369 := [lemma #367]: #368
+#370 := [hypothesis]: #316
+#372 := (or #371 #364)
+#373 := [th-lemma]: #372
+#374 := [unit-resolution #373 #370 #369]: false
+#375 := [lemma #374]: #371
+#320 := (or #202 #316)
+#321 := [def-axiom]: #320
+#376 := [unit-resolution #321 #375]: #202
+#317 := (not #202)
+#318 := (or #317 #315)
+#319 := [def-axiom]: #318
+#377 := [unit-resolution #319 #376]: #315
+#378 := (not #315)
+#379 := (or #378 #354)
+#380 := [th-lemma]: #379
+[unit-resolution #380 #377 #363]: false
+unsat
+WdMJH3tkMv/rps8y9Ukq5Q 86 0
+#2 := false
+#37 := 0::real
+decl uf_2 :: (-> T2 T1 real)
+decl uf_4 :: T1
+#12 := uf_4
+decl uf_3 :: T2
+#5 := uf_3
+#13 := (uf_2 uf_3 uf_4)
+#34 := -1::real
+#140 := (* -1::real #13)
+decl uf_1 :: real
+#4 := uf_1
+#141 := (+ uf_1 #140)
+#143 := (>= #141 0::real)
+#6 := (:var 0 T1)
+#7 := (uf_2 uf_3 #6)
+#127 := (pattern #7)
+#35 := (* -1::real #7)
+#36 := (+ uf_1 #35)
+#47 := (>= #36 0::real)
+#134 := (forall (vars (?x2 T1)) (:pat #127) #47)
+#49 := (forall (vars (?x2 T1)) #47)
+#137 := (iff #49 #134)
+#135 := (iff #47 #47)
+#136 := [refl]: #135
+#138 := [quant-intro #136]: #137
+#67 := (~ #49 #49)
+#58 := (~ #47 #47)
+#66 := [refl]: #58
+#68 := [nnf-pos #66]: #67
+#10 := (<= #7 uf_1)
+#11 := (forall (vars (?x2 T1)) #10)
+#50 := (iff #11 #49)
+#46 := (iff #10 #47)
+#48 := [rewrite]: #46
+#51 := [quant-intro #48]: #50
+#32 := [asserted]: #11
+#52 := [mp #32 #51]: #49
+#69 := [mp~ #52 #68]: #49
+#139 := [mp #69 #138]: #134
+#149 := (not #134)
+#150 := (or #149 #143)
+#151 := [quant-inst]: #150
+#144 := [unit-resolution #151 #139]: #143
+#142 := (<= #141 0::real)
+#38 := (<= #36 0::real)
+#128 := (forall (vars (?x1 T1)) (:pat #127) #38)
+#41 := (forall (vars (?x1 T1)) #38)
+#131 := (iff #41 #128)
+#129 := (iff #38 #38)
+#130 := [refl]: #129
+#132 := [quant-intro #130]: #131
+#62 := (~ #41 #41)
+#64 := (~ #38 #38)
+#65 := [refl]: #64
+#63 := [nnf-pos #65]: #62
+#8 := (<= uf_1 #7)
+#9 := (forall (vars (?x1 T1)) #8)
+#42 := (iff #9 #41)
+#39 := (iff #8 #38)
+#40 := [rewrite]: #39
+#43 := [quant-intro #40]: #42
+#31 := [asserted]: #9
+#44 := [mp #31 #43]: #41
+#61 := [mp~ #44 #63]: #41
+#133 := [mp #61 #132]: #128
+#145 := (not #128)
+#146 := (or #145 #142)
+#147 := [quant-inst]: #146
+#148 := [unit-resolution #147 #133]: #142
+#45 := (= uf_1 #13)
+#55 := (not #45)
+#14 := (= #13 uf_1)
+#15 := (not #14)
+#56 := (iff #15 #55)
+#53 := (iff #14 #45)
+#54 := [rewrite]: #53
+#57 := [monotonicity #54]: #56
+#33 := [asserted]: #15
+#60 := [mp #33 #57]: #55
+#153 := (not #143)
+#152 := (not #142)
+#154 := (or #45 #152 #153)
+#155 := [th-lemma]: #154
+[unit-resolution #155 #60 #148 #144]: false
+unsat
+V+IAyBZU/6QjYs6JkXx8LQ 57 0
+#2 := false
+#4 := 0::real
+decl uf_1 :: (-> T2 real)
+decl uf_2 :: (-> T1 T1 T2)
+decl uf_12 :: (-> T4 T1)
+decl uf_4 :: T4
+#11 := uf_4
+#39 := (uf_12 uf_4)
+decl uf_10 :: T4
+#27 := uf_10
+#38 := (uf_12 uf_10)
+#40 := (uf_2 #38 #39)
+#41 := (uf_1 #40)
+#264 := (>= #41 0::real)
+#266 := (not #264)
+#43 := (= #41 0::real)
+#44 := (not #43)
+#131 := [asserted]: #44
+#272 := (or #43 #266)
+#42 := (<= #41 0::real)
+#130 := [asserted]: #42
+#265 := (not #42)
+#270 := (or #43 #265 #266)
+#271 := [th-lemma]: #270
+#273 := [unit-resolution #271 #130]: #272
+#274 := [unit-resolution #273 #131]: #266
+#6 := (:var 0 T1)
+#5 := (:var 1 T1)
+#7 := (uf_2 #5 #6)
+#241 := (pattern #7)
+#8 := (uf_1 #7)
+#65 := (>= #8 0::real)
+#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65)
+#66 := (forall (vars (?x1 T1) (?x2 T1)) #65)
+#245 := (iff #66 #242)
+#243 := (iff #65 #65)
+#244 := [refl]: #243
+#246 := [quant-intro #244]: #245
+#149 := (~ #66 #66)
+#151 := (~ #65 #65)
+#152 := [refl]: #151
+#150 := [nnf-pos #152]: #149
+#9 := (<= 0::real #8)
+#10 := (forall (vars (?x1 T1) (?x2 T1)) #9)
+#67 := (iff #10 #66)
+#63 := (iff #9 #65)
+#64 := [rewrite]: #63
+#68 := [quant-intro #64]: #67
+#60 := [asserted]: #10
+#69 := [mp #60 #68]: #66
+#147 := [mp~ #69 #150]: #66
+#247 := [mp #147 #246]: #242
+#267 := (not #242)
+#268 := (or #267 #264)
+#269 := [quant-inst]: #268
+[unit-resolution #269 #247 #274]: false
+unsat
+vqiyJ/qjGXZ3iOg6xftiug 15 0
+uf_1 -> val!0
+uf_2 -> val!1
+uf_3 -> val!2
+uf_5 -> val!15
+uf_6 -> val!26
+uf_4 -> {
+  val!0 -> val!12
+  val!1 -> val!13
+  else -> val!13
+}
+uf_7 -> {
+  val!6 -> val!31
+  else -> val!31
+}
+sat
+mrZPJZyTokErrN6SYupisw 9 0
+uf_4 -> val!1
+uf_2 -> val!3
+uf_3 -> val!4
+uf_1 -> {
+  val!5 -> val!6
+  val!4 -> val!7
+  else -> val!7
+}
+sat
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,3473 @@
+
+header {* Kurzweil-Henstock gauge integration in many dimensions. *}
+(*  Author:                     John Harrison
+    Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+
+theory Integration
+  imports Derivative SMT
+begin
+
+declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration.cert"]]
+declare [[smt_record=true]]
+declare [[z3_proofs=true]]
+
+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 conjunctD5: assumes "a \<and> b \<and> c \<and> d \<and> e" shows a b c d e using assms by auto
+
+declare smult_conv_scaleR[simp]
+
+subsection {* Some useful lemmas about intervals. *}
+
+lemma empty_as_interval: "{} = {1..0::real^'n}"
+  apply(rule set_ext,rule) defer unfolding vector_le_def mem_interval
+  using UNIV_witness[where 'a='n] apply(erule_tac exE,rule_tac x=x in allE) by auto
+
+lemma interior_subset_union_intervals: 
+  assumes "i = {a..b::real^'n}" "j = {c..d}" "interior j \<noteq> {}" "i \<subseteq> j \<union> s" "interior(i) \<inter> interior(j) = {}"
+  shows "interior i \<subseteq> interior s" proof-
+  have "{a<..<b} \<inter> {c..d} = {}" using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
+    unfolding assms(1,2) interior_closed_interval by auto
+  moreover have "{a<..<b} \<subseteq> {c..d} \<union> s" apply(rule order_trans,rule interval_open_subset_closed)
+    using assms(4) unfolding assms(1,2) by auto
+  ultimately show ?thesis apply-apply(rule interior_maximal) defer apply(rule open_interior)
+    unfolding assms(1,2) interior_closed_interval by auto qed
+
+lemma inter_interior_unions_intervals: fixes f::"(real^'n) set set"
+  assumes "finite f" "open s" "\<forall>t\<in>f. \<exists>a b. t = {a..b}" "\<forall>t\<in>f. s \<inter> (interior t) = {}"
+  shows "s \<inter> interior(\<Union>f) = {}" proof(rule ccontr,unfold ex_in_conv[THEN sym]) case goal1
+  have lem1:"\<And>x e s U. ball x e \<subseteq> s \<inter> interior U \<longleftrightarrow> ball x e \<subseteq> s \<inter> U" apply rule  defer apply(rule_tac Int_greatest)
+    unfolding open_subset_interior[OF open_ball]  using interior_subset by auto
+  have lem2:"\<And>x s P. \<exists>x\<in>s. P x \<Longrightarrow> \<exists>x\<in>insert x s. P x" by auto
+  have "\<And>f. finite f \<Longrightarrow> (\<forall>t\<in>f. \<exists>a b. t = {a..b}) \<Longrightarrow> (\<exists>x. x \<in> s \<inter> interior (\<Union>f)) \<Longrightarrow> (\<exists>t\<in>f. \<exists>x. \<exists>e>0. ball x e \<subseteq> s \<inter> t)" proof- case goal1
+  thus ?case proof(induct rule:finite_induct) 
+    case empty from this(2) guess x .. hence False unfolding Union_empty interior_empty by auto thus ?case by auto next
+    case (insert i f) guess x using insert(5) .. note x = this
+    then guess e unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior],rule_format] .. note e=this
+    guess a using insert(4)[rule_format,OF insertI1] .. then guess b .. note ab = this
+    show ?case proof(cases "x\<in>i") case False hence "x \<in> UNIV - {a..b}" unfolding ab by auto
+      then guess d unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_interval],rule_format] ..
+      hence "0 < d" "ball x (min d e) \<subseteq> UNIV - i" using e unfolding ab by auto
+      hence "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)" using e unfolding lem1 by auto hence "x \<in> s \<inter> interior (\<Union>f)" using `d>0` e by auto
+      hence "\<exists>t\<in>f. \<exists>x e. 0 < e \<and> ball x e \<subseteq> s \<inter> t" apply-apply(rule insert(3)) using insert(4) by auto thus ?thesis by auto next
+    case True show ?thesis proof(cases "x\<in>{a<..<b}")
+      case True then guess d unfolding open_contains_ball_eq[OF open_interval,rule_format] ..
+      thus ?thesis apply(rule_tac x=i in bexI,rule_tac x=x in exI,rule_tac x="min d e" in exI)
+	unfolding ab using interval_open_subset_closed[of a b] and e by fastsimp+ next
+    case False then obtain k where "x$k \<le> a$k \<or> x$k \<ge> b$k" unfolding mem_interval by(auto simp add:not_less) 
+    hence "x$k = a$k \<or> x$k = b$k" using True unfolding ab and mem_interval apply(erule_tac x=k in allE) by auto
+    hence "\<exists>x. ball x (e/2) \<subseteq> s \<inter> (\<Union>f)" proof(erule_tac disjE)
+      let ?z = "x - (e/2) *\<^sub>R basis k" assume as:"x$k = a$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
+	fix y assume "y \<in> ball ?z (e / 2) \<inter> i" hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
+	hence "\<bar>(?z - y) $ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding vector_dist_norm by auto
+	hence "y$k < a$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps)
+	hence "y \<notin> i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed
+      moreover have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof
+	fix y assume as:"y\<in> ball ?z (e/2)" have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R basis k)"
+	   apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R basis k"])
+	  unfolding norm_scaleR norm_basis by auto
+	also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps)
+	finally show "y\<in>ball x e" unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) qed
+      ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto
+    next let ?z = "x + (e/2) *\<^sub>R basis k" assume as:"x$k = b$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
+	fix y assume "y \<in> ball ?z (e / 2) \<inter> i" hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
+	hence "\<bar>(?z - y) $ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding vector_dist_norm by auto
+	hence "y$k > b$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps)
+	hence "y \<notin> i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed
+      moreover have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof
+	fix y assume as:"y\<in> ball ?z (e/2)" have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R basis k)"
+	   apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R basis k"])
+	  unfolding norm_scaleR norm_basis by auto
+	also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps)
+	finally show "y\<in>ball x e" unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) qed
+      ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto qed 
+    then guess x .. hence "x \<in> s \<inter> interior (\<Union>f)" unfolding lem1[where U="\<Union>f",THEN sym] using centre_in_ball e[THEN conjunct1] by auto
+    thus ?thesis apply-apply(rule lem2,rule insert(3)) using insert(4) by auto qed qed qed qed note * = this
+  guess t using *[OF assms(1,3) goal1]  .. from this(2) guess x .. then guess e ..
+  hence "x \<in> s" "x\<in>interior t" defer using open_subset_interior[OF open_ball, of x e t] by auto
+  thus False using `t\<in>f` assms(4) by auto qed
+subsection {* Bounds on intervals where they exist. *}
+
+definition "interval_upperbound (s::(real^'n) set) = (\<chi> i. Sup {a. \<exists>x\<in>s. x$i = a})"
+
+definition "interval_lowerbound (s::(real^'n) set) = (\<chi> i. Inf {a. \<exists>x\<in>s. x$i = a})"
+
+lemma interval_upperbound[simp]: assumes "\<forall>i. a$i \<le> b$i" shows "interval_upperbound {a..b} = b"
+  using assms unfolding interval_upperbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE)
+  apply(rule Sup_unique) unfolding setle_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer
+  apply(rule,rule) apply(rule_tac x="b$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=b in bexI)
+  unfolding mem_interval using assms by auto
+
+lemma interval_lowerbound[simp]: assumes "\<forall>i. a$i \<le> b$i" shows "interval_lowerbound {a..b} = a"
+  using assms unfolding interval_lowerbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE)
+  apply(rule Inf_unique) unfolding setge_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer
+  apply(rule,rule) apply(rule_tac x="a$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=a in bexI)
+  unfolding mem_interval using assms by auto
+
+lemmas interval_bounds = interval_upperbound interval_lowerbound
+
+lemma interval_bounds'[simp]: assumes "{a..b}\<noteq>{}" shows "interval_upperbound {a..b} = b" "interval_lowerbound {a..b} = a"
+  using assms unfolding interval_ne_empty by auto
+
+lemma interval_upperbound_1[simp]: "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> interval_upperbound {a..b} = (b::real^1)"
+  apply(rule interval_upperbound) by auto
+
+lemma interval_lowerbound_1[simp]: "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> interval_lowerbound {a..b} = (a::real^1)"
+  apply(rule interval_lowerbound) by auto
+
+lemmas interval_bound_1 = interval_upperbound_1 interval_lowerbound_1
+
+subsection {* Content (length, area, volume...) of an interval. *}
+
+definition "content (s::(real^'n) set) =
+       (if s = {} then 0 else (\<Prod>i\<in>UNIV. (interval_upperbound s)$i - (interval_lowerbound s)$i))"
+
+lemma interval_not_empty:"\<forall>i. a$i \<le> b$i \<Longrightarrow> {a..b::real^'n} \<noteq> {}"
+  unfolding interval_eq_empty unfolding not_ex not_less by assumption
+
+lemma content_closed_interval: assumes "\<forall>i. a$i \<le> b$i"
+  shows "content {a..b} = (\<Prod>i\<in>UNIV. b$i - a$i)"
+  using interval_not_empty[OF assms] unfolding content_def interval_upperbound[OF assms] interval_lowerbound[OF assms] by auto
+
+lemma content_closed_interval': assumes "{a..b}\<noteq>{}" shows "content {a..b} = (\<Prod>i\<in>UNIV. b$i - a$i)"
+  apply(rule content_closed_interval) using assms unfolding interval_ne_empty .
+
+lemma content_1:"dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> content {a..b} = dest_vec1 b - dest_vec1 a"
+  using content_closed_interval[of a b] by auto
+
+lemma content_1':"a \<le> b \<Longrightarrow> content {vec1 a..vec1 b} = b - a" using content_1[of "vec a" "vec b"] by auto
+
+lemma content_unit[intro]: "content{0..1::real^'n} = 1" proof-
+  have *:"\<forall>i. 0$i \<le> (1::real^'n::finite)$i" by auto
+  have "0 \<in> {0..1::real^'n::finite}" unfolding mem_interval by auto
+  thus ?thesis unfolding content_def interval_bounds[OF *] using setprod_1 by auto qed
+
+lemma content_pos_le[intro]: "0 \<le> content {a..b}" proof(cases "{a..b}={}")
+  case False hence *:"\<forall>i. a $ i \<le> b $ i" unfolding interval_ne_empty by assumption
+  have "(\<Prod>i\<in>UNIV. interval_upperbound {a..b} $ i - interval_lowerbound {a..b} $ i) \<ge> 0"
+    apply(rule setprod_nonneg) unfolding interval_bounds[OF *] using * apply(erule_tac x=x in allE) by auto
+  thus ?thesis unfolding content_def by(auto simp del:interval_bounds') qed(unfold content_def, auto)
+
+lemma content_pos_lt: assumes "\<forall>i. a$i < b$i" shows "0 < content {a..b}"
+proof- have help_lemma1: "\<forall>i. a$i < b$i \<Longrightarrow> \<forall>i. a$i \<le> ((b$i)::real)" apply(rule,erule_tac x=i in allE) by auto
+  show ?thesis unfolding content_closed_interval[OF help_lemma1[OF assms]] apply(rule setprod_pos)
+    using assms apply(erule_tac x=x in allE) by auto qed
+
+lemma content_pos_lt_1: "dest_vec1 a < dest_vec1 b \<Longrightarrow> 0 < content({a..b})"
+  apply(rule content_pos_lt) by auto
+
+lemma content_eq_0: "content({a..b::real^'n}) = 0 \<longleftrightarrow> (\<exists>i. b$i \<le> a$i)" proof(cases "{a..b} = {}")
+  case True thus ?thesis unfolding content_def if_P[OF True] unfolding interval_eq_empty apply-
+    apply(rule,erule exE) apply(rule_tac x=i in exI) by auto next
+  guess a using UNIV_witness[where 'a='n] .. case False note as=this[unfolded interval_eq_empty not_ex not_less]
+  show ?thesis unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_UNIV]
+    apply(rule) apply(erule_tac[!] exE bexE) unfolding interval_bounds[OF as] apply(rule_tac x=x in exI) defer
+    apply(rule_tac x=i in bexI) using as apply(erule_tac x=i in allE) by auto qed
+
+lemma cond_cases:"(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)" by auto
+
+lemma content_closed_interval_cases:
+  "content {a..b} = (if \<forall>i. a$i \<le> b$i then setprod (\<lambda>i. b$i - a$i) UNIV else 0)" apply(rule cond_cases) 
+  apply(rule content_closed_interval) unfolding content_eq_0 not_all not_le defer apply(erule exE,rule_tac x=x in exI) by auto
+
+lemma content_eq_0_interior: "content {a..b} = 0 \<longleftrightarrow> interior({a..b}) = {}"
+  unfolding content_eq_0 interior_closed_interval interval_eq_empty by auto
+
+lemma content_eq_0_1: "content {a..b::real^1} = 0 \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
+  unfolding content_eq_0 by auto
+
+lemma content_pos_lt_eq: "0 < content {a..b} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
+  apply(rule) defer apply(rule content_pos_lt,assumption) proof- assume "0 < content {a..b}"
+  hence "content {a..b} \<noteq> 0" by auto thus "\<forall>i. a$i < b$i" unfolding content_eq_0 not_ex not_le by auto qed
+
+lemma content_empty[simp]: "content {} = 0" unfolding content_def by auto
+
+lemma content_subset: assumes "{a..b} \<subseteq> {c..d}" shows "content {a..b::real^'n} \<le> content {c..d}" proof(cases "{a..b}={}")
+  case True thus ?thesis using content_pos_le[of c d] by auto next
+  case False hence ab_ne:"\<forall>i. a $ i \<le> b $ i" unfolding interval_ne_empty by auto
+  hence ab_ab:"a\<in>{a..b}" "b\<in>{a..b}" unfolding mem_interval by auto
+  have "{c..d} \<noteq> {}" using assms False by auto
+  hence cd_ne:"\<forall>i. c $ i \<le> d $ i" using assms unfolding interval_ne_empty by auto
+  show ?thesis unfolding content_def unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne]
+    unfolding if_not_P[OF False] if_not_P[OF `{c..d} \<noteq> {}`] apply(rule setprod_mono,rule) proof fix i::'n
+    show "0 \<le> b $ i - a $ i" using ab_ne[THEN spec[where x=i]] by auto
+    show "b $ i - a $ i \<le> d $ i - c $ i"
+      using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(2),of i]
+      using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i] by auto qed qed
+
+lemma content_lt_nz: "0 < content {a..b} \<longleftrightarrow> content {a..b} \<noteq> 0"
+  unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by auto
+
+subsection {* The notion of a gauge --- simply an open set containing the point. *}
+
+definition gauge where "gauge d \<longleftrightarrow> (\<forall>x. x\<in>(d x) \<and> open(d x))"
+
+lemma gaugeI:assumes "\<And>x. x\<in>g x" "\<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" "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)" apply(rule gauge_ball) by auto
+
+lemma gauge_inter: "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" "\<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(meson zero_less_one)
+
+subsection {* Divisions. *}
+
+definition division_of (infixl "division'_of" 40) where
+  "s division_of i \<equiv>
+        finite s \<and>
+        (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = {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>k. k\<in>s \<Longrightarrow> k \<subseteq> i" "\<And>k. k\<in>s \<Longrightarrow>  k \<noteq> {}" "\<And>k. k\<in>s \<Longrightarrow> (\<exists>a b. k = {a..b})"
+  "\<And>k1 k2. \<lbrakk>k1\<in>s; k2\<in>s; k1 \<noteq> k2\<rbrakk> \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}" "\<Union>s = i" using assms unfolding division_of_def by auto
+
+lemma division_ofI:
+  assumes "finite s" "\<And>k. k\<in>s \<Longrightarrow> k \<subseteq> i" "\<And>k. k\<in>s \<Longrightarrow>  k \<noteq> {}" "\<And>k. k\<in>s \<Longrightarrow> (\<exists>a b. k = {a..b})"
+  "\<And>k1 k2. \<lbrakk>k1\<in>s; k2\<in>s; k1 \<noteq> k2\<rbrakk> \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}" "\<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]: "{a..b} \<noteq> {} \<Longrightarrow> {{a..b}} division_of {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 {a..a::real^'n} \<longleftrightarrow> s = {{a..a}}" (is "?l = ?r") proof
+  assume ?r moreover { assume "s = {{a}}" moreover fix k assume "k\<in>s" 
+    ultimately have"\<exists>x y. k = {x..y}" apply(rule_tac x=a in exI)+ unfolding interval_sing[THEN conjunct1] by auto }
+  ultimately show ?l unfolding division_of_def interval_sing[THEN conjunct1] by auto next
+  assume ?l note as=conjunctD4[OF this[unfolded division_of_def interval_sing[THEN conjunct1]]]
+  { fix x assume x:"x\<in>s" have "x={a}" using as(2)[rule_format,OF x] by auto }
+  moreover have "s \<noteq> {}" using as(4) by auto ultimately show ?r unfolding interval_sing[THEN conjunct1] 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 {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. {a..b} \<in> d \<longrightarrow> P {a..b}))"
+  unfolding division_of_def by fastsimp
+
+lemma division_of_subset: assumes "p division_of (\<Union>p)" "q \<subseteq> p" shows "q division_of (\<Union>q)"
+  apply(rule division_ofI) proof- note as=division_ofD[OF assms(1)]
+  show "finite q" apply(rule finite_subset) using as(1) assms(2) by auto
+  { fix k assume "k \<in> q" hence kp:"k\<in>p" using assms(2) by auto show "k\<subseteq>\<Union>q" using `k \<in> q` by auto
+  show "\<exists>a b. k = {a..b}" using as(4)[OF kp] by auto show "k \<noteq> {}" using as(3)[OF kp] by auto }
+  fix k1 k2 assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2" hence *:"k1\<in>p" "k2\<in>p" "k1\<noteq>k2" using assms(2) by auto
+  show "interior k1 \<inter> interior k2 = {}" using as(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 {a..b} = 0" "d division_of {a..b}" shows "\<forall>k\<in>d. content k = 0"
+  unfolding forall_in_division[OF assms(2)] apply(rule,rule,rule) apply(drule division_ofD(2)[OF assms(2)])
+  apply(drule content_subset) unfolding assms(1) proof- case goal1 thus ?case using content_pos_le[of a b] by auto qed
+
+lemma division_inter: assumes "p1 division_of s1" "p2 division_of (s2::(real^'a) set)"
+  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_ext) unfolding * and Union_image_eq UN_iff
+    using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto
+  { 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 thus "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
+  guess a1 using division_ofD(4)[OF assms(1) k(2)] .. then guess b1 .. note ab1=this
+  guess a2 using division_ofD(4)[OF assms(2) k(3)] .. then guess b2 .. note ab2=this
+  show "\<exists>a b. k = {a..b}" unfolding k ab1 ab2 unfolding inter_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" hence 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 *) defer apply(rule_tac[1-4] subset_interior)
+    using division_ofD(5)[OF assms(1) k1(2) k2(2)]
+    using division_ofD(5)[OF assms(2) k1(3) k2(3)] using th by auto qed qed
+
+lemma division_inter_1: assumes "d division_of i" "{a..b::real^'n} \<subseteq> i"
+  shows "{ {a..b} \<inter> k |k. k \<in> d \<and> {a..b} \<inter> k \<noteq> {} } division_of {a..b}" proof(cases "{a..b} = {}")
+  case True show ?thesis unfolding True and division_of_trivial by auto next
+  have *:"{a..b} \<inter> i = {a..b}" using assms(2) by auto 
+  case False show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed
+
+lemma elementary_inter: assumes "p1 division_of s" "p2 division_of (t::(real^'n) set)"
+  shows "\<exists>p. p division_of (s \<inter> t)"
+  by(rule,rule division_inter[OF assms])
+
+lemma elementary_inters: assumes "finite f" "f\<noteq>{}" "\<forall>s\<in>f. \<exists>p. p division_of (s::(real^'n) set)"
+  shows "\<exists>p. p division_of (\<Inter> f)" using assms apply-proof(induct f rule:finite_induct)
+case (insert x f) show ?case proof(cases "f={}")
+  case True thus ?thesis unfolding True using insert by auto next
+  case False guess p using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
+  moreover guess px using insert(5)[rule_format,OF insertI1] .. ultimately
+  show ?thesis unfolding Inter_insert apply(rule_tac elementary_inter) by assumption+ qed qed auto
+
+lemma division_disjoint_union:
+  assumes "p1 division_of s1" "p2 division_of s2" "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)] and 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 subset_interior[OF d1(2)[OF as(1)]] subset_interior[OF d2(2)[OF as(2)]]
+      using assms(3) by blast } moreover
+  { assume as:"k1\<in>p2" "k2\<in>p1" have ?g using subset_interior[OF d1(2)[OF as(2)]] subset_interior[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 = {a..b}" using k d1(4) d2(4) by auto qed
+
+lemma partial_division_extend_1:
+  assumes "{c..d} \<subseteq> {a..b::real^'n}" "{c..d} \<noteq> {}"
+  obtains p where "p division_of {a..b}" "{c..d} \<in> p"
+proof- def n \<equiv> "CARD('n)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by auto
+  guess \<pi> using ex_bij_betw_nat_finite_1[OF finite_UNIV[where 'a='n]] .. note \<pi>=this
+  def \<pi>' \<equiv> "inv_into {1..n} \<pi>"
+  have \<pi>':"bij_betw \<pi>' UNIV {1..n}" using bij_betw_inv_into[OF \<pi>] unfolding \<pi>'_def n_def by auto
+  hence \<pi>'i:"\<And>i. \<pi>' i \<in> {1..n}" unfolding bij_betw_def by auto 
+  have \<pi>\<pi>'[simp]:"\<And>i. \<pi> (\<pi>' i) = i" unfolding \<pi>'_def apply(rule f_inv_into_f) unfolding n_def using \<pi> unfolding bij_betw_def by auto
+  have \<pi>'\<pi>[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi>' (\<pi> i) = i" unfolding \<pi>'_def apply(rule inv_into_f_eq) using \<pi> unfolding n_def bij_betw_def by auto
+  have "{c..d} \<noteq> {}" using assms by auto
+  let ?p1 = "\<lambda>l. {(\<chi> i. if \<pi>' i < l then c$i else a$i) .. (\<chi> i. if \<pi>' i < l then d$i else if \<pi>' i = l then c$\<pi> l else b$i)}"
+  let ?p2 = "\<lambda>l. {(\<chi> i. if \<pi>' i < l then c$i else if \<pi>' i = l then d$\<pi> l else a$i) .. (\<chi> i. if \<pi>' i < l then d$i else b$i)}"
+  let ?p =  "{?p1 l |l. l \<in> {1..n+1}} \<union> {?p2 l |l. l \<in> {1..n+1}}"
+  have abcd:"\<And>i. a $ i \<le> c $ i \<and> c$i \<le> d$i \<and> d $ i \<le> b $ i" using assms unfolding subset_interval interval_eq_empty by(auto simp add:not_le not_less)
+  show ?thesis apply(rule that[of ?p]) apply(rule division_ofI)
+  proof- have "\<And>i. \<pi>' i < Suc n"
+    proof(rule ccontr,unfold not_less) fix i assume "Suc n \<le> \<pi>' i"
+      hence "\<pi>' i \<notin> {1..n}" by auto thus False using \<pi>' unfolding bij_betw_def by auto
+    qed hence "c = (\<chi> i. if \<pi>' i < Suc n then c $ i else a $ i)"
+        "d = (\<chi> i. if \<pi>' i < Suc n then d $ i else if \<pi>' i = n + 1 then c $ \<pi> (n + 1) else b $ i)"
+      unfolding Cart_eq Cart_lambda_beta using \<pi>' unfolding bij_betw_def by auto
+    thus cdp:"{c..d} \<in> ?p" apply-apply(rule UnI1) unfolding mem_Collect_eq apply(rule_tac x="n + 1" in exI) by auto
+    have "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p1 l \<subseteq> {a..b}"  "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p2 l \<subseteq> {a..b}"
+      unfolding subset_eq apply(rule_tac[!] ballI,rule_tac[!] ccontr)
+    proof- fix l assume l:"l\<in>{1..n+1}" fix x assume "x\<notin>{a..b}"
+      then guess i unfolding mem_interval not_all .. note i=this
+      show "x \<in> ?p1 l \<Longrightarrow> False" "x \<in> ?p2 l \<Longrightarrow> False" unfolding mem_interval apply(erule_tac[!] x=i in allE)
+        apply(case_tac[!] "\<pi>' i < l", case_tac[!] "\<pi>' i = l") using abcd[of i] i by auto 
+    qed moreover have "\<And>x. x \<in> {a..b} \<Longrightarrow> x \<in> \<Union>?p"
+    proof- fix x assume x:"x\<in>{a..b}"
+      { presume "x\<notin>{c..d} \<Longrightarrow> x \<in> \<Union>?p" thus "x \<in> \<Union>?p" using cdp by blast }
+      let ?M = "{i. i\<in>{1..n+1} \<and> \<not> (c $ \<pi> i \<le> x $ \<pi> i \<and> x $ \<pi> i \<le> d $ \<pi> i)}"
+      assume "x\<notin>{c..d}" then guess i0 unfolding mem_interval not_all ..
+      hence "\<pi>' i0 \<in> ?M" using \<pi>' unfolding bij_betw_def by(auto intro!:le_SucI)
+      hence M:"finite ?M" "?M \<noteq> {}" by auto
+      def l \<equiv> "Min ?M" note l = Min_less_iff[OF M,unfolded l_def[symmetric]] Min_in[OF M,unfolded mem_Collect_eq l_def[symmetric]]
+        Min_gr_iff[OF M,unfolded l_def[symmetric]]
+      have "x\<in>?p1 l \<or> x\<in>?p2 l" using l(2)[THEN conjunct2] unfolding de_Morgan_conj not_le
+        apply- apply(erule disjE) apply(rule disjI1) defer apply(rule disjI2)
+      proof- assume as:"x $ \<pi> l < c $ \<pi> l"
+        show "x \<in> ?p1 l" unfolding mem_interval Cart_lambda_beta
+        proof case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le by auto
+          thus ?case using as x[unfolded mem_interval,rule_format,of i]
+            apply auto using l(3)[of "\<pi>' i"] by(auto elim!:ballE[where x="\<pi>' i"])
+        qed
+      next assume as:"x $ \<pi> l > d $ \<pi> l"
+        show "x \<in> ?p2 l" unfolding mem_interval Cart_lambda_beta
+        proof case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le by auto
+          thus ?case using as x[unfolded mem_interval,rule_format,of i]
+            apply auto using l(3)[of "\<pi>' i"] by(auto elim!:ballE[where x="\<pi>' i"])
+        qed qed
+      thus "x \<in> \<Union>?p" using l(2) by blast 
+    qed ultimately show "\<Union>?p = {a..b}" apply-apply(rule) defer apply(rule) by(assumption,blast)
+    
+    show "finite ?p" by auto
+    fix k assume k:"k\<in>?p" then obtain l where l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}" by auto
+    show "k\<subseteq>{a..b}" apply(rule,unfold mem_interval,rule,rule) 
+    proof- fix i::'n and x assume "x \<in> k" moreover have "\<pi>' i < l \<or> \<pi>' i = l \<or> \<pi>' i > l" by auto
+      ultimately show "a$i \<le> x$i" "x$i \<le> b$i" using abcd[of i] using l by(auto elim:disjE elim!:allE[where x=i] simp add:vector_le_def)
+    qed have "\<And>l. ?p1 l \<noteq> {}" "\<And>l. ?p2 l \<noteq> {}" unfolding interval_eq_empty not_ex apply(rule_tac[!] allI)
+    proof- case goal1 thus ?case using abcd[of x] by auto
+    next   case goal2 thus ?case using abcd[of x] by auto
+    qed thus "k \<noteq> {}" using k by auto
+    show "\<exists>a b. k = {a..b}" using k by auto
+    fix k' assume k':"k' \<in> ?p" "k \<noteq> k'" then obtain l' where l':"k' = ?p1 l' \<or> k' = ?p2 l'" "l' \<in> {1..n + 1}" by auto
+    { fix k k' l l'
+      assume k:"k\<in>?p" and l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}" 
+      assume k':"k' \<in> ?p" "k \<noteq> k'" and  l':"k' = ?p1 l' \<or> k' = ?p2 l'" "l' \<in> {1..n + 1}" 
+      assume "l \<le> l'" fix x
+      have "x \<notin> interior k \<inter> interior k'" 
+      proof(rule,cases "l' = n+1") assume x:"x \<in> interior k \<inter> interior k'"
+        case True hence "\<And>i. \<pi>' i < l'" using \<pi>'i by(auto simp add:less_Suc_eq_le)
+        hence k':"k' = {c..d}" using l'(1) \<pi>'i by(auto simp add:Cart_nth_inverse)
+        have ln:"l < n + 1" 
+        proof(rule ccontr) case goal1 hence l2:"l = n+1" using l by auto
+          hence "\<And>i. \<pi>' i < l" using \<pi>'i by(auto simp add:less_Suc_eq_le)
+          hence "k = {c..d}" using l(1) \<pi>'i by(auto simp add:Cart_nth_inverse)
+          thus False using `k\<noteq>k'` k' by auto
+        qed have **:"\<pi>' (\<pi> l) = l" using \<pi>'\<pi>[of l] using l ln by auto
+        have "x $ \<pi> l < c $ \<pi> l \<or> d $ \<pi> l < x $ \<pi> l" using l(1) apply-
+        proof(erule disjE)
+          assume as:"k = ?p1 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
+          show ?thesis using *[of "\<pi> l"] using ln unfolding Cart_lambda_beta ** by auto
+        next assume as:"k = ?p2 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
+          show ?thesis using *[of "\<pi> l"] using ln unfolding Cart_lambda_beta ** by auto
+        qed thus False using x unfolding k' unfolding Int_iff interior_closed_interval mem_interval
+          by(auto elim!:allE[where x="\<pi> l"])
+      next case False hence "l < n + 1" using l'(2) using `l\<le>l'` by auto
+        hence ln:"l \<in> {1..n}" "l' \<in> {1..n}" using l l' False by auto
+        note \<pi>l = \<pi>'\<pi>[OF ln(1)] \<pi>'\<pi>[OF ln(2)]
+        assume x:"x \<in> interior k \<inter> interior k'"
+        show False using l(1) l'(1) apply-
+        proof(erule_tac[!] disjE)+
+          assume as:"k = ?p1 l" "k' = ?p1 l'"
+          note * = x[unfolded as Int_iff interior_closed_interval mem_interval]
+          have "l \<noteq> l'" using k'(2)[unfolded as] by auto
+          thus False using * by(smt Cart_lambda_beta \<pi>l)
+        next assume as:"k = ?p2 l" "k' = ?p2 l'"
+          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
+          have "l \<noteq> l'" apply(rule) using k'(2)[unfolded as] by auto
+          thus False using *[of "\<pi> l"] *[of "\<pi> l'"]
+            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` by auto
+        next assume as:"k = ?p1 l" "k' = ?p2 l'"
+          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
+          show False using *[of "\<pi> l"] *[of "\<pi> l'"]
+            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` using abcd[of "\<pi> l'"] by smt 
+        next assume as:"k = ?p2 l" "k' = ?p1 l'"
+          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
+          show False using *[of "\<pi> l"] *[of "\<pi> l'"]
+            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` using abcd[of "\<pi> l'"] by smt
+        qed qed } 
+    from this[OF k l k' l'] this[OF k'(1) l' k _ l] have "\<And>x. x \<notin> interior k \<inter> interior k'"
+      apply - apply(cases "l' \<le> l") using k'(2) by auto            
+    thus "interior k \<inter> interior k' = {}" by auto        
+qed qed
+
+lemma partial_division_extend_interval: assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> {a..b}"
+  obtains q where "p \<subseteq> q" "q division_of {a..b::real^'n}" proof(cases "p = {}")
+  case True guess q apply(rule elementary_interval[of a b]) .
+  thus ?thesis apply- apply(rule that[of q]) unfolding True by auto next
+  case False note p = division_ofD[OF assms(1)]
+  have *:"\<forall>k\<in>p. \<exists>q. q division_of {a..b} \<and> k\<in>q" proof case goal1
+    guess c using p(4)[OF goal1] .. then guess d .. note cd_ = this
+    have *:"{c..d} \<subseteq> {a..b}" "{c..d} \<noteq> {}" using p(2,3)[OF goal1, unfolded cd_] using assms(2) by auto
+    guess q apply(rule partial_division_extend_1[OF *]) . thus ?case unfolding cd_ by auto qed
+  guess q using bchoice[OF *] .. note q = conjunctD2[OF this[rule_format]]
+  have "\<And>x. x\<in>p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})" apply(rule,rule_tac p="q x" in division_of_subset) proof-
+    fix x assume x:"x\<in>p" show "q x division_of \<Union>q x" apply-apply(rule division_ofI)
+      using division_ofD[OF q(1)[OF x]] by auto show "q x - {x} \<subseteq> q x" by auto qed
+  hence "\<exists>d. d division_of \<Inter> ((\<lambda>i. \<Union>(q i - {i})) ` p)" apply- apply(rule elementary_inters)
+    apply(rule finite_imageI[OF p(1)]) unfolding image_is_empty apply(rule False) by auto
+  then guess d .. note d = this
+  show ?thesis apply(rule that[of "d \<union> p"]) proof-
+    have *:"\<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 *:"{a..b} = \<Inter> (\<lambda>i. \<Union>(q i - {i})) ` p \<union> \<Union>p" apply(rule *[OF False]) proof fix i assume i:"i\<in>p"
+      show "\<Union>(q i - {i}) \<union> i = {a..b}" using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto qed
+    show "d \<union> p division_of {a..b}" unfolding * apply(rule division_disjoint_union[OF d assms(1)])
+      apply(rule inter_interior_unions_intervals) apply(rule p open_interior ballI)+ proof(assumption,rule)
+      fix k assume k:"k\<in>p" have *:"\<And>u t s. u \<subseteq> s \<Longrightarrow> s \<inter> t = {} \<Longrightarrow> u \<inter> t = {}" by auto
+      show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<inter> interior k = {}" apply(rule *[of _ "interior (\<Union>(q k - {k}))"])
+	defer apply(subst Int_commute) apply(rule inter_interior_unions_intervals) proof- 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 = {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
+	have *:"\<And>x s. x \<in> s \<Longrightarrow> \<Inter>s \<subseteq> x" by auto show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<subseteq> interior (\<Union>(q k - {k}))"
+	  apply(rule subset_interior *)+ using k by auto qed qed qed auto qed
+
+lemma elementary_bounded[dest]: "p division_of s \<Longrightarrow> bounded (s::(real^'n) set)"
+  unfolding division_of_def by(metis bounded_Union bounded_interval) 
+
+lemma elementary_subset_interval: "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> {a..b::real^'n}"
+  by(meson elementary_bounded bounded_subset_closed_interval)
+
+lemma division_union_intervals_exists: assumes "{a..b::real^'n} \<noteq> {}"
+  obtains p where "(insert {a..b} p) division_of ({a..b} \<union> {c..d})" proof(cases "{c..d} = {}")
+  case True show ?thesis apply(rule that[of "{}"]) unfolding True using assms by auto next
+  case False note false=this show ?thesis proof(cases "{a..b} \<inter> {c..d} = {}")
+  have *:"\<And>a b. {a,b} = {a} \<union> {b}" by auto
+  case True show ?thesis apply(rule that[of "{{c..d}}"]) unfolding * apply(rule division_disjoint_union)
+    using false True assms using interior_subset by auto next
+  case False obtain u v where uv:"{a..b} \<inter> {c..d} = {u..v}" unfolding inter_interval by auto
+  have *:"{u..v} \<subseteq> {c..d}" using uv by auto
+  guess p apply(rule partial_division_extend_1[OF * False[unfolded uv]]) . note p=this division_ofD[OF this(1)]
+  have *:"{a..b} \<union> {c..d} = {a..b} \<union> \<Union>(p - {{u..v}})" "\<And>x s. insert x s = {x} \<union> s" using p(8) unfolding uv[THEN sym] by auto
+  show thesis apply(rule that[of "p - {{u..v}}"]) unfolding *(1) apply(subst *(2)) apply(rule division_disjoint_union)
+    apply(rule,rule assms) apply(rule division_of_subset[of p]) apply(rule division_of_union_self[OF p(1)]) defer
+    unfolding interior_inter[THEN sym] proof-
+    have *:"\<And>cd p uv ab. p \<subseteq> cd \<Longrightarrow> ab \<inter> cd = uv \<Longrightarrow> ab \<inter> p = uv \<inter> p" by auto
+    have "interior ({a..b} \<inter> \<Union>(p - {{u..v}})) = interior({u..v} \<inter> \<Union>(p - {{u..v}}))" 
+      apply(rule arg_cong[of _ _ interior]) apply(rule *[OF _ uv]) using p(8) by auto
+    also have "\<dots> = {}" unfolding interior_inter apply(rule inter_interior_unions_intervals) using p(6) p(7)[OF p(2)] p(3) by auto
+    finally show "interior ({a..b} \<inter> \<Union>(p - {{u..v}})) = {}" by assumption qed auto qed qed
+
+lemma division_of_unions: assumes "finite f"  "\<And>p. p\<in>f \<Longrightarrow> p division_of (\<Union>p)"
+  "\<And>k1 k2. \<lbrakk>k1 \<in> \<Union>f; k2 \<in> \<Union>f; k1 \<noteq> k2\<rbrakk> \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+  shows "\<Union>f division_of \<Union>\<Union>f" apply(rule division_ofI) prefer 5 apply(rule assms(3)|assumption)+
+  apply(rule finite_Union assms(1))+ prefer 3 apply(erule UnionE) apply(rule_tac s=X in division_ofD(3)[OF assms(2)])
+  using division_ofD[OF assms(2)] by auto
+  
+lemma elementary_union_interval: assumes "p division_of \<Union>p"
+  obtains q where "q division_of ({a..b::real^'n} \<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" "{a..b} = {} \<Longrightarrow> thesis" "{a..b} \<noteq> {} \<Longrightarrow> interior {a..b} = {} \<Longrightarrow> thesis"
+    "p\<noteq>{} \<Longrightarrow> interior {a..b}\<noteq>{} \<Longrightarrow> {a..b} \<noteq> {} \<Longrightarrow> thesis"
+  thus thesis by auto
+next assume as:"p={}" guess p apply(rule elementary_interval[of a b]) .
+  thus thesis apply(rule_tac that[of p]) unfolding as by auto 
+next assume as:"{a..b}={}" show thesis apply(rule that) unfolding as using assms by auto
+next assume as:"interior {a..b} = {}" "{a..b} \<noteq> {}"
+  show thesis apply(rule that[of "insert {a..b} p"],rule division_ofI)
+    unfolding finite_insert apply(rule assm(1)) unfolding Union_insert  
+    using assm(2-4) as apply- by(fastsimp dest: assm(5))+
+next assume as:"p \<noteq> {}" "interior {a..b} \<noteq> {}" "{a..b}\<noteq>{}"
+  have "\<forall>k\<in>p. \<exists>q. (insert {a..b} q) division_of ({a..b} \<union> k)" proof case goal1
+    from assm(4)[OF this] guess c .. then guess d ..
+    thus ?case apply-apply(rule division_union_intervals_exists[OF as(3),of c d]) by auto
+  qed from bchoice[OF this] guess q .. note q=division_ofD[OF this[rule_format]]
+  let ?D = "\<Union>{insert {a..b} (q k) | k. k \<in> p}"
+  show thesis apply(rule that[of "?D"]) proof(rule division_ofI)
+    have *:"{insert {a..b} (q k) |k. k \<in> p} = (\<lambda>k. insert {a..b} (q k)) ` p" by auto
+    show "finite ?D" apply(rule finite_Union) unfolding * apply(rule finite_imageI) using assm(1) q(1) by auto
+    show "\<Union>?D = {a..b} \<union> \<Union>p" unfolding * lem1 unfolding lem2[OF as(1), of "{a..b}",THEN sym]
+      using q(6) by auto
+    fix k assume k:"k\<in>?D" thus " k \<subseteq> {a..b} \<union> \<Union>p" using q(2) by auto
+    show "k \<noteq> {}" using q(3) k by auto show "\<exists>a b. k = {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 {a..b} (q x)"  "x\<in>p"  using k  by auto
+    obtain x' where x':"k'\<in>insert {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 apply(rule q(5)) using x x' k' unfolding True by auto
+    next case False 
+      { presume "k = {a..b} \<Longrightarrow> ?thesis" "k' = {a..b} \<Longrightarrow> ?thesis" 
+        "k \<noteq> {a..b} \<Longrightarrow> k' \<noteq> {a..b} \<Longrightarrow> ?thesis"
+        thus ?thesis by auto }
+      { assume as':"k  = {a..b}" show ?thesis apply(rule q(5)) using x' k'(2) unfolding as' by auto }
+      { assume as':"k' = {a..b}" show ?thesis apply(rule q(5)) using x  k'(2) unfolding as' by auto }
+      assume as':"k \<noteq> {a..b}" "k' \<noteq> {a..b}"
+      guess c using q(4)[OF x(2,1)] .. then guess d .. note c_d=this
+      have "interior k  \<inter> interior {a..b} = {}" apply(rule q(5)) using x  k'(2) using as' by auto
+      hence "interior k \<subseteq> interior x" apply-
+        apply(rule interior_subset_union_intervals[OF c_d _ as(2) q(2)[OF x(2,1)]]) by auto moreover
+      guess c using q(4)[OF x'(2,1)] .. then guess d .. note c_d=this
+      have "interior k' \<inter> interior {a..b} = {}" apply(rule q(5)) using x' k'(2) using as' by auto
+      hence "interior k' \<subseteq> interior x'" apply-
+        apply(rule interior_subset_union_intervals[OF c_d _ as(2) q(2)[OF x'(2,1)]]) by auto
+      ultimately show ?thesis using assm(5)[OF x(2) x'(2) False] by auto
+    qed qed } qed
+
+lemma elementary_unions_intervals:
+  assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = {a..b::real^'n}"
+  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
+    fix x F assume as:"finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
+    from this(3) guess p .. note p=this
+    from assms(2)[OF as(4)] guess a .. then guess b .. note ab=this
+    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 ab * by auto
+  qed(insert assms,auto) thus ?thesis apply-apply(erule exE,rule that) by auto qed
+
+lemma elementary_union: assumes "ps division_of s" "pt division_of (t::(real^'n) set)"
+  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
+  hence *:"\<Union>(ps \<union> pt) = s \<union> t" by auto
+  show ?thesis apply-apply(rule elementary_unions_intervals[of "ps\<union>pt"])
+    unfolding * prefer 3 apply(rule_tac p=p in that)
+    using assms[unfolded division_of_def] by auto qed
+
+lemma partial_division_extend: fixes t::"(real^'n) set"
+  assumes "p division_of s" "q division_of t" "s \<subseteq> t"
+  obtains r where "p \<subseteq> r" "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>{a..b}" using elementary_subset_interval[OF assms(2)] by auto
+  guess r1 apply(rule partial_division_extend_interval) apply(rule assms(1)[unfolded divp(6)[THEN sym]])
+    apply(rule subset_trans) by(rule ab assms[unfolded divp(6)[THEN sym]])+  note r1 = this division_ofD[OF this(2)]
+  guess p' apply(rule elementary_unions_intervals[of "r1 - p"]) using r1(3,6) by auto 
+  then obtain r2 where r2:"r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)" 
+    apply- apply(drule elementary_inter[OF _ assms(2)[unfolded divq(6)[THEN sym]]]) by auto
+  { fix x assume x:"x\<in>t" "x\<notin>s"
+    hence "x\<in>\<Union>r1" unfolding r1 using ab by auto
+    then guess r unfolding Union_iff .. note r=this moreover
+    have "r \<notin> p" proof assume "r\<in>p" hence "x\<in>s" using divp(2) r by auto
+      thus False using x by auto qed
+    ultimately have "x\<in>\<Union>(r1 - p)" by auto }
+  hence *:"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)" "open (interior s)" "\<forall>t\<in>r1-p. \<exists>a b. t = {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(rule)
+        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" "open (interior m)" "\<forall>t\<in>p. \<exists>a b. t = {a..b}" using divp by auto
+          show "\<forall>t\<in>p. interior m \<inter> interior t = {}" apply(rule, rule r1(7)) using as using r1 by auto
+        qed thus "interior s \<inter> interior m = {}" unfolding divp by auto
+      qed qed        
+    thus "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}" using interior_subset by auto
+  qed auto qed
+
+subsection {* Tagged (partial) divisions. *}
+
+definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where
+  "(s tagged_partial_division_of i) \<equiv>
+        finite s \<and>
+        (\<forall>x k. (x,k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = {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>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+  "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = {a..b}"
+  "\<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  apply- by blast+ 
+
+definition tagged_division_of (infixr "tagged'_division'_of" 40) where
+  "(s tagged_division_of i) \<equiv>
+        (s tagged_partial_division_of i) \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+
+lemma tagged_division_of_finite[dest]: "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 = {a..b})) \<and>
+        (\<forall>x1 k1 x2 k2. (x1,k1) \<in> s \<and> (x2,k2) \<in> s \<and> ~((x1,k1) = (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>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"  "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = {a..b}"
+  "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2,k2) \<in> s \<Longrightarrow> ~((x1,k1) = (x2,k2)) \<Longrightarrow> (interior(k1) \<inter> interior(k2) = {})"
+  "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  shows "s tagged_division_of i"
+  unfolding tagged_division_of apply(rule) defer apply rule
+  apply(rule allI impI conjI assms)+ apply assumption
+  apply(rule, rule assms, assumption) apply(rule assms, assumption)
+  using assms(1,5-) apply- by blast+
+
+lemma tagged_division_ofD[dest]: assumes "s tagged_division_of i"
+  shows "finite s" "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"  "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = {a..b}"
+  "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2,k2) \<in> s \<Longrightarrow> ~((x1,k1) = (x2,k2)) \<Longrightarrow> (interior(k1) \<inter> interior(k2) = {})"
+  "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)" using assms unfolding tagged_division_of apply- 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
+  thus  "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = {a..b}" using assm apply- by fastsimp+
+  fix k' assume k':"k' \<in> snd ` s" "k \<noteq> k'" from this(1) obtain xk' where xk':"(xk', k') \<in> s" by auto
+  thus "interior k \<inter> interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto
+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
+  thus "k\<noteq>{}" "\<exists>a b. k = {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
+  thus "interior k \<inter> interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto
+qed
+
+lemma tagged_partial_division_subset: assumes "s tagged_partial_division_of i" "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 setsum_over_tagged_division_lemma: fixes d::"(real^'m) set \<Rightarrow> 'a::real_normed_vector"
+  assumes "p tagged_division_of i" "\<And>u v. {u..v} \<noteq> {} \<Longrightarrow> content {u..v} = 0 \<Longrightarrow> d {u..v} = 0"
+  shows "setsum (\<lambda>(x,k). d k) p = setsum d (snd ` p)"
+proof- note assm=tagged_division_ofD[OF assms(1)]
+  have *:"(\<lambda>(x,k). d k) = d \<circ> snd" unfolding o_def apply(rule ext) by auto
+  show ?thesis unfolding * apply(subst eq_commute) proof(rule setsum_reindex_nonzero)
+    show "finite p" using assm by auto
+    fix x y assume as:"x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y" 
+    obtain a b where ab:"snd x = {a..b}" using assm(4)[of "fst x" "snd x"] as(1) by auto
+    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y" unfolding as(4)[THEN sym] using as(1-3) by auto
+    hence "interior (snd x) \<inter> interior (snd y) = {}" apply-apply(rule assm(5)[of "fst x" _ "fst y"]) using as by auto 
+    hence "content {a..b} = 0" unfolding as(4)[THEN sym] ab content_eq_0_interior by auto
+    hence "d {a..b} = 0" apply-apply(rule assms(2)) using assm(2)[of "fst x" "snd x"] as(1) unfolding ab[THEN sym] by auto
+    thus "d (snd x) = 0" 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> {a..b} \<Longrightarrow> {(x,{a..b})} tagged_division_of {a..b}"
+  apply(rule tagged_division_ofI) by auto
+
+lemma tagged_division_union:
+  assumes "p1 tagged_division_of s1"  "p2 tagged_division_of s2" "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)] and 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 = {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) subset_interior by blast
+  show "interior k \<inter> interior k' = {}" apply(cases "(x,k)\<in>p1", case_tac[!] "(x',k')\<in>p1")
+    apply(rule p1(5)) prefer 4 apply(rule *) prefer 6 apply(subst Int_commute,rule *) prefer 8 apply(rule p2(5))
+    using p1(3) p2(3) using xk xk' by auto qed 
+
+lemma tagged_division_unions:
+  assumes "finite iset" "\<forall>i\<in>iset. (pfn(i) tagged_division_of i)"
+  "\<forall>i1 \<in> iset. \<forall>i2 \<in> iset. ~(i1 = 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)" apply(rule finite_Union) 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 = {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] subset_interior by blast
+  show "interior k \<inter> interior k' = {}" apply(cases "i=i'")
+    using assm(5)[OF i _ xk'(2)]  i'(2) using assm(3)[OF i] assm(3)[OF i'] defer apply-apply(rule *) 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] by auto
+
+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] by auto
+
+subsection {* Fine-ness of a partition w.r.t. a gauge. *}
+
+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 {* Gauge integral. Define on compact intervals first, then use a limit. *}
+
+definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) where
+  "(f has_integral_compact_interval y) i \<equiv>
+        (\<forall>e>0. \<exists>d. gauge d \<and>
+          (\<forall>p. p tagged_division_of i \<and> d fine p
+                        \<longrightarrow> norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))"
+
+definition has_integral (infixr "has'_integral" 46) where 
+"((f::(real^'n \<Rightarrow> 'b::real_normed_vector)) has_integral y) i \<equiv>
+        if (\<exists>a b. i = {a..b}) then (f has_integral_compact_interval y) i
+        else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b}
+              \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral_compact_interval z) {a..b} \<and>
+                                       norm(z - y) < e))"
+
+lemma has_integral:
+ "(f has_integral y) ({a..b}) \<longleftrightarrow>
+        (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p
+                        \<longrightarrow> norm(setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
+  unfolding has_integral_def has_integral_compact_interval_def by auto
+
+lemma has_integralD[dest]: assumes
+ "(f has_integral y) ({a..b})" "e>0"
+  obtains d where "gauge d" "\<And>p. p tagged_division_of {a..b} \<Longrightarrow> d fine p
+                        \<Longrightarrow> norm(setsum (\<lambda>(x,k). content(k) *\<^sub>R f(x)) p - y) < e"
+  using assms unfolding has_integral by auto
+
+lemma has_integral_alt:
+ "(f has_integral y) i \<longleftrightarrow>
+      (if (\<exists>a b. i = {a..b}) then (f has_integral y) i
+       else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b}
+                               \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0)
+                                        has_integral z) ({a..b}) \<and>
+                                       norm(z - y) < e)))"
+  unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto
+
+lemma has_integral_altD:
+  assumes "(f has_integral y) i" "\<not> (\<exists>a b. i = {a..b})" "e>0"
+  obtains B where "B>0" "\<forall>a b. ball 0 B \<subseteq> {a..b}\<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) ({a..b}) \<and> norm(z - y) < e)"
+  using assms unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto
+
+definition integrable_on (infixr "integrable'_on" 46) where
+  "(f integrable_on i) \<equiv> \<exists>y. (f has_integral y) i"
+
+definition "integral i f \<equiv> SOME y. (f has_integral y) i"
+
+lemma integrable_integral[dest]:
+ "f integrable_on i \<Longrightarrow> (f has_integral (integral i f)) i"
+  unfolding integrable_on_def integral_def by(rule someI_ex)
+
+lemma has_integral_integrable[intro]: "(f has_integral i) s \<Longrightarrow> f integrable_on s"
+  unfolding integrable_on_def by auto
+
+lemma has_integral_integral:"f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s"
+  by auto
+
+lemma has_integral_vec1: assumes "(f has_integral k) {a..b}"
+  shows "((\<lambda>x. vec1 (f x)) has_integral (vec1 k)) {a..b}"
+proof- have *:"\<And>p. (\<Sum>(x, k)\<in>p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k)"
+    unfolding vec_sub Cart_eq by(auto simp add:vec1_dest_vec1_simps split_beta)
+  show ?thesis using assms unfolding has_integral apply safe
+    apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe)
+    apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed
+
+lemma setsum_content_null:
+  assumes "content({a..b}) = 0" "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
+proof(rule setsum_0',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) guess c .. then guess d .. note c_d=this
+  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 c_d]] content_pos_le[of c d]
+    unfolding assms(1) c_d by auto
+  finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
+qed
+
+subsection {* Some basic combining lemmas. *}
+
+lemma tagged_division_unions_exists:
+  assumes "finite iset" "\<forall>i \<in> iset. \<exists>p. p tagged_division_of i \<and> d fine p"
+  "\<forall>i1\<in>iset. \<forall>i2\<in>iset. ~(i1 = i2) \<longrightarrow> (interior(i1) \<inter> interior(i2) = {})" "(\<Union>iset = i)"
+   obtains p where "p tagged_division_of i" "d fine p"
+proof- guess pfn using bchoice[OF assms(2)] .. note pfn = conjunctD2[OF this[rule_format]]
+  show thesis apply(rule_tac p="\<Union>(pfn ` iset)" in that) unfolding assms(4)[THEN sym]
+    apply(rule tagged_division_unions[OF assms(1) _ assms(3)]) defer 
+    apply(rule fine_unions) using pfn by auto
+qed
+
+subsection {* The set we're concerned with must be closed. *}
+
+lemma division_of_closed: "s division_of i \<Longrightarrow> closed (i::(real^'n) set)"
+  unfolding division_of_def by(fastsimp intro!: closed_Union closed_interval)
+
+subsection {* General bisection principle for intervals; might be useful elsewhere. *}
+
+lemma interval_bisection_step:
+  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "~(P {a..b::real^'n})"
+  obtains c d where "~(P{c..d})"
+  "\<forall>i. a$i \<le> c$i \<and> c$i \<le> d$i \<and> d$i \<le> b$i \<and> 2 * (d$i - c$i) \<le> b$i - a$i"
+proof- have "{a..b} \<noteq> {}" using assms(1,3) by auto
+  note ab=this[unfolded interval_eq_empty not_ex not_less]
+  { fix f have "finite f \<Longrightarrow>
+        (\<forall>s\<in>f. P s) \<Longrightarrow>
+        (\<forall>s\<in>f. \<exists>a b. s = {a..b}) \<Longrightarrow>
+        (\<forall>s\<in>f.\<forall>t\<in>f. ~(s = t) \<longrightarrow> interior(s) \<inter> interior(t) = {}) \<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])
+        apply rule defer apply rule defer apply(rule inter_interior_unions_intervals)
+        using insert by auto
+    qed } note * = this
+  let ?A = "{{c..d} | c d. \<forall>i. (c$i = a$i) \<and> (d$i = (a$i + b$i) / 2) \<or> (c$i = (a$i + b$i) / 2) \<and> (d$i = b$i)}"
+  let ?PP = "\<lambda>c d. \<forall>i. a$i \<le> c$i \<and> c$i \<le> d$i \<and> d$i \<le> b$i \<and> 2 * (d$i - c$i) \<le> b$i - a$i"
+  { presume "\<forall>c d. ?PP c d \<longrightarrow> P {c..d} \<Longrightarrow> False"
+    thus thesis unfolding atomize_not not_all apply-apply(erule exE)+ apply(rule_tac c=x and d=xa in that) by auto }
+  assume as:"\<forall>c d. ?PP c d \<longrightarrow> P {c..d}"
+  have "P (\<Union> ?A)" proof(rule *, rule_tac[2-] ballI, rule_tac[4] ballI, rule_tac[4] impI) 
+    let ?B = "(\<lambda>s.{(\<chi> i. if i \<in> s then a$i else (a$i + b$i) / 2) ..
+      (\<chi> i. if i \<in> s then (a$i + b$i) / 2 else b$i)}) ` {s. s \<subseteq> UNIV}"
+    have "?A \<subseteq> ?B" proof case goal1
+      then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format]
+      have *:"\<And>a b c d. a = c \<Longrightarrow> b = d \<Longrightarrow> {a..b} = {c..d}" by auto
+      show "x\<in>?B" unfolding image_iff apply(rule_tac x="{i. c$i = a$i}" in bexI)
+        unfolding c_d apply(rule * ) unfolding Cart_eq cond_component Cart_lambda_beta
+      proof(rule_tac[1-2] allI) fix i show "c $ i = (if i \<in> {i. c $ i = a $ i} then a $ i else (a $ i + b $ i) / 2)"
+          "d $ i = (if i \<in> {i. c $ i = a $ i} then (a $ i + b $ i) / 2 else b $ i)"
+          using c_d(2)[of i] ab[THEN spec[where x=i]] by(auto simp add:field_simps)
+      qed auto qed
+    thus "finite ?A" apply(rule finite_subset[of _ ?B]) by auto
+    fix s assume "s\<in>?A" then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+)
+    note c_d=this[rule_format]
+    show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 show ?case 
+        using c_d(2)[of i] using ab[THEN spec[where x=i]] by auto qed
+    show "\<exists>a b. s = {a..b}" unfolding c_d by auto
+    fix t assume "t\<in>?A" then guess e unfolding mem_Collect_eq .. then guess f apply- by(erule exE,(erule conjE)+)
+    note e_f=this[rule_format]
+    assume "s \<noteq> t" hence "\<not> (c = e \<and> d = f)" unfolding c_d e_f by auto
+    then obtain i where "c$i \<noteq> e$i \<or> d$i \<noteq> f$i" unfolding de_Morgan_conj Cart_eq by auto
+    hence i:"c$i \<noteq> e$i" "d$i \<noteq> f$i" apply- apply(erule_tac[!] disjE)
+    proof- assume "c$i \<noteq> e$i" thus "d$i \<noteq> f$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
+    next   assume "d$i \<noteq> f$i" thus "c$i \<noteq> e$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
+    qed 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 e_f c_d interior_closed_interval proof(rule *)
+      fix x assume "x\<in>{c<..<d}" "x\<in>{e<..<f}"
+      hence x:"c$i < d$i" "e$i < f$i" "c$i < f$i" "e$i < d$i" unfolding mem_interval apply-apply(erule_tac[!] x=i in allE)+ by auto
+      show False using c_d(2)[of i] apply- apply(erule_tac disjE)
+      proof(erule_tac[!] conjE) assume as:"c $ i = a $ i" "d $ i = (a $ i + b $ i) / 2"
+        show False using e_f(2)[of i] and i x unfolding as by(fastsimp simp add:field_simps)
+      next assume as:"c $ i = (a $ i + b $ i) / 2" "d $ i = b $ i"
+        show False using e_f(2)[of i] and i x unfolding as by(fastsimp simp add:field_simps)
+      qed qed qed
+  also have "\<Union> ?A = {a..b}" proof(rule set_ext,rule)
+    fix x assume "x\<in>\<Union>?A" then guess Y unfolding Union_iff ..
+    from this(1) guess c unfolding mem_Collect_eq .. then guess d ..
+    note c_d = this[THEN conjunct2,rule_format] `x\<in>Y`[unfolded this[THEN conjunct1]]
+    show "x\<in>{a..b}" unfolding mem_interval proof 
+      fix i show "a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
+        using c_d(1)[of i] c_d(2)[unfolded mem_interval,THEN spec[where x=i]] by auto qed
+  next fix x assume x:"x\<in>{a..b}"
+    have "\<forall>i. \<exists>c d. (c = a$i \<and> d = (a$i + b$i) / 2 \<or> c = (a$i + b$i) / 2 \<and> d = b$i) \<and> c\<le>x$i \<and> x$i \<le> d"
+      (is "\<forall>i. \<exists>c d. ?P i c d") unfolding mem_interval proof fix i
+      have "?P i (a$i) ((a $ i + b $ i) / 2) \<or> ?P i ((a $ i + b $ i) / 2) (b$i)"
+        using x[unfolded mem_interval,THEN spec[where x=i]] by auto thus "\<exists>c d. ?P i c d" by blast
+    qed thus "x\<in>\<Union>?A" unfolding Union_iff lambda_skolem unfolding Bex_def mem_Collect_eq
+      apply-apply(erule exE)+ apply(rule_tac x="{xa..xaa}" in exI) unfolding mem_interval by auto
+  qed finally show False using assms by auto qed
+
+lemma interval_bisection:
+  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "\<not> P {a..b::real^'n}"
+  obtains x where "x \<in> {a..b}" "\<forall>e>0. \<exists>c d. x \<in> {c..d} \<and> {c..d} \<subseteq> ball x e \<and> {c..d} \<subseteq> {a..b} \<and> ~P({c..d})"
+proof-
+  have "\<forall>x. \<exists>y. \<not> P {fst x..snd x} \<longrightarrow> (\<not> P {fst y..snd y} \<and> (\<forall>i. fst x$i \<le> fst y$i \<and> fst y$i \<le> snd y$i \<and> snd y$i \<le> snd x$i \<and>
+                           2 * (snd y$i - fst y$i) \<le> snd x$i - fst x$i))" proof case goal1 thus ?case proof-
+      presume "\<not> P {fst x..snd x} \<Longrightarrow> ?thesis"
+      thus ?thesis apply(cases "P {fst x..snd x}") by auto
+    next assume as:"\<not> P {fst x..snd x}" from interval_bisection_step[of P, OF assms(1-2) as] guess c d . 
+      thus ?thesis apply- apply(rule_tac x="(c,d)" in exI) by auto
+    qed qed then guess f apply-apply(drule choice) by(erule exE) note f=this
+  def AB \<equiv> "\<lambda>n. (f ^^ n) (a,b)" def A \<equiv> "\<lambda>n. fst(AB n)" and B \<equiv> "\<lambda>n. snd(AB n)" note ab_def = this AB_def
+  have "A 0 = a" "B 0 = b" "\<And>n. \<not> P {A(Suc n)..B(Suc n)} \<and>
+    (\<forall>i. A(n)$i \<le> A(Suc n)$i \<and> A(Suc n)$i \<le> B(Suc n)$i \<and> B(Suc n)$i \<le> B(n)$i \<and> 
+    2 * (B(Suc n)$i - A(Suc n)$i) \<le> B(n)$i - A(n)$i)" (is "\<And>n. ?P n")
+  proof- show "A 0 = a" "B 0 = b" unfolding ab_def by auto
+    case goal3 note S = ab_def funpow.simps o_def id_apply show ?case
+    proof(induct n) case 0 thus ?case unfolding S apply(rule f[rule_format]) using assms(3) by auto
+    next case (Suc n) show ?case unfolding S apply(rule f[rule_format]) using Suc unfolding S by auto
+    qed qed note AB = this(1-2) conjunctD2[OF this(3),rule_format]
+
+  have interv:"\<And>e. 0 < e \<Longrightarrow> \<exists>n. \<forall>x\<in>{A n..B n}. \<forall>y\<in>{A n..B n}. dist x y < e"
+  proof- case goal1 guess n using real_arch_pow2[of "(setsum (\<lambda>i. b$i - a$i) UNIV) / e"] .. note n=this
+    show ?case apply(rule_tac x=n in exI) proof(rule,rule)
+      fix x y assume xy:"x\<in>{A n..B n}" "y\<in>{A n..B n}"
+      have "dist x y \<le> setsum (\<lambda>i. abs((x - y)$i)) UNIV" unfolding vector_dist_norm by(rule norm_le_l1)
+      also have "\<dots> \<le> setsum (\<lambda>i. B n$i - A n$i) UNIV"
+      proof(rule setsum_mono) fix i show "\<bar>(x - y) $ i\<bar> \<le> B n $ i - A n $ i"
+          using xy[unfolded mem_interval,THEN spec[where x=i]]
+          unfolding vector_minus_component by auto qed
+      also have "\<dots> \<le> setsum (\<lambda>i. b$i - a$i) UNIV / 2^n" unfolding setsum_divide_distrib
+      proof(rule setsum_mono) case goal1 thus ?case
+        proof(induct n) case 0 thus ?case unfolding AB by auto
+        next case (Suc n) have "B (Suc n) $ i - A (Suc n) $ i \<le> (B n $ i - A n $ i) / 2" using AB(4)[of n i] by auto
+          also have "\<dots> \<le> (b $ i - a $ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case .
+        qed qed
+      also have "\<dots> < e" using n using goal1 by(auto simp add:field_simps) finally show "dist x y < e" .
+    qed qed
+  { fix n m ::nat assume "m \<le> n" then guess d unfolding le_Suc_ex_iff .. note d=this
+    have "{A n..B n} \<subseteq> {A m..B m}" unfolding d 
+    proof(induct d) case 0 thus ?case by auto
+    next case (Suc d) show ?case apply(rule subset_trans[OF _ Suc])
+        apply(rule) unfolding mem_interval apply(rule,erule_tac x=i in allE)
+      proof- case goal1 thus ?case using AB(4)[of "m + d" i] by(auto simp add:field_simps)
+      qed qed } note ABsubset = this 
+  have "\<exists>a. \<forall>n. a\<in>{A n..B n}" apply(rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv])
+  proof- fix n show "{A n..B n} \<noteq> {}" apply(cases "0<n") using AB(3)[of "n - 1"] assms(1,3) AB(1-2) by auto qed auto
+  then guess x0 .. note x0=this[rule_format]
+  show thesis proof(rule that[rule_format,of x0])
+    show "x0\<in>{a..b}" using x0[of 0] unfolding AB .
+    fix e assume "0 < (e::real)" from interv[OF this] guess n .. note n=this
+    show "\<exists>c d. x0 \<in> {c..d} \<and> {c..d} \<subseteq> ball x0 e \<and> {c..d} \<subseteq> {a..b} \<and> \<not> P {c..d}"
+      apply(rule_tac x="A n" in exI,rule_tac x="B n" in exI) apply(rule,rule x0) apply rule defer 
+    proof show "\<not> P {A n..B n}" apply(cases "0<n") using AB(3)[of "n - 1"] assms(3) AB(1-2) by auto
+      show "{A n..B n} \<subseteq> ball x0 e" using n using x0[of n] by auto
+      show "{A n..B n} \<subseteq> {a..b}" unfolding AB(1-2)[symmetric] apply(rule ABsubset) by auto
+    qed qed qed 
+
+subsection {* Cousin's lemma. *}
+
+lemma fine_division_exists: assumes "gauge g" 
+  obtains p where "p tagged_division_of {a..b::real^'n}" "g fine p"
+proof- presume "\<not> (\<exists>p. p tagged_division_of {a..b} \<and> g fine p) \<Longrightarrow> False"
+  then guess p unfolding atomize_not not_not .. thus thesis apply-apply(rule that[of p]) by auto
+next assume as:"\<not> (\<exists>p. p tagged_division_of {a..b} \<and> g fine p)"
+  guess x apply(rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p",rule_format,OF _ _ as])
+    apply(rule_tac x="{}" in exI) defer apply(erule conjE exE)+
+  proof- show "{} tagged_division_of {} \<and> g fine {}" unfolding fine_def by auto
+    fix s t p p' assume "p tagged_division_of s" "g fine p" "p' tagged_division_of t" "g fine p'" "interior s \<inter> interior t = {}"
+    thus "\<exists>p. p tagged_division_of s \<union> t \<and> g fine p" apply-apply(rule_tac x="p \<union> p'" in exI) apply rule
+      apply(rule tagged_division_union) prefer 4 apply(rule fine_union) by auto
+  qed note x=this
+  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)] guess c d apply-apply(erule exE conjE)+ . note c_d = this
+  have "g fine {(x, {c..d})}" unfolding fine_def using e using c_d(2) by auto
+  thus False using tagged_division_of_self[OF c_d(1)] using c_d by auto qed
+
+subsection {* Basic theorems about integrals. *}
+
+lemma has_integral_unique: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2"
+proof(rule ccontr) let ?e = "norm(k1 - k2) / 2" assume as:"k1 \<noteq> k2" hence e:"?e > 0" by auto
+  have lem:"\<And>f::real^'n \<Rightarrow> 'a.  \<And> a b k1 k2.
+    (f has_integral k1) ({a..b}) \<Longrightarrow> (f has_integral k2) ({a..b}) \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> False"
+  proof- case goal1 let ?e = "norm(k1 - k2) / 2" from goal1(3) have e:"?e > 0" by auto
+    guess d1 by(rule has_integralD[OF goal1(1) e]) note d1=this
+    guess d2 by(rule has_integralD[OF goal1(2) e]) note d2=this
+    guess p by(rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)],of a b]) note p=this
+    let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
+      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:group_simps norm_minus_commute)
+    also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
+      apply(rule add_strict_mono) apply(rule_tac[!] d2(2) d1(2)) using p unfolding fine_def by auto
+    finally show False by auto
+  qed { presume "\<not> (\<exists>a b. i = {a..b}) \<Longrightarrow> False"
+    thus False apply-apply(cases "\<exists>a b. i = {a..b}")
+      using assms by(auto simp add:has_integral intro:lem[OF _ _ as]) }
+  assume as:"\<not> (\<exists>a b. i = {a..b})"
+  guess B1 by(rule has_integral_altD[OF assms(1) as,OF e]) note B1=this[rule_format]
+  guess B2 by(rule has_integral_altD[OF assms(2) as,OF e]) note B2=this[rule_format]
+  have "\<exists>a b::real^'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> {a..b}" apply(rule bounded_subset_closed_interval)
+    using bounded_Un bounded_ball by auto then guess a b apply-by(erule exE)+
+  note ab=conjunctD2[OF this[unfolded Un_subset_iff]]
+  guess w using B1(2)[OF ab(1)] .. note w=conjunctD2[OF this]
+  guess z using B2(2)[OF ab(2)] .. note z=conjunctD2[OF this]
+  have "z = w" using lem[OF w(1) z(1)] by auto
+  hence "norm (k1 - k2) \<le> norm (z - k2) + norm (w - k1)"
+    using norm_triangle_ineq4[of "k1 - w" "k2 - z"] by(auto simp add: norm_minus_commute) 
+  also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" apply(rule add_strict_mono) by(rule_tac[!] z(2) w(2))
+  finally show False by auto qed
+
+lemma integral_unique[intro]:
+  "(f has_integral y) k \<Longrightarrow> integral k f = y"
+  unfolding integral_def apply(rule some_equality) by(auto intro: has_integral_unique) 
+
+lemma has_integral_is_0: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector" 
+  assumes "\<forall>x\<in>s. f x = 0" shows "(f has_integral 0) s"
+proof- have lem:"\<And>a b. \<And>f::real^'n \<Rightarrow> 'a.
+    (\<forall>x\<in>{a..b}. f(x) = 0) \<Longrightarrow> (f has_integral 0) ({a..b})" unfolding has_integral
+  proof(rule,rule) fix a b e and f::"real^'n \<Rightarrow> 'a"
+    assume as:"\<forall>x\<in>{a..b}. f x = 0" "0 < (e::real)"
+    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e)"
+      apply(rule_tac x="\<lambda>x. ball x 1" in exI)  apply(rule,rule gaugeI) unfolding centre_in_ball defer apply(rule open_ball)
+    proof(rule,rule,erule conjE) case goal1
+      have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) = 0" proof(rule setsum_0',rule)
+        fix x assume x:"x\<in>p" have "f (fst x) = 0" using tagged_division_ofD(2-3)[OF goal1(1), of "fst x" "snd x"] using as x by auto
+        thus "(\<lambda>(x, k). content k *\<^sub>R f x) x = 0" apply(subst surjective_pairing[of x]) unfolding split_conv by auto
+      qed thus ?case using as by auto
+    qed auto qed  { presume "\<not> (\<exists>a b. s = {a..b}) \<Longrightarrow> ?thesis"
+    thus ?thesis apply-apply(cases "\<exists>a b. s = {a..b}")
+      using assms by(auto simp add:has_integral intro:lem) }
+  have *:"(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. 0)" apply(rule ext) using assms by auto
+  assume "\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P *
+  apply(rule,rule,rule_tac x=1 in exI,rule) defer apply(rule,rule,rule)
+  proof- fix e::real and a b assume "e>0"
+    thus "\<exists>z. ((\<lambda>x::real^'n. 0::'a) has_integral z) {a..b} \<and> norm (z - 0) < e"
+      apply(rule_tac x=0 in exI) apply(rule,rule lem) by auto
+  qed auto qed
+
+lemma has_integral_0[simp]: "((\<lambda>x::real^'n. 0) has_integral 0) s"
+  apply(rule has_integral_is_0) by auto 
+
+lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) s \<longleftrightarrow> i = 0"
+  using has_integral_unique[OF has_integral_0] by auto
+
+lemma has_integral_linear: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral y) s" "bounded_linear h" shows "((h o f) has_integral ((h y))) s"
+proof- interpret bounded_linear h using assms(2) . from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format]
+  have lem:"\<And>f::real^'n \<Rightarrow> 'a. \<And> y a b.
+    (f has_integral y) ({a..b}) \<Longrightarrow> ((h o f) has_integral h(y)) ({a..b})"
+  proof(subst has_integral,rule,rule) case goal1
+    from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format]
+    have *:"e / B > 0" apply(rule divide_pos_pos) using goal1(2) B by auto
+    guess g using has_integralD[OF goal1(1) *] . note g=this
+    show ?case apply(rule_tac x=g in exI) apply(rule,rule g(1))
+    proof(rule,rule,erule conjE) fix p assume as:"p tagged_division_of {a..b}" "g fine p" 
+      have *:"\<And>x k. h ((\<lambda>(x, k). content k *\<^sub>R f x) x) = (\<lambda>(x, k). h (content k *\<^sub>R f x)) x" by auto
+      have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = setsum (h \<circ> (\<lambda>(x, k). content k *\<^sub>R f x)) p"
+        unfolding o_def unfolding scaleR[THEN sym] * by simp
+      also have "\<dots> = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" using setsum[of "\<lambda>(x,k). content k *\<^sub>R f x" p] using as by auto
+      finally have *:"(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" .
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) - h y) < e" unfolding * diff[THEN sym]
+        apply(rule le_less_trans[OF B(2)]) using g(2)[OF as] B(1) by(auto simp add:field_simps)
+    qed qed { presume "\<not> (\<exists>a b. s = {a..b}) \<Longrightarrow> ?thesis"
+    thus ?thesis apply-apply(cases "\<exists>a b. s = {a..b}") using assms by(auto simp add:has_integral intro!:lem) }
+  assume as:"\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P
+  proof(rule,rule) fix e::real  assume e:"0<e"
+    have *:"0 < e/B" by(rule divide_pos_pos,rule e,rule B(1))
+    guess M using has_integral_altD[OF assms(1) as *,rule_format] . note M=this
+    show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b} \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) has_integral z) {a..b} \<and> norm (z - h y) < e)"
+      apply(rule_tac x=M in exI) apply(rule,rule M(1))
+    proof(rule,rule,rule) case goal1 guess z using M(2)[OF goal1(1)] .. note z=conjunctD2[OF this]
+      have *:"(\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) = h \<circ> (\<lambda>x. if x \<in> s then f x else 0)"
+        unfolding o_def apply(rule ext) using zero by auto
+      show ?case apply(rule_tac x="h z" in exI,rule) unfolding * apply(rule lem[OF z(1)]) unfolding diff[THEN sym]
+        apply(rule le_less_trans[OF B(2)]) using B(1) z(2) by(auto simp add:field_simps)
+    qed qed qed
+
+lemma has_integral_cmul:
+  shows "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s"
+  unfolding o_def[THEN sym] apply(rule has_integral_linear,assumption)
+  by(rule scaleR.bounded_linear_right)
+
+lemma has_integral_neg:
+  shows "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral (-k)) s"
+  apply(drule_tac c="-1" in has_integral_cmul) by auto
+
+lemma has_integral_add: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector" 
+  assumes "(f has_integral k) s" "(g has_integral l) s"
+  shows "((\<lambda>x. f x + g x) has_integral (k + l)) s"
+proof- have lem:"\<And>f g::real^'n \<Rightarrow> 'a. \<And>a b k l.
+    (f has_integral k) ({a..b}) \<Longrightarrow> (g has_integral l) ({a..b}) \<Longrightarrow>
+     ((\<lambda>x. f(x) + g(x)) has_integral (k + l)) ({a..b})" proof- case goal1
+    show ?case unfolding has_integral proof(rule,rule) fix e::real assume e:"e>0" hence *:"e/2>0" by auto
+      guess d1 using has_integralD[OF goal1(1) *] . note d1=this
+      guess d2 using has_integralD[OF goal1(2) *] . note d2=this
+      show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)"
+        apply(rule_tac x="\<lambda>x. (d1 x) \<inter> (d2 x)" in exI) apply(rule,rule gauge_inter[OF d1(1) d2(1)])
+      proof(rule,rule,erule conjE) fix p assume as:"p tagged_division_of {a..b}" "(\<lambda>x. d1 x \<inter> d2 x) fine p"
+        have *:"(\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) = (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p. content k *\<^sub>R g x)"
+          unfolding scaleR_right_distrib setsum_addf[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,THEN sym]
+          by(rule setsum_cong2,auto)
+        have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) = norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
+          unfolding * by(auto simp add:group_simps) also let ?res = "\<dots>"
+        from as have *:"d1 fine p" "d2 fine p" unfolding fine_inter by auto
+        have "?res < e/2 + e/2" apply(rule le_less_trans[OF norm_triangle_ineq])
+          apply(rule add_strict_mono) using d1(2)[OF as(1) *(1)] and d2(2)[OF as(1) *(2)] by auto
+        finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e" by auto
+      qed qed qed { presume "\<not> (\<exists>a b. s = {a..b}) \<Longrightarrow> ?thesis"
+    thus ?thesis apply-apply(cases "\<exists>a b. s = {a..b}") using assms by(auto simp add:has_integral intro!:lem) }
+  assume as:"\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P
+  proof(rule,rule) case goal1 hence *:"e/2 > 0" by auto
+    from has_integral_altD[OF assms(1) as *] guess B1 . note B1=this[rule_format]
+    from has_integral_altD[OF assms(2) as *] guess B2 . note B2=this[rule_format]
+    show ?case apply(rule_tac x="max B1 B2" in exI) apply(rule,rule min_max.less_supI1,rule B1)
+    proof(rule,rule,rule) fix a b assume "ball 0 (max B1 B2) \<subseteq> {a..b::real^'n}"
+      hence *:"ball 0 B1 \<subseteq> {a..b::real^'n}" "ball 0 B2 \<subseteq> {a..b::real^'n}" by auto
+      guess w using B1(2)[OF *(1)] .. note w=conjunctD2[OF this]
+      guess z using B2(2)[OF *(2)] .. note z=conjunctD2[OF this]
+      have *:"\<And>x. (if x \<in> s then f x + g x else 0) = (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)" by auto
+      show "\<exists>z. ((\<lambda>x. if x \<in> s then f x + g x else 0) has_integral z) {a..b} \<and> norm (z - (k + l)) < e"
+        apply(rule_tac x="w + z" in exI) apply(rule,rule lem[OF w(1) z(1), unfolded *[THEN sym]])
+        using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) by(auto simp add:field_simps)
+    qed qed qed
+
+lemma has_integral_sub:
+  shows "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_integral (k - l)) s"
+  using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding group_simps by auto
+
+lemma integral_0: "integral s (\<lambda>x::real^'n. 0::real^'m) = 0"
+  by(rule integral_unique has_integral_0)+
+
+lemma integral_add:
+  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow>
+   integral s (\<lambda>x. f x + g x) = integral s f + integral s g"
+  apply(rule integral_unique) apply(drule integrable_integral)+
+  apply(rule has_integral_add) by assumption+
+
+lemma integral_cmul:
+  shows "f integrable_on s \<Longrightarrow> integral s (\<lambda>x. c *\<^sub>R f x) = c *\<^sub>R integral s f"
+  apply(rule integral_unique) apply(drule integrable_integral)+
+  apply(rule has_integral_cmul) by assumption+
+
+lemma integral_neg:
+  shows "f integrable_on s \<Longrightarrow> integral s (\<lambda>x. - f x) = - integral s f"
+  apply(rule integral_unique) apply(drule integrable_integral)+
+  apply(rule has_integral_neg) by assumption+
+
+lemma integral_sub:
+  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> integral s (\<lambda>x. f x - g x) = integral s f - integral s g"
+  apply(rule integral_unique) apply(drule integrable_integral)+
+  apply(rule has_integral_sub) by assumption+
+
+lemma integrable_0: "(\<lambda>x. 0) integrable_on s"
+  unfolding integrable_on_def using has_integral_0 by auto
+
+lemma integrable_add:
+  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x + g x) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_add)
+
+lemma integrable_cmul:
+  shows "f integrable_on s \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_cmul)
+
+lemma integrable_neg:
+  shows "f integrable_on s \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_neg)
+
+lemma integrable_sub:
+  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_sub)
+
+lemma integrable_linear:
+  shows "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h o f) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_linear)
+
+lemma integral_linear:
+  shows "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> integral s (h o f) = h(integral s f)"
+  apply(rule has_integral_unique) defer unfolding has_integral_integral 
+  apply(drule has_integral_linear,assumption,assumption) unfolding has_integral_integral[THEN sym]
+  apply(rule integrable_linear) by assumption+
+
+lemma has_integral_setsum:
+  assumes "finite t" "\<forall>a\<in>t. ((f a) has_integral (i a)) s"
+  shows "((\<lambda>x. setsum (\<lambda>a. f a x) t) has_integral (setsum i t)) s"
+proof(insert assms(1) subset_refl[of t],induct rule:finite_subset_induct)
+  case (insert x F) show ?case unfolding setsum_insert[OF insert(1,3)]
+    apply(rule has_integral_add) using insert assms by auto
+qed auto
+
+lemma integral_setsum:
+  shows "finite t \<Longrightarrow> \<forall>a\<in>t. (f a) integrable_on s \<Longrightarrow>
+  integral s (\<lambda>x. setsum (\<lambda>a. f a x) t) = setsum (\<lambda>a. integral s (f a)) t"
+  apply(rule integral_unique) apply(rule has_integral_setsum)
+  using integrable_integral by auto
+
+lemma integrable_setsum:
+  shows "finite t \<Longrightarrow> \<forall>a \<in> t.(f a) integrable_on s \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) t) integrable_on s"
+  unfolding integrable_on_def apply(drule bchoice) using has_integral_setsum[of t] by auto
+
+lemma has_integral_eq:
+  assumes "\<forall>x\<in>s. f x = g x" "(f has_integral k) s" shows "(g has_integral k) s"
+  using has_integral_sub[OF assms(2), of "\<lambda>x. f x - g x" 0]
+  using has_integral_is_0[of s "\<lambda>x. f x - g x"] using assms(1) by auto
+
+lemma integrable_eq:
+  shows "\<forall>x\<in>s. f x = g x \<Longrightarrow> f integrable_on s \<Longrightarrow> g integrable_on s"
+  unfolding integrable_on_def using has_integral_eq[of s f g] by auto
+
+lemma has_integral_eq_eq:
+  shows "\<forall>x\<in>s. f x = g x \<Longrightarrow> ((f has_integral k) s \<longleftrightarrow> (g has_integral k) s)"
+  using has_integral_eq[of s f g] has_integral_eq[of s g f] by auto 
+
+lemma has_integral_null[dest]:
+  assumes "content({a..b}) = 0" shows  "(f has_integral 0) ({a..b})"
+  unfolding has_integral apply(rule,rule,rule_tac x="\<lambda>x. ball x 1" in exI,rule) defer
+proof(rule,rule,erule conjE) fix e::real assume e:"e>0" thus "gauge (\<lambda>x. ball x 1)" by auto
+  fix p assume p:"p tagged_division_of {a..b}" (*"(\<lambda>x. ball x 1) fine p"*)
+  have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) = 0" unfolding norm_eq_zero diff_0_right
+    using setsum_content_null[OF assms(1) p, of f] . 
+  thus "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e" using e by auto qed
+
+lemma has_integral_null_eq[simp]:
+  shows "content({a..b}) = 0 \<Longrightarrow> ((f has_integral i) ({a..b}) \<longleftrightarrow> i = 0)"
+  apply rule apply(rule has_integral_unique,assumption) 
+  apply(drule has_integral_null,assumption)
+  apply(drule has_integral_null) by auto
+
+lemma integral_null[dest]: shows "content({a..b}) = 0 \<Longrightarrow> integral({a..b}) f = 0"
+  by(rule integral_unique,drule has_integral_null)
+
+lemma integrable_on_null[dest]: shows "content({a..b}) = 0 \<Longrightarrow> f integrable_on {a..b}"
+  unfolding integrable_on_def apply(drule has_integral_null) by auto
+
+lemma has_integral_empty[intro]: shows "(f has_integral 0) {}"
+  unfolding empty_as_interval apply(rule has_integral_null) 
+  using content_empty unfolding empty_as_interval .
+
+lemma has_integral_empty_eq[simp]: shows "(f has_integral i) {} \<longleftrightarrow> i = 0"
+  apply(rule,rule has_integral_unique,assumption) by auto
+
+lemma integrable_on_empty[intro]: shows "f integrable_on {}" unfolding integrable_on_def by auto
+
+lemma integral_empty[simp]: shows "integral {} f = 0"
+  apply(rule integral_unique) using has_integral_empty .
+
+lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}"
+  apply(rule has_integral_null) unfolding content_eq_0_interior
+  unfolding interior_closed_interval using interval_sing by auto
+
+lemma integrable_on_refl[intro]: shows "f integrable_on {a..a}" unfolding integrable_on_def by auto
+
+lemma integral_refl: shows "integral {a..a} f = 0" apply(rule integral_unique) by auto
+
+subsection {* Cauchy-type criterion for integrability. *}
+
+lemma integrable_cauchy: fixes f::"real^'n \<Rightarrow> 'a::{real_normed_vector,complete_space}" 
+  shows "f integrable_on {a..b} \<longleftrightarrow>
+  (\<forall>e>0.\<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<and> d fine p1 \<and>
+                            p2 tagged_division_of {a..b} \<and> d fine p2
+                            \<longrightarrow> norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 -
+                                     setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) < e))" (is "?l = (\<forall>e>0. \<exists>d. ?P e d)")
+proof assume ?l
+  then guess y unfolding integrable_on_def has_integral .. note y=this
+  show "\<forall>e>0. \<exists>d. ?P e d" proof(rule,rule) case goal1 hence "e/2 > 0" by auto
+    then guess d apply- apply(drule y[rule_format]) by(erule exE,erule conjE) note d=this[rule_format]
+    show ?case apply(rule_tac x=d in exI,rule,rule d) apply(rule,rule,rule,(erule conjE)+)
+    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b}" "d fine p1" "p2 tagged_division_of {a..b}" "d fine p2"
+      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
+        apply(rule dist_triangle_half_l[where y=y,unfolded vector_dist_norm])
+        using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] .
+    qed qed
+next assume "\<forall>e>0. \<exists>d. ?P e d" hence "\<forall>n::nat. \<exists>d. ?P (inverse(real (n + 1))) d" by auto
+  from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
+  have "\<And>n. gauge (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}})" apply(rule gauge_inters) using d(1) by auto
+  hence "\<forall>n. \<exists>p. p tagged_division_of {a..b} \<and> (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}}) fine p" apply-
+  proof case goal1 from this[of n] show ?case apply(drule_tac fine_division_exists) by auto qed
+  from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]]
+  have dp:"\<And>i n. i\<le>n \<Longrightarrow> d i fine p n" using p(2) unfolding fine_inters by auto
+  have "Cauchy (\<lambda>n. setsum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))"
+  proof(rule CauchyI) case goal1 then guess N unfolding real_arch_inv[of e] .. note N=this
+    show ?case apply(rule_tac x=N in exI)
+    proof(rule,rule,rule,rule) fix m n assume mn:"N \<le> m" "N \<le> n" have *:"N = (N - 1) + 1" using N by auto
+      show "norm ((\<Sum>(x, k)\<in>p m. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p n. content k *\<^sub>R f x)) < e"
+        apply(rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) apply(subst *) apply(rule d(2))
+        using dp p(1) using mn by auto 
+    qed qed
+  then guess y unfolding convergent_eq_cauchy[THEN sym] .. note y=this[unfolded Lim_sequentially,rule_format]
+  show ?l unfolding integrable_on_def has_integral apply(rule_tac x=y in exI)
+  proof(rule,rule) fix e::real assume "e>0" hence *:"e/2 > 0" by auto
+    then guess N1 unfolding real_arch_inv[of "e/2"] .. note N1=this hence N1':"N1 = N1 - 1 + 1" by auto
+    guess N2 using y[OF *] .. note N2=this
+    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e)"
+      apply(rule_tac x="d (N1 + N2)" in exI) apply rule defer 
+    proof(rule,rule,erule conjE) show "gauge (d (N1 + N2))" using d by auto
+      fix q assume as:"q tagged_division_of {a..b}" "d (N1 + N2) fine q"
+      have *:"inverse (real (N1 + N2 + 1)) < e / 2" apply(rule less_trans) using N1 by auto
+      show "norm ((\<Sum>(x, k)\<in>q. content k *\<^sub>R f x) - y) < e" apply(rule norm_triangle_half_r)
+        apply(rule less_trans[OF _ *]) apply(subst N1', rule d(2)[of "p (N1+N2)"]) defer
+        using N2[rule_format,unfolded vector_dist_norm,of "N1+N2"]
+        using as dp[of "N1 - 1 + 1 + N2" "N1 + N2"] using p(1)[of "N1 + N2"] using N1 by auto qed qed qed
+
+subsection {* Additivity of integral on abutting intervals. *}
+
+lemma interval_split:
+  "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
+  "{a..b} \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
+  apply(rule_tac[!] set_ext) unfolding Int_iff mem_interval mem_Collect_eq
+  unfolding Cart_lambda_beta by auto
+
+lemma content_split:
+  "content {a..b::real^'n} = content({a..b} \<inter> {x. x$k \<le> c}) + content({a..b} \<inter> {x. x$k >= c})"
+proof- note simps = interval_split content_closed_interval_cases Cart_lambda_beta vector_le_def
+  { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps by auto }
+  have *:"UNIV = insert k (UNIV - {k})" "\<And>x. finite (UNIV-{x::'n})" "\<And>x. x\<notin>UNIV-{x}" by auto
+  have *:"\<And>X Y Z. (\<Prod>i\<in>UNIV. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>UNIV-{k}. Z i (Y i))"
+    "(\<Prod>i\<in>UNIV. b$i - a$i) = (\<Prod>i\<in>UNIV-{k}. b$i - a$i) * (b$k - a$k)" 
+    apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto
+  assume as:"a\<le>b" moreover have "\<And>x. min (b $ k) c = max (a $ k) c
+    \<Longrightarrow> x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)"
+    by  (auto simp add:field_simps)
+  moreover have "\<not> a $ k \<le> c \<Longrightarrow> \<not> c \<le> b $ k \<Longrightarrow> False"
+    unfolding not_le using as[unfolded vector_le_def,rule_format,of k] by auto
+  ultimately show ?thesis 
+    unfolding simps unfolding *(1)[of "\<lambda>i x. b$i - x"] *(1)[of "\<lambda>i x. x - a$i"] *(2) by(auto)
+qed
+
+lemma division_split_left_inj:
+  assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2"
+  "k1 \<inter> {x::real^'n. x$k \<le> c} = k2 \<inter> {x. x$k \<le> c}"
+  shows "content(k1 \<inter> {x. x$k \<le> c}) = 0"
+proof- note d=division_ofD[OF assms(1)]
+  have *:"\<And>a b::real^'n. \<And> c k. (content({a..b} \<inter> {x. x$k \<le> c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$k \<le> c}) = {})"
+    unfolding interval_split content_eq_0_interior by auto
+  guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
+  guess u2 v2 using d(4)[OF assms(3)] apply-by(erule 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)]])
+    defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
+
+lemma division_split_right_inj:
+  assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2"
+  "k1 \<inter> {x::real^'n. x$k \<ge> c} = k2 \<inter> {x. x$k \<ge> c}"
+  shows "content(k1 \<inter> {x. x$k \<ge> c}) = 0"
+proof- note d=division_ofD[OF assms(1)]
+  have *:"\<And>a b::real^'n. \<And> c k. (content({a..b} \<inter> {x. x$k >= c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$k >= c}) = {})"
+    unfolding interval_split content_eq_0_interior by auto
+  guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
+  guess u2 v2 using d(4)[OF assms(3)] apply-by(erule 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)]])
+    defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
+
+lemma tagged_division_split_left_inj:
+  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$k \<le> c} = k2 \<inter> {x. x$k \<le> c}" 
+  shows "content(k1 \<inter> {x. x$k \<le> c}) = 0"
+proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
+  show ?thesis apply(rule division_split_left_inj[OF division_of_tagged_division[OF assms(1)]])
+    apply(rule_tac[1-2] *) using assms(2-) by auto qed
+
+lemma tagged_division_split_right_inj:
+  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$k \<ge> c} = k2 \<inter> {x. x$k \<ge> c}" 
+  shows "content(k1 \<inter> {x. x$k \<ge> c}) = 0"
+proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
+  show ?thesis apply(rule division_split_right_inj[OF division_of_tagged_division[OF assms(1)]])
+    apply(rule_tac[1-2] *) using assms(2-) by auto qed
+
+lemma division_split:
+  assumes "p division_of {a..b::real^'n}"
+  shows "{l \<inter> {x. x$k \<le> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$k \<le> c} = {})} division_of ({a..b} \<inter> {x. x$k \<le> c})" (is "?p1 division_of ?I1") and 
+        "{l \<inter> {x. x$k \<ge> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$k \<ge> c} = {})} division_of ({a..b} \<inter> {x. x$k \<ge> c})" (is "?p2 division_of ?I2")
+proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms]
+  show "finite ?p1" "finite ?p2" using p(1) by auto show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2" unfolding p(6)[THEN sym] by auto
+  { fix k assume "k\<in>?p1" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this
+    guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this
+    show "k\<subseteq>?I1" "k \<noteq> {}" "\<exists>a b. k = {a..b}" unfolding l
+      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto
+    fix k' assume "k'\<in>?p1" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this
+    assume "k\<noteq>k'" thus "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 apply-by(erule exE,(erule conjE)+) note l=this
+    guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this
+    show "k\<subseteq>?I2" "k \<noteq> {}" "\<exists>a b. k = {a..b}" unfolding l
+      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto
+    fix k' assume "k'\<in>?p2" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this
+    assume "k\<noteq>k'" thus "interior k \<inter> interior k' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto }
+qed
+
+lemma has_integral_split: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral i) ({a..b} \<inter> {x. x$k \<le> c})"  "(f has_integral j) ({a..b} \<inter> {x. x$k \<ge> c})"
+  shows "(f has_integral (i + j)) ({a..b})"
+proof(unfold has_integral,rule,rule) case goal1 hence e:"e/2>0" by auto
+  guess d1 using has_integralD[OF assms(1)[unfolded interval_split] e] . note d1=this[unfolded interval_split[THEN sym]]
+  guess d2 using has_integralD[OF assms(2)[unfolded interval_split] e] . note d2=this[unfolded interval_split[THEN sym]]
+  let ?d = "\<lambda>x. if x$k = c then (d1 x \<inter> d2 x) else ball x (abs(x$k - c)) \<inter> d1 x \<inter> d2 x"
+  show ?case apply(rule_tac x="?d" in exI,rule) defer apply(rule,rule,(erule conjE)+)
+  proof- show "gauge ?d" using d1(1) d2(1) unfolding gauge_def by auto
+    fix p assume "p tagged_division_of {a..b}" "?d fine p" note p = this tagged_division_ofD[OF this(1)]
+    have lem0:"\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$k \<le> c} = {}) \<Longrightarrow> x$k \<le> c"
+         "\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$k \<ge> c} = {}) \<Longrightarrow> x$k \<ge> c"
+    proof- fix x kk assume as:"(x,kk)\<in>p"
+      show "~(kk \<inter> {x. x$k \<le> c} = {}) \<Longrightarrow> x$k \<le> c"
+      proof(rule ccontr) case goal1
+        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $ k - c\<bar>"
+          using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
+        hence "\<exists>y. y \<in> ball x \<bar>x $ k - c\<bar> \<inter> {x. x $ k \<le> c}" using goal1(1) by blast 
+        then guess y .. hence "\<bar>x $ k - y $ k\<bar> < \<bar>x $ k - c\<bar>" "y$k \<le> c" apply-apply(rule le_less_trans)
+          using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:vector_dist_norm)
+        thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
+      qed
+      show "~(kk \<inter> {x. x$k \<ge> c} = {}) \<Longrightarrow> x$k \<ge> c"
+      proof(rule ccontr) case goal1
+        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $ k - c\<bar>"
+          using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
+        hence "\<exists>y. y \<in> ball x \<bar>x $ k - c\<bar> \<inter> {x. x $ k \<ge> c}" using goal1(1) by blast 
+        then guess y .. hence "\<bar>x $ k - y $ k\<bar> < \<bar>x $ k - c\<bar>" "y$k \<ge> c" apply-apply(rule le_less_trans)
+          using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:vector_dist_norm)
+        thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
+      qed
+    qed
+
+    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 lem2: "\<And>f s P f. finite s \<Longrightarrow> finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}"
+    proof- case goal1 thus ?case apply-apply(rule finite_subset[of _ "(\<lambda>(x,k). (x,f k)) ` s"]) by auto qed
+    have lem3: "\<And>g::(real ^ 'n \<Rightarrow> bool) \<Rightarrow> real ^ 'n \<Rightarrow> bool. finite p \<Longrightarrow>
+      setsum (\<lambda>(x,k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> ~(g k = {})}
+               = setsum (\<lambda>(x,k). content k *\<^sub>R f x) ((\<lambda>(x,k). (x,g k)) ` p)"
+      apply(rule setsum_mono_zero_left) prefer 3
+    proof fix g::"(real ^ 'n \<Rightarrow> bool) \<Rightarrow> real ^ 'n \<Rightarrow> bool" and i::"(real^'n) \<times> ((real^'n) set)"
+      assume "i \<in> (\<lambda>(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
+      then obtain x k where xk:"i=(x,g k)" "(x,k)\<in>p" "(x,g k) \<notin> {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}" by auto
+      have "content (g k) = 0" using xk using content_empty by auto
+      thus "(\<lambda>(x, k). content k *\<^sub>R f x) i = 0" unfolding xk split_conv by auto
+    qed auto
+    have lem4:"\<And>g. (\<lambda>(x,l). content (g l) *\<^sub>R f x) = (\<lambda>(x,l). content l *\<^sub>R f x) o (\<lambda>(x,l). (x,g l))" apply(rule ext) by auto
+
+    let ?M1 = "{(x,kk \<inter> {x. x$k \<le> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$k \<le> c} \<noteq> {}}"
+    have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2" apply(rule d1(2),rule tagged_division_ofI)
+      apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
+    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = {a..b} \<inter> {x. x$k \<le> c}" unfolding p(8)[THEN sym] by auto
+      fix x l assume xl:"(x,l)\<in>?M1"
+      then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note xl'=this
+      have "l' \<subseteq> d1 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
+      thus "l \<subseteq> d1 x" unfolding xl' by auto
+      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $ k \<le> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
+        using lem0(1)[OF xl'(3-4)] by auto
+      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k])
+      fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M1"
+      then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note yr'=this
+      assume as:"(x,l) \<noteq> (y,r)" show "interior l \<inter> interior r = {}"
+      proof(cases "l' = r' \<longrightarrow> x' = y'")
+        case False thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      next case True hence "l' \<noteq> r'" using as unfolding xl' yr' by auto
+        thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      qed qed moreover
+
+    let ?M2 = "{(x,kk \<inter> {x. x$k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$k \<ge> c} \<noteq> {}}" 
+    have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2" apply(rule d2(2),rule tagged_division_ofI)
+      apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
+    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = {a..b} \<inter> {x. x$k \<ge> c}" unfolding p(8)[THEN sym] by auto
+      fix x l assume xl:"(x,l)\<in>?M2"
+      then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note xl'=this
+      have "l' \<subseteq> d2 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
+      thus "l \<subseteq> d2 x" unfolding xl' by auto
+      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $ k \<ge> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
+        using lem0(2)[OF xl'(3-4)] by auto
+      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k])
+      fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M2"
+      then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note yr'=this
+      assume as:"(x,l) \<noteq> (y,r)" show "interior l \<inter> interior r = {}"
+      proof(cases "l' = r' \<longrightarrow> x' = y'")
+        case False thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      next case True hence "l' \<noteq> r'" using as unfolding xl' yr' by auto
+        thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      qed qed ultimately
+
+    have "norm (((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2"
+      apply- apply(rule norm_triangle_lt) by auto
+    also { have *:"\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'a) = 0" using scaleR_zero_left by auto
+      have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)
+       = (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)" by auto
+      also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x $ k \<le> c}) *\<^sub>R f x) + (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x $ k}) *\<^sub>R f x) - (i + j)"
+        unfolding lem3[OF p(3)] apply(subst setsum_reindex_nonzero[OF p(3)]) defer apply(subst setsum_reindex_nonzero[OF p(3)])
+        defer unfolding lem4[THEN sym] apply(rule refl) unfolding split_paired_all split_conv apply(rule_tac[!] *)
+      proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) by auto
+      next case   goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) by auto
+      qed also note setsum_addf[THEN sym]
+      also have *:"\<And>x. x\<in>p \<Longrightarrow> (\<lambda>(x, ka). content (ka \<inter> {x. x $ k \<le> c}) *\<^sub>R f x) x + (\<lambda>(x, ka). content (ka \<inter> {x. c \<le> x $ k}) *\<^sub>R f x) x
+        = (\<lambda>(x,ka). content ka *\<^sub>R f x) x" unfolding split_paired_all split_conv
+      proof- fix a b assume "(a,b) \<in> p" from p(6)[OF this] guess u v apply-by(erule exE)+ note uv=this
+        thus "content (b \<inter> {x. x $ k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x $ k}) *\<^sub>R f a = content b *\<^sub>R f a"
+          unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[of u v k c] by auto
+      qed note setsum_cong2[OF this]
+      finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x $ k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x $ k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
+        ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x $ k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x $ k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
+        (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)" by auto }
+    finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed
+
+subsection {* A sort of converse, integrability on subintervals. *}
+
+lemma tagged_division_union_interval:
+  assumes "p1 tagged_division_of ({a..b} \<inter> {x::real^'n. x$k \<le> (c::real)})"  "p2 tagged_division_of ({a..b} \<inter> {x. x$k \<ge> c})"
+  shows "(p1 \<union> p2) tagged_division_of ({a..b})"
+proof- have *:"{a..b} = ({a..b} \<inter> {x. x$k \<le> c}) \<union> ({a..b} \<inter> {x. x$k \<ge> c})" by auto
+  show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms])
+    unfolding interval_split interior_closed_interval
+    by(auto simp add: vector_less_def Cart_lambda_beta elim!:allE[where x=k]) qed
+
+lemma has_integral_separate_sides: fixes f::"real^'m \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral i) ({a..b})" "e>0"
+  obtains d where "gauge d" "(\<forall>p1 p2. p1 tagged_division_of ({a..b} \<inter> {x. x$k \<le> c}) \<and> d fine p1 \<and>
+                                p2 tagged_division_of ({a..b} \<inter> {x. x$k \<ge> c}) \<and> d fine p2
+                                \<longrightarrow> norm((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 +
+                                          setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e)"
+proof- guess d using has_integralD[OF assms] . note d=this
+  show ?thesis apply(rule that[of d]) apply(rule d) apply(rule,rule,rule,(erule conjE)+)
+  proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this
+                   assume "p2 tagged_division_of {a..b} \<inter> {x. c \<le> x $ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this
+    note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this
+    have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) = norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
+      apply(subst setsum_Un_zero) apply(rule p1 p2)+ apply(rule) unfolding split_paired_all split_conv
+    proof- fix a b assume ab:"(a,b) \<in> p1 \<inter> p2"
+      have "(a,b) \<in> p1" using ab by auto from p1(4)[OF this] guess u v apply-by(erule exE)+ note uv =this
+      have "b \<subseteq> {x. x$k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastsimp
+      moreover have "interior {x. x $ k = c} = {}" 
+      proof(rule ccontr) case goal1 then obtain x where x:"x\<in>interior {x. x$k = c}" by auto
+        then guess e unfolding mem_interior .. note e=this
+        have x:"x$k = c" using x interior_subset by fastsimp
+        have *:"\<And>i. \<bar>(x - (x + (\<chi> i. if i = k then e / 2 else 0))) $ i\<bar> = (if i = k then e/2 else 0)" using e by auto
+        have "x + (\<chi> i. if i = k then e/2 else 0) \<in> ball x e" unfolding mem_ball vector_dist_norm 
+          apply(rule le_less_trans[OF norm_le_l1]) unfolding * 
+          unfolding setsum_delta[OF finite_UNIV] using e by auto 
+        hence "x + (\<chi> i. if i = k then e/2 else 0) \<in> {x. x$k = c}" using e by auto
+        thus False unfolding mem_Collect_eq using e x by auto
+      qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule subset_interior) by auto
+      thus "content b *\<^sub>R f a = 0" by auto
+    qed auto
+    also have "\<dots> < e" by(rule d(2) p12 fine_union p1 p2)+
+    finally show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" . qed qed
+
+lemma integrable_split[intro]: fixes f::"real^'n \<Rightarrow> 'a::{real_normed_vector,complete_space}" assumes "f integrable_on {a..b}"
+  shows "f integrable_on ({a..b} \<inter> {x. x$k \<le> c})" (is ?t1) and "f integrable_on ({a..b} \<inter> {x. x$k \<ge> c})" (is ?t2) 
+proof- guess y using assms unfolding integrable_on_def .. note y=this
+  def b' \<equiv> "(\<chi> i. if i = k then min (b$k) c else b$i)::real^'n"
+  and a' \<equiv> "(\<chi> i. if i = k then max (a$k) c else a$i)::real^'n"
+  show ?t1 ?t2 unfolding interval_split integrable_cauchy unfolding interval_split[THEN sym]
+  proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "e/2>0" by auto
+    from has_integral_separate_sides[OF y this,of k c] guess d . note d=this[rule_format]
+    let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<inter> A \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> A \<and> d fine p2 \<longrightarrow>
+                              norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
+    show "?P {x. x $ k \<le> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
+    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c} \<and> d fine p2"
+      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
+      proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
+        show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
+          using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
+          using p using assms by(auto simp add:group_simps)
+      qed qed  
+    show "?P {x. x $ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
+    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p2"
+      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
+      proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
+        show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
+          using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
+          using p using assms by(auto simp add:group_simps) qed qed qed qed
+
+subsection {* Generalized notion of additivity. *}
+
+definition "neutral opp = (SOME x. \<forall>y. opp x y = y \<and> opp y x = y)"
+
+definition operative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ((real^'n) set \<Rightarrow> 'a) \<Rightarrow> bool" where
+  "operative opp f \<equiv> 
+    (\<forall>a b. content {a..b} = 0 \<longrightarrow> f {a..b} = neutral(opp)) \<and>
+    (\<forall>a b c k. f({a..b}) =
+                   opp (f({a..b} \<inter> {x. x$k \<le> c}))
+                       (f({a..b} \<inter> {x. x$k \<ge> c})))"
+
+lemma operativeD[dest]: assumes "operative opp f"
+  shows "\<And>a b. content {a..b} = 0 \<Longrightarrow> f {a..b} = neutral(opp)"
+  "\<And>a b c k. f({a..b}) = opp (f({a..b} \<inter> {x. x$k \<le> c})) (f({a..b} \<inter> {x. x$k \<ge> c}))"
+  using assms unfolding operative_def by auto
+
+lemma operative_trivial:
+ "operative opp f \<Longrightarrow> content({a..b}) = 0 \<Longrightarrow> f({a..b}) = neutral opp"
+  unfolding operative_def by auto
+
+lemma property_empty_interval:
+ "(\<forall>a b. content({a..b}) = 0 \<longrightarrow> P({a..b})) \<Longrightarrow> P {}" 
+  using content_empty unfolding empty_as_interval by auto
+
+lemma operative_empty: "operative opp f \<Longrightarrow> f {} = neutral opp"
+  unfolding operative_def apply(rule property_empty_interval) by auto
+
+subsection {* Using additivity of lifted function to encode definedness. *}
+
+lemma forall_option: "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>x. P(Some x))"
+  by (metis map_of.simps option.nchotomy)
+
+lemma exists_option:
+ "(\<exists>x. P x) \<longleftrightarrow> P None \<or> (\<exists>x. P(Some x))" 
+  by (metis map_of.simps option.nchotomy)
+
+fun lifted where 
+  "lifted (opp::'a\<Rightarrow>'a\<Rightarrow>'b) (Some x) (Some y) = Some(opp x y)" |
+  "lifted opp None _ = (None::'b option)" |
+  "lifted opp _ None = None"
+
+lemma lifted_simp_1[simp]: "lifted opp v None = None"
+  apply(induct v) by auto
+
+definition "monoidal opp \<equiv>  (\<forall>x y. opp x y = opp y x) \<and>
+                   (\<forall>x y z. opp x (opp y z) = opp (opp x y) z) \<and>
+                   (\<forall>x. opp (neutral opp) x = x)"
+
+lemma monoidalI: assumes "\<And>x y. opp x y = opp y x"
+  "\<And>x y z. opp x (opp y z) = opp (opp x y) z"
+  "\<And>x. opp (neutral opp) x = x" shows "monoidal opp"
+  unfolding monoidal_def using assms by fastsimp
+
+lemma monoidal_ac: assumes "monoidal opp"
+  shows "opp (neutral opp) a = a" "opp a (neutral opp) = a" "opp a b = opp b a"
+  "opp (opp a b) c = opp a (opp b c)"  "opp a (opp b c) = opp b (opp a c)"
+  using assms unfolding monoidal_def apply- by metis+
+
+lemma monoidal_simps[simp]: assumes "monoidal opp"
+  shows "opp (neutral opp) a = a" "opp a (neutral opp) = a"
+  using monoidal_ac[OF assms] by auto
+
+lemma neutral_lifted[cong]: assumes "monoidal opp"
+  shows "neutral (lifted opp) = Some(neutral opp)"
+  apply(subst neutral_def) apply(rule some_equality) apply(rule,induct_tac y) prefer 3
+proof- fix x assume "\<forall>y. lifted opp x y = y \<and> lifted opp y x = y"
+  thus "x = Some (neutral opp)" apply(induct x) defer
+    apply rule apply(subst neutral_def) apply(subst eq_commute,rule some_equality)
+    apply(rule,erule_tac x="Some y" in allE) defer apply(erule_tac x="Some x" in allE) by auto
+qed(auto simp add:monoidal_ac[OF assms])
+
+lemma monoidal_lifted[intro]: assumes "monoidal opp" shows "monoidal(lifted opp)"
+  unfolding monoidal_def forall_option neutral_lifted[OF assms] using monoidal_ac[OF assms] by auto
+
+definition "support opp f s = {x. x\<in>s \<and> f x \<noteq> neutral opp}"
+definition "fold' opp e s \<equiv> (if finite s then fold opp e s else e)"
+definition "iterate opp s f \<equiv> fold' (\<lambda>x a. opp (f x) a) (neutral opp) (support opp f s)"
+
+lemma support_subset[intro]:"support opp f s \<subseteq> s" unfolding support_def by auto
+lemma support_empty[simp]:"support opp f {} = {}" using support_subset[of opp f "{}"] by auto
+
+lemma fun_left_comm_monoidal[intro]: assumes "monoidal opp" shows "fun_left_comm opp"
+  unfolding fun_left_comm_def using monoidal_ac[OF assms] by auto
+
+lemma support_clauses:
+  "\<And>f g s. support opp f {} = {}"
+  "\<And>f g s. support opp f (insert x s) = (if f(x) = neutral opp then support opp f s else insert x (support opp f s))"
+  "\<And>f g s. support opp f (s - {x}) = (support opp f s) - {x}"
+  "\<And>f g s. support opp f (s \<union> t) = (support opp f s) \<union> (support opp f t)"
+  "\<And>f g s. support opp f (s \<inter> t) = (support opp f s) \<inter> (support opp f t)"
+  "\<And>f g s. support opp f (s - t) = (support opp f s) - (support opp f t)"
+  "\<And>f g s. support opp g (f ` s) = f ` (support opp (g o f) s)"
+unfolding support_def by auto
+
+lemma finite_support[intro]:"finite s \<Longrightarrow> finite (support opp f s)"
+  unfolding support_def by auto
+
+lemma iterate_empty[simp]:"iterate opp {} f = neutral opp"
+  unfolding iterate_def fold'_def by auto 
+
+lemma iterate_insert[simp]: assumes "monoidal opp" "finite s"
+  shows "iterate opp (insert x s) f = (if x \<in> s then iterate opp s f else opp (f x) (iterate opp s f))" 
+proof(cases "x\<in>s") case True hence *:"insert x s = s" by auto
+  show ?thesis unfolding iterate_def if_P[OF True] * by auto
+next case False note x=this
+  note * = fun_left_comm.fun_left_comm_apply[OF fun_left_comm_monoidal[OF assms(1)]]
+  show ?thesis proof(cases "f x = neutral opp")
+    case True show ?thesis unfolding iterate_def if_not_P[OF x] support_clauses if_P[OF True]
+      unfolding True monoidal_simps[OF assms(1)] by auto
+  next case False show ?thesis unfolding iterate_def fold'_def  if_not_P[OF x] support_clauses if_not_P[OF False]
+      apply(subst fun_left_comm.fold_insert[OF * finite_support])
+      using `finite s` unfolding support_def using False x by auto qed qed 
+
+lemma iterate_some:
+  assumes "monoidal opp"  "finite s"
+  shows "iterate (lifted opp) s (\<lambda>x. Some(f x)) = Some (iterate opp s f)" using assms(2)
+proof(induct s) case empty thus ?case using assms by auto
+next case (insert x F) show ?case apply(subst iterate_insert) prefer 3 apply(subst if_not_P)
+    defer unfolding insert(3) lifted.simps apply rule using assms insert by auto qed
+
+subsection {* Two key instances of additivity. *}
+
+lemma neutral_add[simp]:
+  "neutral op + = (0::_::comm_monoid_add)" unfolding neutral_def 
+  apply(rule some_equality) defer apply(erule_tac x=0 in allE) by auto
+
+lemma operative_content[intro]: "operative (op +) content"
+  unfolding operative_def content_split[THEN sym] neutral_add by auto
+
+lemma neutral_monoid[simp]: "neutral ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a) = 0"
+  unfolding neutral_def apply(rule some_equality) defer
+  apply(erule_tac x=0 in allE) by auto
+
+lemma monoidal_monoid[intro]:
+  shows "monoidal ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a)"
+  unfolding monoidal_def neutral_monoid by(auto simp add: group_simps) 
+
+lemma operative_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
+  unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add
+  apply(rule,rule,rule,rule) defer apply(rule allI)+
+proof- fix a b c k show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) =
+              lifted op + (if f integrable_on {a..b} \<inter> {x. x $ k \<le> c} then Some (integral ({a..b} \<inter> {x. x $ k \<le> c}) f) else None)
+               (if f integrable_on {a..b} \<inter> {x. c \<le> x $ k} then Some (integral ({a..b} \<inter> {x. c \<le> x $ k}) f) else None)"
+  proof(cases "f integrable_on {a..b}") 
+    case True show ?thesis unfolding if_P[OF True]
+      unfolding if_P[OF integrable_split(1)[OF True]] if_P[OF integrable_split(2)[OF True]]
+      unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split) 
+      apply(rule_tac[!] integrable_integral integrable_split)+ using True by assumption+
+  next case False have "(\<not> (f integrable_on {a..b} \<inter> {x. x $ k \<le> c})) \<or> (\<not> ( f integrable_on {a..b} \<inter> {x. c \<le> x $ k}))"
+    proof(rule ccontr) case goal1 hence "f integrable_on {a..b}" apply- unfolding integrable_on_def
+        apply(rule_tac x="integral ({a..b} \<inter> {x. x $ k \<le> c}) f + integral ({a..b} \<inter> {x. x $ k \<ge> c}) f" in exI)
+        apply(rule has_integral_split) apply(rule_tac[!] integrable_integral) by auto
+      thus False using False by auto
+    qed thus ?thesis using False by auto 
+  qed next 
+  fix a b assume as:"content {a..b::real^'n} = 0"
+  thus "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = Some 0"
+    unfolding if_P[OF integrable_on_null[OF as]] using has_integral_null_eq[OF as] by auto qed
+
+subsection {* Points of division of a partition. *}
+
+definition "division_points (k::(real^'n) set) d = 
+    {(j,x). (interval_lowerbound k)$j < x \<and> x < (interval_upperbound k)$j \<and>
+           (\<exists>i\<in>d. (interval_lowerbound i)$j = x \<or> (interval_upperbound i)$j = x)}"
+
+lemma division_points_finite: 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)$j < x \<and> x < (interval_upperbound i)$j \<and>
+           (\<exists>i\<in>d. (interval_lowerbound i)$j = x \<or> (interval_upperbound i)$j = x)}"
+  have *:"division_points i d = \<Union>(?M ` UNIV)"
+    unfolding division_points_def by auto
+  show ?thesis unfolding * using assm by auto qed
+
+lemma division_points_subset:
+  assumes "d division_of {a..b}" "\<forall>i. a$i < b$i"  "a$k < c" "c < b$k"
+  shows "division_points ({a..b} \<inter> {x. x$k \<le> c}) {l \<inter> {x. x$k \<le> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$k \<le> c} = {})}
+                  \<subseteq> division_points ({a..b}) d" (is ?t1) and
+        "division_points ({a..b} \<inter> {x. x$k \<ge> c}) {l \<inter> {x. x$k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$k \<ge> c} = {})}
+                  \<subseteq> division_points ({a..b}) d" (is ?t2)
+proof- note assm = division_ofD[OF assms(1)]
+  have *:"\<forall>i. a$i \<le> b$i"   "\<forall>i. a$i \<le> (\<chi> i. if i = k then min (b $ k) c else b $ i) $ i"
+    "\<forall>i. (\<chi> i. if i = k then max (a $ k) c else a $ i) $ i \<le> b$i"  "min (b $ k) c = c" "max (a $ k) c = c"
+    using assms using less_imp_le by auto
+  show ?t1 unfolding division_points_def interval_split[of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding *
+    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq apply(erule exE conjE)+
+  proof- fix i l x assume as:"a $ fst x < snd x" "snd x < (if fst x = k then c else b $ fst x)"
+      "interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x"  "i = l \<inter> {x. x $ k \<le> c}" "l \<in> d" "l \<inter> {x. x $ k \<le> c} \<noteq> {}"
+    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
+    have *:"\<forall>i. u $ i \<le> (\<chi> i. if i = k then min (v $ k) c else v $ i) $ i" using as(6) unfolding l interval_split interval_ne_empty as .
+    have **:"\<forall>i. u$i \<le> v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
+    show "a $ fst x < snd x \<and> snd x < b $ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x)"
+      using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply-
+      apply(rule,assumption,rule) defer apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
+      apply(case_tac[!] "fst x = k") using assms by auto
+  qed
+  show ?t2 unfolding division_points_def interval_split[of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding *
+    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq apply(erule exE conjE)+
+  proof- fix i l x assume as:"(if fst x = k then c else a $ fst x) < snd x" "snd x < b $ fst x" "interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x"
+      "i = l \<inter> {x. c \<le> x $ k}" "l \<in> d" "l \<inter> {x. c \<le> x $ k} \<noteq> {}"
+    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
+    have *:"\<forall>i. (\<chi> i. if i = k then max (u $ k) c else u $ i) $ i \<le> v $ i" using as(6) unfolding l interval_split interval_ne_empty as .
+    have **:"\<forall>i. u$i \<le> v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
+    show "a $ fst x < snd x \<and> snd x < b $ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x)"
+      using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply-
+      apply rule defer apply(rule,assumption) apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
+      apply(case_tac[!] "fst x = k") using assms by auto qed qed
+
+lemma division_points_psubset:
+  assumes "d division_of {a..b}"  "\<forall>i. a$i < b$i"  "a$k < c" "c < b$k"
+  "l \<in> d" "interval_lowerbound l$k = c \<or> interval_upperbound l$k = c"
+  shows "division_points ({a..b} \<inter> {x. x$k \<le> c}) {l \<inter> {x. x$k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x$k \<le> c} \<noteq> {}} \<subset> division_points ({a..b}) d" (is "?D1 \<subset> ?D") 
+        "division_points ({a..b} \<inter> {x. x$k \<ge> c}) {l \<inter> {x. x$k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x$k \<ge> c} \<noteq> {}} \<subset> division_points ({a..b}) d" (is "?D2 \<subset> ?D") 
+proof- have ab:"\<forall>i. a$i \<le> b$i" using assms(2) by(auto intro!:less_imp_le)
+  guess u v using division_ofD(4)[OF assms(1,5)] apply-by(erule exE)+ note l=this
+  have uv:"\<forall>i. u$i \<le> v$i" "\<forall>i. a$i \<le> u$i \<and> v$i \<le> b$i" using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty
+    unfolding subset_eq apply- defer apply(erule_tac x=u in ballE, erule_tac x=v in ballE) unfolding mem_interval by auto
+  have *:"interval_upperbound ({a..b} \<inter> {x. x $ k \<le> interval_upperbound l $ k}) $ k = interval_upperbound l $ k"
+         "interval_upperbound ({a..b} \<inter> {x. x $ k \<le> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
+    unfolding interval_split apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
+    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab by auto
+  have "\<exists>x. x \<in> ?D - ?D1" using assms(2-) apply-apply(erule disjE)
+    apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
+    apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
+    unfolding division_points_def unfolding interval_bounds[OF ab]
+    apply (auto simp add:interval_bounds) unfolding * by auto
+  thus "?D1 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto
+
+  have *:"interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
+         "interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_upperbound l $ k}) $ k = interval_upperbound l $ k"
+    unfolding interval_split apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
+    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab by auto
+  have "\<exists>x. x \<in> ?D - ?D2" using assms(2-) apply-apply(erule disjE)
+    apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
+    apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
+    unfolding division_points_def unfolding interval_bounds[OF ab]
+    apply (auto simp add:interval_bounds) unfolding * by auto
+  thus "?D2 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto qed
+
+subsection {* Preservation by divisions and tagged divisions. *}
+
+lemma support_support[simp]:"support opp f (support opp f s) = support opp f s"
+  unfolding support_def by auto
+
+lemma iterate_support[simp]: "iterate opp (support opp f s) f = iterate opp s f"
+  unfolding iterate_def support_support by auto
+
+lemma iterate_expand_cases:
+  "iterate opp s f = (if finite(support opp f s) then iterate opp (support opp f s) f else neutral opp)"
+  apply(cases) apply(subst if_P,assumption) unfolding iterate_def support_support fold'_def by auto 
+
+lemma iterate_image: assumes "monoidal opp"  "inj_on f s"
+  shows "iterate opp (f ` s) g = iterate opp s (g \<circ> f)"
+proof- have *:"\<And>s. finite s \<Longrightarrow>  \<forall>x\<in>s. \<forall>y\<in>s. f x = f y \<longrightarrow> x = y \<Longrightarrow>
+     iterate opp (f ` s) g = iterate opp s (g \<circ> f)"
+  proof- case goal1 show ?case using goal1
+    proof(induct s) case empty thus ?case using assms(1) by auto
+    next case (insert x s) show ?case unfolding iterate_insert[OF assms(1) insert(1)]
+        unfolding if_not_P[OF insert(2)] apply(subst insert(3)[THEN sym])
+        unfolding image_insert defer apply(subst iterate_insert[OF assms(1)])
+        apply(rule finite_imageI insert)+ apply(subst if_not_P)
+        unfolding image_iff o_def using insert(2,4) by auto
+    qed qed
+  show ?thesis 
+    apply(cases "finite (support opp g (f ` s))")
+    apply(subst (1) iterate_support[THEN sym],subst (2) iterate_support[THEN sym])
+    unfolding support_clauses apply(rule *)apply(rule finite_imageD,assumption) unfolding inj_on_def[symmetric]
+    apply(rule subset_inj_on[OF assms(2) support_subset])+
+    apply(subst iterate_expand_cases) unfolding support_clauses apply(simp only: if_False)
+    apply(subst iterate_expand_cases) apply(subst if_not_P) by auto qed
+
+
+(* This lemma about iterations comes up in a few places.                     *)
+lemma iterate_nonzero_image_lemma:
+  assumes "monoidal opp" "finite s" "g(a) = neutral opp"
+  "\<forall>x\<in>s. \<forall>y\<in>s. f x = f y \<and> x \<noteq> y \<longrightarrow> g(f x) = neutral opp"
+  shows "iterate opp {f x | x. x \<in> s \<and> f x \<noteq> a} g = iterate opp s (g \<circ> f)"
+proof- have *:"{f x |x. x \<in> s \<and> ~(f x = a)} = f ` {x. x \<in> s \<and> ~(f x = a)}" by auto
+  have **:"support opp (g \<circ> f) {x \<in> s. f x \<noteq> a} = support opp (g \<circ> f) s"
+    unfolding support_def using assms(3) by auto
+  show ?thesis unfolding *
+    apply(subst iterate_support[THEN sym]) unfolding support_clauses
+    apply(subst iterate_image[OF assms(1)]) defer
+    apply(subst(2) iterate_support[THEN sym]) apply(subst **)
+    unfolding inj_on_def using assms(3,4) unfolding support_def by auto qed
+
+lemma iterate_eq_neutral:
+  assumes "monoidal opp"  "\<forall>x \<in> s. (f(x) = neutral opp)"
+  shows "(iterate opp s f = neutral opp)"
+proof- have *:"support opp f s = {}" unfolding support_def using assms(2) by auto
+  show ?thesis apply(subst iterate_support[THEN sym]) 
+    unfolding * using assms(1) by auto qed
+
+lemma iterate_op: assumes "monoidal opp" "finite s"
+  shows "iterate opp s (\<lambda>x. opp (f x) (g x)) = opp (iterate opp s f) (iterate opp s g)" using assms(2)
+proof(induct s) case empty thus ?case unfolding iterate_insert[OF assms(1)] using assms(1) by auto
+next case (insert x F) show ?case unfolding iterate_insert[OF assms(1) insert(1)] if_not_P[OF insert(2)] insert(3)
+    unfolding monoidal_ac[OF assms(1)] by(rule refl) qed
+
+lemma iterate_eq: assumes "monoidal opp" "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
+  shows "iterate opp s f = iterate opp s g"
+proof- have *:"support opp g s = support opp f s"
+    unfolding support_def using assms(2) by auto
+  show ?thesis
+  proof(cases "finite (support opp f s)")
+    case False thus ?thesis apply(subst iterate_expand_cases,subst(2) iterate_expand_cases)
+      unfolding * by auto
+  next def su \<equiv> "support opp f s"
+    case True note support_subset[of opp f s] 
+    thus ?thesis apply- apply(subst iterate_support[THEN sym],subst(2) iterate_support[THEN sym]) unfolding * using True
+      unfolding su_def[symmetric]
+    proof(induct su) case empty show ?case by auto
+    next case (insert x s) show ?case unfolding iterate_insert[OF assms(1) insert(1)] 
+        unfolding if_not_P[OF insert(2)] apply(subst insert(3))
+        defer apply(subst assms(2)[of x]) using insert by auto qed qed qed
+
+lemma nonempty_witness: assumes "s \<noteq> {}" obtains x where "x \<in> s" using assms by auto
+
+lemma operative_division: fixes f::"(real^'n) set \<Rightarrow> 'a"
+  assumes "monoidal opp" "operative opp f" "d division_of {a..b}"
+  shows "iterate opp d f = f {a..b}"
+proof- def C \<equiv> "card (division_points {a..b} d)" thus ?thesis using assms
+  proof(induct C arbitrary:a b d rule:full_nat_induct)
+    case goal1
+    { presume *:"content {a..b} \<noteq> 0 \<Longrightarrow> ?case"
+      thus ?case apply-apply(cases) defer apply assumption
+      proof- assume as:"content {a..b} = 0"
+        show ?case unfolding operativeD(1)[OF assms(2) as] apply(rule iterate_eq_neutral[OF goal1(2)])
+        proof fix x assume x:"x\<in>d"
+          then guess u v apply(drule_tac division_ofD(4)[OF goal1(4)]) by(erule exE)+
+          thus "f x = neutral opp" using division_of_content_0[OF as goal1(4)] 
+            using operativeD(1)[OF assms(2)] x by auto
+        qed qed }
+    assume "content {a..b} \<noteq> 0" note ab = this[unfolded content_lt_nz[THEN sym] content_pos_lt_eq]
+    hence ab':"\<forall>i. a$i \<le> b$i" by (auto intro!: less_imp_le) show ?case 
+    proof(cases "division_points {a..b} d = {}")
+      case True have d':"\<forall>i\<in>d. \<exists>u v. i = {u..v} \<and>
+        (\<forall>j. u$j = a$j \<and> v$j = a$j \<or> u$j = b$j \<and> v$j = b$j \<or> u$j = a$j \<and> v$j = b$j)"
+        unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule)
+        apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) apply(rule)
+      proof- fix u v j assume as:"{u..v} \<in> d" note division_ofD(3)[OF goal1(4) this]
+        hence uv:"\<forall>i. u$i \<le> v$i" "u$j \<le> v$j" unfolding interval_ne_empty by auto
+        have *:"\<And>p r Q. p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> (Q {u..v})" using as by auto
+        have "(j, u$j) \<notin> division_points {a..b} d"
+          "(j, v$j) \<notin> division_points {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$j \<le> u$j" "v$j \<le> b$j" using division_ofD(2,2,3)[OF goal1(4) as] 
+          unfolding subset_eq apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE)
+          unfolding interval_ne_empty mem_interval by auto
+        ultimately show "u$j = a$j \<and> v$j = a$j \<or> u$j = b$j \<and> v$j = b$j \<or> u$j = a$j \<and> v$j = b$j"
+          unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) by auto
+      qed have "(1/2) *\<^sub>R (a+b) \<in> {a..b}" unfolding mem_interval using ab by(auto intro!:less_imp_le)
+      note this[unfolded division_ofD(6)[OF goal1(4),THEN sym] Union_iff]
+      then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] apply-by(erule exE conjE)+ note uv=this
+      have "{a..b} \<in> d"
+      proof- { presume "i = {a..b}" thus ?thesis using i by auto }
+        { presume "u = a" "v = b" thus "i = {a..b}" using uv by auto }
+        show "u = a" "v = b" unfolding Cart_eq
+        proof(rule_tac[!] allI) fix j note i(2)[unfolded uv mem_interval,rule_format,of j]
+          thus "u $ j = a $ j" "v $ j = b $ j" using uv(2)[rule_format,of j] by auto
+        qed qed
+      hence *:"d = insert {a..b} (d - {{a..b}})" by auto
+      have "iterate opp (d - {{a..b}}) f = neutral opp" apply(rule iterate_eq_neutral[OF goal1(2)])
+      proof fix x assume x:"x \<in> d - {{a..b}}" hence "x\<in>d" by auto note d'[rule_format,OF this]
+        then guess u v apply-by(erule exE conjE)+ note uv=this
+        have "u\<noteq>a \<or> v\<noteq>b" using x[unfolded uv] by auto  
+        then obtain j where "u$j \<noteq> a$j \<or> v$j \<noteq> b$j" unfolding Cart_eq by auto
+        hence "u$j = v$j" using uv(2)[rule_format,of j] by auto
+        hence "content {u..v} = 0"  unfolding content_eq_0 apply(rule_tac x=j in exI) by auto
+        thus "f x = neutral opp" unfolding uv(1) by(rule operativeD(1)[OF goal1(3)])
+      qed thus "iterate opp d f = f {a..b}" apply-apply(subst *) 
+        apply(subst iterate_insert[OF goal1(2)]) using goal1(2,4) by auto
+    next case False hence "\<exists>x. x\<in>division_points {a..b} d" by auto
+      then guess k c unfolding split_paired_Ex apply- unfolding division_points_def mem_Collect_eq split_conv
+        by(erule exE conjE)+ note kc=this[unfolded interval_bounds[OF ab']]
+      from this(3) guess j .. note j=this
+      def d1 \<equiv> "{l \<inter> {x. x$k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x$k \<le> c} \<noteq> {}}"
+      def d2 \<equiv> "{l \<inter> {x. x$k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x$k \<ge> c} \<noteq> {}}"
+      def cb \<equiv> "(\<chi> i. if i = k then c else b$i)" and ca \<equiv> "(\<chi> i. if i = k then c else a$i)"
+      note division_points_psubset[OF goal1(4) ab kc(1-2) j]
+      note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
+      hence *:"(iterate opp d1 f) = f ({a..b} \<inter> {x. x$k \<le> c})" "(iterate opp d2 f) = f ({a..b} \<inter> {x. x$k \<ge> c})"
+        apply- unfolding interval_split apply(rule_tac[!] goal1(1)[rule_format])
+        using division_split[OF goal1(4), where k=k and c=c]
+        unfolding interval_split d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono
+        using goal1(2-3) using division_points_finite[OF goal1(4)] by auto
+      have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev")
+        unfolding * apply(rule operativeD(2)) using goal1(3) .
+      also have "iterate opp d1 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$k \<le> c}))"
+        unfolding d1_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
+        unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
+        unfolding empty_as_interval[THEN sym] apply(rule content_empty)
+      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. x $ k \<le> c} = y \<inter> {x. x $ k \<le> c}" "l \<noteq> y" 
+        from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
+        show "f (l \<inter> {x. x $ k \<le> c}) = neutral opp" unfolding l interval_split
+          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_left_inj)
+          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+
+      qed also have "iterate opp d2 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$k \<ge> c}))"
+        unfolding d2_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
+        unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
+        unfolding empty_as_interval[THEN sym] apply(rule content_empty)
+      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x $ k} = y \<inter> {x. c \<le> x $ k}" "l \<noteq> y" 
+        from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
+        show "f (l \<inter> {x. x $ k \<ge> c}) = neutral opp" unfolding l interval_split
+          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_right_inj)
+          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+
+      qed also have *:"\<forall>x\<in>d. f x = opp (f (x \<inter> {x. x $ k \<le> c})) (f (x \<inter> {x. c \<le> x $ k}))"
+        unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) .
+      have "opp (iterate opp d (\<lambda>l. f (l \<inter> {x. x $ k \<le> c}))) (iterate opp d (\<lambda>l. f (l \<inter> {x. c \<le> x $ k})))
+        = iterate opp d f" apply(subst(3) iterate_eq[OF _ *[rule_format]]) prefer 3
+        apply(rule iterate_op[THEN sym]) using goal1 by auto
+      finally show ?thesis by auto
+    qed qed qed 
+
+lemma iterate_image_nonzero: assumes "monoidal opp"
+  "finite s" "\<forall>x\<in>s. \<forall>y\<in>s. ~(x = y) \<and> f x = f y \<longrightarrow> g(f x) = neutral opp"
+  shows "iterate opp (f ` s) g = iterate opp s (g \<circ> f)" using assms
+proof(induct rule:finite_subset_induct[OF assms(2) subset_refl])
+  case goal1 show ?case using assms(1) by auto
+next case goal2 have *:"\<And>x y. y = neutral opp \<Longrightarrow> x = opp y x" using assms(1) by auto
+  show ?case unfolding image_insert apply(subst iterate_insert[OF assms(1)])
+    apply(rule finite_imageI goal2)+
+    apply(cases "f a \<in> f ` F") unfolding if_P if_not_P apply(subst goal2(4)[OF assms(1) goal2(1)]) defer
+    apply(subst iterate_insert[OF assms(1) goal2(1)]) defer
+    apply(subst iterate_insert[OF assms(1) goal2(1)])
+    unfolding if_not_P[OF goal2(3)] defer unfolding image_iff defer apply(erule bexE)
+    apply(rule *) unfolding o_def apply(rule_tac y=x in goal2(7)[rule_format])
+    using goal2 unfolding o_def by auto qed 
+
+lemma operative_tagged_division: assumes "monoidal opp" "operative opp f" "d tagged_division_of {a..b}"
+  shows "iterate(opp) d (\<lambda>(x,l). f l) = f {a..b}"
+proof- have *:"(\<lambda>(x,l). f l) = (f o snd)" unfolding o_def by(rule,auto) note assm = tagged_division_ofD[OF assms(3)]
+  have "iterate(opp) d (\<lambda>(x,l). f l) = iterate opp (snd ` d) f" unfolding *
+    apply(rule iterate_image_nonzero[THEN sym,OF assms(1)]) apply(rule tagged_division_of_finite assms)+ 
+    unfolding Ball_def split_paired_All snd_conv apply(rule,rule,rule,rule,rule,rule,rule,erule conjE)
+  proof- fix a b aa ba assume as:"(a, b) \<in> d" "(aa, ba) \<in> d" "(a, b) \<noteq> (aa, ba)" "b = ba"
+    guess u v using assm(4)[OF as(1)] apply-by(erule exE)+ note uv=this
+    show "f b = neutral opp" unfolding uv apply(rule operativeD(1)[OF assms(2)])
+      unfolding content_eq_0_interior using tagged_division_ofD(5)[OF assms(3) as(1-3)]
+      unfolding as(4)[THEN sym] uv by auto
+  qed also have "\<dots> = f {a..b}" 
+    using operative_division[OF assms(1-2) division_of_tagged_division[OF assms(3)]] .
+  finally show ?thesis . qed
+
+subsection {* Additivity of content. *}
+
+lemma setsum_iterate:assumes "finite s" shows "setsum f s = iterate op + s f"
+proof- have *:"setsum f s = setsum f (support op + f s)"
+    apply(rule setsum_mono_zero_right)
+    unfolding support_def neutral_monoid using assms by auto
+  thus ?thesis unfolding * setsum_def iterate_def fold_image_def fold'_def
+    unfolding neutral_monoid . qed
+
+lemma additive_content_division: assumes "d division_of {a..b}"
+  shows "setsum content d = content({a..b})"
+  unfolding operative_division[OF monoidal_monoid operative_content assms,THEN sym]
+  apply(subst setsum_iterate) using assms by auto
+
+lemma additive_content_tagged_division:
+  assumes "d tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,l). content l) d = content({a..b})"
+  unfolding operative_tagged_division[OF monoidal_monoid operative_content assms,THEN sym]
+  apply(subst setsum_iterate) using assms by auto
+
+subsection {* Finally, the integral of a constant\<forall> *}
+
+lemma has_integral_const[intro]:
+  "((\<lambda>x. c) has_integral (content({a..b::real^'n}) *\<^sub>R c)) ({a..b})"
+  unfolding has_integral apply(rule,rule,rule_tac x="\<lambda>x. ball x 1" in exI)
+  apply(rule,rule gauge_trivial)apply(rule,rule,erule conjE)
+  unfolding split_def apply(subst scaleR_left.setsum[THEN sym, unfolded o_def])
+  defer apply(subst additive_content_tagged_division[unfolded split_def]) apply assumption by auto
+
+subsection {* Bounds on the norm of Riemann sums and the integral itself. *}
+
+lemma dsum_bound: assumes "p division_of {a..b}" "norm(c) \<le> e"
+  shows "norm(setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content({a..b})" (is "?l \<le> ?r")
+  apply(rule order_trans,rule setsum_norm) defer unfolding norm_scaleR setsum_left_distrib[THEN sym]
+  apply(rule order_trans[OF mult_left_mono],rule assms,rule setsum_abs_ge_zero)
+  apply(subst real_mult_commute) apply(rule mult_left_mono)
+  apply(rule order_trans[of _ "setsum content p"]) apply(rule eq_refl,rule setsum_cong2)
+  apply(subst abs_of_nonneg) unfolding additive_content_division[OF assms(1)]
+proof- from order_trans[OF norm_ge_zero[of c] assms(2)] show "0 \<le> e" .
+  fix x assume "x\<in>p" from division_ofD(4)[OF assms(1) this] guess u v apply-by(erule exE)+
+  thus "0 \<le> content x" using content_pos_le by auto
+qed(insert assms,auto)
+
+lemma rsum_bound: assumes "p tagged_division_of {a..b}" "\<forall>x\<in>{a..b}. norm(f x) \<le> e"
+  shows "norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content({a..b})"
+proof(cases "{a..b} = {}") case True
+  show ?thesis using assms(1) unfolding True tagged_division_of_trivial by auto
+next case False show ?thesis
+    apply(rule order_trans,rule setsum_norm) defer unfolding split_def norm_scaleR
+    apply(rule order_trans[OF setsum_mono]) apply(rule mult_left_mono[OF _ abs_ge_zero, of _ e]) defer
+    unfolding setsum_left_distrib[THEN sym] apply(subst real_mult_commute) apply(rule mult_left_mono)
+    apply(rule order_trans[of _ "setsum (content \<circ> snd) p"]) apply(rule eq_refl,rule setsum_cong2)
+    apply(subst o_def, rule abs_of_nonneg)
+  proof- show "setsum (content \<circ> snd) p \<le> content {a..b}" apply(rule eq_refl)
+      unfolding additive_content_tagged_division[OF assms(1),THEN sym] split_def by auto
+    guess w using nonempty_witness[OF False] .
+    thus "e\<ge>0" apply-apply(rule order_trans) defer apply(rule assms(2)[rule_format],assumption) by auto
+    fix xk assume *:"xk\<in>p" guess x k  using surj_pair[of xk] apply-by(erule exE)+ note xk = this *[unfolded this]
+    from tagged_division_ofD(4)[OF assms(1) xk(2)] guess u v apply-by(erule exE)+ note uv=this
+    show "0\<le> content (snd xk)" unfolding xk snd_conv uv by(rule content_pos_le)
+    show "norm (f (fst xk)) \<le> e" unfolding xk fst_conv using tagged_division_ofD(2,3)[OF assms(1) xk(2)] assms(2) by auto
+  qed(insert assms,auto) qed
+
+lemma rsum_diff_bound:
+  assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. norm(f x - g x) \<le> e"
+  shows "norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - setsum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le> e * content({a..b})"
+  apply(rule order_trans[OF _ rsum_bound[OF assms]]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
+  unfolding setsum_subtractf[THEN sym] apply(rule setsum_cong2) unfolding scaleR.diff_right by auto
+
+lemma has_integral_bound: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "0 \<le> B" "(f has_integral i) ({a..b})" "\<forall>x\<in>{a..b}. norm(f x) \<le> B"
+  shows "norm i \<le> B * content {a..b}"
+proof- let ?P = "content {a..b} > 0" { presume "?P \<Longrightarrow> ?thesis"
+    thus ?thesis proof(cases ?P) case False
+      hence *:"content {a..b} = 0" using content_lt_nz by auto
+      hence **:"i = 0" using assms(2) apply(subst has_integral_null_eq[THEN sym]) by auto
+      show ?thesis unfolding * ** using assms(1) by auto
+    qed auto } assume ab:?P
+  { presume "\<not> ?thesis \<Longrightarrow> False" thus ?thesis by auto }
+  assume "\<not> ?thesis" hence *:"norm i - B * content {a..b} > 0" by auto
+  from assms(2)[unfolded has_integral,rule_format,OF *] guess d apply-by(erule exE conjE)+ note d=this[rule_format]
+  from fine_division_exists[OF this(1), of a b] guess p . note p=this
+  have *:"\<And>s B. norm s \<le> B \<Longrightarrow> \<not> (norm (s - i) < norm i - B)"
+  proof- case goal1 thus ?case unfolding not_less
+    using norm_triangle_sub[of i s] unfolding norm_minus_commute by auto
+  qed show False using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto qed
+
+subsection {* Similar theorems about relationship among components. *}
+
+lemma rsum_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
+  assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. (f x)$i \<le> (g x)$i"
+  shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)$i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)$i"
+  unfolding setsum_component apply(rule setsum_mono)
+  apply(rule mp) defer apply assumption apply(induct_tac x,rule) unfolding split_conv
+proof- fix a b assume ab:"(a,b) \<in> p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab]
+  from this(3) guess u v apply-by(erule exE)+ note b=this
+  show "(content b *\<^sub>R f a) $ i \<le> (content b *\<^sub>R g a) $ i" unfolding b
+    unfolding Cart_nth.scaleR real_scaleR_def apply(rule mult_left_mono)
+    defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed
+
+lemma has_integral_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
+  assumes "(f has_integral i) s" "(g has_integral j) s"  "\<forall>x\<in>s. (f x)$k \<le> (g x)$k"
+  shows "i$k \<le> j$k"
+proof- have lem:"\<And>a b g i j. \<And>f::real^'n \<Rightarrow> real^'m. (f has_integral i) ({a..b}) \<Longrightarrow> 
+    (g has_integral j) ({a..b}) \<Longrightarrow> \<forall>x\<in>{a..b}. (f x)$k \<le> (g x)$k \<Longrightarrow> i$k \<le> j$k"
+  proof(rule ccontr) case goal1 hence *:"0 < (i$k - j$k) / 3" by auto
+    guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d1=this[rule_format]
+    guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d2=this[rule_format]
+    guess p using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter .
+    note p = this(1) conjunctD2[OF this(2)]  note le_less_trans[OF component_le_norm, of _ _ k]
+    note this[OF d1(2)[OF conjI[OF p(1,2)]]] this[OF d2(2)[OF conjI[OF p(1,3)]]]
+    thus False unfolding Cart_nth.diff using rsum_component_le[OF p(1) goal1(3)] by smt
+  qed let ?P = "\<exists>a b. s = {a..b}"
+  { presume "\<not> ?P \<Longrightarrow> ?thesis" thus ?thesis proof(cases ?P)
+      case True then guess a b apply-by(erule exE)+ note s=this
+      show ?thesis apply(rule lem) using assms[unfolded s] by auto
+    qed auto } assume as:"\<not> ?P"
+  { presume "\<not> ?thesis \<Longrightarrow> False" thus ?thesis by auto }
+  assume "\<not> i$k \<le> j$k" hence ij:"(i$k - j$k) / 3 > 0" by auto
+  note has_integral_altD[OF _ as this] from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
+  have "bounded (ball 0 B1 \<union> ball (0::real^'n) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+
+  from bounded_subset_closed_interval[OF this] guess a b apply- by(erule exE)+
+  note ab = conjunctD2[OF this[unfolded Un_subset_iff]]
+  guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this]
+  guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this]
+  have *:"\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False" by smt(*SMTSMT*)
+  note le_less_trans[OF component_le_norm[of _ k]] note this[OF w1(2)] this[OF w2(2)] moreover
+  have "w1$k \<le> w2$k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately
+  show False unfolding Cart_nth.diff by(rule *) qed
+
+lemma integral_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
+  assumes "f integrable_on s" "g integrable_on s"  "\<forall>x\<in>s. (f x)$k \<le> (g x)$k"
+  shows "(integral s f)$k \<le> (integral s g)$k"
+  apply(rule has_integral_component_le) using integrable_integral assms by auto
+
+lemma has_integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
+  assumes "(f has_integral i) s"  "(g has_integral j) s" "\<forall>x\<in>s. f x \<le> g x"
+  shows "dest_vec1 i \<le> dest_vec1 j" apply(rule has_integral_component_le[OF assms(1-2)])
+  using assms(3) unfolding vector_le_def by auto
+
+lemma integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
+  assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. f x \<le> g x"
+  shows "dest_vec1(integral s f) \<le> dest_vec1(integral s g)"
+  apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto
+
+lemma has_integral_component_pos: fixes f::"real^'n \<Rightarrow> real^'m"
+  assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> (f x)$k" shows "0 \<le> i$k"
+  using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2) by auto
+
+lemma integral_component_pos: fixes f::"real^'n \<Rightarrow> real^'m"
+  assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> (f x)$k" shows "0 \<le> (integral s f)$k"
+  apply(rule has_integral_component_pos) using assms by auto
+
+lemma has_integral_dest_vec1_pos: fixes f::"real^'n \<Rightarrow> real^1"
+  assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> i"
+  using has_integral_component_pos[OF assms(1), of 1]
+  using assms(2) unfolding vector_le_def by auto
+
+lemma integral_dest_vec1_pos: fixes f::"real^'n \<Rightarrow> real^1"
+  assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> integral s f"
+  apply(rule has_integral_dest_vec1_pos) using assms by auto
+
+lemma has_integral_component_neg: fixes f::"real^'n \<Rightarrow> real^'m"
+  assumes "(f has_integral i) s" "\<forall>x\<in>s. (f x)$k \<le> 0" shows "i$k \<le> 0"
+  using has_integral_component_le[OF assms(1) has_integral_0] assms(2) by auto
+
+lemma has_integral_dest_vec1_neg: fixes f::"real^'n \<Rightarrow> real^1"
+  assumes "(f has_integral i) s" "\<forall>x\<in>s. f x \<le> 0" shows "i \<le> 0"
+  using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto
+
+lemma has_integral_component_lbound:
+  assumes "(f has_integral i) {a..b}"  "\<forall>x\<in>{a..b}. B \<le> f(x)$k" shows "B * content {a..b} \<le> i$k"
+  using has_integral_component_le[OF has_integral_const assms(1),of "(\<chi> i. B)" k] assms(2)
+  unfolding Cart_lambda_beta vector_scaleR_component by(auto simp add:field_simps)
+
+lemma has_integral_component_ubound: 
+  assumes "(f has_integral i) {a..b}" "\<forall>x\<in>{a..b}. f x$k \<le> B"
+  shows "i$k \<le> B * content({a..b})"
+  using has_integral_component_le[OF assms(1) has_integral_const, of k "vec B"]
+  unfolding vec_component Cart_nth.scaleR using assms(2) by(auto simp add:field_simps)
+
+lemma integral_component_lbound:
+  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)$k"
+  shows "B * content({a..b}) \<le> (integral({a..b}) f)$k"
+  apply(rule has_integral_component_lbound) using assms unfolding has_integral_integral by auto
+
+lemma integral_component_ubound:
+  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. f(x)$k \<le> B"
+  shows "(integral({a..b}) f)$k \<le> B * content({a..b})"
+  apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto
+
+subsection {* Uniform limit of integrable functions is integrable. *}
+
+lemma real_arch_invD:
+  "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
+  by(subst(asm) real_arch_inv)
+
+lemma integrable_uniform_limit: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  assumes "\<forall>e>0. \<exists>g. (\<forall>x\<in>{a..b}. norm(f x - g x) \<le> e) \<and> g integrable_on {a..b}"
+  shows "f integrable_on {a..b}"
+proof- { presume *:"content {a..b} > 0 \<Longrightarrow> ?thesis"
+    show ?thesis apply cases apply(rule *,assumption)
+      unfolding content_lt_nz integrable_on_def using has_integral_null by auto }
+  assume as:"content {a..b} > 0"
+  have *:"\<And>P. \<forall>e>(0::real). P e \<Longrightarrow> \<forall>n::nat. P (inverse (real n+1))" by auto
+  from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format]
+  from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\<lambda>x. x"]] guess i .. note i=this[rule_format]
+  
+  have "Cauchy i" unfolding Cauchy_def
+  proof(rule,rule) fix e::real assume "e>0"
+    hence "e / 4 / content {a..b} > 0" using as by(auto simp add:field_simps)
+    then guess M apply-apply(subst(asm) real_arch_inv) by(erule exE conjE)+ note M=this
+    show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (i m) (i n) < e" apply(rule_tac x=M in exI,rule,rule,rule,rule)
+    proof- case goal1 have "e/4>0" using `e>0` by auto note * = i[unfolded has_integral,rule_format,OF this]
+      from *[of m] guess gm apply-by(erule conjE exE)+ note gm=this[rule_format]
+      from *[of n] guess gn apply-by(erule conjE exE)+ note gn=this[rule_format]
+      from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b] guess p . note p=this
+      have lem2:"\<And>s1 s2 i1 i2. norm(s2 - s1) \<le> e/2 \<Longrightarrow> norm(s1 - i1) < e / 4 \<Longrightarrow> norm(s2 - i2) < e / 4 \<Longrightarrow>norm(i1 - i2) < e"
+      proof- case goal1 have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
+          using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
+          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:group_simps)
+        also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+        finally show ?case .
+      qed
+      show ?case unfolding vector_dist_norm apply(rule lem2) defer
+        apply(rule gm(2)[OF conjI[OF p(1)]],rule_tac[2] gn(2)[OF conjI[OF p(1)]])
+        using conjunctD2[OF p(2)[unfolded fine_inter]] apply- apply assumption+ apply(rule order_trans)
+        apply(rule rsum_diff_bound[OF p(1), where e="2 / real M"])
+      proof show "2 / real M * content {a..b} \<le> e / 2" unfolding divide_inverse 
+          using M as by(auto simp add:field_simps)
+        fix x assume x:"x \<in> {a..b}"
+        have "norm (f x - g n x) + norm (f x - g m x) \<le> inverse (real n + 1) + inverse (real m + 1)" 
+            using g(1)[OF x, of n] g(1)[OF x, of m] by auto
+        also have "\<dots> \<le> inverse (real M) + inverse (real M)" apply(rule add_mono)
+          apply(rule_tac[!] le_imp_inverse_le) using goal1 M by auto
+        also have "\<dots> = 2 / real M" unfolding real_divide_def by auto
+        finally show "norm (g n x - g m x) \<le> 2 / real M"
+          using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
+          by(auto simp add:group_simps simp add:norm_minus_commute)
+      qed qed qed
+  from this[unfolded convergent_eq_cauchy[THEN sym]] guess s .. note s=this
+
+  show ?thesis unfolding integrable_on_def apply(rule_tac x=s in exI) unfolding has_integral
+  proof(rule,rule)  
+    case goal1 hence *:"e/3 > 0" by auto
+    from s[unfolded Lim_sequentially,rule_format,OF this] guess N1 .. note N1=this
+    from goal1 as have "e / 3 / content {a..b} > 0" by(auto simp add:field_simps)
+    from real_arch_invD[OF this] guess N2 apply-by(erule exE conjE)+ note N2=this
+    from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format]
+    have lem:"\<And>sf sg i. norm(sf - sg) \<le> e / 3 \<Longrightarrow> norm(i - s) < e / 3 \<Longrightarrow> norm(sg - i) < e / 3 \<Longrightarrow> norm(sf - s) < e"
+    proof- case goal1 have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
+        using norm_triangle_ineq[of "sf - sg" "sg - s"]
+        using norm_triangle_ineq[of "sg -  i" " i - s"] by(auto simp add:group_simps)
+      also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+      finally show ?case .
+    qed
+    show ?case apply(rule_tac x=g' in exI) apply(rule,rule g')
+    proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> g' fine p" note * = g'(2)[OF this]
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e" apply-apply(rule lem[OF _ _ *])
+        apply(rule order_trans,rule rsum_diff_bound[OF p[THEN conjunct1]]) apply(rule,rule g,assumption)
+      proof- have "content {a..b} < e / 3 * (real N2)"
+          using N2 unfolding inverse_eq_divide using as by(auto simp add:field_simps)
+        hence "content {a..b} < e / 3 * (real (N1 + N2) + 1)"
+          apply-apply(rule less_le_trans,assumption) using `e>0` by auto 
+        thus "inverse (real (N1 + N2) + 1) * content {a..b} \<le> e / 3"
+          unfolding inverse_eq_divide by(auto simp add:field_simps)
+        show "norm (i (N1 + N2) - s) < e / 3" by(rule N1[rule_format,unfolded vector_dist_norm],auto)
+      qed qed qed qed
+
+subsection {* Negligible sets. *}
+
+definition "indicator s \<equiv> (\<lambda>x. if x \<in> s then 1 else (0::real))"
+
+lemma dest_vec1_indicator:
+ "indicator s x = (if x \<in> s then 1 else 0)" unfolding indicator_def by auto
+
+lemma indicator_pos_le[intro]: "0 \<le> (indicator s x)" unfolding indicator_def by auto
+
+lemma indicator_le_1[intro]: "(indicator s x) \<le> 1" unfolding indicator_def by auto
+
+lemma dest_vec1_indicator_abs_le_1: "abs(indicator s x) \<le> 1"
+  unfolding indicator_def by auto
+
+definition "negligible (s::(real^'n) set) \<equiv> (\<forall>a b. ((indicator s) has_integral 0) {a..b})"
+
+lemma indicator_simps[simp]:"x\<in>s \<Longrightarrow> indicator s x = 1" "x\<notin>s \<Longrightarrow> indicator s x = 0"
+  unfolding indicator_def by auto
+
+subsection {* Negligibility of hyperplane. *}
+
+lemma vsum_nonzero_image_lemma: 
+  assumes "finite s" "g(a) = 0"
+  "\<forall>x\<in>s. \<forall>y\<in>s. f x = f y \<and> x \<noteq> y \<longrightarrow> g(f x) = 0"
+  shows "setsum g {f x |x. x \<in> s \<and> f x \<noteq> a} = setsum (g o f) s"
+  unfolding setsum_iterate[OF assms(1)] apply(subst setsum_iterate) defer
+  apply(rule iterate_nonzero_image_lemma) apply(rule assms monoidal_monoid)+
+  unfolding assms using neutral_add unfolding neutral_add using assms by auto 
+
+lemma interval_doublesplit: shows "{a..b} \<inter> {x . abs(x$k - c) \<le> (e::real)} =
+  {(\<chi> i. if i = k then max (a$k) (c - e) else a$i) .. (\<chi> i. if i = k then min (b$k) (c + e) else b$i)}"
+proof- have *:"\<And>x c e::real. abs(x - c) \<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 by(rule refl) qed
+
+lemma division_doublesplit: assumes "p division_of {a..b::real^'n}" 
+  shows "{l \<inter> {x. abs(x$k - c) \<le> e} |l. l \<in> p \<and> l \<inter> {x. abs(x$k - c) \<le> e} \<noteq> {}} division_of ({a..b} \<inter> {x. abs(x$k - c) \<le> e})"
+proof- have *:"\<And>x c. abs(x - c) \<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" and k=k,unfolded interval_split]
+  note division_split(2)[OF this, where c="c-e" and k=k] 
+  thus ?thesis apply(rule **) unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit
+    apply(rule set_ext) unfolding mem_Collect_eq apply rule apply(erule conjE exE)+ apply(rule_tac x=la in exI) defer
+    apply(erule conjE exE)+ apply(rule_tac x="l \<inter> {x. c + e \<ge> x $ k}" in exI) apply rule defer apply rule
+    apply(rule_tac x=l in exI) by blast+ qed
+
+lemma content_doublesplit: assumes "0 < e"
+  obtains d where "0 < d" "content({a..b} \<inter> {x. abs(x$k - c) \<le> d}) < e"
+proof(cases "content {a..b} = 0")
+  case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit
+    apply(rule le_less_trans[OF content_subset]) defer apply(subst True)
+    unfolding interval_doublesplit[THEN sym] using assms by auto 
+next case False def d \<equiv> "e / 3 / setprod (\<lambda>i. b$i - a$i) (UNIV - {k})"
+  note False[unfolded content_eq_0 not_ex not_le, rule_format]
+  hence prod0:"0 < setprod (\<lambda>i. b$i - a$i) (UNIV - {k})" apply-apply(rule setprod_pos) by smt
+  hence "d > 0" unfolding d_def using assms by(auto simp add:field_simps) thus ?thesis
+  proof(rule that[of d]) have *:"UNIV = insert k (UNIV - {k})" by auto
+    have **:"{a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} \<noteq> {} \<Longrightarrow> 
+      (\<Prod>i\<in>UNIV - {k}. interval_upperbound ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) $ i - interval_lowerbound ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) $ i)
+      = (\<Prod>i\<in>UNIV - {k}. b$i - a$i)" apply(rule setprod_cong,rule refl)
+      unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds by auto
+    show "content ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms)
+      unfolding if_not_P apply(subst *) apply(subst setprod_insert) unfolding **
+      unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds unfolding Cart_lambda_beta if_P[OF refl]
+    proof- have "(min (b $ k) (c + d) - max (a $ k) (c - d)) \<le> 2 * d" by auto
+      also have "... < e / (\<Prod>i\<in>UNIV - {k}. b $ i - a $ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps)
+      finally show "(min (b $ k) (c + d) - max (a $ k) (c - d)) * (\<Prod>i\<in>UNIV - {k}. b $ i - a $ i) < e"
+        unfolding pos_less_divide_eq[OF prod0] . qed auto qed qed
+
+lemma negligible_standard_hyperplane[intro]: "negligible {x. x$k = (c::real)}" 
+  unfolding negligible_def has_integral apply(rule,rule,rule,rule)
+proof- case goal1 from content_doublesplit[OF this,of a b k c] guess d . note d=this let ?i = "indicator {x. x$k = c}"
+  show ?case apply(rule_tac x="\<lambda>x. ball x d" in exI) apply(rule,rule gauge_ball,rule d)
+  proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> (\<lambda>x. ball x d) fine p"
+    have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x$k - c) \<le> d}) *\<^sub>R ?i x)"
+      apply(rule setsum_cong2) unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
+      apply(cases,rule disjI1,assumption,rule disjI2)
+    proof- fix x l assume as:"(x,l)\<in>p" "?i x \<noteq> 0" hence xk:"x$k = c" unfolding indicator_def apply-by(rule ccontr,auto)
+      show "content l = content (l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d})" apply(rule arg_cong[where f=content])
+        apply(rule set_ext,rule,rule) unfolding mem_Collect_eq
+      proof- fix y assume y:"y\<in>l" note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv]
+        note this[unfolded subset_eq mem_ball vector_dist_norm,rule_format,OF y] note le_less_trans[OF component_le_norm[of _ k] this]
+        thus "\<bar>y $ k - c\<bar> \<le> d" unfolding Cart_nth.diff xk by auto
+      qed auto qed
+    note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]]
+    show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e" unfolding diff_0_right * unfolding real_scaleR_def real_norm_def
+      apply(subst abs_of_nonneg) apply(rule setsum_nonneg,rule) unfolding split_paired_all split_conv
+      apply(rule mult_nonneg_nonneg) apply(drule p'(4)) apply(erule exE)+ apply(rule_tac b=b in back_subst)
+      prefer 2 apply(subst(asm) eq_commute) apply assumption
+      apply(subst interval_doublesplit) apply(rule content_pos_le) apply(rule indicator_pos_le)
+    proof- have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}))"
+        apply(rule setsum_mono) unfolding split_paired_all split_conv 
+        apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit intro!:content_pos_le)
+      also have "... < e" apply(subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]])
+      proof- case goal1 have "content ({u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<le> content {u..v}"
+          unfolding interval_doublesplit apply(rule content_subset) unfolding interval_doublesplit[THEN sym] by auto
+        thus ?case unfolding goal1 unfolding interval_doublesplit using content_pos_le by smt
+      next have *:"setsum content {l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} |l. l \<in> snd ` p \<and> l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} \<noteq> {}} \<ge> 0"
+          apply(rule setsum_nonneg,rule) unfolding mem_Collect_eq image_iff apply(erule exE bexE conjE)+ unfolding split_paired_all 
+        proof- fix x l a b assume as:"x = l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}" "(a, b) \<in> p" "l = snd (a, b)"
+          guess u v using p'(4)[OF as(2)] apply-by(erule exE)+ note * = this
+          show "content x \<ge> 0" unfolding as snd_conv * interval_doublesplit by(rule content_pos_le)
+        qed have **:"norm (1::real) \<le> 1" by auto note division_doublesplit[OF p'',unfolded interval_doublesplit]
+        note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym]]
+        note this[unfolded real_scaleR_def real_norm_def class_semiring.semiring_rules, of k c d] note le_less_trans[OF this d(2)]
+        from this[unfolded abs_of_nonneg[OF *]] show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d})) < e"
+          apply(subst vsum_nonzero_image_lemma[of "snd ` p" content "{}", unfolded o_def,THEN sym])
+          apply(rule finite_imageI p' content_empty)+ unfolding forall_in_division[OF p'']
+        proof(rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v
+          assume as:"{m..n} \<in> snd ` p" "{u..v} \<in> snd ` p" "{m..n} \<noteq> {u..v}"  "{m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} = {u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}"
+          have "({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<inter> ({u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<subseteq> {m..n} \<inter> {u..v}" by blast
+          note subset_interior[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]]
+          hence "interior ({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) = {}" unfolding as Int_absorb by auto
+          thus "content ({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) = 0" unfolding interval_doublesplit content_eq_0_interior[THEN sym] .
+        qed qed
+      finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) * ?i x) < e" .
+    qed qed qed
+
+subsection {* A technical lemma about "refinement" of division. *}
+
+lemma tagged_division_finer: fixes p::"((real^'n) \<times> ((real^'n) set)) set"
+  assumes "p tagged_division_of {a..b}" "gauge d"
+  obtains q where "q tagged_division_of {a..b}" "d fine q" "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
+proof-
+  let ?P = "\<lambda>p. p tagged_partial_division_of {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 {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
+    thus ?thesis apply-apply(rule that[of q]) unfolding tagged_division_ofD[OF assms(1)] by auto
+  } fix p::"((real^'n) \<times> ((real^'n) set)) set" assume as:"finite p"
+  show "?P p" apply(rule,rule) using as proof(induct p) 
+    case empty show ?case apply(rule_tac x="{}" in exI) unfolding fine_def by auto
+  next case (insert xk p) guess x k using surj_pair[of xk] apply- by(erule 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 apply-by(erule exE)+ note uv=this
+
+    have "finite {k. \<exists>x. (x, k) \<in> p}" 
+      apply(rule finite_subset[of _ "snd ` p"],rule) unfolding subset_eq image_iff mem_Collect_eq
+      apply(erule exE,rule_tac x="(xa,x)" in bexI) using p by auto
+    hence int:"interior {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],assumption)      
+      apply(rule p(5))  unfolding uv xk apply(rule insertI1,rule insertI2) apply assumption
+      using insert(2) unfolding uv xk by auto
+
+    show ?case proof(cases "{u..v} \<subseteq> d x")
+      case True thus ?thesis apply(rule_tac x="{(x,{u..v})} \<union> q1" in exI) apply rule
+        unfolding * uv apply(rule tagged_division_union,rule tagged_division_of_self)
+        apply(rule p[unfolded xk uv] insertI1)+  apply(rule q1,rule int) 
+        apply(rule,rule fine_union,subst fine_def) defer apply(rule q1)
+        unfolding Ball_def split_paired_All split_conv apply(rule,rule,rule,rule)
+        apply(erule insertE) defer apply(rule UnI2) apply(drule q1(3)[rule_format]) unfolding xk uv by auto
+    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,rule,rule,rule) apply(erule insertE)
+        apply(rule UnI2) defer apply(drule q1(3)[rule_format])using False unfolding xk uv by auto
+    qed qed qed
+
+subsection {* Hence the main theorem about negligible sets. *}
+
+lemma finite_product_dependent: assumes "finite s" "\<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 by auto qed auto
+
+lemma sum_sum_product: assumes "finite s" "\<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_Un_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 by auto
+    show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}" apply(rule finite_product_dependent) using insert by auto
+  qed(insert insert, auto) qed auto
+
+lemma has_integral_negligible: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "negligible s" "\<forall>x\<in>(t - s). f x = 0"
+  shows "(f has_integral 0) t"
+proof- presume P:"\<And>f::real^'n \<Rightarrow> 'a. \<And>a b. (\<forall>x. ~(x \<in> s) \<longrightarrow> f x = 0) \<Longrightarrow> (f has_integral 0) ({a..b})"
+  let ?f = "(\<lambda>x. if x \<in> t then f x else 0)"
+  show ?thesis apply(rule_tac f="?f" in has_integral_eq) apply(rule) unfolding if_P apply(rule refl)
+    apply(subst has_integral_alt) apply(cases,subst if_P,assumption) unfolding if_not_P
+  proof- assume "\<exists>a b. t = {a..b}" then guess a b apply-by(erule exE)+ note t = this
+    show "(?f has_integral 0) t" unfolding t apply(rule P) using assms(2) unfolding t by auto
+  next show "\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b} \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> t then ?f x else 0) has_integral z) {a..b} \<and> norm (z - 0) < e)"
+      apply(safe,rule_tac x=1 in exI,rule) apply(rule zero_less_one,safe) apply(rule_tac x=0 in exI)
+      apply(rule,rule P) using assms(2) by auto
+  qed
+next fix f::"real^'n \<Rightarrow> 'a" and a b::"real^'n" assume assm:"\<forall>x. x \<notin> s \<longrightarrow> f x = 0" 
+  show "(f has_integral 0) {a..b}" unfolding has_integral
+  proof(safe) case goal1
+    hence "\<And>n. e / 2 / ((real n+1) * (2 ^ n)) > 0" 
+      apply-apply(rule divide_pos_pos) defer apply(rule mult_pos_pos) by(auto simp add:field_simps)
+    note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b] note allI[OF this,of "\<lambda>x. x"] 
+    from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]]
+    show ?case apply(rule_tac x="\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x" in exI) 
+    proof safe show "gauge (\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x)" using d(1) unfolding gauge_def by auto
+      fix p assume as:"p tagged_division_of {a..b}" "(\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x) fine p" 
+      let ?goal = "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
+      { presume "p\<noteq>{} \<Longrightarrow> ?goal" thus ?goal apply(cases "p={}") using goal1 by auto  }
+      assume as':"p \<noteq> {}" from real_arch_simple[of "Sup((\<lambda>(x,k). norm(f x)) ` p)"] guess N ..
+      hence N:"\<forall>x\<in>(\<lambda>(x, k). norm (f x)) ` p. x \<le> real N" apply(subst(asm) Sup_finite_le_iff) using as as' by auto
+      have "\<forall>i. \<exists>q. q tagged_division_of {a..b} \<and> (d i) fine q \<and> (\<forall>(x, k)\<in>p. k \<subseteq> (d i) x \<longrightarrow> (x, k) \<in> q)"
+        apply(rule,rule tagged_division_finer[OF as(1) d(1)]) by auto
+      from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]]
+      have *:"\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> 0" apply(rule setsum_nonneg,safe) 
+        unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) apply(drule tagged_division_ofD(4)[OF q(1)]) by auto
+      have **:"\<And>f g s t. finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow> (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t"
+      proof- case goal1 thus ?case apply-apply(rule setsum_le_included[of s t g snd f]) prefer 4
+          apply safe apply(erule_tac x=x in ballE) apply(erule exE) apply(rule_tac x="(xa,x)" in bexI) by auto qed
+      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
+                     norm(setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x) (q i))) {0..N+1}"
+        unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
+        apply(rule order_trans,rule setsum_norm) defer apply(subst sum_sum_product) prefer 3 
+      proof(rule **,safe) show "finite {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i}" apply(rule finite_product_dependent) using q by auto
+        fix i a b assume as'':"(a,b) \<in> q i" show "0 \<le> (real i + 1) * (content b *\<^sub>R indicator s a)"
+          unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) defer apply(rule mult_nonneg_nonneg)
+          using tagged_division_ofD(4)[OF q(1) as''] by auto
+      next fix i::nat show "finite (q i)" using q by auto
+      next fix x k assume xk:"(x,k) \<in> p" def n \<equiv> "nat \<lfloor>norm (f x)\<rfloor>"
+        have *:"norm (f x) \<in> (\<lambda>(x, k). norm (f x)) ` p" using xk by auto
+        have nfx:"real n \<le> norm(f x)" "norm(f x) \<le> real n + 1" unfolding n_def by auto
+        hence "n \<in> {0..N + 1}" using N[rule_format,OF *] by auto
+        moreover  note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv]
+        note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this] note this[unfolded n_def[symmetric]]
+        moreover have "norm (content k *\<^sub>R f x) \<le> (real n + 1) * (content k * indicator s x)"
+        proof(cases "x\<in>s") case False thus ?thesis using assm by auto
+        next case True have *:"content k \<ge> 0" using tagged_division_ofD(4)[OF as(1) xk] by auto
+          moreover have "content k * norm (f x) \<le> content k * (real n + 1)" apply(rule mult_mono) using nfx * by auto
+          ultimately show ?thesis unfolding abs_mult using nfx True by(auto simp add:field_simps)
+        qed ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le> (real y + 1) * (content k *\<^sub>R indicator s x)"
+          apply(rule_tac x=n in exI,safe) apply(rule_tac x=n in exI,rule_tac x="(x,k)" in exI,safe) by auto
+      qed(insert as, auto)
+      also have "... \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {0..N+1}" apply(rule setsum_mono) 
+      proof- case goal1 thus ?case apply(subst mult_commute, subst pos_le_divide_eq[THEN sym])
+          using d(2)[rule_format,of "q i" i] using q[rule_format] by(auto simp add:field_simps)
+      qed also have "... < e * inverse 2 * 2" unfolding real_divide_def setsum_right_distrib[THEN sym]
+        apply(rule mult_strict_left_mono) unfolding power_inverse atLeastLessThanSuc_atLeastAtMost[THEN sym]
+        apply(subst sumr_geometric) using goal1 by auto
+      finally show "?goal" by auto qed qed qed
+
+lemma has_integral_spike: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "negligible s" "(\<forall>x\<in>(t - s). g x = f x)" "(f has_integral y) t"
+  shows "(g has_integral y) t"
+proof- { fix a b::"real^'n" and f g ::"real^'n \<Rightarrow> 'a" and y::'a
+    assume as:"\<forall>x \<in> {a..b} - s. g x = f x" "(f has_integral y) {a..b}"
+    have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) {a..b}" apply(rule has_integral_add[OF as(2)])
+      apply(rule has_integral_negligible[OF assms(1)]) using as by auto
+    hence "(g has_integral y) {a..b}" by auto } note * = this
+  show ?thesis apply(subst has_integral_alt) using assms(2-) apply-apply(rule cond_cases,safe)
+    apply(rule *, assumption+) apply(subst(asm) has_integral_alt) unfolding if_not_P
+    apply(erule_tac x=e in allE,safe,rule_tac x=B in exI,safe) apply(erule_tac x=a in allE,erule_tac x=b in allE,safe)
+    apply(rule_tac x=z in exI,safe) apply(rule *[where fa2="\<lambda>x. if x\<in>t then f x else 0"]) by auto qed
+
+lemma has_integral_spike_eq:
+  assumes "negligible s" "\<forall>x\<in>(t - s). g x = f x"
+  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
+  apply rule apply(rule_tac[!] has_integral_spike[OF assms(1)]) using assms(2) by auto
+
+lemma integrable_spike: assumes "negligible s" "\<forall>x\<in>(t - s). g x = f x" "f integrable_on t"
+  shows "g integrable_on  t"
+  using assms unfolding integrable_on_def apply-apply(erule exE)
+  apply(rule,rule has_integral_spike) by fastsimp+
+
+lemma integral_spike: assumes "negligible s" "\<forall>x\<in>(t - s). g x = f x"
+  shows "integral t f = integral t g"
+  unfolding integral_def using has_integral_spike_eq[OF assms] by auto
+
+subsection {* Some other trivialities about negligible sets. *}
+
+lemma negligible_subset[intro]: assumes "negligible s" "t \<subseteq> s" shows "negligible t" unfolding negligible_def 
+proof(safe) case goal1 show ?case using assms(1)[unfolded negligible_def,rule_format,of a b]
+    apply-apply(rule has_integral_spike[OF assms(1)]) defer apply assumption
+    using assms(2) unfolding indicator_def by auto qed
+
+lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible(s - t)" using assms by auto
+
+lemma negligible_inter: assumes "negligible s \<or> negligible t" shows "negligible(s \<inter> t)" using assms by auto
+
+lemma negligible_union: assumes "negligible s" "negligible t" shows "negligible (s \<union> t)" unfolding negligible_def 
+proof safe case goal1 note assm = assms[unfolded negligible_def,rule_format,of a b]
+  thus ?case apply(subst has_integral_spike_eq[OF assms(2)])
+    defer apply assumption unfolding indicator_def by auto qed
+
+lemma negligible_union_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> (negligible s \<and> negligible t)"
+  using negligible_union by auto
+
+lemma negligible_sing[intro]: "negligible {a::real^'n}" 
+proof- guess x using UNIV_witness[where 'a='n] ..
+  show ?thesis using negligible_standard_hyperplane[of x "a$x"] by auto qed
+
+lemma negligible_insert[simp]: "negligible(insert a s) \<longleftrightarrow> negligible s"
+  apply(subst insert_is_Un) unfolding negligible_union_eq by auto
+
+lemma negligible_empty[intro]: "negligible {}" by auto
+
+lemma negligible_finite[intro]: assumes "finite s" shows "negligible s"
+  using assms apply(induct s) by auto
+
+lemma negligible_unions[intro]: assumes "finite s" "\<forall>t\<in>s. negligible t" shows "negligible(\<Union>s)"
+  using assms by(induct,auto) 
+
+lemma negligible:  "negligible s \<longleftrightarrow> (\<forall>t::(real^'n) set. (indicator s has_integral 0) t)"
+  apply safe defer apply(subst negligible_def)
+proof- fix t::"(real^'n) set" assume as:"negligible s"
+  have *:"(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)" by(rule ext,auto)  
+  show "(indicator s has_integral 0) t" apply(subst has_integral_alt)
+    apply(cases,subst if_P,assumption) unfolding if_not_P apply(safe,rule as[unfolded negligible_def,rule_format])
+    apply(rule_tac x=1 in exI) apply(safe,rule zero_less_one) apply(rule_tac x=0 in exI)
+    using negligible_subset[OF as,of "s \<inter> t"] unfolding negligible_def indicator_def unfolding * by auto qed auto
+
+subsection {* Finite case of the spike theorem is quite commonly needed. *}
+
+lemma has_integral_spike_finite: assumes "finite s" "\<forall>x\<in>t-s. g x = f x" 
+  "(f has_integral y) t" shows "(g has_integral y) t"
+  apply(rule has_integral_spike) using assms by auto
+
+lemma has_integral_spike_finite_eq: assumes "finite s" "\<forall>x\<in>t-s. g x = f x"
+  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
+  apply rule apply(rule_tac[!] has_integral_spike_finite) using assms by auto
+
+lemma integrable_spike_finite:
+  assumes "finite s" "\<forall>x\<in>t-s. g x = f x" "f integrable_on t" shows "g integrable_on  t"
+  using assms unfolding integrable_on_def apply safe apply(rule_tac x=y in exI)
+  apply(rule has_integral_spike_finite) by auto
+
+subsection {* In particular, the boundary of an interval is negligible. *}
+
+lemma negligible_frontier_interval: "negligible({a..b} - {a<..<b})"
+proof- let ?A = "\<Union>((\<lambda>k. {x. x$k = a$k} \<union> {x. x$k = b$k}) ` UNIV)"
+  have "{a..b} - {a<..<b} \<subseteq> ?A" apply rule unfolding Diff_iff mem_interval not_all
+    apply(erule conjE exE)+ apply(rule_tac X="{x. x $ xa = a $ xa} \<union> {x. x $ xa = b $ xa}" in UnionI)
+    apply(erule_tac[!] x=xa in allE) by auto
+  thus ?thesis apply-apply(rule negligible_subset[of ?A]) apply(rule negligible_unions[OF finite_imageI]) by auto qed
+
+lemma has_integral_spike_interior:
+  assumes "\<forall>x\<in>{a<..<b}. g x = f x" "(f has_integral y) ({a..b})" shows "(g has_integral y) ({a..b})"
+  apply(rule has_integral_spike[OF negligible_frontier_interval _ assms(2)]) using assms(1) by auto
+
+lemma has_integral_spike_interior_eq:
+  assumes "\<forall>x\<in>{a<..<b}. g x = f x" shows "((f has_integral y) ({a..b}) \<longleftrightarrow> (g has_integral y) ({a..b}))"
+  apply rule apply(rule_tac[!] has_integral_spike_interior) using assms by auto
+
+lemma integrable_spike_interior: assumes "\<forall>x\<in>{a<..<b}. g x = f x" "f integrable_on {a..b}" shows "g integrable_on {a..b}"
+  using  assms unfolding integrable_on_def using has_integral_spike_interior[OF assms(1)] by auto
+
+subsection {* Integrability of continuous functions. *}
+
+lemma neutral_and[simp]: "neutral op \<and> = True"
+  unfolding neutral_def apply(rule some_equality) by auto
+
+lemma monoidal_and[intro]: "monoidal op \<and>" unfolding monoidal_def by auto
+
+lemma iterate_and[simp]: assumes "finite s" shows "(iterate op \<and>) s p \<longleftrightarrow> (\<forall>x\<in>s. p x)" using assms
+apply induct unfolding iterate_insert[OF monoidal_and] by auto
+
+lemma operative_division_and: assumes "operative op \<and> P" "d division_of {a..b}"
+  shows "(\<forall>i\<in>d. P i) \<longleftrightarrow> P {a..b}"
+  using operative_division[OF monoidal_and assms] division_of_finite[OF assms(2)] by auto
+
+lemma operative_approximable: assumes "0 \<le> e" fixes f::"real^'n \<Rightarrow> 'a::banach"
+  shows "operative op \<and> (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::real^'n)) \<le> e) \<and> g integrable_on i)" unfolding operative_def neutral_and
+proof safe fix a b::"real^'n" { assume "content {a..b} = 0"
+    thus "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" 
+      apply(rule_tac x=f in exI) using assms by(auto intro!:integrable_on_null) }
+  { fix c k g assume as:"\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}"
+    show "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. x $ k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. x $ k \<le> c}"
+      "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $ k}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. c \<le> x $ k}"
+      apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2)] by auto }
+  fix c k g1 g2 assume as:"\<forall>x\<in>{a..b} \<inter> {x. x $ k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on {a..b} \<inter> {x. x $ k \<le> c}"
+                          "\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $ k}. norm (f x - g2 x) \<le> e" "g2 integrable_on {a..b} \<inter> {x. c \<le> x $ k}"
+  let ?g = "\<lambda>x. if x$k = c then f x else if x$k \<le> c then g1 x else g2 x"
+  show "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" apply(rule_tac x="?g" in exI)
+  proof safe case goal1 thus ?case apply- apply(cases "x$k=c", case_tac "x$k < c") using as assms by auto
+  next case goal2 presume "?g integrable_on {a..b} \<inter> {x. x $ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $ k \<ge> c}"
+    then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this]
+    show ?case unfolding integrable_on_def by auto
+  next show "?g integrable_on {a..b} \<inter> {x. x $ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $ k \<ge> c}"
+      apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using as(2,4) by auto qed qed
+
+lemma approximable_on_division: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  assumes "0 \<le> e" "d division_of {a..b}" "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
+  obtains g where "\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}"
+proof- note * = operative_division[OF monoidal_and operative_approximable[OF assms(1)] assms(2)]
+  note this[unfolded iterate_and[OF division_of_finite[OF assms(2)]]] from assms(3)[unfolded this[of f]]
+  guess g .. thus thesis apply-apply(rule that[of g]) by auto qed
+
+lemma integrable_continuous: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}"
+proof(rule integrable_uniform_limit,safe) fix e::real assume e:"0 < e"
+  from compact_uniformly_continuous[OF assms compact_interval,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d ..
+  note d=conjunctD2[OF this,rule_format]
+  from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this
+  note p' = tagged_division_ofD[OF p(1)]
+  have *:"\<forall>i\<in>snd ` p. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
+  proof(safe,unfold snd_conv) fix x l assume as:"(x,l) \<in> p" 
+    from p'(4)[OF this] guess a b apply-by(erule exE)+ note l=this
+    show "\<exists>g. (\<forall>x\<in>l. norm (f x - g x) \<le> e) \<and> g integrable_on l" apply(rule_tac x="\<lambda>y. f x" in exI)
+    proof safe show "(\<lambda>y. f x) integrable_on l" unfolding integrable_on_def l by(rule,rule has_integral_const)
+      fix y assume y:"y\<in>l" note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this]
+      note d(2)[OF _ _ this[unfolded mem_ball]]
+      thus "norm (f y - f x) \<le> e" using y p'(2-3)[OF as] unfolding vector_dist_norm l norm_minus_commute by fastsimp qed qed
+  from e have "0 \<le> e" by auto from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g .
+  thus "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" by auto qed 
+
+subsection {* Specialization of additivity to one dimension. *}
+
+lemma operative_1_lt: assumes "monoidal opp"
+  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real^1} = neutral opp) \<and>
+                (\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
+  unfolding operative_def content_eq_0_1 forall_1 vector_le_def vector_less_def
+proof safe fix a b c::"real^1" assume as:"\<forall>a b c. f {a..b} = opp (f ({a..b} \<inter> {x. x $ 1 \<le> c})) (f ({a..b} \<inter> {x. c \<le> x $ 1}))" "a $ 1 < c $ 1" "c $ 1 < b $ 1"
+    from this(2-) have "{a..b} \<inter> {x. x $ 1 \<le> c $ 1} = {a..c}" "{a..b} \<inter> {x. x $ 1 \<ge> c $ 1} = {c..b}" by auto
+    thus "opp (f {a..c}) (f {c..b}) = f {a..b}" unfolding as(1)[rule_format,of a b "c$1"] by auto
+next fix a b::"real^1" and c::real
+  assume as:"\<forall>a b. b $ 1 \<le> a $ 1 \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a $ 1 < c $ 1 \<and> c $ 1 < b $ 1 \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}"
+  show "f {a..b} = opp (f ({a..b} \<inter> {x. x $ 1 \<le> c})) (f ({a..b} \<inter> {x. c \<le> x $ 1}))"
+  proof(cases "c \<in> {a$1 .. b$1}")
+    case False hence "c<a$1 \<or> c>b$1" by auto
+    thus ?thesis apply-apply(erule disjE)
+    proof- assume "c<a$1" hence *:"{a..b} \<inter> {x. x $ 1 \<le> c} = {1..0}"  "{a..b} \<inter> {x. c \<le> x $ 1} = {a..b}" by auto
+      show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
+    next   assume "b$1<c" hence *:"{a..b} \<inter> {x. x $ 1 \<le> c} = {a..b}"  "{a..b} \<inter> {x. c \<le> x $ 1} = {1..0}" by auto
+      show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
+    qed
+  next case True hence *:"min (b $ 1) c = c" "max (a $ 1) c = c" by auto
+    show ?thesis unfolding interval_split num1_eq_iff if_True * vec_def[THEN sym]
+    proof(cases "c = a$1 \<or> c = b$1")
+      case False thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})"
+        apply-apply(subst as(2)[rule_format]) using True by auto
+    next case True thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})" apply-
+      proof(erule disjE) assume "c=a$1" hence *:"a = vec1 c" unfolding Cart_eq by auto 
+        hence "f {a..vec1 c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
+        thus ?thesis using assms unfolding * by auto
+      next assume "c=b$1" hence *:"b = vec1 c" unfolding Cart_eq by auto 
+        hence "f {vec1 c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
+        thus ?thesis using assms unfolding * by auto qed qed qed qed
+
+lemma operative_1_le: assumes "monoidal opp"
+  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real^1} = neutral opp) \<and>
+                (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
+unfolding operative_1_lt[OF assms]
+proof safe fix a b c::"real^1" assume as:"\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a < c" "c < b"
+  show "opp (f {a..c}) (f {c..b}) = f {a..b}" apply(rule as(1)[rule_format]) using as(2-) unfolding vector_le_def vector_less_def by auto
+next fix a b c ::"real^1"
+  assume "\<forall>a b. b \<le> a \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a \<le> c" "c \<le> b"
+  note as = this[rule_format]
+  show "opp (f {a..c}) (f {c..b}) = f {a..b}"
+  proof(cases "c = a \<or> c = b")
+    case False thus ?thesis apply-apply(subst as(2)) using as(3-) unfolding vector_le_def vector_less_def Cart_eq by(auto simp del:dest_vec1_eq)
+    next case True thus ?thesis apply-
+      proof(erule disjE) assume *:"c=a" hence "f {a..c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
+        thus ?thesis using assms unfolding * by auto
+      next               assume *:"c=b" hence "f {c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
+        thus ?thesis using assms unfolding * by auto qed qed qed 
+
+subsection {* Special case of additivity we need for the FCT. *}
+
+lemma additive_tagged_division_1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
+  assumes "dest_vec1 a \<le> dest_vec1 b" "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a"
+proof- let ?f = "(\<lambda>k::(real^1) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
+  have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty_1
+    by(auto simp add:not_less interval_bound_1 vector_less_def)
+  have **:"{a..b} \<noteq> {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)]
+  note * = this[unfolded if_not_P[OF **] interval_bound_1[OF assms(1)],THEN sym ]
+  show ?thesis unfolding * apply(subst setsum_iterate[THEN sym]) defer
+    apply(rule setsum_cong2) unfolding split_paired_all split_conv using assms(2) by auto qed
+
+subsection {* A useful lemma allowing us to factor out the content size. *}
+
+lemma has_integral_factor_content:
+  "(f has_integral i) {a..b} \<longleftrightarrow> (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p
+    \<longrightarrow> norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a..b}))"
+proof(cases "content {a..b} = 0")
+  case True show ?thesis unfolding has_integral_null_eq[OF True] apply safe
+    apply(rule,rule,rule gauge_trivial,safe) unfolding setsum_content_null[OF True] True defer 
+    apply(erule_tac x=1 in allE,safe) defer apply(rule fine_division_exists[of _ a b],assumption)
+    apply(erule_tac x=p in allE) unfolding setsum_content_null[OF True] by auto
+next case False note F = this[unfolded content_lt_nz[THEN sym]]
+  let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)"
+  show ?thesis apply(subst has_integral)
+  proof safe fix e::real assume e:"e>0"
+    { assume "\<forall>e>0. ?P e op <" thus "?P (e * content {a..b}) op \<le>" apply(erule_tac x="e * content {a..b}" in allE)
+        apply(erule impE) defer apply(erule exE,rule_tac x=d in exI)
+        using F e by(auto simp add:field_simps intro:mult_pos_pos) }
+    {  assume "\<forall>e>0. ?P (e * content {a..b}) op \<le>" thus "?P e op <" apply(erule_tac x="e / 2 / content {a..b}" in allE)
+        apply(erule impE) defer apply(erule exE,rule_tac x=d in exI)
+        using F e by(auto simp add:field_simps intro:mult_pos_pos) } qed qed
+
+subsection {* Fundamental theorem of calculus. *}
+
+lemma fundamental_theorem_of_calculus: fixes f::"real^1 \<Rightarrow> 'a::banach"
+  assumes "a \<le> b"  "\<forall>x\<in>{a..b}. ((f o vec1) has_vector_derivative f'(vec1 x)) (at x within {a..b})"
+  shows "(f' has_integral (f(vec1 b) - f(vec1 a))) ({vec1 a..vec1 b})"
+unfolding has_integral_factor_content
+proof safe fix e::real assume e:"e>0" have ab:"dest_vec1 (vec1 a) \<le> dest_vec1 (vec1 b)" using assms(1) by auto
+  note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt]
+  have *:"\<And>P Q. \<forall>x\<in>{a..b}. P x \<and> (\<forall>e>0. \<exists>d>0. Q x e d) \<Longrightarrow> \<forall>x. \<exists>(d::real)>0. x\<in>{a..b} \<longrightarrow> Q x e d" using e by blast
+  note this[OF assm,unfolded gauge_existence_lemma] from choice[OF this,unfolded Ball_def[symmetric]]
+  guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
+  show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
+                 norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \<le> e * content {vec1 a..vec1 b})"
+    apply(rule_tac x="\<lambda>x. ball x (d (dest_vec1 x))" in exI,safe)
+    apply(rule gauge_ball_dependent,rule,rule d(1))
+  proof- fix p assume as:"p tagged_division_of {vec1 a..vec1 b}" "(\<lambda>x. ball x (d (dest_vec1 x))) fine p"
+    show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \<le> e * content {vec1 a..vec1 b}" 
+      unfolding content_1[OF ab] additive_tagged_division_1[OF ab as(1),of f,THEN sym]
+      unfolding vector_minus_component[THEN sym] additive_tagged_division_1[OF ab as(1),of "\<lambda>x. x",THEN sym]
+      apply(subst dest_vec1_setsum) unfolding setsum_right_distrib defer unfolding setsum_subtractf[THEN sym] 
+    proof(rule setsum_norm_le,safe) fix x k assume "(x,k)\<in>p"
+      note xk = tagged_division_ofD(2-4)[OF as(1) this] from this(3) guess u v apply-by(erule exE)+ note k=this
+      have *:"dest_vec1 u \<le> dest_vec1 v" using xk unfolding k by auto
+      have ball:"\<forall>xa\<in>k. xa \<in> ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
+      have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \<le> norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)"
+        apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
+        unfolding scaleR.diff_left by(auto simp add:group_simps)
+      also have "... \<le> e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)"
+        apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4
+        apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1])
+        using ball[rule_format,of u] ball[rule_format,of v] 
+        using xk(1-2) unfolding k subset_eq by(auto simp add:vector_dist_norm norm_real)
+      also have "... \<le> e * dest_vec1 (interval_upperbound k - interval_lowerbound k)"
+        unfolding k interval_bound_1[OF *] using xk(1) unfolding k by(auto simp add:vector_dist_norm norm_real field_simps)
+      finally show "norm (content k *\<^sub>R f' x - (f (interval_upperbound k) - f (interval_lowerbound k))) \<le>
+        e * dest_vec1 (interval_upperbound k - interval_lowerbound k)" unfolding k interval_bound_1[OF *] content_1[OF *] .
+    qed(insert as, auto) qed qed
+
+subsection {* Attempt a systematic general set of "offset" results for components. *}
+
+lemma gauge_modify:
+  assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
+  shows "gauge (\<lambda>x y. d (f x) (f y))"
+  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) unfolding mem_def Collect_def by auto
+
+subsection {* Only need trivial subintervals if the interval itself is trivial. *}
+
+lemma division_of_nontrivial: fixes s::"(real^'n) set set"
+  assumes "s division_of {a..b}" "content({a..b}) \<noteq> 0"
+  shows "{k. k \<in> s \<and> content k \<noteq> 0} division_of {a..b}" using assms(1) apply-
+proof(induct "card s" arbitrary:s rule:nat_less_induct)
+  fix s::"(real^'n) set set" assume assm:"s division_of {a..b}"
+    "\<forall>m<card s. \<forall>x. m = card x \<longrightarrow> x division_of {a..b} \<longrightarrow> {k \<in> x. content k \<noteq> 0} division_of {a..b}" 
+  note s = division_ofD[OF assm(1)] let ?thesis = "{k \<in> s. content k \<noteq> 0} division_of {a..b}"
+  { presume *:"{k \<in> s. content k \<noteq> 0} \<noteq> s \<Longrightarrow> ?thesis"
+    show ?thesis apply cases defer apply(rule *,assumption) using assm(1) by auto }
+  assume noteq:"{k \<in> s. content k \<noteq> 0} \<noteq> s"
+  then obtain k where k:"k\<in>s" "content k = 0" by auto
+  from s(4)[OF k(1)] guess c d apply-by(erule exE)+ note k=k this
+  from k have "card s > 0" unfolding card_gt_0_iff using assm(1) by auto
+  hence card:"card (s - {k}) < card s" using assm(1) k(1) apply(subst card_Diff_singleton_if) by auto
+  have *:"closed (\<Union>(s - {k}))" apply(rule closed_Union) defer apply rule apply(drule DiffD1,drule s(4))
+    apply safe apply(rule closed_interval) using assm(1) by auto
+  have "k \<subseteq> \<Union>(s - {k})" apply safe apply(rule *[unfolded closed_limpt,rule_format]) unfolding islimpt_approachable
+  proof safe fix x and e::real assume as:"x\<in>k" "e>0"
+    from k(2)[unfolded k content_eq_0] guess i .. 
+    hence i:"c$i = d$i" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by smt
+    hence xi:"x$i = d$i" using as unfolding k mem_interval by smt
+    def y \<equiv> "(\<chi> j. if j = i then if c$i \<le> (a$i + b$i) / 2 then c$i + min e (b$i - c$i) / 2 else c$i - min e (c$i - a$i) / 2 else x$j)"
+    show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> dist x' x < e" apply(rule_tac x=y in bexI) 
+    proof have "d \<in> {c..d}" using s(3)[OF k(1)] unfolding k interval_eq_empty mem_interval by(fastsimp simp add: not_less)
+      hence "d \<in> {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN spec[where x=i]]
+      hence xyi:"y$i \<noteq> x$i" unfolding y_def unfolding i xi Cart_lambda_beta if_P[OF refl]
+        apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2) using assms(2)[unfolded content_eq_0] by smt+ 
+      thus "y \<noteq> x" unfolding Cart_eq by auto
+      have *:"UNIV = insert i (UNIV - {i})" by auto
+      have "norm (y - x) < e + setsum (\<lambda>i. 0) (UNIV::'n set)" apply(rule le_less_trans[OF norm_le_l1])
+        apply(subst *,subst setsum_insert) prefer 3 apply(rule add_less_le_mono)
+      proof- show "\<bar>(y - x) $ i\<bar> < e" unfolding y_def Cart_lambda_beta vector_minus_component if_P[OF refl]
+          apply(cases) apply(subst if_P,assumption) unfolding if_not_P unfolding i xi using di as(2) by auto
+        show "(\<Sum>i\<in>UNIV - {i}. \<bar>(y - x) $ i\<bar>) \<le> (\<Sum>i\<in>UNIV. 0)" unfolding y_def by auto 
+      qed auto thus "dist y x < e" unfolding vector_dist_norm by auto
+      have "y\<notin>k" unfolding k mem_interval apply rule apply(erule_tac x=i in allE) using xyi unfolding k i xi by auto
+      moreover have "y \<in> \<Union>s" unfolding s mem_interval
+      proof note simps = y_def Cart_lambda_beta if_not_P
+        fix j::'n show "a $ j \<le> y $ j \<and> y $ j \<le> b $ j" 
+        proof(cases "j = i") case False have "x \<in> {a..b}" using s(2)[OF k(1)] as(1) by auto
+          thus ?thesis unfolding simps if_not_P[OF False] unfolding mem_interval by auto
+        next case True note T = this show ?thesis
+          proof(cases "c $ i \<le> (a $ i + b $ i) / 2")
+            case True show ?thesis unfolding simps if_P[OF T] if_P[OF True] unfolding i
+              using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) 
+          next case False thus ?thesis unfolding simps if_P[OF T] if_not_P[OF False] unfolding i
+              using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps)
+          qed qed qed
+      ultimately show "y \<in> \<Union>(s - {k})" by auto
+    qed qed hence "\<Union>(s - {k}) = {a..b}" unfolding s(6)[THEN sym] by auto
+  hence  "{ka \<in> s - {k}. content ka \<noteq> 0} division_of {a..b}" apply-apply(rule assm(2)[rule_format,OF card refl])
+    apply(rule division_ofI) defer apply(rule_tac[1-4] s) using assm(1) by auto
+  moreover have "{ka \<in> s - {k}. content ka \<noteq> 0} = {k \<in> s. content k \<noteq> 0}" using k by auto ultimately show ?thesis by auto qed
+
+subsection {* Integrabibility on subintervals. *}
+
+lemma operative_integrable: fixes f::"real^'n \<Rightarrow> 'a::banach" shows
+  "operative op \<and> (\<lambda>i. f integrable_on i)"
+  unfolding operative_def neutral_and apply safe apply(subst integrable_on_def)
+  unfolding has_integral_null_eq apply(rule,rule refl) apply(rule,assumption)+
+  unfolding integrable_on_def by(auto intro: has_integral_split)
+
+lemma integrable_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach" 
+  assumes "f integrable_on {a..b}" "{c..d} \<subseteq> {a..b}" shows "f integrable_on {c..d}" 
+  apply(cases "{c..d} = {}") defer apply(rule partial_division_extend_1[OF assms(2)],assumption)
+  using operative_division_and[OF operative_integrable,THEN sym,of _ _ _ f] assms(1) by auto
+
+subsection {* Combining adjacent intervals in 1 dimension. *}
+
+lemma has_integral_combine: assumes "(a::real^1) \<le> c" "c \<le> b"
+  "(f has_integral i) {a..c}" "(f has_integral (j::'a::banach)) {c..b}"
+  shows "(f has_integral (i + j)) {a..b}"
+proof- note operative_integral[of f, unfolded operative_1_le[OF monoidal_lifted[OF monoidal_monoid]]]
+  note conjunctD2[OF this,rule_format] note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]]
+  hence "f integrable_on {a..b}" apply- apply(rule ccontr) apply(subst(asm) if_P) defer
+    apply(subst(asm) if_P) using assms(3-) by auto
+  with * show ?thesis apply-apply(subst(asm) if_P) defer apply(subst(asm) if_P) defer apply(subst(asm) if_P)
+    unfolding lifted.simps using assms(3-) by(auto simp add: integrable_on_def integral_unique) qed
+
+lemma integral_combine: fixes f::"real^1 \<Rightarrow> 'a::banach"
+  assumes "a \<le> c" "c \<le> b" "f integrable_on ({a..b})"
+  shows "integral {a..c} f + integral {c..b} f = integral({a..b}) f"
+  apply(rule integral_unique[THEN sym]) apply(rule has_integral_combine[OF assms(1-2)])
+  apply(rule_tac[!] integrable_integral integrable_subinterval[OF assms(3)])+ using assms(1-2) by auto
+
+lemma integrable_combine: fixes f::"real^1 \<Rightarrow> 'a::banach"
+  assumes "a \<le> c" "c \<le> b" "f integrable_on {a..c}" "f integrable_on {c..b}"
+  shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by(fastsimp intro!:has_integral_combine)
+
+subsection {* Reduce integrability to "local" integrability. *}
+
+lemma integrable_on_little_subintervals: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  assumes "\<forall>x\<in>{a..b}. \<exists>d>0. \<forall>u v. x \<in> {u..v} \<and> {u..v} \<subseteq> ball x d \<and> {u..v} \<subseteq> {a..b} \<longrightarrow> f integrable_on {u..v}"
+  shows "f integrable_on {a..b}"
+proof- have "\<forall>x. \<exists>d. x\<in>{a..b} \<longrightarrow> d>0 \<and> (\<forall>u v. x \<in> {u..v} \<and> {u..v} \<subseteq> ball x d \<and> {u..v} \<subseteq> {a..b} \<longrightarrow> f integrable_on {u..v})"
+    using assms by auto note this[unfolded gauge_existence_lemma] from choice[OF this] guess d .. note d=this[rule_format]
+  guess p apply(rule fine_division_exists[OF gauge_ball_dependent,of d a b]) using d by auto note p=this(1-2)
+  note division_of_tagged_division[OF this(1)] note * = operative_division_and[OF operative_integrable,OF this,THEN sym,of f]
+  show ?thesis unfolding * apply safe unfolding snd_conv
+  proof- fix x k assume "(x,k) \<in> p" note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this]
+    thus "f integrable_on k" apply safe apply(rule d[THEN conjunct2,rule_format,of x]) by auto qed qed
+
+subsection {* Second FCT or existence of antiderivative. *}
+
+lemma integrable_const[intro]:"(\<lambda>x. c) integrable_on {a..b}"
+  unfolding integrable_on_def by(rule,rule has_integral_const)
+
+lemma integral_has_vector_derivative: fixes f::"real \<Rightarrow> 'a::banach"
+  assumes "continuous_on {a..b} f" "x \<in> {a..b}"
+  shows "((\<lambda>u. integral {vec a..vec u} (f o dest_vec1)) has_vector_derivative f(x)) (at x within {a..b})"
+  unfolding has_vector_derivative_def has_derivative_within_alt
+apply safe apply(rule scaleR.bounded_linear_left)
+proof- fix e::real assume e:"e>0"
+  note compact_uniformly_continuous[OF assms(1) compact_real_interval,unfolded uniformly_continuous_on_def]
+  from this[rule_format,OF e] guess d apply-by(erule conjE exE)+ note d=this[rule_format]
+  let ?I = "\<lambda>a b. integral {vec1 a..vec1 b} (f \<circ> dest_vec1)"
+  show "\<exists>d>0. \<forall>y\<in>{a..b}. norm (y - x) < d \<longrightarrow> norm (?I a y - ?I a x - (y - x) *\<^sub>R f x) \<le> e * norm (y - x)"
+  proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x")
+      case False have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous)
+        apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
+      hence *:"?I a y - ?I a x = ?I x y" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+        using False unfolding not_less using assms(2) goal1 by auto
+      have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto
+      show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
+        defer apply(rule has_integral_sub) apply(rule integrable_integral)
+        apply(rule integrable_subinterval,rule integrable_continuous) apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+
+      proof- show "{vec1 x..vec1 y} \<subseteq> {vec1 a..vec1 b}" using goal1 assms(2) by auto
+        have *:"y - x = norm(y - x)" using False by auto
+        show "((\<lambda>xa. f x) has_integral (y - x) *\<^sub>R f x) {vec1 x..vec1 y}" apply(subst *) unfolding ** by auto
+        show "\<forall>xa\<in>{vec1 x..vec1 y}. norm (f (dest_vec1 xa) - f x) \<le> e" apply safe apply(rule less_imp_le)
+          apply(rule d(2)[unfolded vector_dist_norm]) using assms(2) using goal1 by auto
+      qed(insert e,auto)
+    next case True have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous)
+        apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
+      hence *:"?I a x - ?I a y = ?I y x" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+        using True using assms(2) goal1 by auto
+      have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto
+      have ***:"\<And>fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto 
+      show ?thesis apply(subst ***) unfolding norm_minus_cancel **
+        apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
+        defer apply(rule has_integral_sub) apply(subst minus_minus[THEN sym]) unfolding minus_minus
+        apply(rule integrable_integral) apply(rule integrable_subinterval,rule integrable_continuous)
+        apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+
+      proof- show "{vec1 y..vec1 x} \<subseteq> {vec1 a..vec1 b}" using goal1 assms(2) by auto
+        have *:"x - y = norm(y - x)" using True by auto
+        show "((\<lambda>xa. f x) has_integral (x - y) *\<^sub>R f x) {vec1 y..vec1 x}" apply(subst *) unfolding ** by auto
+        show "\<forall>xa\<in>{vec1 y..vec1 x}. norm (f (dest_vec1 xa) - f x) \<le> e" apply safe apply(rule less_imp_le)
+          apply(rule d(2)[unfolded vector_dist_norm]) using assms(2) using goal1 by auto
+      qed(insert e,auto) qed qed qed
+
+lemma integral_has_vector_derivative': fixes f::"real^1 \<Rightarrow> 'a::banach"
+  assumes "continuous_on {a..b} f" "x \<in> {a..b}"
+  shows "((\<lambda>u. (integral {a..vec u} f)) has_vector_derivative f x) (at (x$1) within {a$1..b$1})"
+  using integral_has_vector_derivative[OF continuous_on_o_vec1[OF assms(1)], of "x$1"]
+  unfolding o_def vec1_dest_vec1 using assms(2) by auto
+
+lemma antiderivative_continuous: assumes "continuous_on {a..b::real} f"
+  obtains g where "\<forall>x\<in> {a..b}. (g has_vector_derivative (f(x)::_::banach)) (at x within {a..b})"
+  apply(rule that,rule) using integral_has_vector_derivative[OF assms] by auto
+
+subsection {* Combined fundamental theorem of calculus. *}
+
+lemma antiderivative_integral_continuous: fixes f::"real \<Rightarrow> 'a::banach" assumes "continuous_on {a..b} f"
+  obtains g where "\<forall>u\<in>{a..b}. \<forall>v \<in> {a..b}. u \<le> v \<longrightarrow> ((f o dest_vec1) has_integral (g v - g u)) {vec u..vec v}"
+proof- from antiderivative_continuous[OF assms] guess g . note g=this
+  show ?thesis apply(rule that[of g])
+  proof safe case goal1 have "\<forall>x\<in>{u..v}. (g has_vector_derivative f x) (at x within {u..v})"
+      apply(rule,rule has_vector_derivative_within_subset) apply(rule g[rule_format]) using goal1(1-2) by auto
+    thus ?case using fundamental_theorem_of_calculus[OF goal1(3),of "g o dest_vec1" "f o dest_vec1"]
+      unfolding o_def vec1_dest_vec1 by auto qed qed
+
+subsection {* General "twiddling" for interval-to-interval function image. *}
+
+lemma has_integral_twiddle:
+  assumes "0 < r" "\<forall>x. h(g x) = x" "\<forall>x. g(h x) = x" "\<forall>x. continuous (at x) g"
+  "\<forall>u v. \<exists>w z. g ` {u..v} = {w..z}"
+  "\<forall>u v. \<exists>w z. h ` {u..v} = {w..z}"
+  "\<forall>u v. content(g ` {u..v}) = r * content {u..v}"
+  "(f has_integral i) {a..b}"
+  shows "((\<lambda>x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` {a..b})"
+proof- { presume *:"{a..b} \<noteq> {} \<Longrightarrow> ?thesis"
+    show ?thesis apply cases defer apply(rule *,assumption)
+    proof- case goal1 thus ?thesis unfolding goal1 assms(8)[unfolded goal1 has_integral_empty_eq] by auto qed }
+  assume "{a..b} \<noteq> {}" from assms(6)[rule_format,of a b] guess w z apply-by(erule exE)+ note wz=this
+  have inj:"inj g" "inj h" unfolding inj_on_def apply safe apply(rule_tac[!] ccontr)
+    using assms(2) apply(erule_tac x=x in allE) using assms(2) apply(erule_tac x=y in allE) defer
+    using assms(3) apply(erule_tac x=x in allE) using assms(3) apply(erule_tac x=y in allE) by auto
+  show ?thesis unfolding has_integral_def has_integral_compact_interval_def apply(subst if_P) apply(rule,rule,rule wz)
+  proof safe fix e::real assume e:"e>0" hence "e * r > 0" using assms(1) by(rule mult_pos_pos)
+    from assms(8)[unfolded has_integral,rule_format,OF this] guess d apply-by(erule exE conjE)+ note d=this[rule_format]
+    def d' \<equiv> "\<lambda>x y. d (g x) (g y)" have d':"\<And>x. d' x = {y. g y \<in> (d (g x))}" unfolding d'_def by(auto simp add:mem_def)
+    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)"
+    proof(rule_tac x=d' in exI,safe) show "gauge d'" using d(1) unfolding gauge_def d' using continuous_open_preimage_univ[OF assms(4)] by auto
+      fix p assume as:"p tagged_division_of h ` {a..b}" "d' fine p" note p = tagged_division_ofD[OF as(1)] 
+      have "(\<lambda>(x, k). (g x, g ` k)) ` p tagged_division_of {a..b} \<and> d fine (\<lambda>(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of 
+      proof safe show "finite ((\<lambda>(x, k). (g x, g ` k)) ` p)" using as by auto
+        show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p" using as(2) unfolding fine_def d' by auto
+        fix x k assume xk[intro]:"(x,k) \<in> p" show "g x \<in> g ` k" using p(2)[OF xk] by auto
+        show "\<exists>u v. g ` k = {u..v}" using p(4)[OF xk] using assms(5-6) by auto
+        { fix y assume "y \<in> k" thus "g y \<in> {a..b}" "g y \<in> {a..b}" using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"]
+            using assms(2)[rule_format,of y] unfolding inj_image_mem_iff[OF inj(2)] by auto }
+        fix x' k' assume xk':"(x',k') \<in> p" fix z assume "z \<in> interior (g ` k)" "z \<in> interior (g ` k')"
+        hence *:"interior (g ` k) \<inter> interior (g ` k') \<noteq> {}" by auto
+        have same:"(x, k) = (x', k')" apply-apply(rule ccontr,drule p(5)[OF xk xk'])
+        proof- assume as:"interior k \<inter> interior k' = {}" from nonempty_witness[OF *] guess z .
+          hence "z \<in> g ` (interior k \<inter> interior k')" using interior_image_subset[OF assms(4) inj(1)]
+            unfolding image_Int[OF inj(1)] by auto thus False using as by blast
+        qed thus "g x = g x'" by auto
+        { fix z assume "z \<in> k"  thus  "g z \<in> g ` k'" using same by auto }
+        { fix z assume "z \<in> k'" thus  "g z \<in> g ` k"  using same by auto }
+      next fix x assume "x \<in> {a..b}" hence "h x \<in>  \<Union>{k. \<exists>x. (x, k) \<in> p}" using p(6) by auto
+        then guess X unfolding Union_iff .. note X=this from this(1) guess y unfolding mem_Collect_eq ..
+        thus "x \<in> \<Union>{k. \<exists>x. (x, k) \<in> (\<lambda>(x, k). (g x, g ` k)) ` p}" apply-
+          apply(rule_tac X="g ` X" in UnionI) defer apply(rule_tac x="h x" in image_eqI)
+          using X(2) assms(3)[rule_format,of x] by auto
+      qed note ** = d(2)[OF this] have *:"inj_on (\<lambda>(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastsimp
+       have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding group_simps add_left_cancel
+        unfolding setsum_reindex[OF *] apply(subst scaleR_right.setsum) defer apply(rule setsum_cong2) unfolding o_def split_paired_all split_conv
+        apply(drule p(4)) apply safe unfolding assms(7)[rule_format] using p by auto
+      also have "... = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") unfolding scaleR.diff_right scaleR.scaleR_left[THEN sym]
+        unfolding real_scaleR_def using assms(1) by auto finally have *:"?l = ?r" .
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using ** unfolding * unfolding norm_scaleR
+        using assms(1) by(auto simp add:field_simps) qed qed qed
+
+subsection {* Special case of a basic affine transformation. *}
+
+lemma interval_image_affinity_interval: shows "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::real^'n) + c) ` {a..b} = {u..v}"
+  unfolding image_affinity_interval by auto
+
+lemmas Cart_simps = Cart_nth.add Cart_nth.minus Cart_nth.zero Cart_nth.diff Cart_nth.scaleR real_scaleR_def Cart_lambda_beta
+   Cart_eq vector_le_def vector_less_def
+
+lemma setprod_cong2: assumes "\<And>x. x \<in> A \<Longrightarrow> f x = g x" shows "setprod f A = setprod g A"
+  apply(rule setprod_cong) using assms by auto
+
+lemma content_image_affinity_interval: 
+ "content((\<lambda>x::real^'n. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ CARD('n) * content {a..b}" (is "?l = ?r")
+proof- { presume *:"{a..b}\<noteq>{} \<Longrightarrow> ?thesis" show ?thesis apply(cases,rule *,assumption)
+      unfolding not_not using content_empty by auto }
+  assume as:"{a..b}\<noteq>{}" show ?thesis proof(cases "m \<ge> 0")
+    case True show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_P[OF True]
+      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') 
+      defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym]
+      apply(rule setprod_cong2) using True as unfolding interval_ne_empty Cart_simps not_le  
+      by(auto simp add:field_simps intro:mult_left_mono)
+  next case False show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_not_P[OF False]
+      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') 
+      defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym]
+      apply(rule setprod_cong2) using False as unfolding interval_ne_empty Cart_simps not_le 
+      by(auto simp add:field_simps mult_le_cancel_left_neg) qed qed
+
+lemma has_integral_affinity: assumes "(f has_integral i) {a..b::real^'n}" "m \<noteq> 0"
+  shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ CARD('n::finite))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})"
+  apply(rule has_integral_twiddle,safe) unfolding Cart_eq Cart_simps apply(rule zero_less_power)
+  defer apply(insert assms(2), simp add:field_simps) apply(insert assms(2), simp add:field_simps)
+  apply(rule continuous_intros)+ apply(rule interval_image_affinity_interval)+ apply(rule content_image_affinity_interval) using assms by auto
+
+lemma integrable_affinity: assumes "f integrable_on {a..b}" "m \<noteq> 0"
+  shows "(\<lambda>x. f(m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` {a..b})"
+  using assms unfolding integrable_on_def apply safe apply(drule has_integral_affinity) by auto
+
+subsection {* Special case of stretching coordinate axes separately. *}
+
+lemma image_stretch_interval:
+  "(\<lambda>x. \<chi> k. m k * x$k) ` {a..b::real^'n} =
+  (if {a..b} = {} then {} else {(\<chi> k. min (m(k) * a$k) (m(k) * b$k)) ..  (\<chi> k. max (m(k) * a$k) (m(k) * b$k))})" (is "?l = ?r")
+proof(cases "{a..b}={}") case True thus ?thesis unfolding True by auto
+next have *:"\<And>P Q. (\<forall>i. P i) \<and> (\<forall>i. Q i) \<longleftrightarrow> (\<forall>i. P i \<and> Q i)" by auto
+  case False note ab = this[unfolded interval_ne_empty]
+  show ?thesis apply-apply(rule set_ext)
+  proof- fix x::"real^'n" have **:"\<And>P Q. (\<forall>i. P i = Q i) \<Longrightarrow> (\<forall>i. P i) = (\<forall>i. Q i)" by auto
+    show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" unfolding if_not_P[OF False] 
+      unfolding image_iff mem_interval Bex_def Cart_simps Cart_eq *
+      unfolding lambda_skolem[THEN sym,of "\<lambda> i xa. (a $ i \<le> xa \<and> xa \<le> b $ i) \<and> x $ i = m i * xa"]
+    proof(rule **,rule) fix i::'n show "(\<exists>xa. (a $ i \<le> xa \<and> xa \<le> b $ i) \<and> x $ i = m i * xa) =
+        (min (m i * a $ i) (m i * b $ i) \<le> x $ i \<and> x $ i \<le> max (m i * a $ i) (m i * b $ i))"
+      proof(cases "m i = 0") case True thus ?thesis using ab by auto
+      next case False hence "0 < m i \<or> 0 > m i" by auto thus ?thesis apply-
+        proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $ i) (m i * b $ i) = m i * a $ i"
+            "max (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab unfolding min_def max_def by auto
+          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI)
+            using as by(auto simp add:field_simps)
+        next assume as:"0 > m i" hence *:"max (m i * a $ i) (m i * b $ i) = m i * a $ i"
+            "min (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab as unfolding min_def max_def 
+            by(auto simp add:field_simps mult_le_cancel_left_neg intro:real_le_antisym)
+          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI)
+            using as by(auto simp add:field_simps) qed qed qed qed qed 
+
+lemma interval_image_stretch_interval: "\<exists>u v. (\<lambda>x. \<chi> k. m k * x$k) ` {a..b::real^'n} = {u..v}"
+  unfolding image_stretch_interval by auto 
+
+lemma content_image_stretch_interval:
+  "content((\<lambda>x::real^'n. \<chi> k. m k * x$k) ` {a..b}) = abs(setprod m UNIV) * content({a..b})"
+proof(cases "{a..b} = {}") case True thus ?thesis
+    unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto
+next case False hence "(\<lambda>x. \<chi> k. m k * x $ k) ` {a..b} \<noteq> {}" by auto
+  thus ?thesis using False unfolding content_def image_stretch_interval apply- unfolding interval_bounds' if_not_P
+    unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding Cart_lambda_beta
+  proof- fix i::'n have "(m i < 0 \<or> m i > 0) \<or> m i = 0" by auto
+    thus "max (m i * a $ i) (m i * b $ i) - min (m i * a $ i) (m i * b $ i) = \<bar>m i\<bar> * (b $ i - a $ i)"
+      apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] 
+      by(auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) qed qed
+
+lemma has_integral_stretch: assumes "(f has_integral i) {a..b}" "\<forall>k. ~(m k = 0)"
+  shows "((\<lambda>x. f(\<chi> k. m k * x$k)) has_integral
+             ((1/(abs(setprod m UNIV))) *\<^sub>R i)) ((\<lambda>x. \<chi> k. 1/(m k) * x$k) ` {a..b})"
+  apply(rule has_integral_twiddle) unfolding zero_less_abs_iff content_image_stretch_interval
+  unfolding image_stretch_interval empty_as_interval Cart_eq using assms
+proof- show "\<forall>x. continuous (at x) (\<lambda>x. \<chi> k. m k * x $ k)"
+   apply(rule,rule linear_continuous_at) unfolding linear_linear
+   unfolding linear_def Cart_simps Cart_eq by(auto simp add:field_simps) qed auto
+
+lemma integrable_stretch: 
+  assumes "f integrable_on {a..b}" "\<forall>k. ~(m k = 0)"
+  shows "(\<lambda>x. f(\<chi> k. m k * x$k)) integrable_on ((\<lambda>x. \<chi> k. 1/(m k) * x$k) ` {a..b})"
+  using assms unfolding integrable_on_def apply-apply(erule exE) apply(drule has_integral_stretch) by auto
+
+subsection {* even more special cases. *}
+
+lemma uminus_interval_vector[simp]:"uminus ` {a..b} = {-b .. -a::real^'n}"
+  apply(rule set_ext,rule) defer unfolding image_iff
+  apply(rule_tac x="-x" in bexI) by(auto simp add:vector_le_def minus_le_iff le_minus_iff)
+
+lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) {a..b}"
+  shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
+  using has_integral_affinity[OF assms, of "-1" 0] by auto
+
+lemma has_integral_reflect[simp]: "((\<lambda>x. f(-x)) has_integral i) {-b..-a} \<longleftrightarrow> (f has_integral i) ({a..b})"
+  apply rule apply(drule_tac[!] has_integral_reflect_lemma) by auto
+
+lemma integrable_reflect[simp]: "(\<lambda>x. f(-x)) integrable_on {-b..-a} \<longleftrightarrow> f integrable_on {a..b}"
+  unfolding integrable_on_def by auto
+
+lemma integral_reflect[simp]: "integral {-b..-a} (\<lambda>x. f(-x)) = integral ({a..b}) f"
+  unfolding integral_def by auto
+
+subsection {* Stronger form of FCT; quite a tedious proof. *}
+
+(** move this **)
+declare norm_triangle_ineq4[intro] 
+
+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" "p tagged_division_of {vec1 a..vec1 b}"
+  shows "setsum (\<lambda>(x,k). f (dest_vec1 (interval_upperbound k)) - f(dest_vec1 (interval_lowerbound k))) p = f b - f a"
+  using additive_tagged_division_1[OF _ assms(2), of "f o dest_vec1"]
+  unfolding o_def vec1_dest_vec1 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"
+  unfolding split_def by(rule refl)
+
+lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
+  apply(subst(asm)(2) norm_minus_cancel[THEN sym])
+  apply(drule norm_triangle_le) by(auto simp add:group_simps)
+
+lemma fundamental_theorem_of_calculus_interior:
+  assumes"a \<le> b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_vector_derivative f'(x)) (at x)"
+  shows "((f' o dest_vec1) has_integral (f b - f a)) {vec a..vec b}"
+proof- { presume *:"a < b \<Longrightarrow> ?thesis" 
+    show ?thesis proof(cases,rule *,assumption)
+      assume "\<not> a < b" hence "a = b" using assms(1) by auto
+      hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" apply(auto simp add: Cart_simps) by smt
+      show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0_1 using * `a=b` by auto
+    qed } assume ab:"a < b"
+  let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
+                   norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f' \<circ> dest_vec1) x) - (f b - f a)) \<le> e * content {vec1 a..vec1 b})"
+  { presume "\<And>e. e>0 \<Longrightarrow> ?P e" thus ?thesis unfolding has_integral_factor_content by auto }
+  fix e::real assume e:"e>0"
+  note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib]
+  note conjunctD2[OF this] note bounded=this(1) and this(2)
+  from this(2) have "\<forall>x\<in>{a<..<b}. \<exists>d>0. \<forall>y. norm (y - x) < d \<longrightarrow> norm (f y - f x - (y - x) *\<^sub>R f' x) \<le> e/2 * norm (y - x)"
+    apply-apply safe apply(erule_tac x=x in ballE,erule_tac x="e/2" in allE) using e by auto note this[unfolded bgauge_existence_lemma]
+  from choice[OF this] guess d .. note conjunctD2[OF this[rule_format]] note d = this[rule_format]
+  have "bounded (f ` {a..b})" apply(rule compact_imp_bounded compact_continuous_image)+ using compact_real_interval assms by auto
+  from this[unfolded bounded_pos] guess B .. note B = this[rule_format]
+
+  have "\<exists>da. 0 < da \<and> (\<forall>c. a \<le> c \<and> {a..c} \<subseteq> {a..b} \<and> {a..c} \<subseteq> ball a da
+    \<longrightarrow> norm(content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b - a)) / 4)"
+  proof- have "a\<in>{a..b}" using ab by auto
+    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
+    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps)
+    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
+    have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' a) \<le> (e * (b - a)) / 8"
+    proof(cases "f' a = 0") case True
+      thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) 
+    next case False thus ?thesis 
+        apply(rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI)
+        using ab e by(auto simp add:field_simps)
+    qed then guess l .. note l = conjunctD2[OF this]
+    show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+)
+    proof- fix c assume as:"a \<le> c" "{a..c} \<subseteq> {a..b}" "{a..c} \<subseteq> ball a (min k l)" 
+      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval]
+      have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \<le> norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by(rule norm_triangle_ineq4)
+      also have "... \<le> e * (b - a) / 8 + e * (b - a) / 8" 
+      proof(rule add_mono) case goal1 have "\<bar>c - a\<bar> \<le> \<bar>l\<bar>" using as' by auto
+        thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto
+      next case goal2 show ?case apply(rule less_imp_le) apply(cases "a = c") defer
+          apply(rule k(2)[unfolded vector_dist_norm]) using as' e ab by(auto simp add:field_simps)
+      qed finally show "norm (content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto
+    qed qed then guess da .. note da=conjunctD2[OF this,rule_format]
+
+  have "\<exists>db>0. \<forall>c\<le>b. {c..b} \<subseteq> {a..b} \<and> {c..b} \<subseteq> ball b db \<longrightarrow> norm(content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b - a)) / 4"
+  proof- have "b\<in>{a..b}" using ab by auto
+    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
+    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps)
+    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
+    have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' b) \<le> (e * (b - a)) / 8"
+    proof(cases "f' b = 0") case True
+      thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) 
+    next case False thus ?thesis 
+        apply(rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI)
+        using ab e by(auto simp add:field_simps)
+    qed then guess l .. note l = conjunctD2[OF this]
+    show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+)
+    proof- fix c assume as:"c \<le> b" "{c..b} \<subseteq> {a..b}" "{c..b} \<subseteq> ball b (min k l)" 
+      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval]
+      have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \<le> norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by(rule norm_triangle_ineq4)
+      also have "... \<le> e * (b - a) / 8 + e * (b - a) / 8" 
+      proof(rule add_mono) case goal1 have "\<bar>c - b\<bar> \<le> \<bar>l\<bar>" using as' by auto
+        thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto
+      next case goal2 show ?case apply(rule less_imp_le) apply(cases "b = c") defer apply(subst norm_minus_commute)
+          apply(rule k(2)[unfolded vector_dist_norm]) using as' e ab by(auto simp add:field_simps)
+      qed finally show "norm (content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto
+    qed qed then guess db .. note db=conjunctD2[OF this,rule_format]
+
+  let ?d = "(\<lambda>x. ball x (if x=vec1 a then da else if x=vec b then db else d (dest_vec1 x)))"
+  show "?P e" apply(rule_tac x="?d" in exI)
+  proof safe case goal1 show ?case apply(rule gauge_ball_dependent) using ab db(1) da(1) d(1) by auto
+  next case goal2 note as=this let ?A = "{t. fst t \<in> {vec1 a, vec1 b}}" note p = tagged_division_ofD[OF goal2(1)]
+    have pA:"p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"  using goal2 by auto
+    note * = additive_tagged_division_1'[OF assms(1) goal2(1), THEN sym]
+    have **:"\<And>n1 s1 n2 s2::real. n2 \<le> s2 / 2 \<Longrightarrow> n1 - s1 \<le> s2 / 2 \<Longrightarrow> n1 + n2 \<le> s1 + s2" by arith
+    show ?case unfolding content_1'[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[THEN sym] split_minus
+      unfolding setsum_right_distrib apply(subst(2) pA,subst pA) unfolding setsum_Un_disjoint[OF pA(2-)]
+    proof(rule norm_triangle_le,rule **) 
+      case goal1 show ?case apply(rule order_trans,rule setsum_norm_le) apply(rule pA) defer apply(subst divide.setsum)
+      proof(rule order_refl,safe,unfold not_le o_def split_conv fst_conv,rule ccontr) fix x k assume as:"(x,k) \<in> p"
+          "e * (dest_vec1 (interval_upperbound k) - dest_vec1 (interval_lowerbound k)) / 2
+          < norm (content k *\<^sub>R f' (dest_vec1 x) - (f (dest_vec1 (interval_upperbound k)) - f (dest_vec1 (interval_lowerbound k))))"
+        from p(4)[OF this(1)] guess u v apply-by(erule exE)+ note k=this
+        hence "\<forall>i. u$i \<le> v$i" and uv:"{u,v}\<subseteq>{u..v}" using p(2)[OF as(1)] by auto note this(1) this(1)[unfolded forall_1]
+        note result = as(2)[unfolded k interval_bounds[OF this(1)] content_1[OF this(2)]]
+
+        assume as':"x \<noteq> vec1 a" "x \<noteq> vec1 b" hence "x$1 \<in> {a<..<b}" using p(2-3)[OF as(1)] by(auto simp add:Cart_simps) note  * = d(2)[OF this] 
+        have "norm ((v$1 - u$1) *\<^sub>R f' (x$1) - (f (v$1) - f (u$1))) =
+          norm ((f (u$1) - f (x$1) - (u$1 - x$1) *\<^sub>R f' (x$1)) - (f (v$1) - f (x$1) - (v$1 - x$1) *\<^sub>R f' (x$1)))" 
+          apply(rule arg_cong[of _ _ norm]) unfolding scaleR_left.diff by auto 
+        also have "... \<le> e / 2 * norm (u$1 - x$1) + e / 2 * norm (v$1 - x$1)" apply(rule norm_triangle_le_sub)
+          apply(rule add_mono) apply(rule_tac[!] *) using fineD[OF goal2(2) as(1)] as' unfolding k subset_eq
+          apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) using uv by(auto simp add:dist_real)
+        also have "... \<le> e / 2 * norm (v$1 - u$1)" using p(2)[OF as(1)] unfolding k by(auto simp add:field_simps)
+        finally have "e * (dest_vec1 v - dest_vec1 u) / 2 < e * (dest_vec1 v - dest_vec1 u) / 2"
+          apply- apply(rule less_le_trans[OF result]) using uv by auto thus False by auto qed
+
+    next have *:"\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2) / 2 \<Longrightarrow> x - s1 \<le> s2 / 2" by auto
+      case goal2 show ?case apply(rule *) apply(rule setsum_nonneg) apply(rule,unfold split_paired_all split_conv)
+        defer unfolding setsum_Un_disjoint[OF pA(2-),THEN sym] pA(1)[THEN sym] unfolding setsum_right_distrib[THEN sym] 
+        apply(subst additive_tagged_division_1[OF _ as(1)]) unfolding vec1_dest_vec1 apply(rule assms)
+      proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {vec1 a, vec1 b}}" note xk=IntD1[OF this]
+        from p(4)[OF this] guess u v apply-by(erule exE)+ note uv=this
+        with p(2)[OF xk] have "{u..v} \<noteq> {}" by auto
+        thus "0 \<le> e * ((interval_upperbound k)$1 - (interval_lowerbound k)$1)"
+          unfolding uv using e by(auto simp add:field_simps)
+      next have *:"\<And>s f t e. setsum f s = setsum f t \<Longrightarrow> norm(setsum f t) \<le> e \<Longrightarrow> norm(setsum f s) \<le> e" by auto
+        show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R (f' \<circ> dest_vec1) x -
+          (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) \<le> e * (b - a) / 2" 
+          apply(rule *[where t="p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content(snd t) \<noteq> 0}"])
+          apply(rule setsum_mono_zero_right[OF pA(2)]) defer apply(rule) unfolding split_paired_all split_conv o_def
+        proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {vec1 a, vec1 b}} - p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content (snd t) \<noteq> 0}"
+          hence xk:"(x,k)\<in>p" "content k = 0" by auto from p(4)[OF xk(1)] guess u v apply-by(erule exE)+ note uv=this
+          have "k\<noteq>{}" using p(2)[OF xk(1)] by auto hence *:"u = v" using xk unfolding uv content_eq_0_1 interval_eq_empty by auto
+          thus "content k *\<^sub>R (f' (x$1)) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1)) = 0" using xk unfolding uv by auto
+        next have *:"p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content(snd t) \<noteq> 0} = 
+            {t. t\<in>p \<and> fst t = vec1 a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = vec1 b \<and> content(snd t) \<noteq> 0}" by blast
+          have **:"\<And>s f. \<And>e::real. (\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y) \<Longrightarrow> (\<forall>x. x \<in> s \<longrightarrow> norm(f x) \<le> e) \<Longrightarrow> e>0 \<Longrightarrow> norm(setsum f s) \<le> e"
+          proof(case_tac "s={}") case goal2 then obtain x where "x\<in>s" by auto hence *:"s = {x}" using goal2(1) by auto
+            thus ?case using `x\<in>s` goal2(2) by auto
+          qed auto
+          case goal2 show ?case apply(subst *, subst setsum_Un_disjoint) prefer 4 apply(rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) 
+            apply(rule norm_triangle_le,rule add_mono) apply(rule_tac[1-2] **)
+          proof- let ?B = "\<lambda>x. {t \<in> p. fst t = vec1 x \<and> content (snd t) \<noteq> 0}"
+            have pa:"\<And>k. (vec1 a, k) \<in> p \<Longrightarrow> \<exists>v. k = {vec1 a .. v} \<and> vec1 a \<le> v" 
+            proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this
+              have *:"u \<le> v" using p(2)[OF goal1] unfolding uv by auto
+              have u:"u = vec1 a" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
+                have "u \<ge> vec1 a" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "u\<noteq>vec1 a" ultimately
+                have "u > vec1 a" unfolding Cart_simps by auto
+                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps)
+              qed thus ?case apply(rule_tac x=v in exI) unfolding uv using * by auto
+            qed
+            have pb:"\<And>k. (vec1 b, k) \<in> p \<Longrightarrow> \<exists>v. k = {v .. vec1 b} \<and> vec1 b \<ge> v" 
+            proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this
+              have *:"u \<le> v" using p(2)[OF goal1] unfolding uv by auto
+              have u:"v = vec1 b" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
+                have "v \<le> vec1 b" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "v\<noteq>vec1 b" ultimately
+                have "v < vec1 b" unfolding Cart_simps by auto
+                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps)
+              qed thus ?case apply(rule_tac x=u in exI) unfolding uv using * by auto
+            qed
+
+            show "\<forall>x y. x \<in> ?B a \<and> y \<in> ?B a \<longrightarrow> x = y" apply(rule,rule,rule,unfold split_paired_all)
+              unfolding mem_Collect_eq fst_conv snd_conv apply safe
+            proof- fix x k k' assume k:"(vec1 a, k) \<in> p" "(vec1 a, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
+              guess v using pa[OF k(1)] .. note v = conjunctD2[OF this]
+              guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (min (v$1) (v'$1))"
+              have "{vec1 a <..< ?v} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter]
+              moreover have "vec1 ((a + ?v$1)/2) \<in> {vec1 a <..< ?v}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps)
+              ultimately have "vec1 ((a + ?v$1)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
+              hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto
+              { assume "x\<in>k" thus "x\<in>k'" unfolding * . } { assume "x\<in>k'" thus "x\<in>k" unfolding * . }
+            qed 
+            show "\<forall>x y. x \<in> ?B b \<and> y \<in> ?B b \<longrightarrow> x = y" apply(rule,rule,rule,unfold split_paired_all)
+              unfolding mem_Collect_eq fst_conv snd_conv apply safe
+            proof- fix x k k' assume k:"(vec1 b, k) \<in> p" "(vec1 b, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
+              guess v using pb[OF k(1)] .. note v = conjunctD2[OF this]
+              guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (max (v$1) (v'$1))"
+              have "{?v <..< vec1 b} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter]
+              moreover have "vec1 ((b + ?v$1)/2) \<in> {?v <..< vec1 b}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps)
+              ultimately have "vec1 ((b + ?v$1)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
+              hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto
+              { assume "x\<in>k" thus "x\<in>k'" unfolding * . } { assume "x\<in>k'" thus "x\<in>k" unfolding * . }
+            qed
+
+            let ?a = a and ?b = b (* a is something else while proofing the next theorem. *)
+            show "\<forall>x. x \<in> ?B a \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x)
+              \<le> e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1
+            proof- case goal1 guess v using pa[OF goal1(1)] .. note v = conjunctD2[OF this]
+              have "vec1 ?a\<in>{vec1 ?a..v}" using v(2) by auto hence "dest_vec1 v \<le> ?b" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
+              moreover have "{?a..dest_vec1 v} \<subseteq> ball ?a da" using fineD[OF as(2) goal1(1)]
+                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE)
+                by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately
+              show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply-
+                apply(rule da(2)[of "v$1",unfolded vec1_dest_vec1])
+                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto
+            qed
+            show "\<forall>x. x \<in> ?B b \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x)
+              \<le> e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1
+            proof- case goal1 guess v using pb[OF goal1(1)] .. note v = conjunctD2[OF this]
+              have "vec1 ?b\<in>{v..vec1 ?b}" using v(2) by auto hence "dest_vec1 v \<ge> ?a" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
+              moreover have "{dest_vec1 v..?b} \<subseteq> ball ?b db" using fineD[OF as(2) goal1(1)]
+                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE) using ab
+                by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately
+              show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply-
+                apply(rule db(2)[of "v$1",unfolded vec1_dest_vec1])
+                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto
+            qed
+          qed(insert p(1) ab e, auto simp add:field_simps) qed auto qed qed qed qed
+
+subsection {* Stronger form with finite number of exceptional points. *}
+
+lemma fundamental_theorem_of_calculus_interior_strong: fixes f::"real \<Rightarrow> 'a::banach"
+  assumes"finite s" "a \<le> b" "continuous_on {a..b} f"
+  "\<forall>x\<in>{a<..<b} - s. (f has_vector_derivative f'(x)) (at x)"
+  shows "((f' o dest_vec1) has_integral (f b - f a)) {vec a..vec b}" using assms apply- 
+proof(induct "card s" arbitrary:s a b)
+  case 0 show ?case apply(rule fundamental_theorem_of_calculus_interior) using 0 by auto
+next case (Suc n) from this(2) guess c s' apply-apply(subst(asm) eq_commute) unfolding card_Suc_eq
+    apply(subst(asm)(2) eq_commute) by(erule exE conjE)+ note cs = this[rule_format]
+  show ?case proof(cases "c\<in>{a<..<b}")
+    case False thus ?thesis apply- apply(rule Suc(1)[OF cs(3) _ Suc(4,5)]) apply safe defer
+      apply(rule Suc(6)[rule_format]) using Suc(3) unfolding cs by auto
+  next have *:"f b - f a = (f c - f a) + (f b - f c)" by auto
+    case True hence "vec1 a \<le> vec1 c" "vec1 c \<le> vec1 b" by auto
+    thus ?thesis apply(subst *) apply(rule has_integral_combine) apply assumption+
+      apply(rule_tac[!] Suc(1)[OF cs(3)]) using Suc(3) unfolding cs
+    proof- show "continuous_on {a..c} f" "continuous_on {c..b} f"
+        apply(rule_tac[!] continuous_on_subset[OF Suc(5)]) using True by auto
+      let ?P = "\<lambda>i j. \<forall>x\<in>{i<..<j} - s'. (f has_vector_derivative f' x) (at x)"
+      show "?P a c" "?P c b" apply safe apply(rule_tac[!] Suc(6)[rule_format]) using True unfolding cs by auto
+    qed auto qed qed
+
+lemma fundamental_theorem_of_calculus_strong: fixes f::"real \<Rightarrow> 'a::banach"
+  assumes "finite s" "a \<le> b" "continuous_on {a..b} f"
+  "\<forall>x\<in>{a..b} - s. (f has_vector_derivative f'(x)) (at x)"
+  shows "((f' o dest_vec1) has_integral (f(b) - f(a))) {vec1 a..vec1 b}"
+  apply(rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f'])
+  using assms(4) by auto
+
+end
--- a/src/HOL/Multivariate_Analysis/Integration_MV.cert	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3296 +0,0 @@
-tB2Atlor9W4pSnrAz5nHpw 907 0
-#2 := false
-#299 := 0::real
-decl uf_1 :: (-> T3 T2 real)
-decl uf_10 :: (-> T4 T2)
-decl uf_7 :: T4
-#15 := uf_7
-#22 := (uf_10 uf_7)
-decl uf_2 :: (-> T1 T3)
-decl uf_4 :: T1
-#11 := uf_4
-#91 := (uf_2 uf_4)
-#902 := (uf_1 #91 #22)
-#297 := -1::real
-#1084 := (* -1::real #902)
-decl uf_16 :: T1
-#50 := uf_16
-#78 := (uf_2 uf_16)
-#799 := (uf_1 #78 #22)
-#1267 := (+ #799 #1084)
-#1272 := (>= #1267 0::real)
-#1266 := (= #799 #902)
-decl uf_9 :: T3
-#21 := uf_9
-#23 := (uf_1 uf_9 #22)
-#905 := (= #23 #902)
-decl uf_11 :: T3
-#24 := uf_11
-#850 := (uf_1 uf_11 #22)
-#904 := (= #850 #902)
-decl uf_6 :: (-> T2 T4)
-#74 := (uf_6 #22)
-#281 := (= uf_7 #74)
-#922 := (ite #281 #905 #904)
-decl uf_8 :: T3
-#18 := uf_8
-#848 := (uf_1 uf_8 #22)
-#903 := (= #848 #902)
-#60 := 0::int
-decl uf_5 :: (-> T4 int)
-#803 := (uf_5 #74)
-#117 := -1::int
-#813 := (* -1::int #803)
-#16 := (uf_5 uf_7)
-#916 := (+ #16 #813)
-#917 := (<= #916 0::int)
-#925 := (ite #917 #922 #903)
-#6 := (:var 0 T2)
-#19 := (uf_1 uf_8 #6)
-#544 := (pattern #19)
-#25 := (uf_1 uf_11 #6)
-#543 := (pattern #25)
-#92 := (uf_1 #91 #6)
-#542 := (pattern #92)
-#13 := (uf_6 #6)
-#541 := (pattern #13)
-#447 := (= #19 #92)
-#445 := (= #25 #92)
-#444 := (= #23 #92)
-#20 := (= #13 uf_7)
-#446 := (ite #20 #444 #445)
-#120 := (* -1::int #16)
-#14 := (uf_5 #13)
-#121 := (+ #14 #120)
-#119 := (>= #121 0::int)
-#448 := (ite #119 #446 #447)
-#545 := (forall (vars (?x3 T2)) (:pat #541 #542 #543 #544) #448)
-#451 := (forall (vars (?x3 T2)) #448)
-#548 := (iff #451 #545)
-#546 := (iff #448 #448)
-#547 := [refl]: #546
-#549 := [quant-intro #547]: #548
-#26 := (ite #20 #23 #25)
-#127 := (ite #119 #26 #19)
-#368 := (= #92 #127)
-#369 := (forall (vars (?x3 T2)) #368)
-#452 := (iff #369 #451)
-#449 := (iff #368 #448)
-#450 := [rewrite]: #449
-#453 := [quant-intro #450]: #452
-#392 := (~ #369 #369)
-#390 := (~ #368 #368)
-#391 := [refl]: #390
-#366 := [nnf-pos #391]: #392
-decl uf_3 :: (-> T1 T2 real)
-#12 := (uf_3 uf_4 #6)
-#132 := (= #12 #127)
-#135 := (forall (vars (?x3 T2)) #132)
-#370 := (iff #135 #369)
-#4 := (:var 1 T1)
-#8 := (uf_3 #4 #6)
-#5 := (uf_2 #4)
-#7 := (uf_1 #5 #6)
-#9 := (= #7 #8)
-#10 := (forall (vars (?x1 T1) (?x2 T2)) #9)
-#113 := [asserted]: #10
-#371 := [rewrite* #113]: #370
-#17 := (< #14 #16)
-#27 := (ite #17 #19 #26)
-#28 := (= #12 #27)
-#29 := (forall (vars (?x3 T2)) #28)
-#136 := (iff #29 #135)
-#133 := (iff #28 #132)
-#130 := (= #27 #127)
-#118 := (not #119)
-#124 := (ite #118 #19 #26)
-#128 := (= #124 #127)
-#129 := [rewrite]: #128
-#125 := (= #27 #124)
-#122 := (iff #17 #118)
-#123 := [rewrite]: #122
-#126 := [monotonicity #123]: #125
-#131 := [trans #126 #129]: #130
-#134 := [monotonicity #131]: #133
-#137 := [quant-intro #134]: #136
-#114 := [asserted]: #29
-#138 := [mp #114 #137]: #135
-#372 := [mp #138 #371]: #369
-#367 := [mp~ #372 #366]: #369
-#454 := [mp #367 #453]: #451
-#550 := [mp #454 #549]: #545
-#738 := (not #545)
-#928 := (or #738 #925)
-#75 := (= #74 uf_7)
-#906 := (ite #75 #905 #904)
-#907 := (+ #803 #120)
-#908 := (>= #907 0::int)
-#909 := (ite #908 #906 #903)
-#929 := (or #738 #909)
-#931 := (iff #929 #928)
-#933 := (iff #928 #928)
-#934 := [rewrite]: #933
-#926 := (iff #909 #925)
-#923 := (iff #906 #922)
-#283 := (iff #75 #281)
-#284 := [rewrite]: #283
-#924 := [monotonicity #284]: #923
-#920 := (iff #908 #917)
-#910 := (+ #120 #803)
-#913 := (>= #910 0::int)
-#918 := (iff #913 #917)
-#919 := [rewrite]: #918
-#914 := (iff #908 #913)
-#911 := (= #907 #910)
-#912 := [rewrite]: #911
-#915 := [monotonicity #912]: #914
-#921 := [trans #915 #919]: #920
-#927 := [monotonicity #921 #924]: #926
-#932 := [monotonicity #927]: #931
-#935 := [trans #932 #934]: #931
-#930 := [quant-inst]: #929
-#936 := [mp #930 #935]: #928
-#1300 := [unit-resolution #936 #550]: #925
-#989 := (= #16 #803)
-#1277 := (= #803 #16)
-#280 := [asserted]: #75
-#287 := [mp #280 #284]: #281
-#1276 := [symm #287]: #75
-#1278 := [monotonicity #1276]: #1277
-#1301 := [symm #1278]: #989
-#1302 := (not #989)
-#1303 := (or #1302 #917)
-#1304 := [th-lemma]: #1303
-#1305 := [unit-resolution #1304 #1301]: #917
-#950 := (not #917)
-#949 := (not #925)
-#951 := (or #949 #950 #922)
-#952 := [def-axiom]: #951
-#1306 := [unit-resolution #952 #1305 #1300]: #922
-#937 := (not #922)
-#1307 := (or #937 #905)
-#938 := (not #281)
-#939 := (or #937 #938 #905)
-#940 := [def-axiom]: #939
-#1308 := [unit-resolution #940 #287]: #1307
-#1309 := [unit-resolution #1308 #1306]: #905
-#1356 := (= #799 #23)
-#800 := (= #23 #799)
-decl uf_15 :: T4
-#40 := uf_15
-#41 := (uf_5 uf_15)
-#814 := (+ #41 #813)
-#815 := (<= #814 0::int)
-#836 := (not #815)
-#158 := (* -1::int #41)
-#1270 := (+ #16 #158)
-#1265 := (>= #1270 0::int)
-#1339 := (not #1265)
-#1269 := (= #16 #41)
-#1298 := (not #1269)
-#286 := (= uf_7 uf_15)
-#44 := (uf_10 uf_15)
-#72 := (uf_6 #44)
-#73 := (= #72 uf_15)
-#277 := (= uf_15 #72)
-#278 := (iff #73 #277)
-#279 := [rewrite]: #278
-#276 := [asserted]: #73
-#282 := [mp #276 #279]: #277
-#1274 := [symm #282]: #73
-#729 := (= uf_7 #72)
-decl uf_17 :: (-> int T4)
-#611 := (uf_5 #72)
-#991 := (uf_17 #611)
-#1289 := (= #991 #72)
-#992 := (= #72 #991)
-#55 := (:var 0 T4)
-#56 := (uf_5 #55)
-#574 := (pattern #56)
-#57 := (uf_17 #56)
-#177 := (= #55 #57)
-#575 := (forall (vars (?x7 T4)) (:pat #574) #177)
-#195 := (forall (vars (?x7 T4)) #177)
-#578 := (iff #195 #575)
-#576 := (iff #177 #177)
-#577 := [refl]: #576
-#579 := [quant-intro #577]: #578
-#405 := (~ #195 #195)
-#403 := (~ #177 #177)
-#404 := [refl]: #403
-#406 := [nnf-pos #404]: #405
-#58 := (= #57 #55)
-#59 := (forall (vars (?x7 T4)) #58)
-#196 := (iff #59 #195)
-#193 := (iff #58 #177)
-#194 := [rewrite]: #193
-#197 := [quant-intro #194]: #196
-#155 := [asserted]: #59
-#200 := [mp #155 #197]: #195
-#407 := [mp~ #200 #406]: #195
-#580 := [mp #407 #579]: #575
-#995 := (not #575)
-#996 := (or #995 #992)
-#997 := [quant-inst]: #996
-#1273 := [unit-resolution #997 #580]: #992
-#1290 := [symm #1273]: #1289
-#1293 := (= uf_7 #991)
-#993 := (uf_17 #803)
-#1287 := (= #993 #991)
-#1284 := (= #803 #611)
-#987 := (= #41 #611)
-#1279 := (= #611 #41)
-#1280 := [monotonicity #1274]: #1279
-#1281 := [symm #1280]: #987
-#1282 := (= #803 #41)
-#1275 := [hypothesis]: #1269
-#1283 := [trans #1278 #1275]: #1282
-#1285 := [trans #1283 #1281]: #1284
-#1288 := [monotonicity #1285]: #1287
-#1291 := (= uf_7 #993)
-#994 := (= #74 #993)
-#1000 := (or #995 #994)
-#1001 := [quant-inst]: #1000
-#1286 := [unit-resolution #1001 #580]: #994
-#1292 := [trans #287 #1286]: #1291
-#1294 := [trans #1292 #1288]: #1293
-#1295 := [trans #1294 #1290]: #729
-#1296 := [trans #1295 #1274]: #286
-#290 := (not #286)
-#76 := (= uf_15 uf_7)
-#77 := (not #76)
-#291 := (iff #77 #290)
-#288 := (iff #76 #286)
-#289 := [rewrite]: #288
-#292 := [monotonicity #289]: #291
-#285 := [asserted]: #77
-#295 := [mp #285 #292]: #290
-#1297 := [unit-resolution #295 #1296]: false
-#1299 := [lemma #1297]: #1298
-#1342 := (or #1269 #1339)
-#1271 := (<= #1270 0::int)
-#621 := (* -1::int #611)
-#723 := (+ #16 #621)
-#724 := (<= #723 0::int)
-decl uf_12 :: T1
-#30 := uf_12
-#88 := (uf_2 uf_12)
-#771 := (uf_1 #88 #44)
-#45 := (uf_1 uf_9 #44)
-#772 := (= #45 #771)
-#796 := (not #772)
-decl uf_14 :: T1
-#38 := uf_14
-#83 := (uf_2 uf_14)
-#656 := (uf_1 #83 #44)
-#1239 := (= #656 #771)
-#1252 := (not #1239)
-#1324 := (iff #1252 #796)
-#1322 := (iff #1239 #772)
-#1320 := (= #656 #45)
-#661 := (= #45 #656)
-#659 := (uf_1 uf_11 #44)
-#664 := (= #656 #659)
-#667 := (ite #277 #661 #664)
-#657 := (uf_1 uf_8 #44)
-#670 := (= #656 #657)
-#622 := (+ #41 #621)
-#623 := (<= #622 0::int)
-#673 := (ite #623 #667 #670)
-#84 := (uf_1 #83 #6)
-#560 := (pattern #84)
-#467 := (= #19 #84)
-#465 := (= #25 #84)
-#464 := (= #45 #84)
-#43 := (= #13 uf_15)
-#466 := (ite #43 #464 #465)
-#159 := (+ #14 #158)
-#157 := (>= #159 0::int)
-#468 := (ite #157 #466 #467)
-#561 := (forall (vars (?x5 T2)) (:pat #541 #560 #543 #544) #468)
-#471 := (forall (vars (?x5 T2)) #468)
-#564 := (iff #471 #561)
-#562 := (iff #468 #468)
-#563 := [refl]: #562
-#565 := [quant-intro #563]: #564
-#46 := (ite #43 #45 #25)
-#165 := (ite #157 #46 #19)
-#378 := (= #84 #165)
-#379 := (forall (vars (?x5 T2)) #378)
-#472 := (iff #379 #471)
-#469 := (iff #378 #468)
-#470 := [rewrite]: #469
-#473 := [quant-intro #470]: #472
-#359 := (~ #379 #379)
-#361 := (~ #378 #378)
-#358 := [refl]: #361
-#356 := [nnf-pos #358]: #359
-#39 := (uf_3 uf_14 #6)
-#170 := (= #39 #165)
-#173 := (forall (vars (?x5 T2)) #170)
-#380 := (iff #173 #379)
-#381 := [rewrite* #113]: #380
-#42 := (< #14 #41)
-#47 := (ite #42 #19 #46)
-#48 := (= #39 #47)
-#49 := (forall (vars (?x5 T2)) #48)
-#174 := (iff #49 #173)
-#171 := (iff #48 #170)
-#168 := (= #47 #165)
-#156 := (not #157)
-#162 := (ite #156 #19 #46)
-#166 := (= #162 #165)
-#167 := [rewrite]: #166
-#163 := (= #47 #162)
-#160 := (iff #42 #156)
-#161 := [rewrite]: #160
-#164 := [monotonicity #161]: #163
-#169 := [trans #164 #167]: #168
-#172 := [monotonicity #169]: #171
-#175 := [quant-intro #172]: #174
-#116 := [asserted]: #49
-#176 := [mp #116 #175]: #173
-#382 := [mp #176 #381]: #379
-#357 := [mp~ #382 #356]: #379
-#474 := [mp #357 #473]: #471
-#566 := [mp #474 #565]: #561
-#676 := (not #561)
-#677 := (or #676 #673)
-#658 := (= #657 #656)
-#660 := (= #659 #656)
-#662 := (ite #73 #661 #660)
-#612 := (+ #611 #158)
-#613 := (>= #612 0::int)
-#663 := (ite #613 #662 #658)
-#678 := (or #676 #663)
-#680 := (iff #678 #677)
-#682 := (iff #677 #677)
-#683 := [rewrite]: #682
-#674 := (iff #663 #673)
-#671 := (iff #658 #670)
-#672 := [rewrite]: #671
-#668 := (iff #662 #667)
-#665 := (iff #660 #664)
-#666 := [rewrite]: #665
-#669 := [monotonicity #279 #666]: #668
-#626 := (iff #613 #623)
-#615 := (+ #158 #611)
-#618 := (>= #615 0::int)
-#624 := (iff #618 #623)
-#625 := [rewrite]: #624
-#619 := (iff #613 #618)
-#616 := (= #612 #615)
-#617 := [rewrite]: #616
-#620 := [monotonicity #617]: #619
-#627 := [trans #620 #625]: #626
-#675 := [monotonicity #627 #669 #672]: #674
-#681 := [monotonicity #675]: #680
-#684 := [trans #681 #683]: #680
-#679 := [quant-inst]: #678
-#685 := [mp #679 #684]: #677
-#1311 := [unit-resolution #685 #566]: #673
-#1312 := (not #987)
-#1313 := (or #1312 #623)
-#1314 := [th-lemma]: #1313
-#1315 := [unit-resolution #1314 #1281]: #623
-#645 := (not #623)
-#698 := (not #673)
-#699 := (or #698 #645 #667)
-#700 := [def-axiom]: #699
-#1316 := [unit-resolution #700 #1315 #1311]: #667
-#686 := (not #667)
-#1317 := (or #686 #661)
-#687 := (not #277)
-#688 := (or #686 #687 #661)
-#689 := [def-axiom]: #688
-#1318 := [unit-resolution #689 #282]: #1317
-#1319 := [unit-resolution #1318 #1316]: #661
-#1321 := [symm #1319]: #1320
-#1323 := [monotonicity #1321]: #1322
-#1325 := [monotonicity #1323]: #1324
-#1145 := (* -1::real #771)
-#1240 := (+ #656 #1145)
-#1241 := (<= #1240 0::real)
-#1249 := (not #1241)
-#1243 := [hypothesis]: #1241
-decl uf_18 :: T3
-#80 := uf_18
-#1040 := (uf_1 uf_18 #44)
-#1043 := (* -1::real #1040)
-#1156 := (+ #771 #1043)
-#1157 := (>= #1156 0::real)
-#1189 := (not #1157)
-#708 := (uf_1 #91 #44)
-#1168 := (+ #708 #1043)
-#1169 := (<= #1168 0::real)
-#1174 := (or #1157 #1169)
-#1177 := (not #1174)
-#89 := (uf_1 #88 #6)
-#552 := (pattern #89)
-#81 := (uf_1 uf_18 #6)
-#594 := (pattern #81)
-#324 := (* -1::real #92)
-#325 := (+ #81 #324)
-#323 := (>= #325 0::real)
-#317 := (* -1::real #89)
-#318 := (+ #81 #317)
-#319 := (<= #318 0::real)
-#436 := (or #319 #323)
-#437 := (not #436)
-#601 := (forall (vars (?x11 T2)) (:pat #594 #552 #542) #437)
-#440 := (forall (vars (?x11 T2)) #437)
-#604 := (iff #440 #601)
-#602 := (iff #437 #437)
-#603 := [refl]: #602
-#605 := [quant-intro #603]: #604
-#326 := (not #323)
-#320 := (not #319)
-#329 := (and #320 #326)
-#332 := (forall (vars (?x11 T2)) #329)
-#441 := (iff #332 #440)
-#438 := (iff #329 #437)
-#439 := [rewrite]: #438
-#442 := [quant-intro #439]: #441
-#425 := (~ #332 #332)
-#423 := (~ #329 #329)
-#424 := [refl]: #423
-#426 := [nnf-pos #424]: #425
-#306 := (* -1::real #84)
-#307 := (+ #81 #306)
-#305 := (>= #307 0::real)
-#308 := (not #305)
-#301 := (* -1::real #81)
-#79 := (uf_1 #78 #6)
-#302 := (+ #79 #301)
-#300 := (>= #302 0::real)
-#298 := (not #300)
-#311 := (and #298 #308)
-#314 := (forall (vars (?x10 T2)) #311)
-#335 := (and #314 #332)
-#93 := (< #81 #92)
-#90 := (< #89 #81)
-#94 := (and #90 #93)
-#95 := (forall (vars (?x11 T2)) #94)
-#85 := (< #81 #84)
-#82 := (< #79 #81)
-#86 := (and #82 #85)
-#87 := (forall (vars (?x10 T2)) #86)
-#96 := (and #87 #95)
-#336 := (iff #96 #335)
-#333 := (iff #95 #332)
-#330 := (iff #94 #329)
-#327 := (iff #93 #326)
-#328 := [rewrite]: #327
-#321 := (iff #90 #320)
-#322 := [rewrite]: #321
-#331 := [monotonicity #322 #328]: #330
-#334 := [quant-intro #331]: #333
-#315 := (iff #87 #314)
-#312 := (iff #86 #311)
-#309 := (iff #85 #308)
-#310 := [rewrite]: #309
-#303 := (iff #82 #298)
-#304 := [rewrite]: #303
-#313 := [monotonicity #304 #310]: #312
-#316 := [quant-intro #313]: #315
-#337 := [monotonicity #316 #334]: #336
-#293 := [asserted]: #96
-#338 := [mp #293 #337]: #335
-#340 := [and-elim #338]: #332
-#427 := [mp~ #340 #426]: #332
-#443 := [mp #427 #442]: #440
-#606 := [mp #443 #605]: #601
-#1124 := (not #601)
-#1180 := (or #1124 #1177)
-#1142 := (* -1::real #708)
-#1143 := (+ #1040 #1142)
-#1144 := (>= #1143 0::real)
-#1146 := (+ #1040 #1145)
-#1147 := (<= #1146 0::real)
-#1148 := (or #1147 #1144)
-#1149 := (not #1148)
-#1181 := (or #1124 #1149)
-#1183 := (iff #1181 #1180)
-#1185 := (iff #1180 #1180)
-#1186 := [rewrite]: #1185
-#1178 := (iff #1149 #1177)
-#1175 := (iff #1148 #1174)
-#1172 := (iff #1144 #1169)
-#1162 := (+ #1142 #1040)
-#1165 := (>= #1162 0::real)
-#1170 := (iff #1165 #1169)
-#1171 := [rewrite]: #1170
-#1166 := (iff #1144 #1165)
-#1163 := (= #1143 #1162)
-#1164 := [rewrite]: #1163
-#1167 := [monotonicity #1164]: #1166
-#1173 := [trans #1167 #1171]: #1172
-#1160 := (iff #1147 #1157)
-#1150 := (+ #1145 #1040)
-#1153 := (<= #1150 0::real)
-#1158 := (iff #1153 #1157)
-#1159 := [rewrite]: #1158
-#1154 := (iff #1147 #1153)
-#1151 := (= #1146 #1150)
-#1152 := [rewrite]: #1151
-#1155 := [monotonicity #1152]: #1154
-#1161 := [trans #1155 #1159]: #1160
-#1176 := [monotonicity #1161 #1173]: #1175
-#1179 := [monotonicity #1176]: #1178
-#1184 := [monotonicity #1179]: #1183
-#1187 := [trans #1184 #1186]: #1183
-#1182 := [quant-inst]: #1181
-#1188 := [mp #1182 #1187]: #1180
-#1244 := [unit-resolution #1188 #606]: #1177
-#1190 := (or #1174 #1189)
-#1191 := [def-axiom]: #1190
-#1245 := [unit-resolution #1191 #1244]: #1189
-#1054 := (+ #656 #1043)
-#1055 := (<= #1054 0::real)
-#1079 := (not #1055)
-#607 := (uf_1 #78 #44)
-#1044 := (+ #607 #1043)
-#1045 := (>= #1044 0::real)
-#1060 := (or #1045 #1055)
-#1063 := (not #1060)
-#567 := (pattern #79)
-#428 := (or #300 #305)
-#429 := (not #428)
-#595 := (forall (vars (?x10 T2)) (:pat #567 #594 #560) #429)
-#432 := (forall (vars (?x10 T2)) #429)
-#598 := (iff #432 #595)
-#596 := (iff #429 #429)
-#597 := [refl]: #596
-#599 := [quant-intro #597]: #598
-#433 := (iff #314 #432)
-#430 := (iff #311 #429)
-#431 := [rewrite]: #430
-#434 := [quant-intro #431]: #433
-#420 := (~ #314 #314)
-#418 := (~ #311 #311)
-#419 := [refl]: #418
-#421 := [nnf-pos #419]: #420
-#339 := [and-elim #338]: #314
-#422 := [mp~ #339 #421]: #314
-#435 := [mp #422 #434]: #432
-#600 := [mp #435 #599]: #595
-#1066 := (not #595)
-#1067 := (or #1066 #1063)
-#1039 := (* -1::real #656)
-#1041 := (+ #1040 #1039)
-#1042 := (>= #1041 0::real)
-#1046 := (or #1045 #1042)
-#1047 := (not #1046)
-#1068 := (or #1066 #1047)
-#1070 := (iff #1068 #1067)
-#1072 := (iff #1067 #1067)
-#1073 := [rewrite]: #1072
-#1064 := (iff #1047 #1063)
-#1061 := (iff #1046 #1060)
-#1058 := (iff #1042 #1055)
-#1048 := (+ #1039 #1040)
-#1051 := (>= #1048 0::real)
-#1056 := (iff #1051 #1055)
-#1057 := [rewrite]: #1056
-#1052 := (iff #1042 #1051)
-#1049 := (= #1041 #1048)
-#1050 := [rewrite]: #1049
-#1053 := [monotonicity #1050]: #1052
-#1059 := [trans #1053 #1057]: #1058
-#1062 := [monotonicity #1059]: #1061
-#1065 := [monotonicity #1062]: #1064
-#1071 := [monotonicity #1065]: #1070
-#1074 := [trans #1071 #1073]: #1070
-#1069 := [quant-inst]: #1068
-#1075 := [mp #1069 #1074]: #1067
-#1246 := [unit-resolution #1075 #600]: #1063
-#1080 := (or #1060 #1079)
-#1081 := [def-axiom]: #1080
-#1247 := [unit-resolution #1081 #1246]: #1079
-#1248 := [th-lemma #1247 #1245 #1243]: false
-#1250 := [lemma #1248]: #1249
-#1253 := (or #1252 #1241)
-#1254 := [th-lemma]: #1253
-#1310 := [unit-resolution #1254 #1250]: #1252
-#1326 := [mp #1310 #1325]: #796
-#1328 := (or #724 #772)
-decl uf_13 :: T3
-#33 := uf_13
-#609 := (uf_1 uf_13 #44)
-#773 := (= #609 #771)
-#775 := (ite #724 #773 #772)
-#32 := (uf_1 uf_9 #6)
-#553 := (pattern #32)
-#34 := (uf_1 uf_13 #6)
-#551 := (pattern #34)
-#456 := (= #32 #89)
-#455 := (= #34 #89)
-#457 := (ite #119 #455 #456)
-#554 := (forall (vars (?x4 T2)) (:pat #541 #551 #552 #553) #457)
-#460 := (forall (vars (?x4 T2)) #457)
-#557 := (iff #460 #554)
-#555 := (iff #457 #457)
-#556 := [refl]: #555
-#558 := [quant-intro #556]: #557
-#143 := (ite #119 #34 #32)
-#373 := (= #89 #143)
-#374 := (forall (vars (?x4 T2)) #373)
-#461 := (iff #374 #460)
-#458 := (iff #373 #457)
-#459 := [rewrite]: #458
-#462 := [quant-intro #459]: #461
-#362 := (~ #374 #374)
-#364 := (~ #373 #373)
-#365 := [refl]: #364
-#363 := [nnf-pos #365]: #362
-#31 := (uf_3 uf_12 #6)
-#148 := (= #31 #143)
-#151 := (forall (vars (?x4 T2)) #148)
-#375 := (iff #151 #374)
-#376 := [rewrite* #113]: #375
-#35 := (ite #17 #32 #34)
-#36 := (= #31 #35)
-#37 := (forall (vars (?x4 T2)) #36)
-#152 := (iff #37 #151)
-#149 := (iff #36 #148)
-#146 := (= #35 #143)
-#140 := (ite #118 #32 #34)
-#144 := (= #140 #143)
-#145 := [rewrite]: #144
-#141 := (= #35 #140)
-#142 := [monotonicity #123]: #141
-#147 := [trans #142 #145]: #146
-#150 := [monotonicity #147]: #149
-#153 := [quant-intro #150]: #152
-#115 := [asserted]: #37
-#154 := [mp #115 #153]: #151
-#377 := [mp #154 #376]: #374
-#360 := [mp~ #377 #363]: #374
-#463 := [mp #360 #462]: #460
-#559 := [mp #463 #558]: #554
-#778 := (not #554)
-#779 := (or #778 #775)
-#714 := (+ #611 #120)
-#715 := (>= #714 0::int)
-#774 := (ite #715 #773 #772)
-#780 := (or #778 #774)
-#782 := (iff #780 #779)
-#784 := (iff #779 #779)
-#785 := [rewrite]: #784
-#776 := (iff #774 #775)
-#727 := (iff #715 #724)
-#717 := (+ #120 #611)
-#720 := (>= #717 0::int)
-#725 := (iff #720 #724)
-#726 := [rewrite]: #725
-#721 := (iff #715 #720)
-#718 := (= #714 #717)
-#719 := [rewrite]: #718
-#722 := [monotonicity #719]: #721
-#728 := [trans #722 #726]: #727
-#777 := [monotonicity #728]: #776
-#783 := [monotonicity #777]: #782
-#786 := [trans #783 #785]: #782
-#781 := [quant-inst]: #780
-#787 := [mp #781 #786]: #779
-#1327 := [unit-resolution #787 #559]: #775
-#788 := (not #775)
-#791 := (or #788 #724 #772)
-#792 := [def-axiom]: #791
-#1329 := [unit-resolution #792 #1327]: #1328
-#1330 := [unit-resolution #1329 #1326]: #724
-#988 := (>= #622 0::int)
-#1331 := (or #1312 #988)
-#1332 := [th-lemma]: #1331
-#1333 := [unit-resolution #1332 #1281]: #988
-#761 := (not #724)
-#1334 := (not #988)
-#1335 := (or #1271 #1334 #761)
-#1336 := [th-lemma]: #1335
-#1337 := [unit-resolution #1336 #1333 #1330]: #1271
-#1338 := (not #1271)
-#1340 := (or #1269 #1338 #1339)
-#1341 := [th-lemma]: #1340
-#1343 := [unit-resolution #1341 #1337]: #1342
-#1344 := [unit-resolution #1343 #1299]: #1339
-#990 := (>= #916 0::int)
-#1345 := (or #1302 #990)
-#1346 := [th-lemma]: #1345
-#1347 := [unit-resolution #1346 #1301]: #990
-#1348 := (not #990)
-#1349 := (or #836 #1348 #1265)
-#1350 := [th-lemma]: #1349
-#1351 := [unit-resolution #1350 #1347 #1344]: #836
-#1353 := (or #815 #800)
-#801 := (uf_1 uf_13 #22)
-#820 := (= #799 #801)
-#823 := (ite #815 #820 #800)
-#476 := (= #32 #79)
-#475 := (= #34 #79)
-#477 := (ite #157 #475 #476)
-#568 := (forall (vars (?x6 T2)) (:pat #541 #551 #567 #553) #477)
-#480 := (forall (vars (?x6 T2)) #477)
-#571 := (iff #480 #568)
-#569 := (iff #477 #477)
-#570 := [refl]: #569
-#572 := [quant-intro #570]: #571
-#181 := (ite #157 #34 #32)
-#383 := (= #79 #181)
-#384 := (forall (vars (?x6 T2)) #383)
-#481 := (iff #384 #480)
-#478 := (iff #383 #477)
-#479 := [rewrite]: #478
-#482 := [quant-intro #479]: #481
-#352 := (~ #384 #384)
-#354 := (~ #383 #383)
-#355 := [refl]: #354
-#353 := [nnf-pos #355]: #352
-#51 := (uf_3 uf_16 #6)
-#186 := (= #51 #181)
-#189 := (forall (vars (?x6 T2)) #186)
-#385 := (iff #189 #384)
-#386 := [rewrite* #113]: #385
-#52 := (ite #42 #32 #34)
-#53 := (= #51 #52)
-#54 := (forall (vars (?x6 T2)) #53)
-#190 := (iff #54 #189)
-#187 := (iff #53 #186)
-#184 := (= #52 #181)
-#178 := (ite #156 #32 #34)
-#182 := (= #178 #181)
-#183 := [rewrite]: #182
-#179 := (= #52 #178)
-#180 := [monotonicity #161]: #179
-#185 := [trans #180 #183]: #184
-#188 := [monotonicity #185]: #187
-#191 := [quant-intro #188]: #190
-#139 := [asserted]: #54
-#192 := [mp #139 #191]: #189
-#387 := [mp #192 #386]: #384
-#402 := [mp~ #387 #353]: #384
-#483 := [mp #402 #482]: #480
-#573 := [mp #483 #572]: #568
-#634 := (not #568)
-#826 := (or #634 #823)
-#802 := (= #801 #799)
-#804 := (+ #803 #158)
-#805 := (>= #804 0::int)
-#806 := (ite #805 #802 #800)
-#827 := (or #634 #806)
-#829 := (iff #827 #826)
-#831 := (iff #826 #826)
-#832 := [rewrite]: #831
-#824 := (iff #806 #823)
-#821 := (iff #802 #820)
-#822 := [rewrite]: #821
-#818 := (iff #805 #815)
-#807 := (+ #158 #803)
-#810 := (>= #807 0::int)
-#816 := (iff #810 #815)
-#817 := [rewrite]: #816
-#811 := (iff #805 #810)
-#808 := (= #804 #807)
-#809 := [rewrite]: #808
-#812 := [monotonicity #809]: #811
-#819 := [trans #812 #817]: #818
-#825 := [monotonicity #819 #822]: #824
-#830 := [monotonicity #825]: #829
-#833 := [trans #830 #832]: #829
-#828 := [quant-inst]: #827
-#834 := [mp #828 #833]: #826
-#1352 := [unit-resolution #834 #573]: #823
-#835 := (not #823)
-#839 := (or #835 #815 #800)
-#840 := [def-axiom]: #839
-#1354 := [unit-resolution #840 #1352]: #1353
-#1355 := [unit-resolution #1354 #1351]: #800
-#1357 := [symm #1355]: #1356
-#1358 := [trans #1357 #1309]: #1266
-#1359 := (not #1266)
-#1360 := (or #1359 #1272)
-#1361 := [th-lemma]: #1360
-#1362 := [unit-resolution #1361 #1358]: #1272
-#1085 := (uf_1 uf_18 #22)
-#1099 := (* -1::real #1085)
-#1112 := (+ #902 #1099)
-#1113 := (<= #1112 0::real)
-#1137 := (not #1113)
-#960 := (uf_1 #88 #22)
-#1100 := (+ #960 #1099)
-#1101 := (>= #1100 0::real)
-#1118 := (or #1101 #1113)
-#1121 := (not #1118)
-#1125 := (or #1124 #1121)
-#1086 := (+ #1085 #1084)
-#1087 := (>= #1086 0::real)
-#1088 := (* -1::real #960)
-#1089 := (+ #1085 #1088)
-#1090 := (<= #1089 0::real)
-#1091 := (or #1090 #1087)
-#1092 := (not #1091)
-#1126 := (or #1124 #1092)
-#1128 := (iff #1126 #1125)
-#1130 := (iff #1125 #1125)
-#1131 := [rewrite]: #1130
-#1122 := (iff #1092 #1121)
-#1119 := (iff #1091 #1118)
-#1116 := (iff #1087 #1113)
-#1106 := (+ #1084 #1085)
-#1109 := (>= #1106 0::real)
-#1114 := (iff #1109 #1113)
-#1115 := [rewrite]: #1114
-#1110 := (iff #1087 #1109)
-#1107 := (= #1086 #1106)
-#1108 := [rewrite]: #1107
-#1111 := [monotonicity #1108]: #1110
-#1117 := [trans #1111 #1115]: #1116
-#1104 := (iff #1090 #1101)
-#1093 := (+ #1088 #1085)
-#1096 := (<= #1093 0::real)
-#1102 := (iff #1096 #1101)
-#1103 := [rewrite]: #1102
-#1097 := (iff #1090 #1096)
-#1094 := (= #1089 #1093)
-#1095 := [rewrite]: #1094
-#1098 := [monotonicity #1095]: #1097
-#1105 := [trans #1098 #1103]: #1104
-#1120 := [monotonicity #1105 #1117]: #1119
-#1123 := [monotonicity #1120]: #1122
-#1129 := [monotonicity #1123]: #1128
-#1132 := [trans #1129 #1131]: #1128
-#1127 := [quant-inst]: #1126
-#1133 := [mp #1127 #1132]: #1125
-#1363 := [unit-resolution #1133 #606]: #1121
-#1138 := (or #1118 #1137)
-#1139 := [def-axiom]: #1138
-#1364 := [unit-resolution #1139 #1363]: #1137
-#1200 := (+ #799 #1099)
-#1201 := (>= #1200 0::real)
-#1231 := (not #1201)
-#847 := (uf_1 #83 #22)
-#1210 := (+ #847 #1099)
-#1211 := (<= #1210 0::real)
-#1216 := (or #1201 #1211)
-#1219 := (not #1216)
-#1222 := (or #1066 #1219)
-#1197 := (* -1::real #847)
-#1198 := (+ #1085 #1197)
-#1199 := (>= #1198 0::real)
-#1202 := (or #1201 #1199)
-#1203 := (not #1202)
-#1223 := (or #1066 #1203)
-#1225 := (iff #1223 #1222)
-#1227 := (iff #1222 #1222)
-#1228 := [rewrite]: #1227
-#1220 := (iff #1203 #1219)
-#1217 := (iff #1202 #1216)
-#1214 := (iff #1199 #1211)
-#1204 := (+ #1197 #1085)
-#1207 := (>= #1204 0::real)
-#1212 := (iff #1207 #1211)
-#1213 := [rewrite]: #1212
-#1208 := (iff #1199 #1207)
-#1205 := (= #1198 #1204)
-#1206 := [rewrite]: #1205
-#1209 := [monotonicity #1206]: #1208
-#1215 := [trans #1209 #1213]: #1214
-#1218 := [monotonicity #1215]: #1217
-#1221 := [monotonicity #1218]: #1220
-#1226 := [monotonicity #1221]: #1225
-#1229 := [trans #1226 #1228]: #1225
-#1224 := [quant-inst]: #1223
-#1230 := [mp #1224 #1229]: #1222
-#1365 := [unit-resolution #1230 #600]: #1219
-#1232 := (or #1216 #1231)
-#1233 := [def-axiom]: #1232
-#1366 := [unit-resolution #1233 #1365]: #1231
-[th-lemma #1366 #1364 #1362]: false
-unsat
-NQHwTeL311Tq3wf2s5BReA 419 0
-#2 := false
-#194 := 0::real
-decl uf_4 :: (-> T2 T3 real)
-decl uf_6 :: (-> T1 T3)
-decl uf_3 :: T1
-#21 := uf_3
-#25 := (uf_6 uf_3)
-decl uf_5 :: T2
-#24 := uf_5
-#26 := (uf_4 uf_5 #25)
-decl uf_7 :: T2
-#27 := uf_7
-#28 := (uf_4 uf_7 #25)
-decl uf_10 :: T1
-#38 := uf_10
-#42 := (uf_6 uf_10)
-decl uf_9 :: T2
-#33 := uf_9
-#43 := (uf_4 uf_9 #42)
-#41 := (= uf_3 uf_10)
-#44 := (ite #41 #43 #28)
-#9 := 0::int
-decl uf_2 :: (-> T1 int)
-#39 := (uf_2 uf_10)
-#226 := -1::int
-#229 := (* -1::int #39)
-#22 := (uf_2 uf_3)
-#230 := (+ #22 #229)
-#228 := (>= #230 0::int)
-#236 := (ite #228 #44 #26)
-#192 := -1::real
-#244 := (* -1::real #236)
-#642 := (+ #26 #244)
-#643 := (<= #642 0::real)
-#567 := (= #26 #236)
-#227 := (not #228)
-decl uf_1 :: (-> int T1)
-#593 := (uf_1 #39)
-#660 := (= #593 uf_10)
-#594 := (= uf_10 #593)
-#4 := (:var 0 T1)
-#5 := (uf_2 #4)
-#546 := (pattern #5)
-#6 := (uf_1 #5)
-#93 := (= #4 #6)
-#547 := (forall (vars (?x1 T1)) (:pat #546) #93)
-#96 := (forall (vars (?x1 T1)) #93)
-#550 := (iff #96 #547)
-#548 := (iff #93 #93)
-#549 := [refl]: #548
-#551 := [quant-intro #549]: #550
-#448 := (~ #96 #96)
-#450 := (~ #93 #93)
-#451 := [refl]: #450
-#449 := [nnf-pos #451]: #448
-#7 := (= #6 #4)
-#8 := (forall (vars (?x1 T1)) #7)
-#97 := (iff #8 #96)
-#94 := (iff #7 #93)
-#95 := [rewrite]: #94
-#98 := [quant-intro #95]: #97
-#92 := [asserted]: #8
-#101 := [mp #92 #98]: #96
-#446 := [mp~ #101 #449]: #96
-#552 := [mp #446 #551]: #547
-#595 := (not #547)
-#600 := (or #595 #594)
-#601 := [quant-inst]: #600
-#654 := [unit-resolution #601 #552]: #594
-#680 := [symm #654]: #660
-#681 := (= uf_3 #593)
-#591 := (uf_1 #22)
-#658 := (= #591 #593)
-#656 := (= #593 #591)
-#652 := (= #39 #22)
-#647 := (= #22 #39)
-#290 := (<= #230 0::int)
-#70 := (<= #22 #39)
-#388 := (iff #70 #290)
-#389 := [rewrite]: #388
-#341 := [asserted]: #70
-#390 := [mp #341 #389]: #290
-#646 := [hypothesis]: #228
-#648 := [th-lemma #646 #390]: #647
-#653 := [symm #648]: #652
-#657 := [monotonicity #653]: #656
-#659 := [symm #657]: #658
-#592 := (= uf_3 #591)
-#596 := (or #595 #592)
-#597 := [quant-inst]: #596
-#655 := [unit-resolution #597 #552]: #592
-#682 := [trans #655 #659]: #681
-#683 := [trans #682 #680]: #41
-#570 := (not #41)
-decl uf_11 :: T2
-#47 := uf_11
-#59 := (uf_4 uf_11 #42)
-#278 := (ite #41 #26 #59)
-#459 := (* -1::real #278)
-#637 := (+ #26 #459)
-#639 := (>= #637 0::real)
-#585 := (= #26 #278)
-#661 := [hypothesis]: #41
-#587 := (or #570 #585)
-#588 := [def-axiom]: #587
-#662 := [unit-resolution #588 #661]: #585
-#663 := (not #585)
-#664 := (or #663 #639)
-#665 := [th-lemma]: #664
-#666 := [unit-resolution #665 #662]: #639
-decl uf_8 :: T2
-#30 := uf_8
-#56 := (uf_4 uf_8 #42)
-#357 := (* -1::real #56)
-#358 := (+ #43 #357)
-#356 := (>= #358 0::real)
-#355 := (not #356)
-#374 := (* -1::real #59)
-#375 := (+ #56 #374)
-#373 := (>= #375 0::real)
-#376 := (not #373)
-#381 := (and #355 #376)
-#64 := (< #39 #39)
-#67 := (ite #64 #43 #59)
-#68 := (< #56 #67)
-#53 := (uf_4 uf_5 #42)
-#65 := (ite #64 #53 #43)
-#66 := (< #65 #56)
-#69 := (and #66 #68)
-#382 := (iff #69 #381)
-#379 := (iff #68 #376)
-#370 := (< #56 #59)
-#377 := (iff #370 #376)
-#378 := [rewrite]: #377
-#371 := (iff #68 #370)
-#368 := (= #67 #59)
-#363 := (ite false #43 #59)
-#366 := (= #363 #59)
-#367 := [rewrite]: #366
-#364 := (= #67 #363)
-#343 := (iff #64 false)
-#344 := [rewrite]: #343
-#365 := [monotonicity #344]: #364
-#369 := [trans #365 #367]: #368
-#372 := [monotonicity #369]: #371
-#380 := [trans #372 #378]: #379
-#361 := (iff #66 #355)
-#352 := (< #43 #56)
-#359 := (iff #352 #355)
-#360 := [rewrite]: #359
-#353 := (iff #66 #352)
-#350 := (= #65 #43)
-#345 := (ite false #53 #43)
-#348 := (= #345 #43)
-#349 := [rewrite]: #348
-#346 := (= #65 #345)
-#347 := [monotonicity #344]: #346
-#351 := [trans #347 #349]: #350
-#354 := [monotonicity #351]: #353
-#362 := [trans #354 #360]: #361
-#383 := [monotonicity #362 #380]: #382
-#340 := [asserted]: #69
-#384 := [mp #340 #383]: #381
-#385 := [and-elim #384]: #355
-#394 := (* -1::real #53)
-#395 := (+ #43 #394)
-#393 := (>= #395 0::real)
-#54 := (uf_4 uf_7 #42)
-#402 := (* -1::real #54)
-#403 := (+ #53 #402)
-#401 := (>= #403 0::real)
-#397 := (+ #43 #374)
-#398 := (<= #397 0::real)
-#412 := (and #393 #398 #401)
-#73 := (<= #43 #59)
-#72 := (<= #53 #43)
-#74 := (and #72 #73)
-#71 := (<= #54 #53)
-#75 := (and #71 #74)
-#415 := (iff #75 #412)
-#406 := (and #393 #398)
-#409 := (and #401 #406)
-#413 := (iff #409 #412)
-#414 := [rewrite]: #413
-#410 := (iff #75 #409)
-#407 := (iff #74 #406)
-#399 := (iff #73 #398)
-#400 := [rewrite]: #399
-#392 := (iff #72 #393)
-#396 := [rewrite]: #392
-#408 := [monotonicity #396 #400]: #407
-#404 := (iff #71 #401)
-#405 := [rewrite]: #404
-#411 := [monotonicity #405 #408]: #410
-#416 := [trans #411 #414]: #415
-#342 := [asserted]: #75
-#417 := [mp #342 #416]: #412
-#418 := [and-elim #417]: #393
-#650 := (+ #26 #394)
-#651 := (<= #650 0::real)
-#649 := (= #26 #53)
-#671 := (= #53 #26)
-#669 := (= #42 #25)
-#667 := (= #25 #42)
-#668 := [monotonicity #661]: #667
-#670 := [symm #668]: #669
-#672 := [monotonicity #670]: #671
-#673 := [symm #672]: #649
-#674 := (not #649)
-#675 := (or #674 #651)
-#676 := [th-lemma]: #675
-#677 := [unit-resolution #676 #673]: #651
-#462 := (+ #56 #459)
-#465 := (>= #462 0::real)
-#438 := (not #465)
-#316 := (ite #290 #278 #43)
-#326 := (* -1::real #316)
-#327 := (+ #56 #326)
-#325 := (>= #327 0::real)
-#324 := (not #325)
-#439 := (iff #324 #438)
-#466 := (iff #325 #465)
-#463 := (= #327 #462)
-#460 := (= #326 #459)
-#457 := (= #316 #278)
-#1 := true
-#452 := (ite true #278 #43)
-#455 := (= #452 #278)
-#456 := [rewrite]: #455
-#453 := (= #316 #452)
-#444 := (iff #290 true)
-#445 := [iff-true #390]: #444
-#454 := [monotonicity #445]: #453
-#458 := [trans #454 #456]: #457
-#461 := [monotonicity #458]: #460
-#464 := [monotonicity #461]: #463
-#467 := [monotonicity #464]: #466
-#468 := [monotonicity #467]: #439
-#297 := (ite #290 #54 #53)
-#305 := (* -1::real #297)
-#306 := (+ #56 #305)
-#307 := (<= #306 0::real)
-#308 := (not #307)
-#332 := (and #308 #324)
-#58 := (= uf_10 uf_3)
-#60 := (ite #58 #26 #59)
-#52 := (< #39 #22)
-#61 := (ite #52 #43 #60)
-#62 := (< #56 #61)
-#55 := (ite #52 #53 #54)
-#57 := (< #55 #56)
-#63 := (and #57 #62)
-#335 := (iff #63 #332)
-#281 := (ite #52 #43 #278)
-#284 := (< #56 #281)
-#287 := (and #57 #284)
-#333 := (iff #287 #332)
-#330 := (iff #284 #324)
-#321 := (< #56 #316)
-#328 := (iff #321 #324)
-#329 := [rewrite]: #328
-#322 := (iff #284 #321)
-#319 := (= #281 #316)
-#291 := (not #290)
-#313 := (ite #291 #43 #278)
-#317 := (= #313 #316)
-#318 := [rewrite]: #317
-#314 := (= #281 #313)
-#292 := (iff #52 #291)
-#293 := [rewrite]: #292
-#315 := [monotonicity #293]: #314
-#320 := [trans #315 #318]: #319
-#323 := [monotonicity #320]: #322
-#331 := [trans #323 #329]: #330
-#311 := (iff #57 #308)
-#302 := (< #297 #56)
-#309 := (iff #302 #308)
-#310 := [rewrite]: #309
-#303 := (iff #57 #302)
-#300 := (= #55 #297)
-#294 := (ite #291 #53 #54)
-#298 := (= #294 #297)
-#299 := [rewrite]: #298
-#295 := (= #55 #294)
-#296 := [monotonicity #293]: #295
-#301 := [trans #296 #299]: #300
-#304 := [monotonicity #301]: #303
-#312 := [trans #304 #310]: #311
-#334 := [monotonicity #312 #331]: #333
-#288 := (iff #63 #287)
-#285 := (iff #62 #284)
-#282 := (= #61 #281)
-#279 := (= #60 #278)
-#225 := (iff #58 #41)
-#277 := [rewrite]: #225
-#280 := [monotonicity #277]: #279
-#283 := [monotonicity #280]: #282
-#286 := [monotonicity #283]: #285
-#289 := [monotonicity #286]: #288
-#336 := [trans #289 #334]: #335
-#179 := [asserted]: #63
-#337 := [mp #179 #336]: #332
-#339 := [and-elim #337]: #324
-#469 := [mp #339 #468]: #438
-#678 := [th-lemma #469 #677 #418 #385 #666]: false
-#679 := [lemma #678]: #570
-#684 := [unit-resolution #679 #683]: false
-#685 := [lemma #684]: #227
-#577 := (or #228 #567)
-#578 := [def-axiom]: #577
-#645 := [unit-resolution #578 #685]: #567
-#686 := (not #567)
-#687 := (or #686 #643)
-#688 := [th-lemma]: #687
-#689 := [unit-resolution #688 #645]: #643
-#31 := (uf_4 uf_8 #25)
-#245 := (+ #31 #244)
-#246 := (<= #245 0::real)
-#247 := (not #246)
-#34 := (uf_4 uf_9 #25)
-#48 := (uf_4 uf_11 #25)
-#255 := (ite #228 #48 #34)
-#264 := (* -1::real #255)
-#265 := (+ #31 #264)
-#263 := (>= #265 0::real)
-#266 := (not #263)
-#271 := (and #247 #266)
-#40 := (< #22 #39)
-#49 := (ite #40 #34 #48)
-#50 := (< #31 #49)
-#45 := (ite #40 #26 #44)
-#46 := (< #45 #31)
-#51 := (and #46 #50)
-#272 := (iff #51 #271)
-#269 := (iff #50 #266)
-#260 := (< #31 #255)
-#267 := (iff #260 #266)
-#268 := [rewrite]: #267
-#261 := (iff #50 #260)
-#258 := (= #49 #255)
-#252 := (ite #227 #34 #48)
-#256 := (= #252 #255)
-#257 := [rewrite]: #256
-#253 := (= #49 #252)
-#231 := (iff #40 #227)
-#232 := [rewrite]: #231
-#254 := [monotonicity #232]: #253
-#259 := [trans #254 #257]: #258
-#262 := [monotonicity #259]: #261
-#270 := [trans #262 #268]: #269
-#250 := (iff #46 #247)
-#241 := (< #236 #31)
-#248 := (iff #241 #247)
-#249 := [rewrite]: #248
-#242 := (iff #46 #241)
-#239 := (= #45 #236)
-#233 := (ite #227 #26 #44)
-#237 := (= #233 #236)
-#238 := [rewrite]: #237
-#234 := (= #45 #233)
-#235 := [monotonicity #232]: #234
-#240 := [trans #235 #238]: #239
-#243 := [monotonicity #240]: #242
-#251 := [trans #243 #249]: #250
-#273 := [monotonicity #251 #270]: #272
-#178 := [asserted]: #51
-#274 := [mp #178 #273]: #271
-#275 := [and-elim #274]: #247
-#196 := (* -1::real #31)
-#212 := (+ #26 #196)
-#213 := (<= #212 0::real)
-#214 := (not #213)
-#197 := (+ #28 #196)
-#195 := (>= #197 0::real)
-#193 := (not #195)
-#219 := (and #193 #214)
-#23 := (< #22 #22)
-#35 := (ite #23 #34 #26)
-#36 := (< #31 #35)
-#29 := (ite #23 #26 #28)
-#32 := (< #29 #31)
-#37 := (and #32 #36)
-#220 := (iff #37 #219)
-#217 := (iff #36 #214)
-#209 := (< #31 #26)
-#215 := (iff #209 #214)
-#216 := [rewrite]: #215
-#210 := (iff #36 #209)
-#207 := (= #35 #26)
-#202 := (ite false #34 #26)
-#205 := (= #202 #26)
-#206 := [rewrite]: #205
-#203 := (= #35 #202)
-#180 := (iff #23 false)
-#181 := [rewrite]: #180
-#204 := [monotonicity #181]: #203
-#208 := [trans #204 #206]: #207
-#211 := [monotonicity #208]: #210
-#218 := [trans #211 #216]: #217
-#200 := (iff #32 #193)
-#189 := (< #28 #31)
-#198 := (iff #189 #193)
-#199 := [rewrite]: #198
-#190 := (iff #32 #189)
-#187 := (= #29 #28)
-#182 := (ite false #26 #28)
-#185 := (= #182 #28)
-#186 := [rewrite]: #185
-#183 := (= #29 #182)
-#184 := [monotonicity #181]: #183
-#188 := [trans #184 #186]: #187
-#191 := [monotonicity #188]: #190
-#201 := [trans #191 #199]: #200
-#221 := [monotonicity #201 #218]: #220
-#177 := [asserted]: #37
-#222 := [mp #177 #221]: #219
-#224 := [and-elim #222]: #214
-[th-lemma #224 #275 #689]: false
-unsat
-NX/HT1QOfbspC2LtZNKpBA 428 0
-#2 := false
-decl uf_10 :: T1
-#38 := uf_10
-decl uf_3 :: T1
-#21 := uf_3
-#45 := (= uf_3 uf_10)
-decl uf_1 :: (-> int T1)
-decl uf_2 :: (-> T1 int)
-#39 := (uf_2 uf_10)
-#588 := (uf_1 #39)
-#686 := (= #588 uf_10)
-#589 := (= uf_10 #588)
-#4 := (:var 0 T1)
-#5 := (uf_2 #4)
-#541 := (pattern #5)
-#6 := (uf_1 #5)
-#93 := (= #4 #6)
-#542 := (forall (vars (?x1 T1)) (:pat #541) #93)
-#96 := (forall (vars (?x1 T1)) #93)
-#545 := (iff #96 #542)
-#543 := (iff #93 #93)
-#544 := [refl]: #543
-#546 := [quant-intro #544]: #545
-#454 := (~ #96 #96)
-#456 := (~ #93 #93)
-#457 := [refl]: #456
-#455 := [nnf-pos #457]: #454
-#7 := (= #6 #4)
-#8 := (forall (vars (?x1 T1)) #7)
-#97 := (iff #8 #96)
-#94 := (iff #7 #93)
-#95 := [rewrite]: #94
-#98 := [quant-intro #95]: #97
-#92 := [asserted]: #8
-#101 := [mp #92 #98]: #96
-#452 := [mp~ #101 #455]: #96
-#547 := [mp #452 #546]: #542
-#590 := (not #542)
-#595 := (or #590 #589)
-#596 := [quant-inst]: #595
-#680 := [unit-resolution #596 #547]: #589
-#687 := [symm #680]: #686
-#688 := (= uf_3 #588)
-#22 := (uf_2 uf_3)
-#586 := (uf_1 #22)
-#684 := (= #586 #588)
-#682 := (= #588 #586)
-#678 := (= #39 #22)
-#676 := (= #22 #39)
-#9 := 0::int
-#227 := -1::int
-#230 := (* -1::int #39)
-#231 := (+ #22 #230)
-#296 := (<= #231 0::int)
-#70 := (<= #22 #39)
-#393 := (iff #70 #296)
-#394 := [rewrite]: #393
-#347 := [asserted]: #70
-#395 := [mp #347 #394]: #296
-#229 := (>= #231 0::int)
-decl uf_4 :: (-> T2 T3 real)
-decl uf_6 :: (-> T1 T3)
-#25 := (uf_6 uf_3)
-decl uf_7 :: T2
-#27 := uf_7
-#28 := (uf_4 uf_7 #25)
-decl uf_9 :: T2
-#33 := uf_9
-#34 := (uf_4 uf_9 #25)
-#46 := (uf_6 uf_10)
-decl uf_5 :: T2
-#24 := uf_5
-#47 := (uf_4 uf_5 #46)
-#48 := (ite #45 #47 #34)
-#256 := (ite #229 #48 #28)
-#568 := (= #28 #256)
-#648 := (not #568)
-#194 := 0::real
-#192 := -1::real
-#265 := (* -1::real #256)
-#640 := (+ #28 #265)
-#642 := (>= #640 0::real)
-#645 := (not #642)
-#643 := [hypothesis]: #642
-decl uf_8 :: T2
-#30 := uf_8
-#31 := (uf_4 uf_8 #25)
-#266 := (+ #31 #265)
-#264 := (>= #266 0::real)
-#267 := (not #264)
-#26 := (uf_4 uf_5 #25)
-decl uf_11 :: T2
-#41 := uf_11
-#42 := (uf_4 uf_11 #25)
-#237 := (ite #229 #42 #26)
-#245 := (* -1::real #237)
-#246 := (+ #31 #245)
-#247 := (<= #246 0::real)
-#248 := (not #247)
-#272 := (and #248 #267)
-#40 := (< #22 #39)
-#49 := (ite #40 #28 #48)
-#50 := (< #31 #49)
-#43 := (ite #40 #26 #42)
-#44 := (< #43 #31)
-#51 := (and #44 #50)
-#273 := (iff #51 #272)
-#270 := (iff #50 #267)
-#261 := (< #31 #256)
-#268 := (iff #261 #267)
-#269 := [rewrite]: #268
-#262 := (iff #50 #261)
-#259 := (= #49 #256)
-#228 := (not #229)
-#253 := (ite #228 #28 #48)
-#257 := (= #253 #256)
-#258 := [rewrite]: #257
-#254 := (= #49 #253)
-#232 := (iff #40 #228)
-#233 := [rewrite]: #232
-#255 := [monotonicity #233]: #254
-#260 := [trans #255 #258]: #259
-#263 := [monotonicity #260]: #262
-#271 := [trans #263 #269]: #270
-#251 := (iff #44 #248)
-#242 := (< #237 #31)
-#249 := (iff #242 #248)
-#250 := [rewrite]: #249
-#243 := (iff #44 #242)
-#240 := (= #43 #237)
-#234 := (ite #228 #26 #42)
-#238 := (= #234 #237)
-#239 := [rewrite]: #238
-#235 := (= #43 #234)
-#236 := [monotonicity #233]: #235
-#241 := [trans #236 #239]: #240
-#244 := [monotonicity #241]: #243
-#252 := [trans #244 #250]: #251
-#274 := [monotonicity #252 #271]: #273
-#178 := [asserted]: #51
-#275 := [mp #178 #274]: #272
-#277 := [and-elim #275]: #267
-#196 := (* -1::real #31)
-#197 := (+ #28 #196)
-#195 := (>= #197 0::real)
-#193 := (not #195)
-#213 := (* -1::real #34)
-#214 := (+ #31 #213)
-#212 := (>= #214 0::real)
-#215 := (not #212)
-#220 := (and #193 #215)
-#23 := (< #22 #22)
-#35 := (ite #23 #28 #34)
-#36 := (< #31 #35)
-#29 := (ite #23 #26 #28)
-#32 := (< #29 #31)
-#37 := (and #32 #36)
-#221 := (iff #37 #220)
-#218 := (iff #36 #215)
-#209 := (< #31 #34)
-#216 := (iff #209 #215)
-#217 := [rewrite]: #216
-#210 := (iff #36 #209)
-#207 := (= #35 #34)
-#202 := (ite false #28 #34)
-#205 := (= #202 #34)
-#206 := [rewrite]: #205
-#203 := (= #35 #202)
-#180 := (iff #23 false)
-#181 := [rewrite]: #180
-#204 := [monotonicity #181]: #203
-#208 := [trans #204 #206]: #207
-#211 := [monotonicity #208]: #210
-#219 := [trans #211 #217]: #218
-#200 := (iff #32 #193)
-#189 := (< #28 #31)
-#198 := (iff #189 #193)
-#199 := [rewrite]: #198
-#190 := (iff #32 #189)
-#187 := (= #29 #28)
-#182 := (ite false #26 #28)
-#185 := (= #182 #28)
-#186 := [rewrite]: #185
-#183 := (= #29 #182)
-#184 := [monotonicity #181]: #183
-#188 := [trans #184 #186]: #187
-#191 := [monotonicity #188]: #190
-#201 := [trans #191 #199]: #200
-#222 := [monotonicity #201 #219]: #221
-#177 := [asserted]: #37
-#223 := [mp #177 #222]: #220
-#224 := [and-elim #223]: #193
-#644 := [th-lemma #224 #277 #643]: false
-#646 := [lemma #644]: #645
-#647 := [hypothesis]: #568
-#649 := (or #648 #642)
-#650 := [th-lemma]: #649
-#651 := [unit-resolution #650 #647 #646]: false
-#652 := [lemma #651]: #648
-#578 := (or #229 #568)
-#579 := [def-axiom]: #578
-#675 := [unit-resolution #579 #652]: #229
-#677 := [th-lemma #675 #395]: #676
-#679 := [symm #677]: #678
-#683 := [monotonicity #679]: #682
-#685 := [symm #683]: #684
-#587 := (= uf_3 #586)
-#591 := (or #590 #587)
-#592 := [quant-inst]: #591
-#681 := [unit-resolution #592 #547]: #587
-#689 := [trans #681 #685]: #688
-#690 := [trans #689 #687]: #45
-#571 := (not #45)
-#54 := (uf_4 uf_11 #46)
-#279 := (ite #45 #28 #54)
-#465 := (* -1::real #279)
-#632 := (+ #28 #465)
-#633 := (<= #632 0::real)
-#580 := (= #28 #279)
-#656 := [hypothesis]: #45
-#582 := (or #571 #580)
-#583 := [def-axiom]: #582
-#657 := [unit-resolution #583 #656]: #580
-#658 := (not #580)
-#659 := (or #658 #633)
-#660 := [th-lemma]: #659
-#661 := [unit-resolution #660 #657]: #633
-#57 := (uf_4 uf_8 #46)
-#363 := (* -1::real #57)
-#379 := (+ #47 #363)
-#380 := (<= #379 0::real)
-#381 := (not #380)
-#364 := (+ #54 #363)
-#362 := (>= #364 0::real)
-#361 := (not #362)
-#386 := (and #361 #381)
-#59 := (uf_4 uf_7 #46)
-#64 := (< #39 #39)
-#67 := (ite #64 #59 #47)
-#68 := (< #57 #67)
-#65 := (ite #64 #47 #54)
-#66 := (< #65 #57)
-#69 := (and #66 #68)
-#387 := (iff #69 #386)
-#384 := (iff #68 #381)
-#376 := (< #57 #47)
-#382 := (iff #376 #381)
-#383 := [rewrite]: #382
-#377 := (iff #68 #376)
-#374 := (= #67 #47)
-#369 := (ite false #59 #47)
-#372 := (= #369 #47)
-#373 := [rewrite]: #372
-#370 := (= #67 #369)
-#349 := (iff #64 false)
-#350 := [rewrite]: #349
-#371 := [monotonicity #350]: #370
-#375 := [trans #371 #373]: #374
-#378 := [monotonicity #375]: #377
-#385 := [trans #378 #383]: #384
-#367 := (iff #66 #361)
-#358 := (< #54 #57)
-#365 := (iff #358 #361)
-#366 := [rewrite]: #365
-#359 := (iff #66 #358)
-#356 := (= #65 #54)
-#351 := (ite false #47 #54)
-#354 := (= #351 #54)
-#355 := [rewrite]: #354
-#352 := (= #65 #351)
-#353 := [monotonicity #350]: #352
-#357 := [trans #353 #355]: #356
-#360 := [monotonicity #357]: #359
-#368 := [trans #360 #366]: #367
-#388 := [monotonicity #368 #385]: #387
-#346 := [asserted]: #69
-#389 := [mp #346 #388]: #386
-#391 := [and-elim #389]: #381
-#397 := (* -1::real #59)
-#398 := (+ #47 #397)
-#399 := (<= #398 0::real)
-#409 := (* -1::real #54)
-#410 := (+ #47 #409)
-#408 := (>= #410 0::real)
-#60 := (uf_4 uf_9 #46)
-#402 := (* -1::real #60)
-#403 := (+ #59 #402)
-#404 := (<= #403 0::real)
-#418 := (and #399 #404 #408)
-#73 := (<= #59 #60)
-#72 := (<= #47 #59)
-#74 := (and #72 #73)
-#71 := (<= #54 #47)
-#75 := (and #71 #74)
-#421 := (iff #75 #418)
-#412 := (and #399 #404)
-#415 := (and #408 #412)
-#419 := (iff #415 #418)
-#420 := [rewrite]: #419
-#416 := (iff #75 #415)
-#413 := (iff #74 #412)
-#405 := (iff #73 #404)
-#406 := [rewrite]: #405
-#400 := (iff #72 #399)
-#401 := [rewrite]: #400
-#414 := [monotonicity #401 #406]: #413
-#407 := (iff #71 #408)
-#411 := [rewrite]: #407
-#417 := [monotonicity #411 #414]: #416
-#422 := [trans #417 #420]: #421
-#348 := [asserted]: #75
-#423 := [mp #348 #422]: #418
-#424 := [and-elim #423]: #399
-#637 := (+ #28 #397)
-#639 := (>= #637 0::real)
-#636 := (= #28 #59)
-#666 := (= #59 #28)
-#664 := (= #46 #25)
-#662 := (= #25 #46)
-#663 := [monotonicity #656]: #662
-#665 := [symm #663]: #664
-#667 := [monotonicity #665]: #666
-#668 := [symm #667]: #636
-#669 := (not #636)
-#670 := (or #669 #639)
-#671 := [th-lemma]: #670
-#672 := [unit-resolution #671 #668]: #639
-#468 := (+ #57 #465)
-#471 := (<= #468 0::real)
-#444 := (not #471)
-#322 := (ite #296 #279 #47)
-#330 := (* -1::real #322)
-#331 := (+ #57 #330)
-#332 := (<= #331 0::real)
-#333 := (not #332)
-#445 := (iff #333 #444)
-#472 := (iff #332 #471)
-#469 := (= #331 #468)
-#466 := (= #330 #465)
-#463 := (= #322 #279)
-#1 := true
-#458 := (ite true #279 #47)
-#461 := (= #458 #279)
-#462 := [rewrite]: #461
-#459 := (= #322 #458)
-#450 := (iff #296 true)
-#451 := [iff-true #395]: #450
-#460 := [monotonicity #451]: #459
-#464 := [trans #460 #462]: #463
-#467 := [monotonicity #464]: #466
-#470 := [monotonicity #467]: #469
-#473 := [monotonicity #470]: #472
-#474 := [monotonicity #473]: #445
-#303 := (ite #296 #60 #59)
-#313 := (* -1::real #303)
-#314 := (+ #57 #313)
-#312 := (>= #314 0::real)
-#311 := (not #312)
-#338 := (and #311 #333)
-#52 := (< #39 #22)
-#61 := (ite #52 #59 #60)
-#62 := (< #57 #61)
-#53 := (= uf_10 uf_3)
-#55 := (ite #53 #28 #54)
-#56 := (ite #52 #47 #55)
-#58 := (< #56 #57)
-#63 := (and #58 #62)
-#341 := (iff #63 #338)
-#282 := (ite #52 #47 #279)
-#285 := (< #282 #57)
-#291 := (and #62 #285)
-#339 := (iff #291 #338)
-#336 := (iff #285 #333)
-#327 := (< #322 #57)
-#334 := (iff #327 #333)
-#335 := [rewrite]: #334
-#328 := (iff #285 #327)
-#325 := (= #282 #322)
-#297 := (not #296)
-#319 := (ite #297 #47 #279)
-#323 := (= #319 #322)
-#324 := [rewrite]: #323
-#320 := (= #282 #319)
-#298 := (iff #52 #297)
-#299 := [rewrite]: #298
-#321 := [monotonicity #299]: #320
-#326 := [trans #321 #324]: #325
-#329 := [monotonicity #326]: #328
-#337 := [trans #329 #335]: #336
-#317 := (iff #62 #311)
-#308 := (< #57 #303)
-#315 := (iff #308 #311)
-#316 := [rewrite]: #315
-#309 := (iff #62 #308)
-#306 := (= #61 #303)
-#300 := (ite #297 #59 #60)
-#304 := (= #300 #303)
-#305 := [rewrite]: #304
-#301 := (= #61 #300)
-#302 := [monotonicity #299]: #301
-#307 := [trans #302 #305]: #306
-#310 := [monotonicity #307]: #309
-#318 := [trans #310 #316]: #317
-#340 := [monotonicity #318 #337]: #339
-#294 := (iff #63 #291)
-#288 := (and #285 #62)
-#292 := (iff #288 #291)
-#293 := [rewrite]: #292
-#289 := (iff #63 #288)
-#286 := (iff #58 #285)
-#283 := (= #56 #282)
-#280 := (= #55 #279)
-#226 := (iff #53 #45)
-#278 := [rewrite]: #226
-#281 := [monotonicity #278]: #280
-#284 := [monotonicity #281]: #283
-#287 := [monotonicity #284]: #286
-#290 := [monotonicity #287]: #289
-#295 := [trans #290 #293]: #294
-#342 := [trans #295 #340]: #341
-#179 := [asserted]: #63
-#343 := [mp #179 #342]: #338
-#345 := [and-elim #343]: #333
-#475 := [mp #345 #474]: #444
-#673 := [th-lemma #475 #672 #424 #391 #661]: false
-#674 := [lemma #673]: #571
-[unit-resolution #674 #690]: false
-unsat
-IL2powemHjRpCJYwmXFxyw 211 0
-#2 := false
-#33 := 0::real
-decl uf_11 :: (-> T5 T6 real)
-decl uf_15 :: T6
-#28 := uf_15
-decl uf_16 :: T5
-#30 := uf_16
-#31 := (uf_11 uf_16 uf_15)
-decl uf_12 :: (-> T7 T8 T5)
-decl uf_14 :: T8
-#26 := uf_14
-decl uf_13 :: (-> T1 T7)
-decl uf_8 :: T1
-#16 := uf_8
-#25 := (uf_13 uf_8)
-#27 := (uf_12 #25 uf_14)
-#29 := (uf_11 #27 uf_15)
-#73 := -1::real
-#84 := (* -1::real #29)
-#85 := (+ #84 #31)
-#74 := (* -1::real #31)
-#75 := (+ #29 #74)
-#112 := (>= #75 0::real)
-#119 := (ite #112 #75 #85)
-#127 := (* -1::real #119)
-decl uf_17 :: T5
-#37 := uf_17
-#38 := (uf_11 uf_17 uf_15)
-#102 := -1/3::real
-#103 := (* -1/3::real #38)
-#128 := (+ #103 #127)
-#100 := 1/3::real
-#101 := (* 1/3::real #31)
-#129 := (+ #101 #128)
-#130 := (<= #129 0::real)
-#131 := (not #130)
-#40 := 3::real
-#39 := (- #31 #38)
-#41 := (/ #39 3::real)
-#32 := (- #29 #31)
-#35 := (- #32)
-#34 := (< #32 0::real)
-#36 := (ite #34 #35 #32)
-#42 := (< #36 #41)
-#136 := (iff #42 #131)
-#104 := (+ #101 #103)
-#78 := (< #75 0::real)
-#90 := (ite #78 #85 #75)
-#109 := (< #90 #104)
-#134 := (iff #109 #131)
-#124 := (< #119 #104)
-#132 := (iff #124 #131)
-#133 := [rewrite]: #132
-#125 := (iff #109 #124)
-#122 := (= #90 #119)
-#113 := (not #112)
-#116 := (ite #113 #85 #75)
-#120 := (= #116 #119)
-#121 := [rewrite]: #120
-#117 := (= #90 #116)
-#114 := (iff #78 #113)
-#115 := [rewrite]: #114
-#118 := [monotonicity #115]: #117
-#123 := [trans #118 #121]: #122
-#126 := [monotonicity #123]: #125
-#135 := [trans #126 #133]: #134
-#110 := (iff #42 #109)
-#107 := (= #41 #104)
-#93 := (* -1::real #38)
-#94 := (+ #31 #93)
-#97 := (/ #94 3::real)
-#105 := (= #97 #104)
-#106 := [rewrite]: #105
-#98 := (= #41 #97)
-#95 := (= #39 #94)
-#96 := [rewrite]: #95
-#99 := [monotonicity #96]: #98
-#108 := [trans #99 #106]: #107
-#91 := (= #36 #90)
-#76 := (= #32 #75)
-#77 := [rewrite]: #76
-#88 := (= #35 #85)
-#81 := (- #75)
-#86 := (= #81 #85)
-#87 := [rewrite]: #86
-#82 := (= #35 #81)
-#83 := [monotonicity #77]: #82
-#89 := [trans #83 #87]: #88
-#79 := (iff #34 #78)
-#80 := [monotonicity #77]: #79
-#92 := [monotonicity #80 #89 #77]: #91
-#111 := [monotonicity #92 #108]: #110
-#137 := [trans #111 #135]: #136
-#72 := [asserted]: #42
-#138 := [mp #72 #137]: #131
-decl uf_1 :: T1
-#4 := uf_1
-#43 := (uf_13 uf_1)
-#44 := (uf_12 #43 uf_14)
-#45 := (uf_11 #44 uf_15)
-#149 := (* -1::real #45)
-#150 := (+ #38 #149)
-#140 := (+ #93 #45)
-#161 := (<= #150 0::real)
-#168 := (ite #161 #140 #150)
-#176 := (* -1::real #168)
-#177 := (+ #103 #176)
-#178 := (+ #101 #177)
-#179 := (<= #178 0::real)
-#180 := (not #179)
-#46 := (- #45 #38)
-#48 := (- #46)
-#47 := (< #46 0::real)
-#49 := (ite #47 #48 #46)
-#50 := (< #49 #41)
-#185 := (iff #50 #180)
-#143 := (< #140 0::real)
-#155 := (ite #143 #150 #140)
-#158 := (< #155 #104)
-#183 := (iff #158 #180)
-#173 := (< #168 #104)
-#181 := (iff #173 #180)
-#182 := [rewrite]: #181
-#174 := (iff #158 #173)
-#171 := (= #155 #168)
-#162 := (not #161)
-#165 := (ite #162 #150 #140)
-#169 := (= #165 #168)
-#170 := [rewrite]: #169
-#166 := (= #155 #165)
-#163 := (iff #143 #162)
-#164 := [rewrite]: #163
-#167 := [monotonicity #164]: #166
-#172 := [trans #167 #170]: #171
-#175 := [monotonicity #172]: #174
-#184 := [trans #175 #182]: #183
-#159 := (iff #50 #158)
-#156 := (= #49 #155)
-#141 := (= #46 #140)
-#142 := [rewrite]: #141
-#153 := (= #48 #150)
-#146 := (- #140)
-#151 := (= #146 #150)
-#152 := [rewrite]: #151
-#147 := (= #48 #146)
-#148 := [monotonicity #142]: #147
-#154 := [trans #148 #152]: #153
-#144 := (iff #47 #143)
-#145 := [monotonicity #142]: #144
-#157 := [monotonicity #145 #154 #142]: #156
-#160 := [monotonicity #157 #108]: #159
-#186 := [trans #160 #184]: #185
-#139 := [asserted]: #50
-#187 := [mp #139 #186]: #180
-#299 := (+ #140 #176)
-#300 := (<= #299 0::real)
-#290 := (= #140 #168)
-#329 := [hypothesis]: #162
-#191 := (+ #29 #149)
-#192 := (<= #191 0::real)
-#51 := (<= #29 #45)
-#193 := (iff #51 #192)
-#194 := [rewrite]: #193
-#188 := [asserted]: #51
-#195 := [mp #188 #194]: #192
-#298 := (+ #75 #127)
-#301 := (<= #298 0::real)
-#284 := (= #75 #119)
-#302 := [hypothesis]: #113
-#296 := (+ #85 #127)
-#297 := (<= #296 0::real)
-#285 := (= #85 #119)
-#288 := (or #112 #285)
-#289 := [def-axiom]: #288
-#303 := [unit-resolution #289 #302]: #285
-#304 := (not #285)
-#305 := (or #304 #297)
-#306 := [th-lemma]: #305
-#307 := [unit-resolution #306 #303]: #297
-#315 := (not #290)
-#310 := (not #300)
-#311 := (or #310 #112)
-#308 := [hypothesis]: #300
-#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false
-#312 := [lemma #309]: #311
-#322 := [unit-resolution #312 #302]: #310
-#316 := (or #315 #300)
-#313 := [hypothesis]: #310
-#314 := [hypothesis]: #290
-#317 := [th-lemma]: #316
-#318 := [unit-resolution #317 #314 #313]: false
-#319 := [lemma #318]: #316
-#323 := [unit-resolution #319 #322]: #315
-#292 := (or #162 #290)
-#293 := [def-axiom]: #292
-#324 := [unit-resolution #293 #323]: #162
-#325 := [th-lemma #324 #307 #138 #302 #195]: false
-#326 := [lemma #325]: #112
-#286 := (or #113 #284)
-#287 := [def-axiom]: #286
-#330 := [unit-resolution #287 #326]: #284
-#331 := (not #284)
-#332 := (or #331 #301)
-#333 := [th-lemma]: #332
-#334 := [unit-resolution #333 #330]: #301
-#335 := [th-lemma #326 #334 #195 #329 #138]: false
-#336 := [lemma #335]: #161
-#327 := [unit-resolution #293 #336]: #290
-#328 := [unit-resolution #319 #327]: #300
-[th-lemma #326 #334 #195 #328 #187 #138]: false
-unsat
-GX51o3DUO/UBS3eNP2P9kA 285 0
-#2 := false
-#7 := 0::real
-decl uf_4 :: real
-#16 := uf_4
-#40 := -1::real
-#116 := (* -1::real uf_4)
-decl uf_3 :: real
-#11 := uf_3
-#117 := (+ uf_3 #116)
-#128 := (<= #117 0::real)
-#129 := (not #128)
-#220 := 2/3::real
-#221 := (* 2/3::real uf_3)
-#222 := (+ #221 #116)
-decl uf_2 :: real
-#5 := uf_2
-#67 := 1/3::real
-#68 := (* 1/3::real uf_2)
-#233 := (+ #68 #222)
-#243 := (<= #233 0::real)
-#268 := (not #243)
-#287 := [hypothesis]: #268
-#41 := (* -1::real uf_2)
-decl uf_1 :: real
-#4 := uf_1
-#42 := (+ uf_1 #41)
-#79 := (>= #42 0::real)
-#80 := (not #79)
-#297 := (or #80 #243)
-#158 := (+ uf_1 #116)
-#159 := (<= #158 0::real)
-#22 := (<= uf_1 uf_4)
-#160 := (iff #22 #159)
-#161 := [rewrite]: #160
-#155 := [asserted]: #22
-#162 := [mp #155 #161]: #159
-#200 := (* 1/3::real uf_3)
-#198 := -4/3::real
-#199 := (* -4/3::real uf_2)
-#201 := (+ #199 #200)
-#202 := (+ uf_1 #201)
-#203 := (>= #202 0::real)
-#258 := (not #203)
-#292 := [hypothesis]: #79
-#293 := (or #80 #258)
-#69 := -1/3::real
-#70 := (* -1/3::real uf_3)
-#186 := -2/3::real
-#187 := (* -2/3::real uf_2)
-#188 := (+ #187 #70)
-#189 := (+ uf_1 #188)
-#204 := (<= #189 0::real)
-#205 := (ite #79 #203 #204)
-#210 := (not #205)
-#51 := (* -1::real uf_1)
-#52 := (+ #51 uf_2)
-#86 := (ite #79 #42 #52)
-#94 := (* -1::real #86)
-#95 := (+ #70 #94)
-#96 := (+ #68 #95)
-#97 := (<= #96 0::real)
-#98 := (not #97)
-#211 := (iff #98 #210)
-#208 := (iff #97 #205)
-#182 := 4/3::real
-#183 := (* 4/3::real uf_2)
-#184 := (+ #183 #70)
-#185 := (+ #51 #184)
-#190 := (ite #79 #185 #189)
-#195 := (<= #190 0::real)
-#206 := (iff #195 #205)
-#207 := [rewrite]: #206
-#196 := (iff #97 #195)
-#193 := (= #96 #190)
-#172 := (+ #41 #70)
-#173 := (+ uf_1 #172)
-#170 := (+ uf_2 #70)
-#171 := (+ #51 #170)
-#174 := (ite #79 #171 #173)
-#179 := (+ #68 #174)
-#191 := (= #179 #190)
-#192 := [rewrite]: #191
-#180 := (= #96 #179)
-#177 := (= #95 #174)
-#164 := (ite #79 #52 #42)
-#167 := (+ #70 #164)
-#175 := (= #167 #174)
-#176 := [rewrite]: #175
-#168 := (= #95 #167)
-#156 := (= #94 #164)
-#165 := [rewrite]: #156
-#169 := [monotonicity #165]: #168
-#178 := [trans #169 #176]: #177
-#181 := [monotonicity #178]: #180
-#194 := [trans #181 #192]: #193
-#197 := [monotonicity #194]: #196
-#209 := [trans #197 #207]: #208
-#212 := [monotonicity #209]: #211
-#13 := 3::real
-#12 := (- uf_2 uf_3)
-#14 := (/ #12 3::real)
-#6 := (- uf_1 uf_2)
-#9 := (- #6)
-#8 := (< #6 0::real)
-#10 := (ite #8 #9 #6)
-#15 := (< #10 #14)
-#103 := (iff #15 #98)
-#71 := (+ #68 #70)
-#45 := (< #42 0::real)
-#57 := (ite #45 #52 #42)
-#76 := (< #57 #71)
-#101 := (iff #76 #98)
-#91 := (< #86 #71)
-#99 := (iff #91 #98)
-#100 := [rewrite]: #99
-#92 := (iff #76 #91)
-#89 := (= #57 #86)
-#83 := (ite #80 #52 #42)
-#87 := (= #83 #86)
-#88 := [rewrite]: #87
-#84 := (= #57 #83)
-#81 := (iff #45 #80)
-#82 := [rewrite]: #81
-#85 := [monotonicity #82]: #84
-#90 := [trans #85 #88]: #89
-#93 := [monotonicity #90]: #92
-#102 := [trans #93 #100]: #101
-#77 := (iff #15 #76)
-#74 := (= #14 #71)
-#60 := (* -1::real uf_3)
-#61 := (+ uf_2 #60)
-#64 := (/ #61 3::real)
-#72 := (= #64 #71)
-#73 := [rewrite]: #72
-#65 := (= #14 #64)
-#62 := (= #12 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#75 := [trans #66 #73]: #74
-#58 := (= #10 #57)
-#43 := (= #6 #42)
-#44 := [rewrite]: #43
-#55 := (= #9 #52)
-#48 := (- #42)
-#53 := (= #48 #52)
-#54 := [rewrite]: #53
-#49 := (= #9 #48)
-#50 := [monotonicity #44]: #49
-#56 := [trans #50 #54]: #55
-#46 := (iff #8 #45)
-#47 := [monotonicity #44]: #46
-#59 := [monotonicity #47 #56 #44]: #58
-#78 := [monotonicity #59 #75]: #77
-#104 := [trans #78 #102]: #103
-#39 := [asserted]: #15
-#105 := [mp #39 #104]: #98
-#213 := [mp #105 #212]: #210
-#259 := (or #205 #80 #258)
-#260 := [def-axiom]: #259
-#294 := [unit-resolution #260 #213]: #293
-#295 := [unit-resolution #294 #292]: #258
-#296 := [th-lemma #287 #292 #295 #162]: false
-#298 := [lemma #296]: #297
-#299 := [unit-resolution #298 #287]: #80
-#261 := (not #204)
-#281 := (or #79 #261)
-#262 := (or #205 #79 #261)
-#263 := [def-axiom]: #262
-#282 := [unit-resolution #263 #213]: #281
-#300 := [unit-resolution #282 #299]: #261
-#290 := (or #79 #204 #243)
-#276 := [hypothesis]: #261
-#288 := [hypothesis]: #80
-#289 := [th-lemma #288 #276 #162 #287]: false
-#291 := [lemma #289]: #290
-#301 := [unit-resolution #291 #300 #299 #287]: false
-#302 := [lemma #301]: #243
-#303 := (or #129 #268)
-#223 := (* -4/3::real uf_3)
-#224 := (+ #223 uf_4)
-#234 := (+ #68 #224)
-#244 := (<= #234 0::real)
-#245 := (ite #128 #243 #244)
-#250 := (not #245)
-#107 := (+ #60 uf_4)
-#135 := (ite #128 #107 #117)
-#143 := (* -1::real #135)
-#144 := (+ #70 #143)
-#145 := (+ #68 #144)
-#146 := (<= #145 0::real)
-#147 := (not #146)
-#251 := (iff #147 #250)
-#248 := (iff #146 #245)
-#235 := (ite #128 #233 #234)
-#240 := (<= #235 0::real)
-#246 := (iff #240 #245)
-#247 := [rewrite]: #246
-#241 := (iff #146 #240)
-#238 := (= #145 #235)
-#225 := (ite #128 #222 #224)
-#230 := (+ #68 #225)
-#236 := (= #230 #235)
-#237 := [rewrite]: #236
-#231 := (= #145 #230)
-#228 := (= #144 #225)
-#214 := (ite #128 #117 #107)
-#217 := (+ #70 #214)
-#226 := (= #217 #225)
-#227 := [rewrite]: #226
-#218 := (= #144 #217)
-#215 := (= #143 #214)
-#216 := [rewrite]: #215
-#219 := [monotonicity #216]: #218
-#229 := [trans #219 #227]: #228
-#232 := [monotonicity #229]: #231
-#239 := [trans #232 #237]: #238
-#242 := [monotonicity #239]: #241
-#249 := [trans #242 #247]: #248
-#252 := [monotonicity #249]: #251
-#17 := (- uf_4 uf_3)
-#19 := (- #17)
-#18 := (< #17 0::real)
-#20 := (ite #18 #19 #17)
-#21 := (< #20 #14)
-#152 := (iff #21 #147)
-#110 := (< #107 0::real)
-#122 := (ite #110 #117 #107)
-#125 := (< #122 #71)
-#150 := (iff #125 #147)
-#140 := (< #135 #71)
-#148 := (iff #140 #147)
-#149 := [rewrite]: #148
-#141 := (iff #125 #140)
-#138 := (= #122 #135)
-#132 := (ite #129 #117 #107)
-#136 := (= #132 #135)
-#137 := [rewrite]: #136
-#133 := (= #122 #132)
-#130 := (iff #110 #129)
-#131 := [rewrite]: #130
-#134 := [monotonicity #131]: #133
-#139 := [trans #134 #137]: #138
-#142 := [monotonicity #139]: #141
-#151 := [trans #142 #149]: #150
-#126 := (iff #21 #125)
-#123 := (= #20 #122)
-#108 := (= #17 #107)
-#109 := [rewrite]: #108
-#120 := (= #19 #117)
-#113 := (- #107)
-#118 := (= #113 #117)
-#119 := [rewrite]: #118
-#114 := (= #19 #113)
-#115 := [monotonicity #109]: #114
-#121 := [trans #115 #119]: #120
-#111 := (iff #18 #110)
-#112 := [monotonicity #109]: #111
-#124 := [monotonicity #112 #121 #109]: #123
-#127 := [monotonicity #124 #75]: #126
-#153 := [trans #127 #151]: #152
-#106 := [asserted]: #21
-#154 := [mp #106 #153]: #147
-#253 := [mp #154 #252]: #250
-#269 := (or #245 #129 #268)
-#270 := [def-axiom]: #269
-#304 := [unit-resolution #270 #253]: #303
-#305 := [unit-resolution #304 #302]: #129
-#271 := (not #244)
-#306 := (or #128 #271)
-#272 := (or #245 #128 #271)
-#273 := [def-axiom]: #272
-#307 := [unit-resolution #273 #253]: #306
-#308 := [unit-resolution #307 #305]: #271
-#285 := (or #128 #244)
-#274 := [hypothesis]: #271
-#275 := [hypothesis]: #129
-#278 := (or #204 #128 #244)
-#277 := [th-lemma #276 #275 #274 #162]: false
-#279 := [lemma #277]: #278
-#280 := [unit-resolution #279 #275 #274]: #204
-#283 := [unit-resolution #282 #280]: #79
-#284 := [th-lemma #275 #274 #283 #162]: false
-#286 := [lemma #284]: #285
-[unit-resolution #286 #308 #305]: false
-unsat
-cebG074uorSr8ODzgTmcKg 97 0
-#2 := false
-#18 := 0::real
-decl uf_1 :: (-> T2 T1 real)
-decl uf_5 :: T1
-#11 := uf_5
-decl uf_2 :: T2
-#4 := uf_2
-#20 := (uf_1 uf_2 uf_5)
-#42 := -1::real
-#53 := (* -1::real #20)
-decl uf_3 :: T2
-#7 := uf_3
-#19 := (uf_1 uf_3 uf_5)
-#54 := (+ #19 #53)
-#63 := (<= #54 0::real)
-#21 := (- #19 #20)
-#22 := (< 0::real #21)
-#23 := (not #22)
-#74 := (iff #23 #63)
-#57 := (< 0::real #54)
-#60 := (not #57)
-#72 := (iff #60 #63)
-#64 := (not #63)
-#67 := (not #64)
-#70 := (iff #67 #63)
-#71 := [rewrite]: #70
-#68 := (iff #60 #67)
-#65 := (iff #57 #64)
-#66 := [rewrite]: #65
-#69 := [monotonicity #66]: #68
-#73 := [trans #69 #71]: #72
-#61 := (iff #23 #60)
-#58 := (iff #22 #57)
-#55 := (= #21 #54)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#62 := [monotonicity #59]: #61
-#75 := [trans #62 #73]: #74
-#41 := [asserted]: #23
-#76 := [mp #41 #75]: #63
-#5 := (:var 0 T1)
-#8 := (uf_1 uf_3 #5)
-#141 := (pattern #8)
-#6 := (uf_1 uf_2 #5)
-#140 := (pattern #6)
-#45 := (* -1::real #8)
-#46 := (+ #6 #45)
-#44 := (>= #46 0::real)
-#43 := (not #44)
-#142 := (forall (vars (?x1 T1)) (:pat #140 #141) #43)
-#49 := (forall (vars (?x1 T1)) #43)
-#145 := (iff #49 #142)
-#143 := (iff #43 #43)
-#144 := [refl]: #143
-#146 := [quant-intro #144]: #145
-#80 := (~ #49 #49)
-#82 := (~ #43 #43)
-#83 := [refl]: #82
-#81 := [nnf-pos #83]: #80
-#9 := (< #6 #8)
-#10 := (forall (vars (?x1 T1)) #9)
-#50 := (iff #10 #49)
-#47 := (iff #9 #43)
-#48 := [rewrite]: #47
-#51 := [quant-intro #48]: #50
-#39 := [asserted]: #10
-#52 := [mp #39 #51]: #49
-#79 := [mp~ #52 #81]: #49
-#147 := [mp #79 #146]: #142
-#164 := (not #142)
-#165 := (or #164 #64)
-#148 := (* -1::real #19)
-#149 := (+ #20 #148)
-#150 := (>= #149 0::real)
-#151 := (not #150)
-#166 := (or #164 #151)
-#168 := (iff #166 #165)
-#170 := (iff #165 #165)
-#171 := [rewrite]: #170
-#162 := (iff #151 #64)
-#160 := (iff #150 #63)
-#152 := (+ #148 #20)
-#155 := (>= #152 0::real)
-#158 := (iff #155 #63)
-#159 := [rewrite]: #158
-#156 := (iff #150 #155)
-#153 := (= #149 #152)
-#154 := [rewrite]: #153
-#157 := [monotonicity #154]: #156
-#161 := [trans #157 #159]: #160
-#163 := [monotonicity #161]: #162
-#169 := [monotonicity #163]: #168
-#172 := [trans #169 #171]: #168
-#167 := [quant-inst]: #166
-#173 := [mp #167 #172]: #165
-[unit-resolution #173 #147 #76]: false
-unsat
-DKRtrJ2XceCkITuNwNViRw 57 0
-#2 := false
-#4 := 0::real
-decl uf_1 :: (-> T2 real)
-decl uf_2 :: (-> T1 T1 T2)
-decl uf_12 :: (-> T4 T1)
-decl uf_4 :: T4
-#11 := uf_4
-#39 := (uf_12 uf_4)
-decl uf_10 :: T4
-#27 := uf_10
-#38 := (uf_12 uf_10)
-#40 := (uf_2 #38 #39)
-#41 := (uf_1 #40)
-#264 := (>= #41 0::real)
-#266 := (not #264)
-#43 := (= #41 0::real)
-#44 := (not #43)
-#131 := [asserted]: #44
-#272 := (or #43 #266)
-#42 := (<= #41 0::real)
-#130 := [asserted]: #42
-#265 := (not #42)
-#270 := (or #43 #265 #266)
-#271 := [th-lemma]: #270
-#273 := [unit-resolution #271 #130]: #272
-#274 := [unit-resolution #273 #131]: #266
-#6 := (:var 0 T1)
-#5 := (:var 1 T1)
-#7 := (uf_2 #5 #6)
-#241 := (pattern #7)
-#8 := (uf_1 #7)
-#65 := (>= #8 0::real)
-#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65)
-#66 := (forall (vars (?x1 T1) (?x2 T1)) #65)
-#245 := (iff #66 #242)
-#243 := (iff #65 #65)
-#244 := [refl]: #243
-#246 := [quant-intro #244]: #245
-#149 := (~ #66 #66)
-#151 := (~ #65 #65)
-#152 := [refl]: #151
-#150 := [nnf-pos #152]: #149
-#9 := (<= 0::real #8)
-#10 := (forall (vars (?x1 T1) (?x2 T1)) #9)
-#67 := (iff #10 #66)
-#63 := (iff #9 #65)
-#64 := [rewrite]: #63
-#68 := [quant-intro #64]: #67
-#60 := [asserted]: #10
-#69 := [mp #60 #68]: #66
-#147 := [mp~ #69 #150]: #66
-#247 := [mp #147 #246]: #242
-#267 := (not #242)
-#268 := (or #267 #264)
-#269 := [quant-inst]: #268
-[unit-resolution #269 #247 #274]: false
-unsat
-97KJAJfUio+nGchEHWvgAw 91 0
-#2 := false
-#38 := 0::real
-decl uf_1 :: (-> T1 T2 real)
-decl uf_3 :: T2
-#5 := uf_3
-decl uf_4 :: T1
-#7 := uf_4
-#8 := (uf_1 uf_4 uf_3)
-#35 := -1::real
-#36 := (* -1::real #8)
-decl uf_2 :: T1
-#4 := uf_2
-#6 := (uf_1 uf_2 uf_3)
-#37 := (+ #6 #36)
-#130 := (>= #37 0::real)
-#155 := (not #130)
-#43 := (= #6 #8)
-#55 := (not #43)
-#15 := (= #8 #6)
-#16 := (not #15)
-#56 := (iff #16 #55)
-#53 := (iff #15 #43)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#34 := [asserted]: #16
-#60 := [mp #34 #57]: #55
-#158 := (or #43 #155)
-#39 := (<= #37 0::real)
-#9 := (<= #6 #8)
-#40 := (iff #9 #39)
-#41 := [rewrite]: #40
-#32 := [asserted]: #9
-#42 := [mp #32 #41]: #39
-#154 := (not #39)
-#156 := (or #43 #154 #155)
-#157 := [th-lemma]: #156
-#159 := [unit-resolution #157 #42]: #158
-#160 := [unit-resolution #159 #60]: #155
-#10 := (:var 0 T2)
-#12 := (uf_1 uf_2 #10)
-#123 := (pattern #12)
-#11 := (uf_1 uf_4 #10)
-#122 := (pattern #11)
-#44 := (* -1::real #12)
-#45 := (+ #11 #44)
-#46 := (<= #45 0::real)
-#124 := (forall (vars (?x1 T2)) (:pat #122 #123) #46)
-#49 := (forall (vars (?x1 T2)) #46)
-#127 := (iff #49 #124)
-#125 := (iff #46 #46)
-#126 := [refl]: #125
-#128 := [quant-intro #126]: #127
-#62 := (~ #49 #49)
-#64 := (~ #46 #46)
-#65 := [refl]: #64
-#63 := [nnf-pos #65]: #62
-#13 := (<= #11 #12)
-#14 := (forall (vars (?x1 T2)) #13)
-#50 := (iff #14 #49)
-#47 := (iff #13 #46)
-#48 := [rewrite]: #47
-#51 := [quant-intro #48]: #50
-#33 := [asserted]: #14
-#52 := [mp #33 #51]: #49
-#61 := [mp~ #52 #63]: #49
-#129 := [mp #61 #128]: #124
-#144 := (not #124)
-#145 := (or #144 #130)
-#131 := (* -1::real #6)
-#132 := (+ #8 #131)
-#133 := (<= #132 0::real)
-#146 := (or #144 #133)
-#148 := (iff #146 #145)
-#150 := (iff #145 #145)
-#151 := [rewrite]: #150
-#142 := (iff #133 #130)
-#134 := (+ #131 #8)
-#137 := (<= #134 0::real)
-#140 := (iff #137 #130)
-#141 := [rewrite]: #140
-#138 := (iff #133 #137)
-#135 := (= #132 #134)
-#136 := [rewrite]: #135
-#139 := [monotonicity #136]: #138
-#143 := [trans #139 #141]: #142
-#149 := [monotonicity #143]: #148
-#152 := [trans #149 #151]: #148
-#147 := [quant-inst]: #146
-#153 := [mp #147 #152]: #145
-[unit-resolution #153 #129 #160]: false
-unsat
-flJYbeWfe+t2l/zsRqdujA 149 0
-#2 := false
-#19 := 0::real
-decl uf_1 :: (-> T1 T2 real)
-decl uf_3 :: T2
-#5 := uf_3
-decl uf_4 :: T1
-#7 := uf_4
-#8 := (uf_1 uf_4 uf_3)
-#44 := -1::real
-#156 := (* -1::real #8)
-decl uf_2 :: T1
-#4 := uf_2
-#6 := (uf_1 uf_2 uf_3)
-#203 := (+ #6 #156)
-#205 := (>= #203 0::real)
-#9 := (= #6 #8)
-#40 := [asserted]: #9
-#208 := (not #9)
-#209 := (or #208 #205)
-#210 := [th-lemma]: #209
-#211 := [unit-resolution #210 #40]: #205
-decl uf_5 :: T1
-#12 := uf_5
-#22 := (uf_1 uf_5 uf_3)
-#160 := (* -1::real #22)
-#161 := (+ #6 #160)
-#207 := (>= #161 0::real)
-#222 := (not #207)
-#206 := (= #6 #22)
-#216 := (not #206)
-#62 := (= #8 #22)
-#70 := (not #62)
-#217 := (iff #70 #216)
-#214 := (iff #62 #206)
-#212 := (iff #206 #62)
-#213 := [monotonicity #40]: #212
-#215 := [symm #213]: #214
-#218 := [monotonicity #215]: #217
-#23 := (= #22 #8)
-#24 := (not #23)
-#71 := (iff #24 #70)
-#68 := (iff #23 #62)
-#69 := [rewrite]: #68
-#72 := [monotonicity #69]: #71
-#43 := [asserted]: #24
-#75 := [mp #43 #72]: #70
-#219 := [mp #75 #218]: #216
-#225 := (or #206 #222)
-#162 := (<= #161 0::real)
-#172 := (+ #8 #160)
-#173 := (>= #172 0::real)
-#178 := (not #173)
-#163 := (not #162)
-#181 := (or #163 #178)
-#184 := (not #181)
-#10 := (:var 0 T2)
-#15 := (uf_1 uf_4 #10)
-#149 := (pattern #15)
-#13 := (uf_1 uf_5 #10)
-#148 := (pattern #13)
-#11 := (uf_1 uf_2 #10)
-#147 := (pattern #11)
-#50 := (* -1::real #15)
-#51 := (+ #13 #50)
-#52 := (<= #51 0::real)
-#76 := (not #52)
-#45 := (* -1::real #13)
-#46 := (+ #11 #45)
-#47 := (<= #46 0::real)
-#78 := (not #47)
-#73 := (or #78 #76)
-#83 := (not #73)
-#150 := (forall (vars (?x1 T2)) (:pat #147 #148 #149) #83)
-#86 := (forall (vars (?x1 T2)) #83)
-#153 := (iff #86 #150)
-#151 := (iff #83 #83)
-#152 := [refl]: #151
-#154 := [quant-intro #152]: #153
-#55 := (and #47 #52)
-#58 := (forall (vars (?x1 T2)) #55)
-#87 := (iff #58 #86)
-#84 := (iff #55 #83)
-#85 := [rewrite]: #84
-#88 := [quant-intro #85]: #87
-#79 := (~ #58 #58)
-#81 := (~ #55 #55)
-#82 := [refl]: #81
-#80 := [nnf-pos #82]: #79
-#16 := (<= #13 #15)
-#14 := (<= #11 #13)
-#17 := (and #14 #16)
-#18 := (forall (vars (?x1 T2)) #17)
-#59 := (iff #18 #58)
-#56 := (iff #17 #55)
-#53 := (iff #16 #52)
-#54 := [rewrite]: #53
-#48 := (iff #14 #47)
-#49 := [rewrite]: #48
-#57 := [monotonicity #49 #54]: #56
-#60 := [quant-intro #57]: #59
-#41 := [asserted]: #18
-#61 := [mp #41 #60]: #58
-#77 := [mp~ #61 #80]: #58
-#89 := [mp #77 #88]: #86
-#155 := [mp #89 #154]: #150
-#187 := (not #150)
-#188 := (or #187 #184)
-#157 := (+ #22 #156)
-#158 := (<= #157 0::real)
-#159 := (not #158)
-#164 := (or #163 #159)
-#165 := (not #164)
-#189 := (or #187 #165)
-#191 := (iff #189 #188)
-#193 := (iff #188 #188)
-#194 := [rewrite]: #193
-#185 := (iff #165 #184)
-#182 := (iff #164 #181)
-#179 := (iff #159 #178)
-#176 := (iff #158 #173)
-#166 := (+ #156 #22)
-#169 := (<= #166 0::real)
-#174 := (iff #169 #173)
-#175 := [rewrite]: #174
-#170 := (iff #158 #169)
-#167 := (= #157 #166)
-#168 := [rewrite]: #167
-#171 := [monotonicity #168]: #170
-#177 := [trans #171 #175]: #176
-#180 := [monotonicity #177]: #179
-#183 := [monotonicity #180]: #182
-#186 := [monotonicity #183]: #185
-#192 := [monotonicity #186]: #191
-#195 := [trans #192 #194]: #191
-#190 := [quant-inst]: #189
-#196 := [mp #190 #195]: #188
-#220 := [unit-resolution #196 #155]: #184
-#197 := (or #181 #162)
-#198 := [def-axiom]: #197
-#221 := [unit-resolution #198 #220]: #162
-#223 := (or #206 #163 #222)
-#224 := [th-lemma]: #223
-#226 := [unit-resolution #224 #221]: #225
-#227 := [unit-resolution #226 #219]: #222
-#199 := (or #181 #173)
-#200 := [def-axiom]: #199
-#228 := [unit-resolution #200 #220]: #173
-[th-lemma #228 #227 #211]: false
-unsat
-rbrrQuQfaijtLkQizgEXnQ 222 0
-#2 := false
-#4 := 0::real
-decl uf_2 :: (-> T2 T1 real)
-decl uf_5 :: T1
-#15 := uf_5
-decl uf_3 :: T2
-#7 := uf_3
-#20 := (uf_2 uf_3 uf_5)
-decl uf_6 :: T2
-#17 := uf_6
-#18 := (uf_2 uf_6 uf_5)
-#59 := -1::real
-#73 := (* -1::real #18)
-#106 := (+ #73 #20)
-decl uf_1 :: real
-#5 := uf_1
-#78 := (* -1::real #20)
-#79 := (+ #18 #78)
-#144 := (+ uf_1 #79)
-#145 := (<= #144 0::real)
-#148 := (ite #145 uf_1 #106)
-#279 := (* -1::real #148)
-#280 := (+ uf_1 #279)
-#281 := (<= #280 0::real)
-#289 := (not #281)
-#72 := 1/2::real
-#151 := (* 1/2::real #148)
-#248 := (<= #151 0::real)
-#162 := (= #151 0::real)
-#24 := 2::real
-#27 := (- #20 #18)
-#28 := (<= uf_1 #27)
-#29 := (ite #28 uf_1 #27)
-#30 := (/ #29 2::real)
-#31 := (+ #18 #30)
-#32 := (= #31 #18)
-#33 := (not #32)
-#34 := (not #33)
-#165 := (iff #34 #162)
-#109 := (<= uf_1 #106)
-#112 := (ite #109 uf_1 #106)
-#118 := (* 1/2::real #112)
-#123 := (+ #18 #118)
-#129 := (= #18 #123)
-#163 := (iff #129 #162)
-#154 := (+ #18 #151)
-#157 := (= #18 #154)
-#160 := (iff #157 #162)
-#161 := [rewrite]: #160
-#158 := (iff #129 #157)
-#155 := (= #123 #154)
-#152 := (= #118 #151)
-#149 := (= #112 #148)
-#146 := (iff #109 #145)
-#147 := [rewrite]: #146
-#150 := [monotonicity #147]: #149
-#153 := [monotonicity #150]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#164 := [trans #159 #161]: #163
-#142 := (iff #34 #129)
-#134 := (not #129)
-#137 := (not #134)
-#140 := (iff #137 #129)
-#141 := [rewrite]: #140
-#138 := (iff #34 #137)
-#135 := (iff #33 #134)
-#132 := (iff #32 #129)
-#126 := (= #123 #18)
-#130 := (iff #126 #129)
-#131 := [rewrite]: #130
-#127 := (iff #32 #126)
-#124 := (= #31 #123)
-#121 := (= #30 #118)
-#115 := (/ #112 2::real)
-#119 := (= #115 #118)
-#120 := [rewrite]: #119
-#116 := (= #30 #115)
-#113 := (= #29 #112)
-#107 := (= #27 #106)
-#108 := [rewrite]: #107
-#110 := (iff #28 #109)
-#111 := [monotonicity #108]: #110
-#114 := [monotonicity #111 #108]: #113
-#117 := [monotonicity #114]: #116
-#122 := [trans #117 #120]: #121
-#125 := [monotonicity #122]: #124
-#128 := [monotonicity #125]: #127
-#133 := [trans #128 #131]: #132
-#136 := [monotonicity #133]: #135
-#139 := [monotonicity #136]: #138
-#143 := [trans #139 #141]: #142
-#166 := [trans #143 #164]: #165
-#105 := [asserted]: #34
-#167 := [mp #105 #166]: #162
-#283 := (not #162)
-#284 := (or #283 #248)
-#285 := [th-lemma]: #284
-#286 := [unit-resolution #285 #167]: #248
-#287 := [hypothesis]: #281
-#53 := (<= uf_1 0::real)
-#54 := (not #53)
-#6 := (< 0::real uf_1)
-#55 := (iff #6 #54)
-#56 := [rewrite]: #55
-#50 := [asserted]: #6
-#57 := [mp #50 #56]: #54
-#288 := [th-lemma #57 #287 #286]: false
-#290 := [lemma #288]: #289
-#241 := (= uf_1 #148)
-#242 := (= #106 #148)
-#299 := (not #242)
-#282 := (+ #106 #279)
-#291 := (<= #282 0::real)
-#296 := (not #291)
-decl uf_4 :: T2
-#10 := uf_4
-#16 := (uf_2 uf_4 uf_5)
-#260 := (+ #16 #78)
-#261 := (>= #260 0::real)
-#266 := (not #261)
-#8 := (:var 0 T1)
-#11 := (uf_2 uf_4 #8)
-#234 := (pattern #11)
-#9 := (uf_2 uf_3 #8)
-#233 := (pattern #9)
-#60 := (* -1::real #11)
-#61 := (+ #9 #60)
-#62 := (<= #61 0::real)
-#179 := (not #62)
-#235 := (forall (vars (?x1 T1)) (:pat #233 #234) #179)
-#178 := (forall (vars (?x1 T1)) #179)
-#238 := (iff #178 #235)
-#236 := (iff #179 #179)
-#237 := [refl]: #236
-#239 := [quant-intro #237]: #238
-#65 := (exists (vars (?x1 T1)) #62)
-#68 := (not #65)
-#175 := (~ #68 #178)
-#180 := (~ #179 #179)
-#177 := [refl]: #180
-#176 := [nnf-neg #177]: #175
-#12 := (<= #9 #11)
-#13 := (exists (vars (?x1 T1)) #12)
-#14 := (not #13)
-#69 := (iff #14 #68)
-#66 := (iff #13 #65)
-#63 := (iff #12 #62)
-#64 := [rewrite]: #63
-#67 := [quant-intro #64]: #66
-#70 := [monotonicity #67]: #69
-#51 := [asserted]: #14
-#71 := [mp #51 #70]: #68
-#173 := [mp~ #71 #176]: #178
-#240 := [mp #173 #239]: #235
-#269 := (not #235)
-#270 := (or #269 #266)
-#250 := (* -1::real #16)
-#251 := (+ #20 #250)
-#252 := (<= #251 0::real)
-#253 := (not #252)
-#271 := (or #269 #253)
-#273 := (iff #271 #270)
-#275 := (iff #270 #270)
-#276 := [rewrite]: #275
-#267 := (iff #253 #266)
-#264 := (iff #252 #261)
-#254 := (+ #250 #20)
-#257 := (<= #254 0::real)
-#262 := (iff #257 #261)
-#263 := [rewrite]: #262
-#258 := (iff #252 #257)
-#255 := (= #251 #254)
-#256 := [rewrite]: #255
-#259 := [monotonicity #256]: #258
-#265 := [trans #259 #263]: #264
-#268 := [monotonicity #265]: #267
-#274 := [monotonicity #268]: #273
-#277 := [trans #274 #276]: #273
-#272 := [quant-inst]: #271
-#278 := [mp #272 #277]: #270
-#293 := [unit-resolution #278 #240]: #266
-#90 := (* 1/2::real #20)
-#102 := (+ #73 #90)
-#89 := (* 1/2::real #16)
-#103 := (+ #89 #102)
-#100 := (>= #103 0::real)
-#23 := (+ #16 #20)
-#25 := (/ #23 2::real)
-#26 := (<= #18 #25)
-#98 := (iff #26 #100)
-#91 := (+ #89 #90)
-#94 := (<= #18 #91)
-#97 := (iff #94 #100)
-#99 := [rewrite]: #97
-#95 := (iff #26 #94)
-#92 := (= #25 #91)
-#93 := [rewrite]: #92
-#96 := [monotonicity #93]: #95
-#101 := [trans #96 #99]: #98
-#58 := [asserted]: #26
-#104 := [mp #58 #101]: #100
-#294 := [hypothesis]: #291
-#295 := [th-lemma #294 #104 #293 #286]: false
-#297 := [lemma #295]: #296
-#298 := [hypothesis]: #242
-#300 := (or #299 #291)
-#301 := [th-lemma]: #300
-#302 := [unit-resolution #301 #298 #297]: false
-#303 := [lemma #302]: #299
-#246 := (or #145 #242)
-#247 := [def-axiom]: #246
-#304 := [unit-resolution #247 #303]: #145
-#243 := (not #145)
-#244 := (or #243 #241)
-#245 := [def-axiom]: #244
-#305 := [unit-resolution #245 #304]: #241
-#306 := (not #241)
-#307 := (or #306 #281)
-#308 := [th-lemma]: #307
-[unit-resolution #308 #305 #290]: false
-unsat
-hwh3oeLAWt56hnKIa8Wuow 248 0
-#2 := false
-#4 := 0::real
-decl uf_2 :: (-> T2 T1 real)
-decl uf_5 :: T1
-#15 := uf_5
-decl uf_6 :: T2
-#17 := uf_6
-#18 := (uf_2 uf_6 uf_5)
-decl uf_4 :: T2
-#10 := uf_4
-#16 := (uf_2 uf_4 uf_5)
-#66 := -1::real
-#137 := (* -1::real #16)
-#138 := (+ #137 #18)
-decl uf_1 :: real
-#5 := uf_1
-#80 := (* -1::real #18)
-#81 := (+ #16 #80)
-#201 := (+ uf_1 #81)
-#202 := (<= #201 0::real)
-#205 := (ite #202 uf_1 #138)
-#352 := (* -1::real #205)
-#353 := (+ uf_1 #352)
-#354 := (<= #353 0::real)
-#362 := (not #354)
-#79 := 1/2::real
-#244 := (* 1/2::real #205)
-#322 := (<= #244 0::real)
-#245 := (= #244 0::real)
-#158 := -1/2::real
-#208 := (* -1/2::real #205)
-#211 := (+ #18 #208)
-decl uf_3 :: T2
-#7 := uf_3
-#20 := (uf_2 uf_3 uf_5)
-#117 := (+ #80 #20)
-#85 := (* -1::real #20)
-#86 := (+ #18 #85)
-#188 := (+ uf_1 #86)
-#189 := (<= #188 0::real)
-#192 := (ite #189 uf_1 #117)
-#195 := (* 1/2::real #192)
-#198 := (+ #18 #195)
-#97 := (* 1/2::real #20)
-#109 := (+ #80 #97)
-#96 := (* 1/2::real #16)
-#110 := (+ #96 #109)
-#107 := (>= #110 0::real)
-#214 := (ite #107 #198 #211)
-#217 := (= #18 #214)
-#248 := (iff #217 #245)
-#241 := (= #18 #211)
-#246 := (iff #241 #245)
-#247 := [rewrite]: #246
-#242 := (iff #217 #241)
-#239 := (= #214 #211)
-#234 := (ite false #198 #211)
-#237 := (= #234 #211)
-#238 := [rewrite]: #237
-#235 := (= #214 #234)
-#232 := (iff #107 false)
-#104 := (not #107)
-#24 := 2::real
-#23 := (+ #16 #20)
-#25 := (/ #23 2::real)
-#26 := (< #25 #18)
-#108 := (iff #26 #104)
-#98 := (+ #96 #97)
-#101 := (< #98 #18)
-#106 := (iff #101 #104)
-#105 := [rewrite]: #106
-#102 := (iff #26 #101)
-#99 := (= #25 #98)
-#100 := [rewrite]: #99
-#103 := [monotonicity #100]: #102
-#111 := [trans #103 #105]: #108
-#65 := [asserted]: #26
-#112 := [mp #65 #111]: #104
-#233 := [iff-false #112]: #232
-#236 := [monotonicity #233]: #235
-#240 := [trans #236 #238]: #239
-#243 := [monotonicity #240]: #242
-#249 := [trans #243 #247]: #248
-#33 := (- #18 #16)
-#34 := (<= uf_1 #33)
-#35 := (ite #34 uf_1 #33)
-#36 := (/ #35 2::real)
-#37 := (- #18 #36)
-#28 := (- #20 #18)
-#29 := (<= uf_1 #28)
-#30 := (ite #29 uf_1 #28)
-#31 := (/ #30 2::real)
-#32 := (+ #18 #31)
-#27 := (<= #18 #25)
-#38 := (ite #27 #32 #37)
-#39 := (= #38 #18)
-#40 := (not #39)
-#41 := (not #40)
-#220 := (iff #41 #217)
-#141 := (<= uf_1 #138)
-#144 := (ite #141 uf_1 #138)
-#159 := (* -1/2::real #144)
-#160 := (+ #18 #159)
-#120 := (<= uf_1 #117)
-#123 := (ite #120 uf_1 #117)
-#129 := (* 1/2::real #123)
-#134 := (+ #18 #129)
-#114 := (<= #18 #98)
-#165 := (ite #114 #134 #160)
-#171 := (= #18 #165)
-#218 := (iff #171 #217)
-#215 := (= #165 #214)
-#212 := (= #160 #211)
-#209 := (= #159 #208)
-#206 := (= #144 #205)
-#203 := (iff #141 #202)
-#204 := [rewrite]: #203
-#207 := [monotonicity #204]: #206
-#210 := [monotonicity #207]: #209
-#213 := [monotonicity #210]: #212
-#199 := (= #134 #198)
-#196 := (= #129 #195)
-#193 := (= #123 #192)
-#190 := (iff #120 #189)
-#191 := [rewrite]: #190
-#194 := [monotonicity #191]: #193
-#197 := [monotonicity #194]: #196
-#200 := [monotonicity #197]: #199
-#187 := (iff #114 #107)
-#186 := [rewrite]: #187
-#216 := [monotonicity #186 #200 #213]: #215
-#219 := [monotonicity #216]: #218
-#184 := (iff #41 #171)
-#176 := (not #171)
-#179 := (not #176)
-#182 := (iff #179 #171)
-#183 := [rewrite]: #182
-#180 := (iff #41 #179)
-#177 := (iff #40 #176)
-#174 := (iff #39 #171)
-#168 := (= #165 #18)
-#172 := (iff #168 #171)
-#173 := [rewrite]: #172
-#169 := (iff #39 #168)
-#166 := (= #38 #165)
-#163 := (= #37 #160)
-#150 := (* 1/2::real #144)
-#155 := (- #18 #150)
-#161 := (= #155 #160)
-#162 := [rewrite]: #161
-#156 := (= #37 #155)
-#153 := (= #36 #150)
-#147 := (/ #144 2::real)
-#151 := (= #147 #150)
-#152 := [rewrite]: #151
-#148 := (= #36 #147)
-#145 := (= #35 #144)
-#139 := (= #33 #138)
-#140 := [rewrite]: #139
-#142 := (iff #34 #141)
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143 #140]: #145
-#149 := [monotonicity #146]: #148
-#154 := [trans #149 #152]: #153
-#157 := [monotonicity #154]: #156
-#164 := [trans #157 #162]: #163
-#135 := (= #32 #134)
-#132 := (= #31 #129)
-#126 := (/ #123 2::real)
-#130 := (= #126 #129)
-#131 := [rewrite]: #130
-#127 := (= #31 #126)
-#124 := (= #30 #123)
-#118 := (= #28 #117)
-#119 := [rewrite]: #118
-#121 := (iff #29 #120)
-#122 := [monotonicity #119]: #121
-#125 := [monotonicity #122 #119]: #124
-#128 := [monotonicity #125]: #127
-#133 := [trans #128 #131]: #132
-#136 := [monotonicity #133]: #135
-#115 := (iff #27 #114)
-#116 := [monotonicity #100]: #115
-#167 := [monotonicity #116 #136 #164]: #166
-#170 := [monotonicity #167]: #169
-#175 := [trans #170 #173]: #174
-#178 := [monotonicity #175]: #177
-#181 := [monotonicity #178]: #180
-#185 := [trans #181 #183]: #184
-#221 := [trans #185 #219]: #220
-#113 := [asserted]: #41
-#222 := [mp #113 #221]: #217
-#250 := [mp #222 #249]: #245
-#356 := (not #245)
-#357 := (or #356 #322)
-#358 := [th-lemma]: #357
-#359 := [unit-resolution #358 #250]: #322
-#360 := [hypothesis]: #354
-#60 := (<= uf_1 0::real)
-#61 := (not #60)
-#6 := (< 0::real uf_1)
-#62 := (iff #6 #61)
-#63 := [rewrite]: #62
-#57 := [asserted]: #6
-#64 := [mp #57 #63]: #61
-#361 := [th-lemma #64 #360 #359]: false
-#363 := [lemma #361]: #362
-#315 := (= uf_1 #205)
-#316 := (= #138 #205)
-#371 := (not #316)
-#355 := (+ #138 #352)
-#364 := (<= #355 0::real)
-#368 := (not #364)
-#87 := (<= #86 0::real)
-#82 := (<= #81 0::real)
-#90 := (and #82 #87)
-#21 := (<= #18 #20)
-#19 := (<= #16 #18)
-#22 := (and #19 #21)
-#91 := (iff #22 #90)
-#88 := (iff #21 #87)
-#89 := [rewrite]: #88
-#83 := (iff #19 #82)
-#84 := [rewrite]: #83
-#92 := [monotonicity #84 #89]: #91
-#59 := [asserted]: #22
-#93 := [mp #59 #92]: #90
-#95 := [and-elim #93]: #87
-#366 := [hypothesis]: #364
-#367 := [th-lemma #366 #95 #112 #359]: false
-#369 := [lemma #367]: #368
-#370 := [hypothesis]: #316
-#372 := (or #371 #364)
-#373 := [th-lemma]: #372
-#374 := [unit-resolution #373 #370 #369]: false
-#375 := [lemma #374]: #371
-#320 := (or #202 #316)
-#321 := [def-axiom]: #320
-#376 := [unit-resolution #321 #375]: #202
-#317 := (not #202)
-#318 := (or #317 #315)
-#319 := [def-axiom]: #318
-#377 := [unit-resolution #319 #376]: #315
-#378 := (not #315)
-#379 := (or #378 #354)
-#380 := [th-lemma]: #379
-[unit-resolution #380 #377 #363]: false
-unsat
-WdMJH3tkMv/rps8y9Ukq5Q 86 0
-#2 := false
-#37 := 0::real
-decl uf_2 :: (-> T2 T1 real)
-decl uf_4 :: T1
-#12 := uf_4
-decl uf_3 :: T2
-#5 := uf_3
-#13 := (uf_2 uf_3 uf_4)
-#34 := -1::real
-#140 := (* -1::real #13)
-decl uf_1 :: real
-#4 := uf_1
-#141 := (+ uf_1 #140)
-#143 := (>= #141 0::real)
-#6 := (:var 0 T1)
-#7 := (uf_2 uf_3 #6)
-#127 := (pattern #7)
-#35 := (* -1::real #7)
-#36 := (+ uf_1 #35)
-#47 := (>= #36 0::real)
-#134 := (forall (vars (?x2 T1)) (:pat #127) #47)
-#49 := (forall (vars (?x2 T1)) #47)
-#137 := (iff #49 #134)
-#135 := (iff #47 #47)
-#136 := [refl]: #135
-#138 := [quant-intro #136]: #137
-#67 := (~ #49 #49)
-#58 := (~ #47 #47)
-#66 := [refl]: #58
-#68 := [nnf-pos #66]: #67
-#10 := (<= #7 uf_1)
-#11 := (forall (vars (?x2 T1)) #10)
-#50 := (iff #11 #49)
-#46 := (iff #10 #47)
-#48 := [rewrite]: #46
-#51 := [quant-intro #48]: #50
-#32 := [asserted]: #11
-#52 := [mp #32 #51]: #49
-#69 := [mp~ #52 #68]: #49
-#139 := [mp #69 #138]: #134
-#149 := (not #134)
-#150 := (or #149 #143)
-#151 := [quant-inst]: #150
-#144 := [unit-resolution #151 #139]: #143
-#142 := (<= #141 0::real)
-#38 := (<= #36 0::real)
-#128 := (forall (vars (?x1 T1)) (:pat #127) #38)
-#41 := (forall (vars (?x1 T1)) #38)
-#131 := (iff #41 #128)
-#129 := (iff #38 #38)
-#130 := [refl]: #129
-#132 := [quant-intro #130]: #131
-#62 := (~ #41 #41)
-#64 := (~ #38 #38)
-#65 := [refl]: #64
-#63 := [nnf-pos #65]: #62
-#8 := (<= uf_1 #7)
-#9 := (forall (vars (?x1 T1)) #8)
-#42 := (iff #9 #41)
-#39 := (iff #8 #38)
-#40 := [rewrite]: #39
-#43 := [quant-intro #40]: #42
-#31 := [asserted]: #9
-#44 := [mp #31 #43]: #41
-#61 := [mp~ #44 #63]: #41
-#133 := [mp #61 #132]: #128
-#145 := (not #128)
-#146 := (or #145 #142)
-#147 := [quant-inst]: #146
-#148 := [unit-resolution #147 #133]: #142
-#45 := (= uf_1 #13)
-#55 := (not #45)
-#14 := (= #13 uf_1)
-#15 := (not #14)
-#56 := (iff #15 #55)
-#53 := (iff #14 #45)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#33 := [asserted]: #15
-#60 := [mp #33 #57]: #55
-#153 := (not #143)
-#152 := (not #142)
-#154 := (or #45 #152 #153)
-#155 := [th-lemma]: #154
-[unit-resolution #155 #60 #148 #144]: false
-unsat
-V+IAyBZU/6QjYs6JkXx8LQ 57 0
-#2 := false
-#4 := 0::real
-decl uf_1 :: (-> T2 real)
-decl uf_2 :: (-> T1 T1 T2)
-decl uf_12 :: (-> T4 T1)
-decl uf_4 :: T4
-#11 := uf_4
-#39 := (uf_12 uf_4)
-decl uf_10 :: T4
-#27 := uf_10
-#38 := (uf_12 uf_10)
-#40 := (uf_2 #38 #39)
-#41 := (uf_1 #40)
-#264 := (>= #41 0::real)
-#266 := (not #264)
-#43 := (= #41 0::real)
-#44 := (not #43)
-#131 := [asserted]: #44
-#272 := (or #43 #266)
-#42 := (<= #41 0::real)
-#130 := [asserted]: #42
-#265 := (not #42)
-#270 := (or #43 #265 #266)
-#271 := [th-lemma]: #270
-#273 := [unit-resolution #271 #130]: #272
-#274 := [unit-resolution #273 #131]: #266
-#6 := (:var 0 T1)
-#5 := (:var 1 T1)
-#7 := (uf_2 #5 #6)
-#241 := (pattern #7)
-#8 := (uf_1 #7)
-#65 := (>= #8 0::real)
-#242 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #241) #65)
-#66 := (forall (vars (?x1 T1) (?x2 T1)) #65)
-#245 := (iff #66 #242)
-#243 := (iff #65 #65)
-#244 := [refl]: #243
-#246 := [quant-intro #244]: #245
-#149 := (~ #66 #66)
-#151 := (~ #65 #65)
-#152 := [refl]: #151
-#150 := [nnf-pos #152]: #149
-#9 := (<= 0::real #8)
-#10 := (forall (vars (?x1 T1) (?x2 T1)) #9)
-#67 := (iff #10 #66)
-#63 := (iff #9 #65)
-#64 := [rewrite]: #63
-#68 := [quant-intro #64]: #67
-#60 := [asserted]: #10
-#69 := [mp #60 #68]: #66
-#147 := [mp~ #69 #150]: #66
-#247 := [mp #147 #246]: #242
-#267 := (not #242)
-#268 := (or #267 #264)
-#269 := [quant-inst]: #268
-[unit-resolution #269 #247 #274]: false
-unsat
-vqiyJ/qjGXZ3iOg6xftiug 15 0
-uf_1 -> val!0
-uf_2 -> val!1
-uf_3 -> val!2
-uf_5 -> val!15
-uf_6 -> val!26
-uf_4 -> {
-  val!0 -> val!12
-  val!1 -> val!13
-  else -> val!13
-}
-uf_7 -> {
-  val!6 -> val!31
-  else -> val!31
-}
-sat
-mrZPJZyTokErrN6SYupisw 9 0
-uf_4 -> val!1
-uf_2 -> val!3
-uf_3 -> val!4
-uf_1 -> {
-  val!5 -> val!6
-  val!4 -> val!7
-  else -> val!7
-}
-sat
--- a/src/HOL/Multivariate_Analysis/Integration_MV.thy	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3473 +0,0 @@
-
-header {* Kurzweil-Henstock gauge integration in many dimensions. *}
-(*  Author:                     John Harrison
-    Translation from HOL light: Robert Himmelmann, TU Muenchen *)
-
-theory Integration_MV
-  imports Derivative SMT
-begin
-
-declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration_MV.cert"]]
-declare [[smt_record=true]]
-declare [[z3_proofs=true]]
-
-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 conjunctD5: assumes "a \<and> b \<and> c \<and> d \<and> e" shows a b c d e using assms by auto
-
-declare smult_conv_scaleR[simp]
-
-subsection {* Some useful lemmas about intervals. *}
-
-lemma empty_as_interval: "{} = {1..0::real^'n}"
-  apply(rule set_ext,rule) defer unfolding vector_le_def mem_interval
-  using UNIV_witness[where 'a='n] apply(erule_tac exE,rule_tac x=x in allE) by auto
-
-lemma interior_subset_union_intervals: 
-  assumes "i = {a..b::real^'n}" "j = {c..d}" "interior j \<noteq> {}" "i \<subseteq> j \<union> s" "interior(i) \<inter> interior(j) = {}"
-  shows "interior i \<subseteq> interior s" proof-
-  have "{a<..<b} \<inter> {c..d} = {}" using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
-    unfolding assms(1,2) interior_closed_interval by auto
-  moreover have "{a<..<b} \<subseteq> {c..d} \<union> s" apply(rule order_trans,rule interval_open_subset_closed)
-    using assms(4) unfolding assms(1,2) by auto
-  ultimately show ?thesis apply-apply(rule interior_maximal) defer apply(rule open_interior)
-    unfolding assms(1,2) interior_closed_interval by auto qed
-
-lemma inter_interior_unions_intervals: fixes f::"(real^'n) set set"
-  assumes "finite f" "open s" "\<forall>t\<in>f. \<exists>a b. t = {a..b}" "\<forall>t\<in>f. s \<inter> (interior t) = {}"
-  shows "s \<inter> interior(\<Union>f) = {}" proof(rule ccontr,unfold ex_in_conv[THEN sym]) case goal1
-  have lem1:"\<And>x e s U. ball x e \<subseteq> s \<inter> interior U \<longleftrightarrow> ball x e \<subseteq> s \<inter> U" apply rule  defer apply(rule_tac Int_greatest)
-    unfolding open_subset_interior[OF open_ball]  using interior_subset by auto
-  have lem2:"\<And>x s P. \<exists>x\<in>s. P x \<Longrightarrow> \<exists>x\<in>insert x s. P x" by auto
-  have "\<And>f. finite f \<Longrightarrow> (\<forall>t\<in>f. \<exists>a b. t = {a..b}) \<Longrightarrow> (\<exists>x. x \<in> s \<inter> interior (\<Union>f)) \<Longrightarrow> (\<exists>t\<in>f. \<exists>x. \<exists>e>0. ball x e \<subseteq> s \<inter> t)" proof- case goal1
-  thus ?case proof(induct rule:finite_induct) 
-    case empty from this(2) guess x .. hence False unfolding Union_empty interior_empty by auto thus ?case by auto next
-    case (insert i f) guess x using insert(5) .. note x = this
-    then guess e unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior],rule_format] .. note e=this
-    guess a using insert(4)[rule_format,OF insertI1] .. then guess b .. note ab = this
-    show ?case proof(cases "x\<in>i") case False hence "x \<in> UNIV - {a..b}" unfolding ab by auto
-      then guess d unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_interval],rule_format] ..
-      hence "0 < d" "ball x (min d e) \<subseteq> UNIV - i" using e unfolding ab by auto
-      hence "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)" using e unfolding lem1 by auto hence "x \<in> s \<inter> interior (\<Union>f)" using `d>0` e by auto
-      hence "\<exists>t\<in>f. \<exists>x e. 0 < e \<and> ball x e \<subseteq> s \<inter> t" apply-apply(rule insert(3)) using insert(4) by auto thus ?thesis by auto next
-    case True show ?thesis proof(cases "x\<in>{a<..<b}")
-      case True then guess d unfolding open_contains_ball_eq[OF open_interval,rule_format] ..
-      thus ?thesis apply(rule_tac x=i in bexI,rule_tac x=x in exI,rule_tac x="min d e" in exI)
-	unfolding ab using interval_open_subset_closed[of a b] and e by fastsimp+ next
-    case False then obtain k where "x$k \<le> a$k \<or> x$k \<ge> b$k" unfolding mem_interval by(auto simp add:not_less) 
-    hence "x$k = a$k \<or> x$k = b$k" using True unfolding ab and mem_interval apply(erule_tac x=k in allE) by auto
-    hence "\<exists>x. ball x (e/2) \<subseteq> s \<inter> (\<Union>f)" proof(erule_tac disjE)
-      let ?z = "x - (e/2) *\<^sub>R basis k" assume as:"x$k = a$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
-	fix y assume "y \<in> ball ?z (e / 2) \<inter> i" hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
-	hence "\<bar>(?z - y) $ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding vector_dist_norm by auto
-	hence "y$k < a$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps)
-	hence "y \<notin> i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed
-      moreover have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof
-	fix y assume as:"y\<in> ball ?z (e/2)" have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R basis k)"
-	   apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R basis k"])
-	  unfolding norm_scaleR norm_basis by auto
-	also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps)
-	finally show "y\<in>ball x e" unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) qed
-      ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto
-    next let ?z = "x + (e/2) *\<^sub>R basis k" assume as:"x$k = b$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
-	fix y assume "y \<in> ball ?z (e / 2) \<inter> i" hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
-	hence "\<bar>(?z - y) $ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding vector_dist_norm by auto
-	hence "y$k > b$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps)
-	hence "y \<notin> i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed
-      moreover have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof
-	fix y assume as:"y\<in> ball ?z (e/2)" have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R basis k)"
-	   apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R basis k"])
-	  unfolding norm_scaleR norm_basis by auto
-	also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps)
-	finally show "y\<in>ball x e" unfolding mem_ball vector_dist_norm using e by(auto simp add:field_simps) qed
-      ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto qed 
-    then guess x .. hence "x \<in> s \<inter> interior (\<Union>f)" unfolding lem1[where U="\<Union>f",THEN sym] using centre_in_ball e[THEN conjunct1] by auto
-    thus ?thesis apply-apply(rule lem2,rule insert(3)) using insert(4) by auto qed qed qed qed note * = this
-  guess t using *[OF assms(1,3) goal1]  .. from this(2) guess x .. then guess e ..
-  hence "x \<in> s" "x\<in>interior t" defer using open_subset_interior[OF open_ball, of x e t] by auto
-  thus False using `t\<in>f` assms(4) by auto qed
-subsection {* Bounds on intervals where they exist. *}
-
-definition "interval_upperbound (s::(real^'n) set) = (\<chi> i. Sup {a. \<exists>x\<in>s. x$i = a})"
-
-definition "interval_lowerbound (s::(real^'n) set) = (\<chi> i. Inf {a. \<exists>x\<in>s. x$i = a})"
-
-lemma interval_upperbound[simp]: assumes "\<forall>i. a$i \<le> b$i" shows "interval_upperbound {a..b} = b"
-  using assms unfolding interval_upperbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE)
-  apply(rule Sup_unique) unfolding setle_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer
-  apply(rule,rule) apply(rule_tac x="b$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=b in bexI)
-  unfolding mem_interval using assms by auto
-
-lemma interval_lowerbound[simp]: assumes "\<forall>i. a$i \<le> b$i" shows "interval_lowerbound {a..b} = a"
-  using assms unfolding interval_lowerbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE)
-  apply(rule Inf_unique) unfolding setge_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer
-  apply(rule,rule) apply(rule_tac x="a$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=a in bexI)
-  unfolding mem_interval using assms by auto
-
-lemmas interval_bounds = interval_upperbound interval_lowerbound
-
-lemma interval_bounds'[simp]: assumes "{a..b}\<noteq>{}" shows "interval_upperbound {a..b} = b" "interval_lowerbound {a..b} = a"
-  using assms unfolding interval_ne_empty by auto
-
-lemma interval_upperbound_1[simp]: "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> interval_upperbound {a..b} = (b::real^1)"
-  apply(rule interval_upperbound) by auto
-
-lemma interval_lowerbound_1[simp]: "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> interval_lowerbound {a..b} = (a::real^1)"
-  apply(rule interval_lowerbound) by auto
-
-lemmas interval_bound_1 = interval_upperbound_1 interval_lowerbound_1
-
-subsection {* Content (length, area, volume...) of an interval. *}
-
-definition "content (s::(real^'n) set) =
-       (if s = {} then 0 else (\<Prod>i\<in>UNIV. (interval_upperbound s)$i - (interval_lowerbound s)$i))"
-
-lemma interval_not_empty:"\<forall>i. a$i \<le> b$i \<Longrightarrow> {a..b::real^'n} \<noteq> {}"
-  unfolding interval_eq_empty unfolding not_ex not_less by assumption
-
-lemma content_closed_interval: assumes "\<forall>i. a$i \<le> b$i"
-  shows "content {a..b} = (\<Prod>i\<in>UNIV. b$i - a$i)"
-  using interval_not_empty[OF assms] unfolding content_def interval_upperbound[OF assms] interval_lowerbound[OF assms] by auto
-
-lemma content_closed_interval': assumes "{a..b}\<noteq>{}" shows "content {a..b} = (\<Prod>i\<in>UNIV. b$i - a$i)"
-  apply(rule content_closed_interval) using assms unfolding interval_ne_empty .
-
-lemma content_1:"dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> content {a..b} = dest_vec1 b - dest_vec1 a"
-  using content_closed_interval[of a b] by auto
-
-lemma content_1':"a \<le> b \<Longrightarrow> content {vec1 a..vec1 b} = b - a" using content_1[of "vec a" "vec b"] by auto
-
-lemma content_unit[intro]: "content{0..1::real^'n} = 1" proof-
-  have *:"\<forall>i. 0$i \<le> (1::real^'n::finite)$i" by auto
-  have "0 \<in> {0..1::real^'n::finite}" unfolding mem_interval by auto
-  thus ?thesis unfolding content_def interval_bounds[OF *] using setprod_1 by auto qed
-
-lemma content_pos_le[intro]: "0 \<le> content {a..b}" proof(cases "{a..b}={}")
-  case False hence *:"\<forall>i. a $ i \<le> b $ i" unfolding interval_ne_empty by assumption
-  have "(\<Prod>i\<in>UNIV. interval_upperbound {a..b} $ i - interval_lowerbound {a..b} $ i) \<ge> 0"
-    apply(rule setprod_nonneg) unfolding interval_bounds[OF *] using * apply(erule_tac x=x in allE) by auto
-  thus ?thesis unfolding content_def by(auto simp del:interval_bounds') qed(unfold content_def, auto)
-
-lemma content_pos_lt: assumes "\<forall>i. a$i < b$i" shows "0 < content {a..b}"
-proof- have help_lemma1: "\<forall>i. a$i < b$i \<Longrightarrow> \<forall>i. a$i \<le> ((b$i)::real)" apply(rule,erule_tac x=i in allE) by auto
-  show ?thesis unfolding content_closed_interval[OF help_lemma1[OF assms]] apply(rule setprod_pos)
-    using assms apply(erule_tac x=x in allE) by auto qed
-
-lemma content_pos_lt_1: "dest_vec1 a < dest_vec1 b \<Longrightarrow> 0 < content({a..b})"
-  apply(rule content_pos_lt) by auto
-
-lemma content_eq_0: "content({a..b::real^'n}) = 0 \<longleftrightarrow> (\<exists>i. b$i \<le> a$i)" proof(cases "{a..b} = {}")
-  case True thus ?thesis unfolding content_def if_P[OF True] unfolding interval_eq_empty apply-
-    apply(rule,erule exE) apply(rule_tac x=i in exI) by auto next
-  guess a using UNIV_witness[where 'a='n] .. case False note as=this[unfolded interval_eq_empty not_ex not_less]
-  show ?thesis unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_UNIV]
-    apply(rule) apply(erule_tac[!] exE bexE) unfolding interval_bounds[OF as] apply(rule_tac x=x in exI) defer
-    apply(rule_tac x=i in bexI) using as apply(erule_tac x=i in allE) by auto qed
-
-lemma cond_cases:"(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)" by auto
-
-lemma content_closed_interval_cases:
-  "content {a..b} = (if \<forall>i. a$i \<le> b$i then setprod (\<lambda>i. b$i - a$i) UNIV else 0)" apply(rule cond_cases) 
-  apply(rule content_closed_interval) unfolding content_eq_0 not_all not_le defer apply(erule exE,rule_tac x=x in exI) by auto
-
-lemma content_eq_0_interior: "content {a..b} = 0 \<longleftrightarrow> interior({a..b}) = {}"
-  unfolding content_eq_0 interior_closed_interval interval_eq_empty by auto
-
-lemma content_eq_0_1: "content {a..b::real^1} = 0 \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
-  unfolding content_eq_0 by auto
-
-lemma content_pos_lt_eq: "0 < content {a..b} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
-  apply(rule) defer apply(rule content_pos_lt,assumption) proof- assume "0 < content {a..b}"
-  hence "content {a..b} \<noteq> 0" by auto thus "\<forall>i. a$i < b$i" unfolding content_eq_0 not_ex not_le by auto qed
-
-lemma content_empty[simp]: "content {} = 0" unfolding content_def by auto
-
-lemma content_subset: assumes "{a..b} \<subseteq> {c..d}" shows "content {a..b::real^'n} \<le> content {c..d}" proof(cases "{a..b}={}")
-  case True thus ?thesis using content_pos_le[of c d] by auto next
-  case False hence ab_ne:"\<forall>i. a $ i \<le> b $ i" unfolding interval_ne_empty by auto
-  hence ab_ab:"a\<in>{a..b}" "b\<in>{a..b}" unfolding mem_interval by auto
-  have "{c..d} \<noteq> {}" using assms False by auto
-  hence cd_ne:"\<forall>i. c $ i \<le> d $ i" using assms unfolding interval_ne_empty by auto
-  show ?thesis unfolding content_def unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne]
-    unfolding if_not_P[OF False] if_not_P[OF `{c..d} \<noteq> {}`] apply(rule setprod_mono,rule) proof fix i::'n
-    show "0 \<le> b $ i - a $ i" using ab_ne[THEN spec[where x=i]] by auto
-    show "b $ i - a $ i \<le> d $ i - c $ i"
-      using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(2),of i]
-      using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i] by auto qed qed
-
-lemma content_lt_nz: "0 < content {a..b} \<longleftrightarrow> content {a..b} \<noteq> 0"
-  unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by auto
-
-subsection {* The notion of a gauge --- simply an open set containing the point. *}
-
-definition gauge where "gauge d \<longleftrightarrow> (\<forall>x. x\<in>(d x) \<and> open(d x))"
-
-lemma gaugeI:assumes "\<And>x. x\<in>g x" "\<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" "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)" apply(rule gauge_ball) by auto
-
-lemma gauge_inter: "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" "\<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(meson zero_less_one)
-
-subsection {* Divisions. *}
-
-definition division_of (infixl "division'_of" 40) where
-  "s division_of i \<equiv>
-        finite s \<and>
-        (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = {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>k. k\<in>s \<Longrightarrow> k \<subseteq> i" "\<And>k. k\<in>s \<Longrightarrow>  k \<noteq> {}" "\<And>k. k\<in>s \<Longrightarrow> (\<exists>a b. k = {a..b})"
-  "\<And>k1 k2. \<lbrakk>k1\<in>s; k2\<in>s; k1 \<noteq> k2\<rbrakk> \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}" "\<Union>s = i" using assms unfolding division_of_def by auto
-
-lemma division_ofI:
-  assumes "finite s" "\<And>k. k\<in>s \<Longrightarrow> k \<subseteq> i" "\<And>k. k\<in>s \<Longrightarrow>  k \<noteq> {}" "\<And>k. k\<in>s \<Longrightarrow> (\<exists>a b. k = {a..b})"
-  "\<And>k1 k2. \<lbrakk>k1\<in>s; k2\<in>s; k1 \<noteq> k2\<rbrakk> \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}" "\<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]: "{a..b} \<noteq> {} \<Longrightarrow> {{a..b}} division_of {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 {a..a::real^'n} \<longleftrightarrow> s = {{a..a}}" (is "?l = ?r") proof
-  assume ?r moreover { assume "s = {{a}}" moreover fix k assume "k\<in>s" 
-    ultimately have"\<exists>x y. k = {x..y}" apply(rule_tac x=a in exI)+ unfolding interval_sing[THEN conjunct1] by auto }
-  ultimately show ?l unfolding division_of_def interval_sing[THEN conjunct1] by auto next
-  assume ?l note as=conjunctD4[OF this[unfolded division_of_def interval_sing[THEN conjunct1]]]
-  { fix x assume x:"x\<in>s" have "x={a}" using as(2)[rule_format,OF x] by auto }
-  moreover have "s \<noteq> {}" using as(4) by auto ultimately show ?r unfolding interval_sing[THEN conjunct1] 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 {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. {a..b} \<in> d \<longrightarrow> P {a..b}))"
-  unfolding division_of_def by fastsimp
-
-lemma division_of_subset: assumes "p division_of (\<Union>p)" "q \<subseteq> p" shows "q division_of (\<Union>q)"
-  apply(rule division_ofI) proof- note as=division_ofD[OF assms(1)]
-  show "finite q" apply(rule finite_subset) using as(1) assms(2) by auto
-  { fix k assume "k \<in> q" hence kp:"k\<in>p" using assms(2) by auto show "k\<subseteq>\<Union>q" using `k \<in> q` by auto
-  show "\<exists>a b. k = {a..b}" using as(4)[OF kp] by auto show "k \<noteq> {}" using as(3)[OF kp] by auto }
-  fix k1 k2 assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2" hence *:"k1\<in>p" "k2\<in>p" "k1\<noteq>k2" using assms(2) by auto
-  show "interior k1 \<inter> interior k2 = {}" using as(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 {a..b} = 0" "d division_of {a..b}" shows "\<forall>k\<in>d. content k = 0"
-  unfolding forall_in_division[OF assms(2)] apply(rule,rule,rule) apply(drule division_ofD(2)[OF assms(2)])
-  apply(drule content_subset) unfolding assms(1) proof- case goal1 thus ?case using content_pos_le[of a b] by auto qed
-
-lemma division_inter: assumes "p1 division_of s1" "p2 division_of (s2::(real^'a) set)"
-  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_ext) unfolding * and Union_image_eq UN_iff
-    using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto
-  { 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 thus "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
-  guess a1 using division_ofD(4)[OF assms(1) k(2)] .. then guess b1 .. note ab1=this
-  guess a2 using division_ofD(4)[OF assms(2) k(3)] .. then guess b2 .. note ab2=this
-  show "\<exists>a b. k = {a..b}" unfolding k ab1 ab2 unfolding inter_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" hence 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 *) defer apply(rule_tac[1-4] subset_interior)
-    using division_ofD(5)[OF assms(1) k1(2) k2(2)]
-    using division_ofD(5)[OF assms(2) k1(3) k2(3)] using th by auto qed qed
-
-lemma division_inter_1: assumes "d division_of i" "{a..b::real^'n} \<subseteq> i"
-  shows "{ {a..b} \<inter> k |k. k \<in> d \<and> {a..b} \<inter> k \<noteq> {} } division_of {a..b}" proof(cases "{a..b} = {}")
-  case True show ?thesis unfolding True and division_of_trivial by auto next
-  have *:"{a..b} \<inter> i = {a..b}" using assms(2) by auto 
-  case False show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed
-
-lemma elementary_inter: assumes "p1 division_of s" "p2 division_of (t::(real^'n) set)"
-  shows "\<exists>p. p division_of (s \<inter> t)"
-  by(rule,rule division_inter[OF assms])
-
-lemma elementary_inters: assumes "finite f" "f\<noteq>{}" "\<forall>s\<in>f. \<exists>p. p division_of (s::(real^'n) set)"
-  shows "\<exists>p. p division_of (\<Inter> f)" using assms apply-proof(induct f rule:finite_induct)
-case (insert x f) show ?case proof(cases "f={}")
-  case True thus ?thesis unfolding True using insert by auto next
-  case False guess p using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
-  moreover guess px using insert(5)[rule_format,OF insertI1] .. ultimately
-  show ?thesis unfolding Inter_insert apply(rule_tac elementary_inter) by assumption+ qed qed auto
-
-lemma division_disjoint_union:
-  assumes "p1 division_of s1" "p2 division_of s2" "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)] and 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 subset_interior[OF d1(2)[OF as(1)]] subset_interior[OF d2(2)[OF as(2)]]
-      using assms(3) by blast } moreover
-  { assume as:"k1\<in>p2" "k2\<in>p1" have ?g using subset_interior[OF d1(2)[OF as(2)]] subset_interior[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 = {a..b}" using k d1(4) d2(4) by auto qed
-
-lemma partial_division_extend_1:
-  assumes "{c..d} \<subseteq> {a..b::real^'n}" "{c..d} \<noteq> {}"
-  obtains p where "p division_of {a..b}" "{c..d} \<in> p"
-proof- def n \<equiv> "CARD('n)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by auto
-  guess \<pi> using ex_bij_betw_nat_finite_1[OF finite_UNIV[where 'a='n]] .. note \<pi>=this
-  def \<pi>' \<equiv> "inv_into {1..n} \<pi>"
-  have \<pi>':"bij_betw \<pi>' UNIV {1..n}" using bij_betw_inv_into[OF \<pi>] unfolding \<pi>'_def n_def by auto
-  hence \<pi>'i:"\<And>i. \<pi>' i \<in> {1..n}" unfolding bij_betw_def by auto 
-  have \<pi>\<pi>'[simp]:"\<And>i. \<pi> (\<pi>' i) = i" unfolding \<pi>'_def apply(rule f_inv_into_f) unfolding n_def using \<pi> unfolding bij_betw_def by auto
-  have \<pi>'\<pi>[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi>' (\<pi> i) = i" unfolding \<pi>'_def apply(rule inv_into_f_eq) using \<pi> unfolding n_def bij_betw_def by auto
-  have "{c..d} \<noteq> {}" using assms by auto
-  let ?p1 = "\<lambda>l. {(\<chi> i. if \<pi>' i < l then c$i else a$i) .. (\<chi> i. if \<pi>' i < l then d$i else if \<pi>' i = l then c$\<pi> l else b$i)}"
-  let ?p2 = "\<lambda>l. {(\<chi> i. if \<pi>' i < l then c$i else if \<pi>' i = l then d$\<pi> l else a$i) .. (\<chi> i. if \<pi>' i < l then d$i else b$i)}"
-  let ?p =  "{?p1 l |l. l \<in> {1..n+1}} \<union> {?p2 l |l. l \<in> {1..n+1}}"
-  have abcd:"\<And>i. a $ i \<le> c $ i \<and> c$i \<le> d$i \<and> d $ i \<le> b $ i" using assms unfolding subset_interval interval_eq_empty by(auto simp add:not_le not_less)
-  show ?thesis apply(rule that[of ?p]) apply(rule division_ofI)
-  proof- have "\<And>i. \<pi>' i < Suc n"
-    proof(rule ccontr,unfold not_less) fix i assume "Suc n \<le> \<pi>' i"
-      hence "\<pi>' i \<notin> {1..n}" by auto thus False using \<pi>' unfolding bij_betw_def by auto
-    qed hence "c = (\<chi> i. if \<pi>' i < Suc n then c $ i else a $ i)"
-        "d = (\<chi> i. if \<pi>' i < Suc n then d $ i else if \<pi>' i = n + 1 then c $ \<pi> (n + 1) else b $ i)"
-      unfolding Cart_eq Cart_lambda_beta using \<pi>' unfolding bij_betw_def by auto
-    thus cdp:"{c..d} \<in> ?p" apply-apply(rule UnI1) unfolding mem_Collect_eq apply(rule_tac x="n + 1" in exI) by auto
-    have "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p1 l \<subseteq> {a..b}"  "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p2 l \<subseteq> {a..b}"
-      unfolding subset_eq apply(rule_tac[!] ballI,rule_tac[!] ccontr)
-    proof- fix l assume l:"l\<in>{1..n+1}" fix x assume "x\<notin>{a..b}"
-      then guess i unfolding mem_interval not_all .. note i=this
-      show "x \<in> ?p1 l \<Longrightarrow> False" "x \<in> ?p2 l \<Longrightarrow> False" unfolding mem_interval apply(erule_tac[!] x=i in allE)
-        apply(case_tac[!] "\<pi>' i < l", case_tac[!] "\<pi>' i = l") using abcd[of i] i by auto 
-    qed moreover have "\<And>x. x \<in> {a..b} \<Longrightarrow> x \<in> \<Union>?p"
-    proof- fix x assume x:"x\<in>{a..b}"
-      { presume "x\<notin>{c..d} \<Longrightarrow> x \<in> \<Union>?p" thus "x \<in> \<Union>?p" using cdp by blast }
-      let ?M = "{i. i\<in>{1..n+1} \<and> \<not> (c $ \<pi> i \<le> x $ \<pi> i \<and> x $ \<pi> i \<le> d $ \<pi> i)}"
-      assume "x\<notin>{c..d}" then guess i0 unfolding mem_interval not_all ..
-      hence "\<pi>' i0 \<in> ?M" using \<pi>' unfolding bij_betw_def by(auto intro!:le_SucI)
-      hence M:"finite ?M" "?M \<noteq> {}" by auto
-      def l \<equiv> "Min ?M" note l = Min_less_iff[OF M,unfolded l_def[symmetric]] Min_in[OF M,unfolded mem_Collect_eq l_def[symmetric]]
-        Min_gr_iff[OF M,unfolded l_def[symmetric]]
-      have "x\<in>?p1 l \<or> x\<in>?p2 l" using l(2)[THEN conjunct2] unfolding de_Morgan_conj not_le
-        apply- apply(erule disjE) apply(rule disjI1) defer apply(rule disjI2)
-      proof- assume as:"x $ \<pi> l < c $ \<pi> l"
-        show "x \<in> ?p1 l" unfolding mem_interval Cart_lambda_beta
-        proof case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le by auto
-          thus ?case using as x[unfolded mem_interval,rule_format,of i]
-            apply auto using l(3)[of "\<pi>' i"] by(auto elim!:ballE[where x="\<pi>' i"])
-        qed
-      next assume as:"x $ \<pi> l > d $ \<pi> l"
-        show "x \<in> ?p2 l" unfolding mem_interval Cart_lambda_beta
-        proof case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le by auto
-          thus ?case using as x[unfolded mem_interval,rule_format,of i]
-            apply auto using l(3)[of "\<pi>' i"] by(auto elim!:ballE[where x="\<pi>' i"])
-        qed qed
-      thus "x \<in> \<Union>?p" using l(2) by blast 
-    qed ultimately show "\<Union>?p = {a..b}" apply-apply(rule) defer apply(rule) by(assumption,blast)
-    
-    show "finite ?p" by auto
-    fix k assume k:"k\<in>?p" then obtain l where l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}" by auto
-    show "k\<subseteq>{a..b}" apply(rule,unfold mem_interval,rule,rule) 
-    proof- fix i::'n and x assume "x \<in> k" moreover have "\<pi>' i < l \<or> \<pi>' i = l \<or> \<pi>' i > l" by auto
-      ultimately show "a$i \<le> x$i" "x$i \<le> b$i" using abcd[of i] using l by(auto elim:disjE elim!:allE[where x=i] simp add:vector_le_def)
-    qed have "\<And>l. ?p1 l \<noteq> {}" "\<And>l. ?p2 l \<noteq> {}" unfolding interval_eq_empty not_ex apply(rule_tac[!] allI)
-    proof- case goal1 thus ?case using abcd[of x] by auto
-    next   case goal2 thus ?case using abcd[of x] by auto
-    qed thus "k \<noteq> {}" using k by auto
-    show "\<exists>a b. k = {a..b}" using k by auto
-    fix k' assume k':"k' \<in> ?p" "k \<noteq> k'" then obtain l' where l':"k' = ?p1 l' \<or> k' = ?p2 l'" "l' \<in> {1..n + 1}" by auto
-    { fix k k' l l'
-      assume k:"k\<in>?p" and l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}" 
-      assume k':"k' \<in> ?p" "k \<noteq> k'" and  l':"k' = ?p1 l' \<or> k' = ?p2 l'" "l' \<in> {1..n + 1}" 
-      assume "l \<le> l'" fix x
-      have "x \<notin> interior k \<inter> interior k'" 
-      proof(rule,cases "l' = n+1") assume x:"x \<in> interior k \<inter> interior k'"
-        case True hence "\<And>i. \<pi>' i < l'" using \<pi>'i by(auto simp add:less_Suc_eq_le)
-        hence k':"k' = {c..d}" using l'(1) \<pi>'i by(auto simp add:Cart_nth_inverse)
-        have ln:"l < n + 1" 
-        proof(rule ccontr) case goal1 hence l2:"l = n+1" using l by auto
-          hence "\<And>i. \<pi>' i < l" using \<pi>'i by(auto simp add:less_Suc_eq_le)
-          hence "k = {c..d}" using l(1) \<pi>'i by(auto simp add:Cart_nth_inverse)
-          thus False using `k\<noteq>k'` k' by auto
-        qed have **:"\<pi>' (\<pi> l) = l" using \<pi>'\<pi>[of l] using l ln by auto
-        have "x $ \<pi> l < c $ \<pi> l \<or> d $ \<pi> l < x $ \<pi> l" using l(1) apply-
-        proof(erule disjE)
-          assume as:"k = ?p1 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show ?thesis using *[of "\<pi> l"] using ln unfolding Cart_lambda_beta ** by auto
-        next assume as:"k = ?p2 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show ?thesis using *[of "\<pi> l"] using ln unfolding Cart_lambda_beta ** by auto
-        qed thus False using x unfolding k' unfolding Int_iff interior_closed_interval mem_interval
-          by(auto elim!:allE[where x="\<pi> l"])
-      next case False hence "l < n + 1" using l'(2) using `l\<le>l'` by auto
-        hence ln:"l \<in> {1..n}" "l' \<in> {1..n}" using l l' False by auto
-        note \<pi>l = \<pi>'\<pi>[OF ln(1)] \<pi>'\<pi>[OF ln(2)]
-        assume x:"x \<in> interior k \<inter> interior k'"
-        show False using l(1) l'(1) apply-
-        proof(erule_tac[!] disjE)+
-          assume as:"k = ?p1 l" "k' = ?p1 l'"
-          note * = x[unfolded as Int_iff interior_closed_interval mem_interval]
-          have "l \<noteq> l'" using k'(2)[unfolded as] by auto
-          thus False using * by(smt Cart_lambda_beta \<pi>l)
-        next assume as:"k = ?p2 l" "k' = ?p2 l'"
-          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          have "l \<noteq> l'" apply(rule) using k'(2)[unfolded as] by auto
-          thus False using *[of "\<pi> l"] *[of "\<pi> l'"]
-            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` by auto
-        next assume as:"k = ?p1 l" "k' = ?p2 l'"
-          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show False using *[of "\<pi> l"] *[of "\<pi> l'"]
-            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` using abcd[of "\<pi> l'"] by smt 
-        next assume as:"k = ?p2 l" "k' = ?p1 l'"
-          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show False using *[of "\<pi> l"] *[of "\<pi> l'"]
-            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` using abcd[of "\<pi> l'"] by smt
-        qed qed } 
-    from this[OF k l k' l'] this[OF k'(1) l' k _ l] have "\<And>x. x \<notin> interior k \<inter> interior k'"
-      apply - apply(cases "l' \<le> l") using k'(2) by auto            
-    thus "interior k \<inter> interior k' = {}" by auto        
-qed qed
-
-lemma partial_division_extend_interval: assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> {a..b}"
-  obtains q where "p \<subseteq> q" "q division_of {a..b::real^'n}" proof(cases "p = {}")
-  case True guess q apply(rule elementary_interval[of a b]) .
-  thus ?thesis apply- apply(rule that[of q]) unfolding True by auto next
-  case False note p = division_ofD[OF assms(1)]
-  have *:"\<forall>k\<in>p. \<exists>q. q division_of {a..b} \<and> k\<in>q" proof case goal1
-    guess c using p(4)[OF goal1] .. then guess d .. note cd_ = this
-    have *:"{c..d} \<subseteq> {a..b}" "{c..d} \<noteq> {}" using p(2,3)[OF goal1, unfolded cd_] using assms(2) by auto
-    guess q apply(rule partial_division_extend_1[OF *]) . thus ?case unfolding cd_ by auto qed
-  guess q using bchoice[OF *] .. note q = conjunctD2[OF this[rule_format]]
-  have "\<And>x. x\<in>p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})" apply(rule,rule_tac p="q x" in division_of_subset) proof-
-    fix x assume x:"x\<in>p" show "q x division_of \<Union>q x" apply-apply(rule division_ofI)
-      using division_ofD[OF q(1)[OF x]] by auto show "q x - {x} \<subseteq> q x" by auto qed
-  hence "\<exists>d. d division_of \<Inter> ((\<lambda>i. \<Union>(q i - {i})) ` p)" apply- apply(rule elementary_inters)
-    apply(rule finite_imageI[OF p(1)]) unfolding image_is_empty apply(rule False) by auto
-  then guess d .. note d = this
-  show ?thesis apply(rule that[of "d \<union> p"]) proof-
-    have *:"\<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 *:"{a..b} = \<Inter> (\<lambda>i. \<Union>(q i - {i})) ` p \<union> \<Union>p" apply(rule *[OF False]) proof fix i assume i:"i\<in>p"
-      show "\<Union>(q i - {i}) \<union> i = {a..b}" using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto qed
-    show "d \<union> p division_of {a..b}" unfolding * apply(rule division_disjoint_union[OF d assms(1)])
-      apply(rule inter_interior_unions_intervals) apply(rule p open_interior ballI)+ proof(assumption,rule)
-      fix k assume k:"k\<in>p" have *:"\<And>u t s. u \<subseteq> s \<Longrightarrow> s \<inter> t = {} \<Longrightarrow> u \<inter> t = {}" by auto
-      show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<inter> interior k = {}" apply(rule *[of _ "interior (\<Union>(q k - {k}))"])
-	defer apply(subst Int_commute) apply(rule inter_interior_unions_intervals) proof- 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 = {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
-	have *:"\<And>x s. x \<in> s \<Longrightarrow> \<Inter>s \<subseteq> x" by auto show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<subseteq> interior (\<Union>(q k - {k}))"
-	  apply(rule subset_interior *)+ using k by auto qed qed qed auto qed
-
-lemma elementary_bounded[dest]: "p division_of s \<Longrightarrow> bounded (s::(real^'n) set)"
-  unfolding division_of_def by(metis bounded_Union bounded_interval) 
-
-lemma elementary_subset_interval: "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> {a..b::real^'n}"
-  by(meson elementary_bounded bounded_subset_closed_interval)
-
-lemma division_union_intervals_exists: assumes "{a..b::real^'n} \<noteq> {}"
-  obtains p where "(insert {a..b} p) division_of ({a..b} \<union> {c..d})" proof(cases "{c..d} = {}")
-  case True show ?thesis apply(rule that[of "{}"]) unfolding True using assms by auto next
-  case False note false=this show ?thesis proof(cases "{a..b} \<inter> {c..d} = {}")
-  have *:"\<And>a b. {a,b} = {a} \<union> {b}" by auto
-  case True show ?thesis apply(rule that[of "{{c..d}}"]) unfolding * apply(rule division_disjoint_union)
-    using false True assms using interior_subset by auto next
-  case False obtain u v where uv:"{a..b} \<inter> {c..d} = {u..v}" unfolding inter_interval by auto
-  have *:"{u..v} \<subseteq> {c..d}" using uv by auto
-  guess p apply(rule partial_division_extend_1[OF * False[unfolded uv]]) . note p=this division_ofD[OF this(1)]
-  have *:"{a..b} \<union> {c..d} = {a..b} \<union> \<Union>(p - {{u..v}})" "\<And>x s. insert x s = {x} \<union> s" using p(8) unfolding uv[THEN sym] by auto
-  show thesis apply(rule that[of "p - {{u..v}}"]) unfolding *(1) apply(subst *(2)) apply(rule division_disjoint_union)
-    apply(rule,rule assms) apply(rule division_of_subset[of p]) apply(rule division_of_union_self[OF p(1)]) defer
-    unfolding interior_inter[THEN sym] proof-
-    have *:"\<And>cd p uv ab. p \<subseteq> cd \<Longrightarrow> ab \<inter> cd = uv \<Longrightarrow> ab \<inter> p = uv \<inter> p" by auto
-    have "interior ({a..b} \<inter> \<Union>(p - {{u..v}})) = interior({u..v} \<inter> \<Union>(p - {{u..v}}))" 
-      apply(rule arg_cong[of _ _ interior]) apply(rule *[OF _ uv]) using p(8) by auto
-    also have "\<dots> = {}" unfolding interior_inter apply(rule inter_interior_unions_intervals) using p(6) p(7)[OF p(2)] p(3) by auto
-    finally show "interior ({a..b} \<inter> \<Union>(p - {{u..v}})) = {}" by assumption qed auto qed qed
-
-lemma division_of_unions: assumes "finite f"  "\<And>p. p\<in>f \<Longrightarrow> p division_of (\<Union>p)"
-  "\<And>k1 k2. \<lbrakk>k1 \<in> \<Union>f; k2 \<in> \<Union>f; k1 \<noteq> k2\<rbrakk> \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-  shows "\<Union>f division_of \<Union>\<Union>f" apply(rule division_ofI) prefer 5 apply(rule assms(3)|assumption)+
-  apply(rule finite_Union assms(1))+ prefer 3 apply(erule UnionE) apply(rule_tac s=X in division_ofD(3)[OF assms(2)])
-  using division_ofD[OF assms(2)] by auto
-  
-lemma elementary_union_interval: assumes "p division_of \<Union>p"
-  obtains q where "q division_of ({a..b::real^'n} \<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" "{a..b} = {} \<Longrightarrow> thesis" "{a..b} \<noteq> {} \<Longrightarrow> interior {a..b} = {} \<Longrightarrow> thesis"
-    "p\<noteq>{} \<Longrightarrow> interior {a..b}\<noteq>{} \<Longrightarrow> {a..b} \<noteq> {} \<Longrightarrow> thesis"
-  thus thesis by auto
-next assume as:"p={}" guess p apply(rule elementary_interval[of a b]) .
-  thus thesis apply(rule_tac that[of p]) unfolding as by auto 
-next assume as:"{a..b}={}" show thesis apply(rule that) unfolding as using assms by auto
-next assume as:"interior {a..b} = {}" "{a..b} \<noteq> {}"
-  show thesis apply(rule that[of "insert {a..b} p"],rule division_ofI)
-    unfolding finite_insert apply(rule assm(1)) unfolding Union_insert  
-    using assm(2-4) as apply- by(fastsimp dest: assm(5))+
-next assume as:"p \<noteq> {}" "interior {a..b} \<noteq> {}" "{a..b}\<noteq>{}"
-  have "\<forall>k\<in>p. \<exists>q. (insert {a..b} q) division_of ({a..b} \<union> k)" proof case goal1
-    from assm(4)[OF this] guess c .. then guess d ..
-    thus ?case apply-apply(rule division_union_intervals_exists[OF as(3),of c d]) by auto
-  qed from bchoice[OF this] guess q .. note q=division_ofD[OF this[rule_format]]
-  let ?D = "\<Union>{insert {a..b} (q k) | k. k \<in> p}"
-  show thesis apply(rule that[of "?D"]) proof(rule division_ofI)
-    have *:"{insert {a..b} (q k) |k. k \<in> p} = (\<lambda>k. insert {a..b} (q k)) ` p" by auto
-    show "finite ?D" apply(rule finite_Union) unfolding * apply(rule finite_imageI) using assm(1) q(1) by auto
-    show "\<Union>?D = {a..b} \<union> \<Union>p" unfolding * lem1 unfolding lem2[OF as(1), of "{a..b}",THEN sym]
-      using q(6) by auto
-    fix k assume k:"k\<in>?D" thus " k \<subseteq> {a..b} \<union> \<Union>p" using q(2) by auto
-    show "k \<noteq> {}" using q(3) k by auto show "\<exists>a b. k = {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 {a..b} (q x)"  "x\<in>p"  using k  by auto
-    obtain x' where x':"k'\<in>insert {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 apply(rule q(5)) using x x' k' unfolding True by auto
-    next case False 
-      { presume "k = {a..b} \<Longrightarrow> ?thesis" "k' = {a..b} \<Longrightarrow> ?thesis" 
-        "k \<noteq> {a..b} \<Longrightarrow> k' \<noteq> {a..b} \<Longrightarrow> ?thesis"
-        thus ?thesis by auto }
-      { assume as':"k  = {a..b}" show ?thesis apply(rule q(5)) using x' k'(2) unfolding as' by auto }
-      { assume as':"k' = {a..b}" show ?thesis apply(rule q(5)) using x  k'(2) unfolding as' by auto }
-      assume as':"k \<noteq> {a..b}" "k' \<noteq> {a..b}"
-      guess c using q(4)[OF x(2,1)] .. then guess d .. note c_d=this
-      have "interior k  \<inter> interior {a..b} = {}" apply(rule q(5)) using x  k'(2) using as' by auto
-      hence "interior k \<subseteq> interior x" apply-
-        apply(rule interior_subset_union_intervals[OF c_d _ as(2) q(2)[OF x(2,1)]]) by auto moreover
-      guess c using q(4)[OF x'(2,1)] .. then guess d .. note c_d=this
-      have "interior k' \<inter> interior {a..b} = {}" apply(rule q(5)) using x' k'(2) using as' by auto
-      hence "interior k' \<subseteq> interior x'" apply-
-        apply(rule interior_subset_union_intervals[OF c_d _ as(2) q(2)[OF x'(2,1)]]) by auto
-      ultimately show ?thesis using assm(5)[OF x(2) x'(2) False] by auto
-    qed qed } qed
-
-lemma elementary_unions_intervals:
-  assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = {a..b::real^'n}"
-  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
-    fix x F assume as:"finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
-    from this(3) guess p .. note p=this
-    from assms(2)[OF as(4)] guess a .. then guess b .. note ab=this
-    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 ab * by auto
-  qed(insert assms,auto) thus ?thesis apply-apply(erule exE,rule that) by auto qed
-
-lemma elementary_union: assumes "ps division_of s" "pt division_of (t::(real^'n) set)"
-  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
-  hence *:"\<Union>(ps \<union> pt) = s \<union> t" by auto
-  show ?thesis apply-apply(rule elementary_unions_intervals[of "ps\<union>pt"])
-    unfolding * prefer 3 apply(rule_tac p=p in that)
-    using assms[unfolded division_of_def] by auto qed
-
-lemma partial_division_extend: fixes t::"(real^'n) set"
-  assumes "p division_of s" "q division_of t" "s \<subseteq> t"
-  obtains r where "p \<subseteq> r" "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>{a..b}" using elementary_subset_interval[OF assms(2)] by auto
-  guess r1 apply(rule partial_division_extend_interval) apply(rule assms(1)[unfolded divp(6)[THEN sym]])
-    apply(rule subset_trans) by(rule ab assms[unfolded divp(6)[THEN sym]])+  note r1 = this division_ofD[OF this(2)]
-  guess p' apply(rule elementary_unions_intervals[of "r1 - p"]) using r1(3,6) by auto 
-  then obtain r2 where r2:"r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)" 
-    apply- apply(drule elementary_inter[OF _ assms(2)[unfolded divq(6)[THEN sym]]]) by auto
-  { fix x assume x:"x\<in>t" "x\<notin>s"
-    hence "x\<in>\<Union>r1" unfolding r1 using ab by auto
-    then guess r unfolding Union_iff .. note r=this moreover
-    have "r \<notin> p" proof assume "r\<in>p" hence "x\<in>s" using divp(2) r by auto
-      thus False using x by auto qed
-    ultimately have "x\<in>\<Union>(r1 - p)" by auto }
-  hence *:"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)" "open (interior s)" "\<forall>t\<in>r1-p. \<exists>a b. t = {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(rule)
-        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" "open (interior m)" "\<forall>t\<in>p. \<exists>a b. t = {a..b}" using divp by auto
-          show "\<forall>t\<in>p. interior m \<inter> interior t = {}" apply(rule, rule r1(7)) using as using r1 by auto
-        qed thus "interior s \<inter> interior m = {}" unfolding divp by auto
-      qed qed        
-    thus "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}" using interior_subset by auto
-  qed auto qed
-
-subsection {* Tagged (partial) divisions. *}
-
-definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where
-  "(s tagged_partial_division_of i) \<equiv>
-        finite s \<and>
-        (\<forall>x k. (x,k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = {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>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-  "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = {a..b}"
-  "\<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  apply- by blast+ 
-
-definition tagged_division_of (infixr "tagged'_division'_of" 40) where
-  "(s tagged_division_of i) \<equiv>
-        (s tagged_partial_division_of i) \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-
-lemma tagged_division_of_finite[dest]: "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 = {a..b})) \<and>
-        (\<forall>x1 k1 x2 k2. (x1,k1) \<in> s \<and> (x2,k2) \<in> s \<and> ~((x1,k1) = (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>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"  "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = {a..b}"
-  "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2,k2) \<in> s \<Longrightarrow> ~((x1,k1) = (x2,k2)) \<Longrightarrow> (interior(k1) \<inter> interior(k2) = {})"
-  "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  shows "s tagged_division_of i"
-  unfolding tagged_division_of apply(rule) defer apply rule
-  apply(rule allI impI conjI assms)+ apply assumption
-  apply(rule, rule assms, assumption) apply(rule assms, assumption)
-  using assms(1,5-) apply- by blast+
-
-lemma tagged_division_ofD[dest]: assumes "s tagged_division_of i"
-  shows "finite s" "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"  "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = {a..b}"
-  "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2,k2) \<in> s \<Longrightarrow> ~((x1,k1) = (x2,k2)) \<Longrightarrow> (interior(k1) \<inter> interior(k2) = {})"
-  "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)" using assms unfolding tagged_division_of apply- 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
-  thus  "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = {a..b}" using assm apply- by fastsimp+
-  fix k' assume k':"k' \<in> snd ` s" "k \<noteq> k'" from this(1) obtain xk' where xk':"(xk', k') \<in> s" by auto
-  thus "interior k \<inter> interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto
-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
-  thus "k\<noteq>{}" "\<exists>a b. k = {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
-  thus "interior k \<inter> interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto
-qed
-
-lemma tagged_partial_division_subset: assumes "s tagged_partial_division_of i" "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 setsum_over_tagged_division_lemma: fixes d::"(real^'m) set \<Rightarrow> 'a::real_normed_vector"
-  assumes "p tagged_division_of i" "\<And>u v. {u..v} \<noteq> {} \<Longrightarrow> content {u..v} = 0 \<Longrightarrow> d {u..v} = 0"
-  shows "setsum (\<lambda>(x,k). d k) p = setsum d (snd ` p)"
-proof- note assm=tagged_division_ofD[OF assms(1)]
-  have *:"(\<lambda>(x,k). d k) = d \<circ> snd" unfolding o_def apply(rule ext) by auto
-  show ?thesis unfolding * apply(subst eq_commute) proof(rule setsum_reindex_nonzero)
-    show "finite p" using assm by auto
-    fix x y assume as:"x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y" 
-    obtain a b where ab:"snd x = {a..b}" using assm(4)[of "fst x" "snd x"] as(1) by auto
-    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y" unfolding as(4)[THEN sym] using as(1-3) by auto
-    hence "interior (snd x) \<inter> interior (snd y) = {}" apply-apply(rule assm(5)[of "fst x" _ "fst y"]) using as by auto 
-    hence "content {a..b} = 0" unfolding as(4)[THEN sym] ab content_eq_0_interior by auto
-    hence "d {a..b} = 0" apply-apply(rule assms(2)) using assm(2)[of "fst x" "snd x"] as(1) unfolding ab[THEN sym] by auto
-    thus "d (snd x) = 0" 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> {a..b} \<Longrightarrow> {(x,{a..b})} tagged_division_of {a..b}"
-  apply(rule tagged_division_ofI) by auto
-
-lemma tagged_division_union:
-  assumes "p1 tagged_division_of s1"  "p2 tagged_division_of s2" "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)] and 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 = {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) subset_interior by blast
-  show "interior k \<inter> interior k' = {}" apply(cases "(x,k)\<in>p1", case_tac[!] "(x',k')\<in>p1")
-    apply(rule p1(5)) prefer 4 apply(rule *) prefer 6 apply(subst Int_commute,rule *) prefer 8 apply(rule p2(5))
-    using p1(3) p2(3) using xk xk' by auto qed 
-
-lemma tagged_division_unions:
-  assumes "finite iset" "\<forall>i\<in>iset. (pfn(i) tagged_division_of i)"
-  "\<forall>i1 \<in> iset. \<forall>i2 \<in> iset. ~(i1 = 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)" apply(rule finite_Union) 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 = {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] subset_interior by blast
-  show "interior k \<inter> interior k' = {}" apply(cases "i=i'")
-    using assm(5)[OF i _ xk'(2)]  i'(2) using assm(3)[OF i] assm(3)[OF i'] defer apply-apply(rule *) 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] by auto
-
-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] by auto
-
-subsection {* Fine-ness of a partition w.r.t. a gauge. *}
-
-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 {* Gauge integral. Define on compact intervals first, then use a limit. *}
-
-definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) where
-  "(f has_integral_compact_interval y) i \<equiv>
-        (\<forall>e>0. \<exists>d. gauge d \<and>
-          (\<forall>p. p tagged_division_of i \<and> d fine p
-                        \<longrightarrow> norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))"
-
-definition has_integral (infixr "has'_integral" 46) where 
-"((f::(real^'n \<Rightarrow> 'b::real_normed_vector)) has_integral y) i \<equiv>
-        if (\<exists>a b. i = {a..b}) then (f has_integral_compact_interval y) i
-        else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b}
-              \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral_compact_interval z) {a..b} \<and>
-                                       norm(z - y) < e))"
-
-lemma has_integral:
- "(f has_integral y) ({a..b}) \<longleftrightarrow>
-        (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p
-                        \<longrightarrow> norm(setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
-  unfolding has_integral_def has_integral_compact_interval_def by auto
-
-lemma has_integralD[dest]: assumes
- "(f has_integral y) ({a..b})" "e>0"
-  obtains d where "gauge d" "\<And>p. p tagged_division_of {a..b} \<Longrightarrow> d fine p
-                        \<Longrightarrow> norm(setsum (\<lambda>(x,k). content(k) *\<^sub>R f(x)) p - y) < e"
-  using assms unfolding has_integral by auto
-
-lemma has_integral_alt:
- "(f has_integral y) i \<longleftrightarrow>
-      (if (\<exists>a b. i = {a..b}) then (f has_integral y) i
-       else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b}
-                               \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0)
-                                        has_integral z) ({a..b}) \<and>
-                                       norm(z - y) < e)))"
-  unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto
-
-lemma has_integral_altD:
-  assumes "(f has_integral y) i" "\<not> (\<exists>a b. i = {a..b})" "e>0"
-  obtains B where "B>0" "\<forall>a b. ball 0 B \<subseteq> {a..b}\<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) ({a..b}) \<and> norm(z - y) < e)"
-  using assms unfolding has_integral unfolding has_integral_compact_interval_def has_integral_def by auto
-
-definition integrable_on (infixr "integrable'_on" 46) where
-  "(f integrable_on i) \<equiv> \<exists>y. (f has_integral y) i"
-
-definition "integral i f \<equiv> SOME y. (f has_integral y) i"
-
-lemma integrable_integral[dest]:
- "f integrable_on i \<Longrightarrow> (f has_integral (integral i f)) i"
-  unfolding integrable_on_def integral_def by(rule someI_ex)
-
-lemma has_integral_integrable[intro]: "(f has_integral i) s \<Longrightarrow> f integrable_on s"
-  unfolding integrable_on_def by auto
-
-lemma has_integral_integral:"f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s"
-  by auto
-
-lemma has_integral_vec1: assumes "(f has_integral k) {a..b}"
-  shows "((\<lambda>x. vec1 (f x)) has_integral (vec1 k)) {a..b}"
-proof- have *:"\<And>p. (\<Sum>(x, k)\<in>p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k)"
-    unfolding vec_sub Cart_eq by(auto simp add:vec1_dest_vec1_simps split_beta)
-  show ?thesis using assms unfolding has_integral apply safe
-    apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe)
-    apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed
-
-lemma setsum_content_null:
-  assumes "content({a..b}) = 0" "p tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
-proof(rule setsum_0',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) guess c .. then guess d .. note c_d=this
-  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 c_d]] content_pos_le[of c d]
-    unfolding assms(1) c_d by auto
-  finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
-qed
-
-subsection {* Some basic combining lemmas. *}
-
-lemma tagged_division_unions_exists:
-  assumes "finite iset" "\<forall>i \<in> iset. \<exists>p. p tagged_division_of i \<and> d fine p"
-  "\<forall>i1\<in>iset. \<forall>i2\<in>iset. ~(i1 = i2) \<longrightarrow> (interior(i1) \<inter> interior(i2) = {})" "(\<Union>iset = i)"
-   obtains p where "p tagged_division_of i" "d fine p"
-proof- guess pfn using bchoice[OF assms(2)] .. note pfn = conjunctD2[OF this[rule_format]]
-  show thesis apply(rule_tac p="\<Union>(pfn ` iset)" in that) unfolding assms(4)[THEN sym]
-    apply(rule tagged_division_unions[OF assms(1) _ assms(3)]) defer 
-    apply(rule fine_unions) using pfn by auto
-qed
-
-subsection {* The set we're concerned with must be closed. *}
-
-lemma division_of_closed: "s division_of i \<Longrightarrow> closed (i::(real^'n) set)"
-  unfolding division_of_def by(fastsimp intro!: closed_Union closed_interval)
-
-subsection {* General bisection principle for intervals; might be useful elsewhere. *}
-
-lemma interval_bisection_step:
-  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "~(P {a..b::real^'n})"
-  obtains c d where "~(P{c..d})"
-  "\<forall>i. a$i \<le> c$i \<and> c$i \<le> d$i \<and> d$i \<le> b$i \<and> 2 * (d$i - c$i) \<le> b$i - a$i"
-proof- have "{a..b} \<noteq> {}" using assms(1,3) by auto
-  note ab=this[unfolded interval_eq_empty not_ex not_less]
-  { fix f have "finite f \<Longrightarrow>
-        (\<forall>s\<in>f. P s) \<Longrightarrow>
-        (\<forall>s\<in>f. \<exists>a b. s = {a..b}) \<Longrightarrow>
-        (\<forall>s\<in>f.\<forall>t\<in>f. ~(s = t) \<longrightarrow> interior(s) \<inter> interior(t) = {}) \<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])
-        apply rule defer apply rule defer apply(rule inter_interior_unions_intervals)
-        using insert by auto
-    qed } note * = this
-  let ?A = "{{c..d} | c d. \<forall>i. (c$i = a$i) \<and> (d$i = (a$i + b$i) / 2) \<or> (c$i = (a$i + b$i) / 2) \<and> (d$i = b$i)}"
-  let ?PP = "\<lambda>c d. \<forall>i. a$i \<le> c$i \<and> c$i \<le> d$i \<and> d$i \<le> b$i \<and> 2 * (d$i - c$i) \<le> b$i - a$i"
-  { presume "\<forall>c d. ?PP c d \<longrightarrow> P {c..d} \<Longrightarrow> False"
-    thus thesis unfolding atomize_not not_all apply-apply(erule exE)+ apply(rule_tac c=x and d=xa in that) by auto }
-  assume as:"\<forall>c d. ?PP c d \<longrightarrow> P {c..d}"
-  have "P (\<Union> ?A)" proof(rule *, rule_tac[2-] ballI, rule_tac[4] ballI, rule_tac[4] impI) 
-    let ?B = "(\<lambda>s.{(\<chi> i. if i \<in> s then a$i else (a$i + b$i) / 2) ..
-      (\<chi> i. if i \<in> s then (a$i + b$i) / 2 else b$i)}) ` {s. s \<subseteq> UNIV}"
-    have "?A \<subseteq> ?B" proof case goal1
-      then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format]
-      have *:"\<And>a b c d. a = c \<Longrightarrow> b = d \<Longrightarrow> {a..b} = {c..d}" by auto
-      show "x\<in>?B" unfolding image_iff apply(rule_tac x="{i. c$i = a$i}" in bexI)
-        unfolding c_d apply(rule * ) unfolding Cart_eq cond_component Cart_lambda_beta
-      proof(rule_tac[1-2] allI) fix i show "c $ i = (if i \<in> {i. c $ i = a $ i} then a $ i else (a $ i + b $ i) / 2)"
-          "d $ i = (if i \<in> {i. c $ i = a $ i} then (a $ i + b $ i) / 2 else b $ i)"
-          using c_d(2)[of i] ab[THEN spec[where x=i]] by(auto simp add:field_simps)
-      qed auto qed
-    thus "finite ?A" apply(rule finite_subset[of _ ?B]) by auto
-    fix s assume "s\<in>?A" then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+)
-    note c_d=this[rule_format]
-    show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 show ?case 
-        using c_d(2)[of i] using ab[THEN spec[where x=i]] by auto qed
-    show "\<exists>a b. s = {a..b}" unfolding c_d by auto
-    fix t assume "t\<in>?A" then guess e unfolding mem_Collect_eq .. then guess f apply- by(erule exE,(erule conjE)+)
-    note e_f=this[rule_format]
-    assume "s \<noteq> t" hence "\<not> (c = e \<and> d = f)" unfolding c_d e_f by auto
-    then obtain i where "c$i \<noteq> e$i \<or> d$i \<noteq> f$i" unfolding de_Morgan_conj Cart_eq by auto
-    hence i:"c$i \<noteq> e$i" "d$i \<noteq> f$i" apply- apply(erule_tac[!] disjE)
-    proof- assume "c$i \<noteq> e$i" thus "d$i \<noteq> f$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
-    next   assume "d$i \<noteq> f$i" thus "c$i \<noteq> e$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
-    qed 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 e_f c_d interior_closed_interval proof(rule *)
-      fix x assume "x\<in>{c<..<d}" "x\<in>{e<..<f}"
-      hence x:"c$i < d$i" "e$i < f$i" "c$i < f$i" "e$i < d$i" unfolding mem_interval apply-apply(erule_tac[!] x=i in allE)+ by auto
-      show False using c_d(2)[of i] apply- apply(erule_tac disjE)
-      proof(erule_tac[!] conjE) assume as:"c $ i = a $ i" "d $ i = (a $ i + b $ i) / 2"
-        show False using e_f(2)[of i] and i x unfolding as by(fastsimp simp add:field_simps)
-      next assume as:"c $ i = (a $ i + b $ i) / 2" "d $ i = b $ i"
-        show False using e_f(2)[of i] and i x unfolding as by(fastsimp simp add:field_simps)
-      qed qed qed
-  also have "\<Union> ?A = {a..b}" proof(rule set_ext,rule)
-    fix x assume "x\<in>\<Union>?A" then guess Y unfolding Union_iff ..
-    from this(1) guess c unfolding mem_Collect_eq .. then guess d ..
-    note c_d = this[THEN conjunct2,rule_format] `x\<in>Y`[unfolded this[THEN conjunct1]]
-    show "x\<in>{a..b}" unfolding mem_interval proof 
-      fix i show "a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
-        using c_d(1)[of i] c_d(2)[unfolded mem_interval,THEN spec[where x=i]] by auto qed
-  next fix x assume x:"x\<in>{a..b}"
-    have "\<forall>i. \<exists>c d. (c = a$i \<and> d = (a$i + b$i) / 2 \<or> c = (a$i + b$i) / 2 \<and> d = b$i) \<and> c\<le>x$i \<and> x$i \<le> d"
-      (is "\<forall>i. \<exists>c d. ?P i c d") unfolding mem_interval proof fix i
-      have "?P i (a$i) ((a $ i + b $ i) / 2) \<or> ?P i ((a $ i + b $ i) / 2) (b$i)"
-        using x[unfolded mem_interval,THEN spec[where x=i]] by auto thus "\<exists>c d. ?P i c d" by blast
-    qed thus "x\<in>\<Union>?A" unfolding Union_iff lambda_skolem unfolding Bex_def mem_Collect_eq
-      apply-apply(erule exE)+ apply(rule_tac x="{xa..xaa}" in exI) unfolding mem_interval by auto
-  qed finally show False using assms by auto qed
-
-lemma interval_bisection:
-  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "\<not> P {a..b::real^'n}"
-  obtains x where "x \<in> {a..b}" "\<forall>e>0. \<exists>c d. x \<in> {c..d} \<and> {c..d} \<subseteq> ball x e \<and> {c..d} \<subseteq> {a..b} \<and> ~P({c..d})"
-proof-
-  have "\<forall>x. \<exists>y. \<not> P {fst x..snd x} \<longrightarrow> (\<not> P {fst y..snd y} \<and> (\<forall>i. fst x$i \<le> fst y$i \<and> fst y$i \<le> snd y$i \<and> snd y$i \<le> snd x$i \<and>
-                           2 * (snd y$i - fst y$i) \<le> snd x$i - fst x$i))" proof case goal1 thus ?case proof-
-      presume "\<not> P {fst x..snd x} \<Longrightarrow> ?thesis"
-      thus ?thesis apply(cases "P {fst x..snd x}") by auto
-    next assume as:"\<not> P {fst x..snd x}" from interval_bisection_step[of P, OF assms(1-2) as] guess c d . 
-      thus ?thesis apply- apply(rule_tac x="(c,d)" in exI) by auto
-    qed qed then guess f apply-apply(drule choice) by(erule exE) note f=this
-  def AB \<equiv> "\<lambda>n. (f ^^ n) (a,b)" def A \<equiv> "\<lambda>n. fst(AB n)" and B \<equiv> "\<lambda>n. snd(AB n)" note ab_def = this AB_def
-  have "A 0 = a" "B 0 = b" "\<And>n. \<not> P {A(Suc n)..B(Suc n)} \<and>
-    (\<forall>i. A(n)$i \<le> A(Suc n)$i \<and> A(Suc n)$i \<le> B(Suc n)$i \<and> B(Suc n)$i \<le> B(n)$i \<and> 
-    2 * (B(Suc n)$i - A(Suc n)$i) \<le> B(n)$i - A(n)$i)" (is "\<And>n. ?P n")
-  proof- show "A 0 = a" "B 0 = b" unfolding ab_def by auto
-    case goal3 note S = ab_def funpow.simps o_def id_apply show ?case
-    proof(induct n) case 0 thus ?case unfolding S apply(rule f[rule_format]) using assms(3) by auto
-    next case (Suc n) show ?case unfolding S apply(rule f[rule_format]) using Suc unfolding S by auto
-    qed qed note AB = this(1-2) conjunctD2[OF this(3),rule_format]
-
-  have interv:"\<And>e. 0 < e \<Longrightarrow> \<exists>n. \<forall>x\<in>{A n..B n}. \<forall>y\<in>{A n..B n}. dist x y < e"
-  proof- case goal1 guess n using real_arch_pow2[of "(setsum (\<lambda>i. b$i - a$i) UNIV) / e"] .. note n=this
-    show ?case apply(rule_tac x=n in exI) proof(rule,rule)
-      fix x y assume xy:"x\<in>{A n..B n}" "y\<in>{A n..B n}"
-      have "dist x y \<le> setsum (\<lambda>i. abs((x - y)$i)) UNIV" unfolding vector_dist_norm by(rule norm_le_l1)
-      also have "\<dots> \<le> setsum (\<lambda>i. B n$i - A n$i) UNIV"
-      proof(rule setsum_mono) fix i show "\<bar>(x - y) $ i\<bar> \<le> B n $ i - A n $ i"
-          using xy[unfolded mem_interval,THEN spec[where x=i]]
-          unfolding vector_minus_component by auto qed
-      also have "\<dots> \<le> setsum (\<lambda>i. b$i - a$i) UNIV / 2^n" unfolding setsum_divide_distrib
-      proof(rule setsum_mono) case goal1 thus ?case
-        proof(induct n) case 0 thus ?case unfolding AB by auto
-        next case (Suc n) have "B (Suc n) $ i - A (Suc n) $ i \<le> (B n $ i - A n $ i) / 2" using AB(4)[of n i] by auto
-          also have "\<dots> \<le> (b $ i - a $ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case .
-        qed qed
-      also have "\<dots> < e" using n using goal1 by(auto simp add:field_simps) finally show "dist x y < e" .
-    qed qed
-  { fix n m ::nat assume "m \<le> n" then guess d unfolding le_Suc_ex_iff .. note d=this
-    have "{A n..B n} \<subseteq> {A m..B m}" unfolding d 
-    proof(induct d) case 0 thus ?case by auto
-    next case (Suc d) show ?case apply(rule subset_trans[OF _ Suc])
-        apply(rule) unfolding mem_interval apply(rule,erule_tac x=i in allE)
-      proof- case goal1 thus ?case using AB(4)[of "m + d" i] by(auto simp add:field_simps)
-      qed qed } note ABsubset = this 
-  have "\<exists>a. \<forall>n. a\<in>{A n..B n}" apply(rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv])
-  proof- fix n show "{A n..B n} \<noteq> {}" apply(cases "0<n") using AB(3)[of "n - 1"] assms(1,3) AB(1-2) by auto qed auto
-  then guess x0 .. note x0=this[rule_format]
-  show thesis proof(rule that[rule_format,of x0])
-    show "x0\<in>{a..b}" using x0[of 0] unfolding AB .
-    fix e assume "0 < (e::real)" from interv[OF this] guess n .. note n=this
-    show "\<exists>c d. x0 \<in> {c..d} \<and> {c..d} \<subseteq> ball x0 e \<and> {c..d} \<subseteq> {a..b} \<and> \<not> P {c..d}"
-      apply(rule_tac x="A n" in exI,rule_tac x="B n" in exI) apply(rule,rule x0) apply rule defer 
-    proof show "\<not> P {A n..B n}" apply(cases "0<n") using AB(3)[of "n - 1"] assms(3) AB(1-2) by auto
-      show "{A n..B n} \<subseteq> ball x0 e" using n using x0[of n] by auto
-      show "{A n..B n} \<subseteq> {a..b}" unfolding AB(1-2)[symmetric] apply(rule ABsubset) by auto
-    qed qed qed 
-
-subsection {* Cousin's lemma. *}
-
-lemma fine_division_exists: assumes "gauge g" 
-  obtains p where "p tagged_division_of {a..b::real^'n}" "g fine p"
-proof- presume "\<not> (\<exists>p. p tagged_division_of {a..b} \<and> g fine p) \<Longrightarrow> False"
-  then guess p unfolding atomize_not not_not .. thus thesis apply-apply(rule that[of p]) by auto
-next assume as:"\<not> (\<exists>p. p tagged_division_of {a..b} \<and> g fine p)"
-  guess x apply(rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p",rule_format,OF _ _ as])
-    apply(rule_tac x="{}" in exI) defer apply(erule conjE exE)+
-  proof- show "{} tagged_division_of {} \<and> g fine {}" unfolding fine_def by auto
-    fix s t p p' assume "p tagged_division_of s" "g fine p" "p' tagged_division_of t" "g fine p'" "interior s \<inter> interior t = {}"
-    thus "\<exists>p. p tagged_division_of s \<union> t \<and> g fine p" apply-apply(rule_tac x="p \<union> p'" in exI) apply rule
-      apply(rule tagged_division_union) prefer 4 apply(rule fine_union) by auto
-  qed note x=this
-  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)] guess c d apply-apply(erule exE conjE)+ . note c_d = this
-  have "g fine {(x, {c..d})}" unfolding fine_def using e using c_d(2) by auto
-  thus False using tagged_division_of_self[OF c_d(1)] using c_d by auto qed
-
-subsection {* Basic theorems about integrals. *}
-
-lemma has_integral_unique: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2"
-proof(rule ccontr) let ?e = "norm(k1 - k2) / 2" assume as:"k1 \<noteq> k2" hence e:"?e > 0" by auto
-  have lem:"\<And>f::real^'n \<Rightarrow> 'a.  \<And> a b k1 k2.
-    (f has_integral k1) ({a..b}) \<Longrightarrow> (f has_integral k2) ({a..b}) \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> False"
-  proof- case goal1 let ?e = "norm(k1 - k2) / 2" from goal1(3) have e:"?e > 0" by auto
-    guess d1 by(rule has_integralD[OF goal1(1) e]) note d1=this
-    guess d2 by(rule has_integralD[OF goal1(2) e]) note d2=this
-    guess p by(rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)],of a b]) note p=this
-    let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
-      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:group_simps norm_minus_commute)
-    also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
-      apply(rule add_strict_mono) apply(rule_tac[!] d2(2) d1(2)) using p unfolding fine_def by auto
-    finally show False by auto
-  qed { presume "\<not> (\<exists>a b. i = {a..b}) \<Longrightarrow> False"
-    thus False apply-apply(cases "\<exists>a b. i = {a..b}")
-      using assms by(auto simp add:has_integral intro:lem[OF _ _ as]) }
-  assume as:"\<not> (\<exists>a b. i = {a..b})"
-  guess B1 by(rule has_integral_altD[OF assms(1) as,OF e]) note B1=this[rule_format]
-  guess B2 by(rule has_integral_altD[OF assms(2) as,OF e]) note B2=this[rule_format]
-  have "\<exists>a b::real^'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> {a..b}" apply(rule bounded_subset_closed_interval)
-    using bounded_Un bounded_ball by auto then guess a b apply-by(erule exE)+
-  note ab=conjunctD2[OF this[unfolded Un_subset_iff]]
-  guess w using B1(2)[OF ab(1)] .. note w=conjunctD2[OF this]
-  guess z using B2(2)[OF ab(2)] .. note z=conjunctD2[OF this]
-  have "z = w" using lem[OF w(1) z(1)] by auto
-  hence "norm (k1 - k2) \<le> norm (z - k2) + norm (w - k1)"
-    using norm_triangle_ineq4[of "k1 - w" "k2 - z"] by(auto simp add: norm_minus_commute) 
-  also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" apply(rule add_strict_mono) by(rule_tac[!] z(2) w(2))
-  finally show False by auto qed
-
-lemma integral_unique[intro]:
-  "(f has_integral y) k \<Longrightarrow> integral k f = y"
-  unfolding integral_def apply(rule some_equality) by(auto intro: has_integral_unique) 
-
-lemma has_integral_is_0: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector" 
-  assumes "\<forall>x\<in>s. f x = 0" shows "(f has_integral 0) s"
-proof- have lem:"\<And>a b. \<And>f::real^'n \<Rightarrow> 'a.
-    (\<forall>x\<in>{a..b}. f(x) = 0) \<Longrightarrow> (f has_integral 0) ({a..b})" unfolding has_integral
-  proof(rule,rule) fix a b e and f::"real^'n \<Rightarrow> 'a"
-    assume as:"\<forall>x\<in>{a..b}. f x = 0" "0 < (e::real)"
-    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e)"
-      apply(rule_tac x="\<lambda>x. ball x 1" in exI)  apply(rule,rule gaugeI) unfolding centre_in_ball defer apply(rule open_ball)
-    proof(rule,rule,erule conjE) case goal1
-      have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) = 0" proof(rule setsum_0',rule)
-        fix x assume x:"x\<in>p" have "f (fst x) = 0" using tagged_division_ofD(2-3)[OF goal1(1), of "fst x" "snd x"] using as x by auto
-        thus "(\<lambda>(x, k). content k *\<^sub>R f x) x = 0" apply(subst surjective_pairing[of x]) unfolding split_conv by auto
-      qed thus ?case using as by auto
-    qed auto qed  { presume "\<not> (\<exists>a b. s = {a..b}) \<Longrightarrow> ?thesis"
-    thus ?thesis apply-apply(cases "\<exists>a b. s = {a..b}")
-      using assms by(auto simp add:has_integral intro:lem) }
-  have *:"(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. 0)" apply(rule ext) using assms by auto
-  assume "\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P *
-  apply(rule,rule,rule_tac x=1 in exI,rule) defer apply(rule,rule,rule)
-  proof- fix e::real and a b assume "e>0"
-    thus "\<exists>z. ((\<lambda>x::real^'n. 0::'a) has_integral z) {a..b} \<and> norm (z - 0) < e"
-      apply(rule_tac x=0 in exI) apply(rule,rule lem) by auto
-  qed auto qed
-
-lemma has_integral_0[simp]: "((\<lambda>x::real^'n. 0) has_integral 0) s"
-  apply(rule has_integral_is_0) by auto 
-
-lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) s \<longleftrightarrow> i = 0"
-  using has_integral_unique[OF has_integral_0] by auto
-
-lemma has_integral_linear: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral y) s" "bounded_linear h" shows "((h o f) has_integral ((h y))) s"
-proof- interpret bounded_linear h using assms(2) . from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format]
-  have lem:"\<And>f::real^'n \<Rightarrow> 'a. \<And> y a b.
-    (f has_integral y) ({a..b}) \<Longrightarrow> ((h o f) has_integral h(y)) ({a..b})"
-  proof(subst has_integral,rule,rule) case goal1
-    from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format]
-    have *:"e / B > 0" apply(rule divide_pos_pos) using goal1(2) B by auto
-    guess g using has_integralD[OF goal1(1) *] . note g=this
-    show ?case apply(rule_tac x=g in exI) apply(rule,rule g(1))
-    proof(rule,rule,erule conjE) fix p assume as:"p tagged_division_of {a..b}" "g fine p" 
-      have *:"\<And>x k. h ((\<lambda>(x, k). content k *\<^sub>R f x) x) = (\<lambda>(x, k). h (content k *\<^sub>R f x)) x" by auto
-      have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = setsum (h \<circ> (\<lambda>(x, k). content k *\<^sub>R f x)) p"
-        unfolding o_def unfolding scaleR[THEN sym] * by simp
-      also have "\<dots> = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" using setsum[of "\<lambda>(x,k). content k *\<^sub>R f x" p] using as by auto
-      finally have *:"(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" .
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) - h y) < e" unfolding * diff[THEN sym]
-        apply(rule le_less_trans[OF B(2)]) using g(2)[OF as] B(1) by(auto simp add:field_simps)
-    qed qed { presume "\<not> (\<exists>a b. s = {a..b}) \<Longrightarrow> ?thesis"
-    thus ?thesis apply-apply(cases "\<exists>a b. s = {a..b}") using assms by(auto simp add:has_integral intro!:lem) }
-  assume as:"\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P
-  proof(rule,rule) fix e::real  assume e:"0<e"
-    have *:"0 < e/B" by(rule divide_pos_pos,rule e,rule B(1))
-    guess M using has_integral_altD[OF assms(1) as *,rule_format] . note M=this
-    show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b} \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) has_integral z) {a..b} \<and> norm (z - h y) < e)"
-      apply(rule_tac x=M in exI) apply(rule,rule M(1))
-    proof(rule,rule,rule) case goal1 guess z using M(2)[OF goal1(1)] .. note z=conjunctD2[OF this]
-      have *:"(\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) = h \<circ> (\<lambda>x. if x \<in> s then f x else 0)"
-        unfolding o_def apply(rule ext) using zero by auto
-      show ?case apply(rule_tac x="h z" in exI,rule) unfolding * apply(rule lem[OF z(1)]) unfolding diff[THEN sym]
-        apply(rule le_less_trans[OF B(2)]) using B(1) z(2) by(auto simp add:field_simps)
-    qed qed qed
-
-lemma has_integral_cmul:
-  shows "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s"
-  unfolding o_def[THEN sym] apply(rule has_integral_linear,assumption)
-  by(rule scaleR.bounded_linear_right)
-
-lemma has_integral_neg:
-  shows "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral (-k)) s"
-  apply(drule_tac c="-1" in has_integral_cmul) by auto
-
-lemma has_integral_add: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector" 
-  assumes "(f has_integral k) s" "(g has_integral l) s"
-  shows "((\<lambda>x. f x + g x) has_integral (k + l)) s"
-proof- have lem:"\<And>f g::real^'n \<Rightarrow> 'a. \<And>a b k l.
-    (f has_integral k) ({a..b}) \<Longrightarrow> (g has_integral l) ({a..b}) \<Longrightarrow>
-     ((\<lambda>x. f(x) + g(x)) has_integral (k + l)) ({a..b})" proof- case goal1
-    show ?case unfolding has_integral proof(rule,rule) fix e::real assume e:"e>0" hence *:"e/2>0" by auto
-      guess d1 using has_integralD[OF goal1(1) *] . note d1=this
-      guess d2 using has_integralD[OF goal1(2) *] . note d2=this
-      show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)"
-        apply(rule_tac x="\<lambda>x. (d1 x) \<inter> (d2 x)" in exI) apply(rule,rule gauge_inter[OF d1(1) d2(1)])
-      proof(rule,rule,erule conjE) fix p assume as:"p tagged_division_of {a..b}" "(\<lambda>x. d1 x \<inter> d2 x) fine p"
-        have *:"(\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) = (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p. content k *\<^sub>R g x)"
-          unfolding scaleR_right_distrib setsum_addf[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,THEN sym]
-          by(rule setsum_cong2,auto)
-        have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) = norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
-          unfolding * by(auto simp add:group_simps) also let ?res = "\<dots>"
-        from as have *:"d1 fine p" "d2 fine p" unfolding fine_inter by auto
-        have "?res < e/2 + e/2" apply(rule le_less_trans[OF norm_triangle_ineq])
-          apply(rule add_strict_mono) using d1(2)[OF as(1) *(1)] and d2(2)[OF as(1) *(2)] by auto
-        finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e" by auto
-      qed qed qed { presume "\<not> (\<exists>a b. s = {a..b}) \<Longrightarrow> ?thesis"
-    thus ?thesis apply-apply(cases "\<exists>a b. s = {a..b}") using assms by(auto simp add:has_integral intro!:lem) }
-  assume as:"\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P
-  proof(rule,rule) case goal1 hence *:"e/2 > 0" by auto
-    from has_integral_altD[OF assms(1) as *] guess B1 . note B1=this[rule_format]
-    from has_integral_altD[OF assms(2) as *] guess B2 . note B2=this[rule_format]
-    show ?case apply(rule_tac x="max B1 B2" in exI) apply(rule,rule min_max.less_supI1,rule B1)
-    proof(rule,rule,rule) fix a b assume "ball 0 (max B1 B2) \<subseteq> {a..b::real^'n}"
-      hence *:"ball 0 B1 \<subseteq> {a..b::real^'n}" "ball 0 B2 \<subseteq> {a..b::real^'n}" by auto
-      guess w using B1(2)[OF *(1)] .. note w=conjunctD2[OF this]
-      guess z using B2(2)[OF *(2)] .. note z=conjunctD2[OF this]
-      have *:"\<And>x. (if x \<in> s then f x + g x else 0) = (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)" by auto
-      show "\<exists>z. ((\<lambda>x. if x \<in> s then f x + g x else 0) has_integral z) {a..b} \<and> norm (z - (k + l)) < e"
-        apply(rule_tac x="w + z" in exI) apply(rule,rule lem[OF w(1) z(1), unfolded *[THEN sym]])
-        using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) by(auto simp add:field_simps)
-    qed qed qed
-
-lemma has_integral_sub:
-  shows "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_integral (k - l)) s"
-  using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding group_simps by auto
-
-lemma integral_0: "integral s (\<lambda>x::real^'n. 0::real^'m) = 0"
-  by(rule integral_unique has_integral_0)+
-
-lemma integral_add:
-  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow>
-   integral s (\<lambda>x. f x + g x) = integral s f + integral s g"
-  apply(rule integral_unique) apply(drule integrable_integral)+
-  apply(rule has_integral_add) by assumption+
-
-lemma integral_cmul:
-  shows "f integrable_on s \<Longrightarrow> integral s (\<lambda>x. c *\<^sub>R f x) = c *\<^sub>R integral s f"
-  apply(rule integral_unique) apply(drule integrable_integral)+
-  apply(rule has_integral_cmul) by assumption+
-
-lemma integral_neg:
-  shows "f integrable_on s \<Longrightarrow> integral s (\<lambda>x. - f x) = - integral s f"
-  apply(rule integral_unique) apply(drule integrable_integral)+
-  apply(rule has_integral_neg) by assumption+
-
-lemma integral_sub:
-  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> integral s (\<lambda>x. f x - g x) = integral s f - integral s g"
-  apply(rule integral_unique) apply(drule integrable_integral)+
-  apply(rule has_integral_sub) by assumption+
-
-lemma integrable_0: "(\<lambda>x. 0) integrable_on s"
-  unfolding integrable_on_def using has_integral_0 by auto
-
-lemma integrable_add:
-  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x + g x) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_add)
-
-lemma integrable_cmul:
-  shows "f integrable_on s \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_cmul)
-
-lemma integrable_neg:
-  shows "f integrable_on s \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_neg)
-
-lemma integrable_sub:
-  shows "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_sub)
-
-lemma integrable_linear:
-  shows "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h o f) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_linear)
-
-lemma integral_linear:
-  shows "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> integral s (h o f) = h(integral s f)"
-  apply(rule has_integral_unique) defer unfolding has_integral_integral 
-  apply(drule has_integral_linear,assumption,assumption) unfolding has_integral_integral[THEN sym]
-  apply(rule integrable_linear) by assumption+
-
-lemma has_integral_setsum:
-  assumes "finite t" "\<forall>a\<in>t. ((f a) has_integral (i a)) s"
-  shows "((\<lambda>x. setsum (\<lambda>a. f a x) t) has_integral (setsum i t)) s"
-proof(insert assms(1) subset_refl[of t],induct rule:finite_subset_induct)
-  case (insert x F) show ?case unfolding setsum_insert[OF insert(1,3)]
-    apply(rule has_integral_add) using insert assms by auto
-qed auto
-
-lemma integral_setsum:
-  shows "finite t \<Longrightarrow> \<forall>a\<in>t. (f a) integrable_on s \<Longrightarrow>
-  integral s (\<lambda>x. setsum (\<lambda>a. f a x) t) = setsum (\<lambda>a. integral s (f a)) t"
-  apply(rule integral_unique) apply(rule has_integral_setsum)
-  using integrable_integral by auto
-
-lemma integrable_setsum:
-  shows "finite t \<Longrightarrow> \<forall>a \<in> t.(f a) integrable_on s \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) t) integrable_on s"
-  unfolding integrable_on_def apply(drule bchoice) using has_integral_setsum[of t] by auto
-
-lemma has_integral_eq:
-  assumes "\<forall>x\<in>s. f x = g x" "(f has_integral k) s" shows "(g has_integral k) s"
-  using has_integral_sub[OF assms(2), of "\<lambda>x. f x - g x" 0]
-  using has_integral_is_0[of s "\<lambda>x. f x - g x"] using assms(1) by auto
-
-lemma integrable_eq:
-  shows "\<forall>x\<in>s. f x = g x \<Longrightarrow> f integrable_on s \<Longrightarrow> g integrable_on s"
-  unfolding integrable_on_def using has_integral_eq[of s f g] by auto
-
-lemma has_integral_eq_eq:
-  shows "\<forall>x\<in>s. f x = g x \<Longrightarrow> ((f has_integral k) s \<longleftrightarrow> (g has_integral k) s)"
-  using has_integral_eq[of s f g] has_integral_eq[of s g f] by auto 
-
-lemma has_integral_null[dest]:
-  assumes "content({a..b}) = 0" shows  "(f has_integral 0) ({a..b})"
-  unfolding has_integral apply(rule,rule,rule_tac x="\<lambda>x. ball x 1" in exI,rule) defer
-proof(rule,rule,erule conjE) fix e::real assume e:"e>0" thus "gauge (\<lambda>x. ball x 1)" by auto
-  fix p assume p:"p tagged_division_of {a..b}" (*"(\<lambda>x. ball x 1) fine p"*)
-  have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) = 0" unfolding norm_eq_zero diff_0_right
-    using setsum_content_null[OF assms(1) p, of f] . 
-  thus "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e" using e by auto qed
-
-lemma has_integral_null_eq[simp]:
-  shows "content({a..b}) = 0 \<Longrightarrow> ((f has_integral i) ({a..b}) \<longleftrightarrow> i = 0)"
-  apply rule apply(rule has_integral_unique,assumption) 
-  apply(drule has_integral_null,assumption)
-  apply(drule has_integral_null) by auto
-
-lemma integral_null[dest]: shows "content({a..b}) = 0 \<Longrightarrow> integral({a..b}) f = 0"
-  by(rule integral_unique,drule has_integral_null)
-
-lemma integrable_on_null[dest]: shows "content({a..b}) = 0 \<Longrightarrow> f integrable_on {a..b}"
-  unfolding integrable_on_def apply(drule has_integral_null) by auto
-
-lemma has_integral_empty[intro]: shows "(f has_integral 0) {}"
-  unfolding empty_as_interval apply(rule has_integral_null) 
-  using content_empty unfolding empty_as_interval .
-
-lemma has_integral_empty_eq[simp]: shows "(f has_integral i) {} \<longleftrightarrow> i = 0"
-  apply(rule,rule has_integral_unique,assumption) by auto
-
-lemma integrable_on_empty[intro]: shows "f integrable_on {}" unfolding integrable_on_def by auto
-
-lemma integral_empty[simp]: shows "integral {} f = 0"
-  apply(rule integral_unique) using has_integral_empty .
-
-lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}"
-  apply(rule has_integral_null) unfolding content_eq_0_interior
-  unfolding interior_closed_interval using interval_sing by auto
-
-lemma integrable_on_refl[intro]: shows "f integrable_on {a..a}" unfolding integrable_on_def by auto
-
-lemma integral_refl: shows "integral {a..a} f = 0" apply(rule integral_unique) by auto
-
-subsection {* Cauchy-type criterion for integrability. *}
-
-lemma integrable_cauchy: fixes f::"real^'n \<Rightarrow> 'a::{real_normed_vector,complete_space}" 
-  shows "f integrable_on {a..b} \<longleftrightarrow>
-  (\<forall>e>0.\<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<and> d fine p1 \<and>
-                            p2 tagged_division_of {a..b} \<and> d fine p2
-                            \<longrightarrow> norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 -
-                                     setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) < e))" (is "?l = (\<forall>e>0. \<exists>d. ?P e d)")
-proof assume ?l
-  then guess y unfolding integrable_on_def has_integral .. note y=this
-  show "\<forall>e>0. \<exists>d. ?P e d" proof(rule,rule) case goal1 hence "e/2 > 0" by auto
-    then guess d apply- apply(drule y[rule_format]) by(erule exE,erule conjE) note d=this[rule_format]
-    show ?case apply(rule_tac x=d in exI,rule,rule d) apply(rule,rule,rule,(erule conjE)+)
-    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b}" "d fine p1" "p2 tagged_division_of {a..b}" "d fine p2"
-      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
-        apply(rule dist_triangle_half_l[where y=y,unfolded vector_dist_norm])
-        using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] .
-    qed qed
-next assume "\<forall>e>0. \<exists>d. ?P e d" hence "\<forall>n::nat. \<exists>d. ?P (inverse(real (n + 1))) d" by auto
-  from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
-  have "\<And>n. gauge (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}})" apply(rule gauge_inters) using d(1) by auto
-  hence "\<forall>n. \<exists>p. p tagged_division_of {a..b} \<and> (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}}) fine p" apply-
-  proof case goal1 from this[of n] show ?case apply(drule_tac fine_division_exists) by auto qed
-  from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]]
-  have dp:"\<And>i n. i\<le>n \<Longrightarrow> d i fine p n" using p(2) unfolding fine_inters by auto
-  have "Cauchy (\<lambda>n. setsum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))"
-  proof(rule CauchyI) case goal1 then guess N unfolding real_arch_inv[of e] .. note N=this
-    show ?case apply(rule_tac x=N in exI)
-    proof(rule,rule,rule,rule) fix m n assume mn:"N \<le> m" "N \<le> n" have *:"N = (N - 1) + 1" using N by auto
-      show "norm ((\<Sum>(x, k)\<in>p m. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p n. content k *\<^sub>R f x)) < e"
-        apply(rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) apply(subst *) apply(rule d(2))
-        using dp p(1) using mn by auto 
-    qed qed
-  then guess y unfolding convergent_eq_cauchy[THEN sym] .. note y=this[unfolded Lim_sequentially,rule_format]
-  show ?l unfolding integrable_on_def has_integral apply(rule_tac x=y in exI)
-  proof(rule,rule) fix e::real assume "e>0" hence *:"e/2 > 0" by auto
-    then guess N1 unfolding real_arch_inv[of "e/2"] .. note N1=this hence N1':"N1 = N1 - 1 + 1" by auto
-    guess N2 using y[OF *] .. note N2=this
-    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e)"
-      apply(rule_tac x="d (N1 + N2)" in exI) apply rule defer 
-    proof(rule,rule,erule conjE) show "gauge (d (N1 + N2))" using d by auto
-      fix q assume as:"q tagged_division_of {a..b}" "d (N1 + N2) fine q"
-      have *:"inverse (real (N1 + N2 + 1)) < e / 2" apply(rule less_trans) using N1 by auto
-      show "norm ((\<Sum>(x, k)\<in>q. content k *\<^sub>R f x) - y) < e" apply(rule norm_triangle_half_r)
-        apply(rule less_trans[OF _ *]) apply(subst N1', rule d(2)[of "p (N1+N2)"]) defer
-        using N2[rule_format,unfolded vector_dist_norm,of "N1+N2"]
-        using as dp[of "N1 - 1 + 1 + N2" "N1 + N2"] using p(1)[of "N1 + N2"] using N1 by auto qed qed qed
-
-subsection {* Additivity of integral on abutting intervals. *}
-
-lemma interval_split:
-  "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
-  "{a..b} \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
-  apply(rule_tac[!] set_ext) unfolding Int_iff mem_interval mem_Collect_eq
-  unfolding Cart_lambda_beta by auto
-
-lemma content_split:
-  "content {a..b::real^'n} = content({a..b} \<inter> {x. x$k \<le> c}) + content({a..b} \<inter> {x. x$k >= c})"
-proof- note simps = interval_split content_closed_interval_cases Cart_lambda_beta vector_le_def
-  { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps by auto }
-  have *:"UNIV = insert k (UNIV - {k})" "\<And>x. finite (UNIV-{x::'n})" "\<And>x. x\<notin>UNIV-{x}" by auto
-  have *:"\<And>X Y Z. (\<Prod>i\<in>UNIV. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>UNIV-{k}. Z i (Y i))"
-    "(\<Prod>i\<in>UNIV. b$i - a$i) = (\<Prod>i\<in>UNIV-{k}. b$i - a$i) * (b$k - a$k)" 
-    apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto
-  assume as:"a\<le>b" moreover have "\<And>x. min (b $ k) c = max (a $ k) c
-    \<Longrightarrow> x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)"
-    by  (auto simp add:field_simps)
-  moreover have "\<not> a $ k \<le> c \<Longrightarrow> \<not> c \<le> b $ k \<Longrightarrow> False"
-    unfolding not_le using as[unfolded vector_le_def,rule_format,of k] by auto
-  ultimately show ?thesis 
-    unfolding simps unfolding *(1)[of "\<lambda>i x. b$i - x"] *(1)[of "\<lambda>i x. x - a$i"] *(2) by(auto)
-qed
-
-lemma division_split_left_inj:
-  assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2"
-  "k1 \<inter> {x::real^'n. x$k \<le> c} = k2 \<inter> {x. x$k \<le> c}"
-  shows "content(k1 \<inter> {x. x$k \<le> c}) = 0"
-proof- note d=division_ofD[OF assms(1)]
-  have *:"\<And>a b::real^'n. \<And> c k. (content({a..b} \<inter> {x. x$k \<le> c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$k \<le> c}) = {})"
-    unfolding interval_split content_eq_0_interior by auto
-  guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
-  guess u2 v2 using d(4)[OF assms(3)] apply-by(erule 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)]])
-    defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
-
-lemma division_split_right_inj:
-  assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2"
-  "k1 \<inter> {x::real^'n. x$k \<ge> c} = k2 \<inter> {x. x$k \<ge> c}"
-  shows "content(k1 \<inter> {x. x$k \<ge> c}) = 0"
-proof- note d=division_ofD[OF assms(1)]
-  have *:"\<And>a b::real^'n. \<And> c k. (content({a..b} \<inter> {x. x$k >= c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$k >= c}) = {})"
-    unfolding interval_split content_eq_0_interior by auto
-  guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
-  guess u2 v2 using d(4)[OF assms(3)] apply-by(erule 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)]])
-    defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
-
-lemma tagged_division_split_left_inj:
-  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$k \<le> c} = k2 \<inter> {x. x$k \<le> c}" 
-  shows "content(k1 \<inter> {x. x$k \<le> c}) = 0"
-proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
-  show ?thesis apply(rule division_split_left_inj[OF division_of_tagged_division[OF assms(1)]])
-    apply(rule_tac[1-2] *) using assms(2-) by auto qed
-
-lemma tagged_division_split_right_inj:
-  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$k \<ge> c} = k2 \<inter> {x. x$k \<ge> c}" 
-  shows "content(k1 \<inter> {x. x$k \<ge> c}) = 0"
-proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
-  show ?thesis apply(rule division_split_right_inj[OF division_of_tagged_division[OF assms(1)]])
-    apply(rule_tac[1-2] *) using assms(2-) by auto qed
-
-lemma division_split:
-  assumes "p division_of {a..b::real^'n}"
-  shows "{l \<inter> {x. x$k \<le> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$k \<le> c} = {})} division_of ({a..b} \<inter> {x. x$k \<le> c})" (is "?p1 division_of ?I1") and 
-        "{l \<inter> {x. x$k \<ge> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$k \<ge> c} = {})} division_of ({a..b} \<inter> {x. x$k \<ge> c})" (is "?p2 division_of ?I2")
-proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms]
-  show "finite ?p1" "finite ?p2" using p(1) by auto show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2" unfolding p(6)[THEN sym] by auto
-  { fix k assume "k\<in>?p1" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this
-    guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this
-    show "k\<subseteq>?I1" "k \<noteq> {}" "\<exists>a b. k = {a..b}" unfolding l
-      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto
-    fix k' assume "k'\<in>?p1" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this
-    assume "k\<noteq>k'" thus "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 apply-by(erule exE,(erule conjE)+) note l=this
-    guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this
-    show "k\<subseteq>?I2" "k \<noteq> {}" "\<exists>a b. k = {a..b}" unfolding l
-      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto
-    fix k' assume "k'\<in>?p2" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this
-    assume "k\<noteq>k'" thus "interior k \<inter> interior k' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto }
-qed
-
-lemma has_integral_split: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral i) ({a..b} \<inter> {x. x$k \<le> c})"  "(f has_integral j) ({a..b} \<inter> {x. x$k \<ge> c})"
-  shows "(f has_integral (i + j)) ({a..b})"
-proof(unfold has_integral,rule,rule) case goal1 hence e:"e/2>0" by auto
-  guess d1 using has_integralD[OF assms(1)[unfolded interval_split] e] . note d1=this[unfolded interval_split[THEN sym]]
-  guess d2 using has_integralD[OF assms(2)[unfolded interval_split] e] . note d2=this[unfolded interval_split[THEN sym]]
-  let ?d = "\<lambda>x. if x$k = c then (d1 x \<inter> d2 x) else ball x (abs(x$k - c)) \<inter> d1 x \<inter> d2 x"
-  show ?case apply(rule_tac x="?d" in exI,rule) defer apply(rule,rule,(erule conjE)+)
-  proof- show "gauge ?d" using d1(1) d2(1) unfolding gauge_def by auto
-    fix p assume "p tagged_division_of {a..b}" "?d fine p" note p = this tagged_division_ofD[OF this(1)]
-    have lem0:"\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$k \<le> c} = {}) \<Longrightarrow> x$k \<le> c"
-         "\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$k \<ge> c} = {}) \<Longrightarrow> x$k \<ge> c"
-    proof- fix x kk assume as:"(x,kk)\<in>p"
-      show "~(kk \<inter> {x. x$k \<le> c} = {}) \<Longrightarrow> x$k \<le> c"
-      proof(rule ccontr) case goal1
-        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $ k - c\<bar>"
-          using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
-        hence "\<exists>y. y \<in> ball x \<bar>x $ k - c\<bar> \<inter> {x. x $ k \<le> c}" using goal1(1) by blast 
-        then guess y .. hence "\<bar>x $ k - y $ k\<bar> < \<bar>x $ k - c\<bar>" "y$k \<le> c" apply-apply(rule le_less_trans)
-          using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:vector_dist_norm)
-        thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
-      qed
-      show "~(kk \<inter> {x. x$k \<ge> c} = {}) \<Longrightarrow> x$k \<ge> c"
-      proof(rule ccontr) case goal1
-        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $ k - c\<bar>"
-          using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
-        hence "\<exists>y. y \<in> ball x \<bar>x $ k - c\<bar> \<inter> {x. x $ k \<ge> c}" using goal1(1) by blast 
-        then guess y .. hence "\<bar>x $ k - y $ k\<bar> < \<bar>x $ k - c\<bar>" "y$k \<ge> c" apply-apply(rule le_less_trans)
-          using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:vector_dist_norm)
-        thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
-      qed
-    qed
-
-    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 lem2: "\<And>f s P f. finite s \<Longrightarrow> finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}"
-    proof- case goal1 thus ?case apply-apply(rule finite_subset[of _ "(\<lambda>(x,k). (x,f k)) ` s"]) by auto qed
-    have lem3: "\<And>g::(real ^ 'n \<Rightarrow> bool) \<Rightarrow> real ^ 'n \<Rightarrow> bool. finite p \<Longrightarrow>
-      setsum (\<lambda>(x,k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> ~(g k = {})}
-               = setsum (\<lambda>(x,k). content k *\<^sub>R f x) ((\<lambda>(x,k). (x,g k)) ` p)"
-      apply(rule setsum_mono_zero_left) prefer 3
-    proof fix g::"(real ^ 'n \<Rightarrow> bool) \<Rightarrow> real ^ 'n \<Rightarrow> bool" and i::"(real^'n) \<times> ((real^'n) set)"
-      assume "i \<in> (\<lambda>(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
-      then obtain x k where xk:"i=(x,g k)" "(x,k)\<in>p" "(x,g k) \<notin> {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}" by auto
-      have "content (g k) = 0" using xk using content_empty by auto
-      thus "(\<lambda>(x, k). content k *\<^sub>R f x) i = 0" unfolding xk split_conv by auto
-    qed auto
-    have lem4:"\<And>g. (\<lambda>(x,l). content (g l) *\<^sub>R f x) = (\<lambda>(x,l). content l *\<^sub>R f x) o (\<lambda>(x,l). (x,g l))" apply(rule ext) by auto
-
-    let ?M1 = "{(x,kk \<inter> {x. x$k \<le> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$k \<le> c} \<noteq> {}}"
-    have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2" apply(rule d1(2),rule tagged_division_ofI)
-      apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
-    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = {a..b} \<inter> {x. x$k \<le> c}" unfolding p(8)[THEN sym] by auto
-      fix x l assume xl:"(x,l)\<in>?M1"
-      then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note xl'=this
-      have "l' \<subseteq> d1 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
-      thus "l \<subseteq> d1 x" unfolding xl' by auto
-      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $ k \<le> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
-        using lem0(1)[OF xl'(3-4)] by auto
-      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k])
-      fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M1"
-      then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note yr'=this
-      assume as:"(x,l) \<noteq> (y,r)" show "interior l \<inter> interior r = {}"
-      proof(cases "l' = r' \<longrightarrow> x' = y'")
-        case False thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      next case True hence "l' \<noteq> r'" using as unfolding xl' yr' by auto
-        thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      qed qed moreover
-
-    let ?M2 = "{(x,kk \<inter> {x. x$k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$k \<ge> c} \<noteq> {}}" 
-    have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2" apply(rule d2(2),rule tagged_division_ofI)
-      apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
-    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = {a..b} \<inter> {x. x$k \<ge> c}" unfolding p(8)[THEN sym] by auto
-      fix x l assume xl:"(x,l)\<in>?M2"
-      then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note xl'=this
-      have "l' \<subseteq> d2 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
-      thus "l \<subseteq> d2 x" unfolding xl' by auto
-      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $ k \<ge> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
-        using lem0(2)[OF xl'(3-4)] by auto
-      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k])
-      fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M2"
-      then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note yr'=this
-      assume as:"(x,l) \<noteq> (y,r)" show "interior l \<inter> interior r = {}"
-      proof(cases "l' = r' \<longrightarrow> x' = y'")
-        case False thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      next case True hence "l' \<noteq> r'" using as unfolding xl' yr' by auto
-        thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      qed qed ultimately
-
-    have "norm (((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2"
-      apply- apply(rule norm_triangle_lt) by auto
-    also { have *:"\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'a) = 0" using scaleR_zero_left by auto
-      have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)
-       = (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)" by auto
-      also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x $ k \<le> c}) *\<^sub>R f x) + (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x $ k}) *\<^sub>R f x) - (i + j)"
-        unfolding lem3[OF p(3)] apply(subst setsum_reindex_nonzero[OF p(3)]) defer apply(subst setsum_reindex_nonzero[OF p(3)])
-        defer unfolding lem4[THEN sym] apply(rule refl) unfolding split_paired_all split_conv apply(rule_tac[!] *)
-      proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) by auto
-      next case   goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) by auto
-      qed also note setsum_addf[THEN sym]
-      also have *:"\<And>x. x\<in>p \<Longrightarrow> (\<lambda>(x, ka). content (ka \<inter> {x. x $ k \<le> c}) *\<^sub>R f x) x + (\<lambda>(x, ka). content (ka \<inter> {x. c \<le> x $ k}) *\<^sub>R f x) x
-        = (\<lambda>(x,ka). content ka *\<^sub>R f x) x" unfolding split_paired_all split_conv
-      proof- fix a b assume "(a,b) \<in> p" from p(6)[OF this] guess u v apply-by(erule exE)+ note uv=this
-        thus "content (b \<inter> {x. x $ k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x $ k}) *\<^sub>R f a = content b *\<^sub>R f a"
-          unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[of u v k c] by auto
-      qed note setsum_cong2[OF this]
-      finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x $ k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x $ k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
-        ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x $ k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x $ k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
-        (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)" by auto }
-    finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed
-
-subsection {* A sort of converse, integrability on subintervals. *}
-
-lemma tagged_division_union_interval:
-  assumes "p1 tagged_division_of ({a..b} \<inter> {x::real^'n. x$k \<le> (c::real)})"  "p2 tagged_division_of ({a..b} \<inter> {x. x$k \<ge> c})"
-  shows "(p1 \<union> p2) tagged_division_of ({a..b})"
-proof- have *:"{a..b} = ({a..b} \<inter> {x. x$k \<le> c}) \<union> ({a..b} \<inter> {x. x$k \<ge> c})" by auto
-  show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms])
-    unfolding interval_split interior_closed_interval
-    by(auto simp add: vector_less_def Cart_lambda_beta elim!:allE[where x=k]) qed
-
-lemma has_integral_separate_sides: fixes f::"real^'m \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral i) ({a..b})" "e>0"
-  obtains d where "gauge d" "(\<forall>p1 p2. p1 tagged_division_of ({a..b} \<inter> {x. x$k \<le> c}) \<and> d fine p1 \<and>
-                                p2 tagged_division_of ({a..b} \<inter> {x. x$k \<ge> c}) \<and> d fine p2
-                                \<longrightarrow> norm((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 +
-                                          setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e)"
-proof- guess d using has_integralD[OF assms] . note d=this
-  show ?thesis apply(rule that[of d]) apply(rule d) apply(rule,rule,rule,(erule conjE)+)
-  proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this
-                   assume "p2 tagged_division_of {a..b} \<inter> {x. c \<le> x $ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this
-    note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this
-    have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) = norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
-      apply(subst setsum_Un_zero) apply(rule p1 p2)+ apply(rule) unfolding split_paired_all split_conv
-    proof- fix a b assume ab:"(a,b) \<in> p1 \<inter> p2"
-      have "(a,b) \<in> p1" using ab by auto from p1(4)[OF this] guess u v apply-by(erule exE)+ note uv =this
-      have "b \<subseteq> {x. x$k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastsimp
-      moreover have "interior {x. x $ k = c} = {}" 
-      proof(rule ccontr) case goal1 then obtain x where x:"x\<in>interior {x. x$k = c}" by auto
-        then guess e unfolding mem_interior .. note e=this
-        have x:"x$k = c" using x interior_subset by fastsimp
-        have *:"\<And>i. \<bar>(x - (x + (\<chi> i. if i = k then e / 2 else 0))) $ i\<bar> = (if i = k then e/2 else 0)" using e by auto
-        have "x + (\<chi> i. if i = k then e/2 else 0) \<in> ball x e" unfolding mem_ball vector_dist_norm 
-          apply(rule le_less_trans[OF norm_le_l1]) unfolding * 
-          unfolding setsum_delta[OF finite_UNIV] using e by auto 
-        hence "x + (\<chi> i. if i = k then e/2 else 0) \<in> {x. x$k = c}" using e by auto
-        thus False unfolding mem_Collect_eq using e x by auto
-      qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule subset_interior) by auto
-      thus "content b *\<^sub>R f a = 0" by auto
-    qed auto
-    also have "\<dots> < e" by(rule d(2) p12 fine_union p1 p2)+
-    finally show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" . qed qed
-
-lemma integrable_split[intro]: fixes f::"real^'n \<Rightarrow> 'a::{real_normed_vector,complete_space}" assumes "f integrable_on {a..b}"
-  shows "f integrable_on ({a..b} \<inter> {x. x$k \<le> c})" (is ?t1) and "f integrable_on ({a..b} \<inter> {x. x$k \<ge> c})" (is ?t2) 
-proof- guess y using assms unfolding integrable_on_def .. note y=this
-  def b' \<equiv> "(\<chi> i. if i = k then min (b$k) c else b$i)::real^'n"
-  and a' \<equiv> "(\<chi> i. if i = k then max (a$k) c else a$i)::real^'n"
-  show ?t1 ?t2 unfolding interval_split integrable_cauchy unfolding interval_split[THEN sym]
-  proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "e/2>0" by auto
-    from has_integral_separate_sides[OF y this,of k c] guess d . note d=this[rule_format]
-    let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<inter> A \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> A \<and> d fine p2 \<longrightarrow>
-                              norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
-    show "?P {x. x $ k \<le> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
-    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c} \<and> d fine p2"
-      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
-      proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
-        show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
-          using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
-          using p using assms by(auto simp add:group_simps)
-      qed qed  
-    show "?P {x. x $ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
-    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p2"
-      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
-      proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
-        show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
-          using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
-          using p using assms by(auto simp add:group_simps) qed qed qed qed
-
-subsection {* Generalized notion of additivity. *}
-
-definition "neutral opp = (SOME x. \<forall>y. opp x y = y \<and> opp y x = y)"
-
-definition operative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ((real^'n) set \<Rightarrow> 'a) \<Rightarrow> bool" where
-  "operative opp f \<equiv> 
-    (\<forall>a b. content {a..b} = 0 \<longrightarrow> f {a..b} = neutral(opp)) \<and>
-    (\<forall>a b c k. f({a..b}) =
-                   opp (f({a..b} \<inter> {x. x$k \<le> c}))
-                       (f({a..b} \<inter> {x. x$k \<ge> c})))"
-
-lemma operativeD[dest]: assumes "operative opp f"
-  shows "\<And>a b. content {a..b} = 0 \<Longrightarrow> f {a..b} = neutral(opp)"
-  "\<And>a b c k. f({a..b}) = opp (f({a..b} \<inter> {x. x$k \<le> c})) (f({a..b} \<inter> {x. x$k \<ge> c}))"
-  using assms unfolding operative_def by auto
-
-lemma operative_trivial:
- "operative opp f \<Longrightarrow> content({a..b}) = 0 \<Longrightarrow> f({a..b}) = neutral opp"
-  unfolding operative_def by auto
-
-lemma property_empty_interval:
- "(\<forall>a b. content({a..b}) = 0 \<longrightarrow> P({a..b})) \<Longrightarrow> P {}" 
-  using content_empty unfolding empty_as_interval by auto
-
-lemma operative_empty: "operative opp f \<Longrightarrow> f {} = neutral opp"
-  unfolding operative_def apply(rule property_empty_interval) by auto
-
-subsection {* Using additivity of lifted function to encode definedness. *}
-
-lemma forall_option: "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>x. P(Some x))"
-  by (metis map_of.simps option.nchotomy)
-
-lemma exists_option:
- "(\<exists>x. P x) \<longleftrightarrow> P None \<or> (\<exists>x. P(Some x))" 
-  by (metis map_of.simps option.nchotomy)
-
-fun lifted where 
-  "lifted (opp::'a\<Rightarrow>'a\<Rightarrow>'b) (Some x) (Some y) = Some(opp x y)" |
-  "lifted opp None _ = (None::'b option)" |
-  "lifted opp _ None = None"
-
-lemma lifted_simp_1[simp]: "lifted opp v None = None"
-  apply(induct v) by auto
-
-definition "monoidal opp \<equiv>  (\<forall>x y. opp x y = opp y x) \<and>
-                   (\<forall>x y z. opp x (opp y z) = opp (opp x y) z) \<and>
-                   (\<forall>x. opp (neutral opp) x = x)"
-
-lemma monoidalI: assumes "\<And>x y. opp x y = opp y x"
-  "\<And>x y z. opp x (opp y z) = opp (opp x y) z"
-  "\<And>x. opp (neutral opp) x = x" shows "monoidal opp"
-  unfolding monoidal_def using assms by fastsimp
-
-lemma monoidal_ac: assumes "monoidal opp"
-  shows "opp (neutral opp) a = a" "opp a (neutral opp) = a" "opp a b = opp b a"
-  "opp (opp a b) c = opp a (opp b c)"  "opp a (opp b c) = opp b (opp a c)"
-  using assms unfolding monoidal_def apply- by metis+
-
-lemma monoidal_simps[simp]: assumes "monoidal opp"
-  shows "opp (neutral opp) a = a" "opp a (neutral opp) = a"
-  using monoidal_ac[OF assms] by auto
-
-lemma neutral_lifted[cong]: assumes "monoidal opp"
-  shows "neutral (lifted opp) = Some(neutral opp)"
-  apply(subst neutral_def) apply(rule some_equality) apply(rule,induct_tac y) prefer 3
-proof- fix x assume "\<forall>y. lifted opp x y = y \<and> lifted opp y x = y"
-  thus "x = Some (neutral opp)" apply(induct x) defer
-    apply rule apply(subst neutral_def) apply(subst eq_commute,rule some_equality)
-    apply(rule,erule_tac x="Some y" in allE) defer apply(erule_tac x="Some x" in allE) by auto
-qed(auto simp add:monoidal_ac[OF assms])
-
-lemma monoidal_lifted[intro]: assumes "monoidal opp" shows "monoidal(lifted opp)"
-  unfolding monoidal_def forall_option neutral_lifted[OF assms] using monoidal_ac[OF assms] by auto
-
-definition "support opp f s = {x. x\<in>s \<and> f x \<noteq> neutral opp}"
-definition "fold' opp e s \<equiv> (if finite s then fold opp e s else e)"
-definition "iterate opp s f \<equiv> fold' (\<lambda>x a. opp (f x) a) (neutral opp) (support opp f s)"
-
-lemma support_subset[intro]:"support opp f s \<subseteq> s" unfolding support_def by auto
-lemma support_empty[simp]:"support opp f {} = {}" using support_subset[of opp f "{}"] by auto
-
-lemma fun_left_comm_monoidal[intro]: assumes "monoidal opp" shows "fun_left_comm opp"
-  unfolding fun_left_comm_def using monoidal_ac[OF assms] by auto
-
-lemma support_clauses:
-  "\<And>f g s. support opp f {} = {}"
-  "\<And>f g s. support opp f (insert x s) = (if f(x) = neutral opp then support opp f s else insert x (support opp f s))"
-  "\<And>f g s. support opp f (s - {x}) = (support opp f s) - {x}"
-  "\<And>f g s. support opp f (s \<union> t) = (support opp f s) \<union> (support opp f t)"
-  "\<And>f g s. support opp f (s \<inter> t) = (support opp f s) \<inter> (support opp f t)"
-  "\<And>f g s. support opp f (s - t) = (support opp f s) - (support opp f t)"
-  "\<And>f g s. support opp g (f ` s) = f ` (support opp (g o f) s)"
-unfolding support_def by auto
-
-lemma finite_support[intro]:"finite s \<Longrightarrow> finite (support opp f s)"
-  unfolding support_def by auto
-
-lemma iterate_empty[simp]:"iterate opp {} f = neutral opp"
-  unfolding iterate_def fold'_def by auto 
-
-lemma iterate_insert[simp]: assumes "monoidal opp" "finite s"
-  shows "iterate opp (insert x s) f = (if x \<in> s then iterate opp s f else opp (f x) (iterate opp s f))" 
-proof(cases "x\<in>s") case True hence *:"insert x s = s" by auto
-  show ?thesis unfolding iterate_def if_P[OF True] * by auto
-next case False note x=this
-  note * = fun_left_comm.fun_left_comm_apply[OF fun_left_comm_monoidal[OF assms(1)]]
-  show ?thesis proof(cases "f x = neutral opp")
-    case True show ?thesis unfolding iterate_def if_not_P[OF x] support_clauses if_P[OF True]
-      unfolding True monoidal_simps[OF assms(1)] by auto
-  next case False show ?thesis unfolding iterate_def fold'_def  if_not_P[OF x] support_clauses if_not_P[OF False]
-      apply(subst fun_left_comm.fold_insert[OF * finite_support])
-      using `finite s` unfolding support_def using False x by auto qed qed 
-
-lemma iterate_some:
-  assumes "monoidal opp"  "finite s"
-  shows "iterate (lifted opp) s (\<lambda>x. Some(f x)) = Some (iterate opp s f)" using assms(2)
-proof(induct s) case empty thus ?case using assms by auto
-next case (insert x F) show ?case apply(subst iterate_insert) prefer 3 apply(subst if_not_P)
-    defer unfolding insert(3) lifted.simps apply rule using assms insert by auto qed
-
-subsection {* Two key instances of additivity. *}
-
-lemma neutral_add[simp]:
-  "neutral op + = (0::_::comm_monoid_add)" unfolding neutral_def 
-  apply(rule some_equality) defer apply(erule_tac x=0 in allE) by auto
-
-lemma operative_content[intro]: "operative (op +) content"
-  unfolding operative_def content_split[THEN sym] neutral_add by auto
-
-lemma neutral_monoid[simp]: "neutral ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a) = 0"
-  unfolding neutral_def apply(rule some_equality) defer
-  apply(erule_tac x=0 in allE) by auto
-
-lemma monoidal_monoid[intro]:
-  shows "monoidal ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a)"
-  unfolding monoidal_def neutral_monoid by(auto simp add: group_simps) 
-
-lemma operative_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
-  unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add
-  apply(rule,rule,rule,rule) defer apply(rule allI)+
-proof- fix a b c k show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) =
-              lifted op + (if f integrable_on {a..b} \<inter> {x. x $ k \<le> c} then Some (integral ({a..b} \<inter> {x. x $ k \<le> c}) f) else None)
-               (if f integrable_on {a..b} \<inter> {x. c \<le> x $ k} then Some (integral ({a..b} \<inter> {x. c \<le> x $ k}) f) else None)"
-  proof(cases "f integrable_on {a..b}") 
-    case True show ?thesis unfolding if_P[OF True]
-      unfolding if_P[OF integrable_split(1)[OF True]] if_P[OF integrable_split(2)[OF True]]
-      unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split) 
-      apply(rule_tac[!] integrable_integral integrable_split)+ using True by assumption+
-  next case False have "(\<not> (f integrable_on {a..b} \<inter> {x. x $ k \<le> c})) \<or> (\<not> ( f integrable_on {a..b} \<inter> {x. c \<le> x $ k}))"
-    proof(rule ccontr) case goal1 hence "f integrable_on {a..b}" apply- unfolding integrable_on_def
-        apply(rule_tac x="integral ({a..b} \<inter> {x. x $ k \<le> c}) f + integral ({a..b} \<inter> {x. x $ k \<ge> c}) f" in exI)
-        apply(rule has_integral_split) apply(rule_tac[!] integrable_integral) by auto
-      thus False using False by auto
-    qed thus ?thesis using False by auto 
-  qed next 
-  fix a b assume as:"content {a..b::real^'n} = 0"
-  thus "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = Some 0"
-    unfolding if_P[OF integrable_on_null[OF as]] using has_integral_null_eq[OF as] by auto qed
-
-subsection {* Points of division of a partition. *}
-
-definition "division_points (k::(real^'n) set) d = 
-    {(j,x). (interval_lowerbound k)$j < x \<and> x < (interval_upperbound k)$j \<and>
-           (\<exists>i\<in>d. (interval_lowerbound i)$j = x \<or> (interval_upperbound i)$j = x)}"
-
-lemma division_points_finite: 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)$j < x \<and> x < (interval_upperbound i)$j \<and>
-           (\<exists>i\<in>d. (interval_lowerbound i)$j = x \<or> (interval_upperbound i)$j = x)}"
-  have *:"division_points i d = \<Union>(?M ` UNIV)"
-    unfolding division_points_def by auto
-  show ?thesis unfolding * using assm by auto qed
-
-lemma division_points_subset:
-  assumes "d division_of {a..b}" "\<forall>i. a$i < b$i"  "a$k < c" "c < b$k"
-  shows "division_points ({a..b} \<inter> {x. x$k \<le> c}) {l \<inter> {x. x$k \<le> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$k \<le> c} = {})}
-                  \<subseteq> division_points ({a..b}) d" (is ?t1) and
-        "division_points ({a..b} \<inter> {x. x$k \<ge> c}) {l \<inter> {x. x$k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$k \<ge> c} = {})}
-                  \<subseteq> division_points ({a..b}) d" (is ?t2)
-proof- note assm = division_ofD[OF assms(1)]
-  have *:"\<forall>i. a$i \<le> b$i"   "\<forall>i. a$i \<le> (\<chi> i. if i = k then min (b $ k) c else b $ i) $ i"
-    "\<forall>i. (\<chi> i. if i = k then max (a $ k) c else a $ i) $ i \<le> b$i"  "min (b $ k) c = c" "max (a $ k) c = c"
-    using assms using less_imp_le by auto
-  show ?t1 unfolding division_points_def interval_split[of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding *
-    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq apply(erule exE conjE)+
-  proof- fix i l x assume as:"a $ fst x < snd x" "snd x < (if fst x = k then c else b $ fst x)"
-      "interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x"  "i = l \<inter> {x. x $ k \<le> c}" "l \<in> d" "l \<inter> {x. x $ k \<le> c} \<noteq> {}"
-    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
-    have *:"\<forall>i. u $ i \<le> (\<chi> i. if i = k then min (v $ k) c else v $ i) $ i" using as(6) unfolding l interval_split interval_ne_empty as .
-    have **:"\<forall>i. u$i \<le> v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
-    show "a $ fst x < snd x \<and> snd x < b $ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x)"
-      using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply-
-      apply(rule,assumption,rule) defer apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
-      apply(case_tac[!] "fst x = k") using assms by auto
-  qed
-  show ?t2 unfolding division_points_def interval_split[of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding *
-    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq apply(erule exE conjE)+
-  proof- fix i l x assume as:"(if fst x = k then c else a $ fst x) < snd x" "snd x < b $ fst x" "interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x"
-      "i = l \<inter> {x. c \<le> x $ k}" "l \<in> d" "l \<inter> {x. c \<le> x $ k} \<noteq> {}"
-    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
-    have *:"\<forall>i. (\<chi> i. if i = k then max (u $ k) c else u $ i) $ i \<le> v $ i" using as(6) unfolding l interval_split interval_ne_empty as .
-    have **:"\<forall>i. u$i \<le> v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
-    show "a $ fst x < snd x \<and> snd x < b $ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x)"
-      using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply-
-      apply rule defer apply(rule,assumption) apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
-      apply(case_tac[!] "fst x = k") using assms by auto qed qed
-
-lemma division_points_psubset:
-  assumes "d division_of {a..b}"  "\<forall>i. a$i < b$i"  "a$k < c" "c < b$k"
-  "l \<in> d" "interval_lowerbound l$k = c \<or> interval_upperbound l$k = c"
-  shows "division_points ({a..b} \<inter> {x. x$k \<le> c}) {l \<inter> {x. x$k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x$k \<le> c} \<noteq> {}} \<subset> division_points ({a..b}) d" (is "?D1 \<subset> ?D") 
-        "division_points ({a..b} \<inter> {x. x$k \<ge> c}) {l \<inter> {x. x$k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x$k \<ge> c} \<noteq> {}} \<subset> division_points ({a..b}) d" (is "?D2 \<subset> ?D") 
-proof- have ab:"\<forall>i. a$i \<le> b$i" using assms(2) by(auto intro!:less_imp_le)
-  guess u v using division_ofD(4)[OF assms(1,5)] apply-by(erule exE)+ note l=this
-  have uv:"\<forall>i. u$i \<le> v$i" "\<forall>i. a$i \<le> u$i \<and> v$i \<le> b$i" using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty
-    unfolding subset_eq apply- defer apply(erule_tac x=u in ballE, erule_tac x=v in ballE) unfolding mem_interval by auto
-  have *:"interval_upperbound ({a..b} \<inter> {x. x $ k \<le> interval_upperbound l $ k}) $ k = interval_upperbound l $ k"
-         "interval_upperbound ({a..b} \<inter> {x. x $ k \<le> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
-    unfolding interval_split apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
-    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab by auto
-  have "\<exists>x. x \<in> ?D - ?D1" using assms(2-) apply-apply(erule disjE)
-    apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
-    apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
-    unfolding division_points_def unfolding interval_bounds[OF ab]
-    apply (auto simp add:interval_bounds) unfolding * by auto
-  thus "?D1 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto
-
-  have *:"interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
-         "interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_upperbound l $ k}) $ k = interval_upperbound l $ k"
-    unfolding interval_split apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
-    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab by auto
-  have "\<exists>x. x \<in> ?D - ?D2" using assms(2-) apply-apply(erule disjE)
-    apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
-    apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
-    unfolding division_points_def unfolding interval_bounds[OF ab]
-    apply (auto simp add:interval_bounds) unfolding * by auto
-  thus "?D2 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto qed
-
-subsection {* Preservation by divisions and tagged divisions. *}
-
-lemma support_support[simp]:"support opp f (support opp f s) = support opp f s"
-  unfolding support_def by auto
-
-lemma iterate_support[simp]: "iterate opp (support opp f s) f = iterate opp s f"
-  unfolding iterate_def support_support by auto
-
-lemma iterate_expand_cases:
-  "iterate opp s f = (if finite(support opp f s) then iterate opp (support opp f s) f else neutral opp)"
-  apply(cases) apply(subst if_P,assumption) unfolding iterate_def support_support fold'_def by auto 
-
-lemma iterate_image: assumes "monoidal opp"  "inj_on f s"
-  shows "iterate opp (f ` s) g = iterate opp s (g \<circ> f)"
-proof- have *:"\<And>s. finite s \<Longrightarrow>  \<forall>x\<in>s. \<forall>y\<in>s. f x = f y \<longrightarrow> x = y \<Longrightarrow>
-     iterate opp (f ` s) g = iterate opp s (g \<circ> f)"
-  proof- case goal1 show ?case using goal1
-    proof(induct s) case empty thus ?case using assms(1) by auto
-    next case (insert x s) show ?case unfolding iterate_insert[OF assms(1) insert(1)]
-        unfolding if_not_P[OF insert(2)] apply(subst insert(3)[THEN sym])
-        unfolding image_insert defer apply(subst iterate_insert[OF assms(1)])
-        apply(rule finite_imageI insert)+ apply(subst if_not_P)
-        unfolding image_iff o_def using insert(2,4) by auto
-    qed qed
-  show ?thesis 
-    apply(cases "finite (support opp g (f ` s))")
-    apply(subst (1) iterate_support[THEN sym],subst (2) iterate_support[THEN sym])
-    unfolding support_clauses apply(rule *)apply(rule finite_imageD,assumption) unfolding inj_on_def[symmetric]
-    apply(rule subset_inj_on[OF assms(2) support_subset])+
-    apply(subst iterate_expand_cases) unfolding support_clauses apply(simp only: if_False)
-    apply(subst iterate_expand_cases) apply(subst if_not_P) by auto qed
-
-
-(* This lemma about iterations comes up in a few places.                     *)
-lemma iterate_nonzero_image_lemma:
-  assumes "monoidal opp" "finite s" "g(a) = neutral opp"
-  "\<forall>x\<in>s. \<forall>y\<in>s. f x = f y \<and> x \<noteq> y \<longrightarrow> g(f x) = neutral opp"
-  shows "iterate opp {f x | x. x \<in> s \<and> f x \<noteq> a} g = iterate opp s (g \<circ> f)"
-proof- have *:"{f x |x. x \<in> s \<and> ~(f x = a)} = f ` {x. x \<in> s \<and> ~(f x = a)}" by auto
-  have **:"support opp (g \<circ> f) {x \<in> s. f x \<noteq> a} = support opp (g \<circ> f) s"
-    unfolding support_def using assms(3) by auto
-  show ?thesis unfolding *
-    apply(subst iterate_support[THEN sym]) unfolding support_clauses
-    apply(subst iterate_image[OF assms(1)]) defer
-    apply(subst(2) iterate_support[THEN sym]) apply(subst **)
-    unfolding inj_on_def using assms(3,4) unfolding support_def by auto qed
-
-lemma iterate_eq_neutral:
-  assumes "monoidal opp"  "\<forall>x \<in> s. (f(x) = neutral opp)"
-  shows "(iterate opp s f = neutral opp)"
-proof- have *:"support opp f s = {}" unfolding support_def using assms(2) by auto
-  show ?thesis apply(subst iterate_support[THEN sym]) 
-    unfolding * using assms(1) by auto qed
-
-lemma iterate_op: assumes "monoidal opp" "finite s"
-  shows "iterate opp s (\<lambda>x. opp (f x) (g x)) = opp (iterate opp s f) (iterate opp s g)" using assms(2)
-proof(induct s) case empty thus ?case unfolding iterate_insert[OF assms(1)] using assms(1) by auto
-next case (insert x F) show ?case unfolding iterate_insert[OF assms(1) insert(1)] if_not_P[OF insert(2)] insert(3)
-    unfolding monoidal_ac[OF assms(1)] by(rule refl) qed
-
-lemma iterate_eq: assumes "monoidal opp" "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
-  shows "iterate opp s f = iterate opp s g"
-proof- have *:"support opp g s = support opp f s"
-    unfolding support_def using assms(2) by auto
-  show ?thesis
-  proof(cases "finite (support opp f s)")
-    case False thus ?thesis apply(subst iterate_expand_cases,subst(2) iterate_expand_cases)
-      unfolding * by auto
-  next def su \<equiv> "support opp f s"
-    case True note support_subset[of opp f s] 
-    thus ?thesis apply- apply(subst iterate_support[THEN sym],subst(2) iterate_support[THEN sym]) unfolding * using True
-      unfolding su_def[symmetric]
-    proof(induct su) case empty show ?case by auto
-    next case (insert x s) show ?case unfolding iterate_insert[OF assms(1) insert(1)] 
-        unfolding if_not_P[OF insert(2)] apply(subst insert(3))
-        defer apply(subst assms(2)[of x]) using insert by auto qed qed qed
-
-lemma nonempty_witness: assumes "s \<noteq> {}" obtains x where "x \<in> s" using assms by auto
-
-lemma operative_division: fixes f::"(real^'n) set \<Rightarrow> 'a"
-  assumes "monoidal opp" "operative opp f" "d division_of {a..b}"
-  shows "iterate opp d f = f {a..b}"
-proof- def C \<equiv> "card (division_points {a..b} d)" thus ?thesis using assms
-  proof(induct C arbitrary:a b d rule:full_nat_induct)
-    case goal1
-    { presume *:"content {a..b} \<noteq> 0 \<Longrightarrow> ?case"
-      thus ?case apply-apply(cases) defer apply assumption
-      proof- assume as:"content {a..b} = 0"
-        show ?case unfolding operativeD(1)[OF assms(2) as] apply(rule iterate_eq_neutral[OF goal1(2)])
-        proof fix x assume x:"x\<in>d"
-          then guess u v apply(drule_tac division_ofD(4)[OF goal1(4)]) by(erule exE)+
-          thus "f x = neutral opp" using division_of_content_0[OF as goal1(4)] 
-            using operativeD(1)[OF assms(2)] x by auto
-        qed qed }
-    assume "content {a..b} \<noteq> 0" note ab = this[unfolded content_lt_nz[THEN sym] content_pos_lt_eq]
-    hence ab':"\<forall>i. a$i \<le> b$i" by (auto intro!: less_imp_le) show ?case 
-    proof(cases "division_points {a..b} d = {}")
-      case True have d':"\<forall>i\<in>d. \<exists>u v. i = {u..v} \<and>
-        (\<forall>j. u$j = a$j \<and> v$j = a$j \<or> u$j = b$j \<and> v$j = b$j \<or> u$j = a$j \<and> v$j = b$j)"
-        unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule)
-        apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) apply(rule)
-      proof- fix u v j assume as:"{u..v} \<in> d" note division_ofD(3)[OF goal1(4) this]
-        hence uv:"\<forall>i. u$i \<le> v$i" "u$j \<le> v$j" unfolding interval_ne_empty by auto
-        have *:"\<And>p r Q. p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> (Q {u..v})" using as by auto
-        have "(j, u$j) \<notin> division_points {a..b} d"
-          "(j, v$j) \<notin> division_points {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$j \<le> u$j" "v$j \<le> b$j" using division_ofD(2,2,3)[OF goal1(4) as] 
-          unfolding subset_eq apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE)
-          unfolding interval_ne_empty mem_interval by auto
-        ultimately show "u$j = a$j \<and> v$j = a$j \<or> u$j = b$j \<and> v$j = b$j \<or> u$j = a$j \<and> v$j = b$j"
-          unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) by auto
-      qed have "(1/2) *\<^sub>R (a+b) \<in> {a..b}" unfolding mem_interval using ab by(auto intro!:less_imp_le)
-      note this[unfolded division_ofD(6)[OF goal1(4),THEN sym] Union_iff]
-      then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] apply-by(erule exE conjE)+ note uv=this
-      have "{a..b} \<in> d"
-      proof- { presume "i = {a..b}" thus ?thesis using i by auto }
-        { presume "u = a" "v = b" thus "i = {a..b}" using uv by auto }
-        show "u = a" "v = b" unfolding Cart_eq
-        proof(rule_tac[!] allI) fix j note i(2)[unfolded uv mem_interval,rule_format,of j]
-          thus "u $ j = a $ j" "v $ j = b $ j" using uv(2)[rule_format,of j] by auto
-        qed qed
-      hence *:"d = insert {a..b} (d - {{a..b}})" by auto
-      have "iterate opp (d - {{a..b}}) f = neutral opp" apply(rule iterate_eq_neutral[OF goal1(2)])
-      proof fix x assume x:"x \<in> d - {{a..b}}" hence "x\<in>d" by auto note d'[rule_format,OF this]
-        then guess u v apply-by(erule exE conjE)+ note uv=this
-        have "u\<noteq>a \<or> v\<noteq>b" using x[unfolded uv] by auto  
-        then obtain j where "u$j \<noteq> a$j \<or> v$j \<noteq> b$j" unfolding Cart_eq by auto
-        hence "u$j = v$j" using uv(2)[rule_format,of j] by auto
-        hence "content {u..v} = 0"  unfolding content_eq_0 apply(rule_tac x=j in exI) by auto
-        thus "f x = neutral opp" unfolding uv(1) by(rule operativeD(1)[OF goal1(3)])
-      qed thus "iterate opp d f = f {a..b}" apply-apply(subst *) 
-        apply(subst iterate_insert[OF goal1(2)]) using goal1(2,4) by auto
-    next case False hence "\<exists>x. x\<in>division_points {a..b} d" by auto
-      then guess k c unfolding split_paired_Ex apply- unfolding division_points_def mem_Collect_eq split_conv
-        by(erule exE conjE)+ note kc=this[unfolded interval_bounds[OF ab']]
-      from this(3) guess j .. note j=this
-      def d1 \<equiv> "{l \<inter> {x. x$k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x$k \<le> c} \<noteq> {}}"
-      def d2 \<equiv> "{l \<inter> {x. x$k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x$k \<ge> c} \<noteq> {}}"
-      def cb \<equiv> "(\<chi> i. if i = k then c else b$i)" and ca \<equiv> "(\<chi> i. if i = k then c else a$i)"
-      note division_points_psubset[OF goal1(4) ab kc(1-2) j]
-      note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
-      hence *:"(iterate opp d1 f) = f ({a..b} \<inter> {x. x$k \<le> c})" "(iterate opp d2 f) = f ({a..b} \<inter> {x. x$k \<ge> c})"
-        apply- unfolding interval_split apply(rule_tac[!] goal1(1)[rule_format])
-        using division_split[OF goal1(4), where k=k and c=c]
-        unfolding interval_split d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono
-        using goal1(2-3) using division_points_finite[OF goal1(4)] by auto
-      have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev")
-        unfolding * apply(rule operativeD(2)) using goal1(3) .
-      also have "iterate opp d1 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$k \<le> c}))"
-        unfolding d1_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
-        unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
-        unfolding empty_as_interval[THEN sym] apply(rule content_empty)
-      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. x $ k \<le> c} = y \<inter> {x. x $ k \<le> c}" "l \<noteq> y" 
-        from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
-        show "f (l \<inter> {x. x $ k \<le> c}) = neutral opp" unfolding l interval_split
-          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_left_inj)
-          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+
-      qed also have "iterate opp d2 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$k \<ge> c}))"
-        unfolding d2_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
-        unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
-        unfolding empty_as_interval[THEN sym] apply(rule content_empty)
-      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x $ k} = y \<inter> {x. c \<le> x $ k}" "l \<noteq> y" 
-        from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
-        show "f (l \<inter> {x. x $ k \<ge> c}) = neutral opp" unfolding l interval_split
-          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_right_inj)
-          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+
-      qed also have *:"\<forall>x\<in>d. f x = opp (f (x \<inter> {x. x $ k \<le> c})) (f (x \<inter> {x. c \<le> x $ k}))"
-        unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) .
-      have "opp (iterate opp d (\<lambda>l. f (l \<inter> {x. x $ k \<le> c}))) (iterate opp d (\<lambda>l. f (l \<inter> {x. c \<le> x $ k})))
-        = iterate opp d f" apply(subst(3) iterate_eq[OF _ *[rule_format]]) prefer 3
-        apply(rule iterate_op[THEN sym]) using goal1 by auto
-      finally show ?thesis by auto
-    qed qed qed 
-
-lemma iterate_image_nonzero: assumes "monoidal opp"
-  "finite s" "\<forall>x\<in>s. \<forall>y\<in>s. ~(x = y) \<and> f x = f y \<longrightarrow> g(f x) = neutral opp"
-  shows "iterate opp (f ` s) g = iterate opp s (g \<circ> f)" using assms
-proof(induct rule:finite_subset_induct[OF assms(2) subset_refl])
-  case goal1 show ?case using assms(1) by auto
-next case goal2 have *:"\<And>x y. y = neutral opp \<Longrightarrow> x = opp y x" using assms(1) by auto
-  show ?case unfolding image_insert apply(subst iterate_insert[OF assms(1)])
-    apply(rule finite_imageI goal2)+
-    apply(cases "f a \<in> f ` F") unfolding if_P if_not_P apply(subst goal2(4)[OF assms(1) goal2(1)]) defer
-    apply(subst iterate_insert[OF assms(1) goal2(1)]) defer
-    apply(subst iterate_insert[OF assms(1) goal2(1)])
-    unfolding if_not_P[OF goal2(3)] defer unfolding image_iff defer apply(erule bexE)
-    apply(rule *) unfolding o_def apply(rule_tac y=x in goal2(7)[rule_format])
-    using goal2 unfolding o_def by auto qed 
-
-lemma operative_tagged_division: assumes "monoidal opp" "operative opp f" "d tagged_division_of {a..b}"
-  shows "iterate(opp) d (\<lambda>(x,l). f l) = f {a..b}"
-proof- have *:"(\<lambda>(x,l). f l) = (f o snd)" unfolding o_def by(rule,auto) note assm = tagged_division_ofD[OF assms(3)]
-  have "iterate(opp) d (\<lambda>(x,l). f l) = iterate opp (snd ` d) f" unfolding *
-    apply(rule iterate_image_nonzero[THEN sym,OF assms(1)]) apply(rule tagged_division_of_finite assms)+ 
-    unfolding Ball_def split_paired_All snd_conv apply(rule,rule,rule,rule,rule,rule,rule,erule conjE)
-  proof- fix a b aa ba assume as:"(a, b) \<in> d" "(aa, ba) \<in> d" "(a, b) \<noteq> (aa, ba)" "b = ba"
-    guess u v using assm(4)[OF as(1)] apply-by(erule exE)+ note uv=this
-    show "f b = neutral opp" unfolding uv apply(rule operativeD(1)[OF assms(2)])
-      unfolding content_eq_0_interior using tagged_division_ofD(5)[OF assms(3) as(1-3)]
-      unfolding as(4)[THEN sym] uv by auto
-  qed also have "\<dots> = f {a..b}" 
-    using operative_division[OF assms(1-2) division_of_tagged_division[OF assms(3)]] .
-  finally show ?thesis . qed
-
-subsection {* Additivity of content. *}
-
-lemma setsum_iterate:assumes "finite s" shows "setsum f s = iterate op + s f"
-proof- have *:"setsum f s = setsum f (support op + f s)"
-    apply(rule setsum_mono_zero_right)
-    unfolding support_def neutral_monoid using assms by auto
-  thus ?thesis unfolding * setsum_def iterate_def fold_image_def fold'_def
-    unfolding neutral_monoid . qed
-
-lemma additive_content_division: assumes "d division_of {a..b}"
-  shows "setsum content d = content({a..b})"
-  unfolding operative_division[OF monoidal_monoid operative_content assms,THEN sym]
-  apply(subst setsum_iterate) using assms by auto
-
-lemma additive_content_tagged_division:
-  assumes "d tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,l). content l) d = content({a..b})"
-  unfolding operative_tagged_division[OF monoidal_monoid operative_content assms,THEN sym]
-  apply(subst setsum_iterate) using assms by auto
-
-subsection {* Finally, the integral of a constant\<forall> *}
-
-lemma has_integral_const[intro]:
-  "((\<lambda>x. c) has_integral (content({a..b::real^'n}) *\<^sub>R c)) ({a..b})"
-  unfolding has_integral apply(rule,rule,rule_tac x="\<lambda>x. ball x 1" in exI)
-  apply(rule,rule gauge_trivial)apply(rule,rule,erule conjE)
-  unfolding split_def apply(subst scaleR_left.setsum[THEN sym, unfolded o_def])
-  defer apply(subst additive_content_tagged_division[unfolded split_def]) apply assumption by auto
-
-subsection {* Bounds on the norm of Riemann sums and the integral itself. *}
-
-lemma dsum_bound: assumes "p division_of {a..b}" "norm(c) \<le> e"
-  shows "norm(setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content({a..b})" (is "?l \<le> ?r")
-  apply(rule order_trans,rule setsum_norm) defer unfolding norm_scaleR setsum_left_distrib[THEN sym]
-  apply(rule order_trans[OF mult_left_mono],rule assms,rule setsum_abs_ge_zero)
-  apply(subst real_mult_commute) apply(rule mult_left_mono)
-  apply(rule order_trans[of _ "setsum content p"]) apply(rule eq_refl,rule setsum_cong2)
-  apply(subst abs_of_nonneg) unfolding additive_content_division[OF assms(1)]
-proof- from order_trans[OF norm_ge_zero[of c] assms(2)] show "0 \<le> e" .
-  fix x assume "x\<in>p" from division_ofD(4)[OF assms(1) this] guess u v apply-by(erule exE)+
-  thus "0 \<le> content x" using content_pos_le by auto
-qed(insert assms,auto)
-
-lemma rsum_bound: assumes "p tagged_division_of {a..b}" "\<forall>x\<in>{a..b}. norm(f x) \<le> e"
-  shows "norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content({a..b})"
-proof(cases "{a..b} = {}") case True
-  show ?thesis using assms(1) unfolding True tagged_division_of_trivial by auto
-next case False show ?thesis
-    apply(rule order_trans,rule setsum_norm) defer unfolding split_def norm_scaleR
-    apply(rule order_trans[OF setsum_mono]) apply(rule mult_left_mono[OF _ abs_ge_zero, of _ e]) defer
-    unfolding setsum_left_distrib[THEN sym] apply(subst real_mult_commute) apply(rule mult_left_mono)
-    apply(rule order_trans[of _ "setsum (content \<circ> snd) p"]) apply(rule eq_refl,rule setsum_cong2)
-    apply(subst o_def, rule abs_of_nonneg)
-  proof- show "setsum (content \<circ> snd) p \<le> content {a..b}" apply(rule eq_refl)
-      unfolding additive_content_tagged_division[OF assms(1),THEN sym] split_def by auto
-    guess w using nonempty_witness[OF False] .
-    thus "e\<ge>0" apply-apply(rule order_trans) defer apply(rule assms(2)[rule_format],assumption) by auto
-    fix xk assume *:"xk\<in>p" guess x k  using surj_pair[of xk] apply-by(erule exE)+ note xk = this *[unfolded this]
-    from tagged_division_ofD(4)[OF assms(1) xk(2)] guess u v apply-by(erule exE)+ note uv=this
-    show "0\<le> content (snd xk)" unfolding xk snd_conv uv by(rule content_pos_le)
-    show "norm (f (fst xk)) \<le> e" unfolding xk fst_conv using tagged_division_ofD(2,3)[OF assms(1) xk(2)] assms(2) by auto
-  qed(insert assms,auto) qed
-
-lemma rsum_diff_bound:
-  assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. norm(f x - g x) \<le> e"
-  shows "norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - setsum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le> e * content({a..b})"
-  apply(rule order_trans[OF _ rsum_bound[OF assms]]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
-  unfolding setsum_subtractf[THEN sym] apply(rule setsum_cong2) unfolding scaleR.diff_right by auto
-
-lemma has_integral_bound: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "0 \<le> B" "(f has_integral i) ({a..b})" "\<forall>x\<in>{a..b}. norm(f x) \<le> B"
-  shows "norm i \<le> B * content {a..b}"
-proof- let ?P = "content {a..b} > 0" { presume "?P \<Longrightarrow> ?thesis"
-    thus ?thesis proof(cases ?P) case False
-      hence *:"content {a..b} = 0" using content_lt_nz by auto
-      hence **:"i = 0" using assms(2) apply(subst has_integral_null_eq[THEN sym]) by auto
-      show ?thesis unfolding * ** using assms(1) by auto
-    qed auto } assume ab:?P
-  { presume "\<not> ?thesis \<Longrightarrow> False" thus ?thesis by auto }
-  assume "\<not> ?thesis" hence *:"norm i - B * content {a..b} > 0" by auto
-  from assms(2)[unfolded has_integral,rule_format,OF *] guess d apply-by(erule exE conjE)+ note d=this[rule_format]
-  from fine_division_exists[OF this(1), of a b] guess p . note p=this
-  have *:"\<And>s B. norm s \<le> B \<Longrightarrow> \<not> (norm (s - i) < norm i - B)"
-  proof- case goal1 thus ?case unfolding not_less
-    using norm_triangle_sub[of i s] unfolding norm_minus_commute by auto
-  qed show False using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto qed
-
-subsection {* Similar theorems about relationship among components. *}
-
-lemma rsum_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. (f x)$i \<le> (g x)$i"
-  shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)$i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)$i"
-  unfolding setsum_component apply(rule setsum_mono)
-  apply(rule mp) defer apply assumption apply(induct_tac x,rule) unfolding split_conv
-proof- fix a b assume ab:"(a,b) \<in> p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab]
-  from this(3) guess u v apply-by(erule exE)+ note b=this
-  show "(content b *\<^sub>R f a) $ i \<le> (content b *\<^sub>R g a) $ i" unfolding b
-    unfolding Cart_nth.scaleR real_scaleR_def apply(rule mult_left_mono)
-    defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed
-
-lemma has_integral_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "(f has_integral i) s" "(g has_integral j) s"  "\<forall>x\<in>s. (f x)$k \<le> (g x)$k"
-  shows "i$k \<le> j$k"
-proof- have lem:"\<And>a b g i j. \<And>f::real^'n \<Rightarrow> real^'m. (f has_integral i) ({a..b}) \<Longrightarrow> 
-    (g has_integral j) ({a..b}) \<Longrightarrow> \<forall>x\<in>{a..b}. (f x)$k \<le> (g x)$k \<Longrightarrow> i$k \<le> j$k"
-  proof(rule ccontr) case goal1 hence *:"0 < (i$k - j$k) / 3" by auto
-    guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d1=this[rule_format]
-    guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d2=this[rule_format]
-    guess p using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter .
-    note p = this(1) conjunctD2[OF this(2)]  note le_less_trans[OF component_le_norm, of _ _ k]
-    note this[OF d1(2)[OF conjI[OF p(1,2)]]] this[OF d2(2)[OF conjI[OF p(1,3)]]]
-    thus False unfolding Cart_nth.diff using rsum_component_le[OF p(1) goal1(3)] by smt
-  qed let ?P = "\<exists>a b. s = {a..b}"
-  { presume "\<not> ?P \<Longrightarrow> ?thesis" thus ?thesis proof(cases ?P)
-      case True then guess a b apply-by(erule exE)+ note s=this
-      show ?thesis apply(rule lem) using assms[unfolded s] by auto
-    qed auto } assume as:"\<not> ?P"
-  { presume "\<not> ?thesis \<Longrightarrow> False" thus ?thesis by auto }
-  assume "\<not> i$k \<le> j$k" hence ij:"(i$k - j$k) / 3 > 0" by auto
-  note has_integral_altD[OF _ as this] from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
-  have "bounded (ball 0 B1 \<union> ball (0::real^'n) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+
-  from bounded_subset_closed_interval[OF this] guess a b apply- by(erule exE)+
-  note ab = conjunctD2[OF this[unfolded Un_subset_iff]]
-  guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this]
-  guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this]
-  have *:"\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False" by smt(*SMTSMT*)
-  note le_less_trans[OF component_le_norm[of _ k]] note this[OF w1(2)] this[OF w2(2)] moreover
-  have "w1$k \<le> w2$k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately
-  show False unfolding Cart_nth.diff by(rule *) qed
-
-lemma integral_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "f integrable_on s" "g integrable_on s"  "\<forall>x\<in>s. (f x)$k \<le> (g x)$k"
-  shows "(integral s f)$k \<le> (integral s g)$k"
-  apply(rule has_integral_component_le) using integrable_integral assms by auto
-
-lemma has_integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
-  assumes "(f has_integral i) s"  "(g has_integral j) s" "\<forall>x\<in>s. f x \<le> g x"
-  shows "dest_vec1 i \<le> dest_vec1 j" apply(rule has_integral_component_le[OF assms(1-2)])
-  using assms(3) unfolding vector_le_def by auto
-
-lemma integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
-  assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. f x \<le> g x"
-  shows "dest_vec1(integral s f) \<le> dest_vec1(integral s g)"
-  apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto
-
-lemma has_integral_component_pos: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> (f x)$k" shows "0 \<le> i$k"
-  using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2) by auto
-
-lemma integral_component_pos: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> (f x)$k" shows "0 \<le> (integral s f)$k"
-  apply(rule has_integral_component_pos) using assms by auto
-
-lemma has_integral_dest_vec1_pos: fixes f::"real^'n \<Rightarrow> real^1"
-  assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> i"
-  using has_integral_component_pos[OF assms(1), of 1]
-  using assms(2) unfolding vector_le_def by auto
-
-lemma integral_dest_vec1_pos: fixes f::"real^'n \<Rightarrow> real^1"
-  assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> integral s f"
-  apply(rule has_integral_dest_vec1_pos) using assms by auto
-
-lemma has_integral_component_neg: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "(f has_integral i) s" "\<forall>x\<in>s. (f x)$k \<le> 0" shows "i$k \<le> 0"
-  using has_integral_component_le[OF assms(1) has_integral_0] assms(2) by auto
-
-lemma has_integral_dest_vec1_neg: fixes f::"real^'n \<Rightarrow> real^1"
-  assumes "(f has_integral i) s" "\<forall>x\<in>s. f x \<le> 0" shows "i \<le> 0"
-  using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto
-
-lemma has_integral_component_lbound:
-  assumes "(f has_integral i) {a..b}"  "\<forall>x\<in>{a..b}. B \<le> f(x)$k" shows "B * content {a..b} \<le> i$k"
-  using has_integral_component_le[OF has_integral_const assms(1),of "(\<chi> i. B)" k] assms(2)
-  unfolding Cart_lambda_beta vector_scaleR_component by(auto simp add:field_simps)
-
-lemma has_integral_component_ubound: 
-  assumes "(f has_integral i) {a..b}" "\<forall>x\<in>{a..b}. f x$k \<le> B"
-  shows "i$k \<le> B * content({a..b})"
-  using has_integral_component_le[OF assms(1) has_integral_const, of k "vec B"]
-  unfolding vec_component Cart_nth.scaleR using assms(2) by(auto simp add:field_simps)
-
-lemma integral_component_lbound:
-  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)$k"
-  shows "B * content({a..b}) \<le> (integral({a..b}) f)$k"
-  apply(rule has_integral_component_lbound) using assms unfolding has_integral_integral by auto
-
-lemma integral_component_ubound:
-  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. f(x)$k \<le> B"
-  shows "(integral({a..b}) f)$k \<le> B * content({a..b})"
-  apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto
-
-subsection {* Uniform limit of integrable functions is integrable. *}
-
-lemma real_arch_invD:
-  "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
-  by(subst(asm) real_arch_inv)
-
-lemma integrable_uniform_limit: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  assumes "\<forall>e>0. \<exists>g. (\<forall>x\<in>{a..b}. norm(f x - g x) \<le> e) \<and> g integrable_on {a..b}"
-  shows "f integrable_on {a..b}"
-proof- { presume *:"content {a..b} > 0 \<Longrightarrow> ?thesis"
-    show ?thesis apply cases apply(rule *,assumption)
-      unfolding content_lt_nz integrable_on_def using has_integral_null by auto }
-  assume as:"content {a..b} > 0"
-  have *:"\<And>P. \<forall>e>(0::real). P e \<Longrightarrow> \<forall>n::nat. P (inverse (real n+1))" by auto
-  from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format]
-  from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\<lambda>x. x"]] guess i .. note i=this[rule_format]
-  
-  have "Cauchy i" unfolding Cauchy_def
-  proof(rule,rule) fix e::real assume "e>0"
-    hence "e / 4 / content {a..b} > 0" using as by(auto simp add:field_simps)
-    then guess M apply-apply(subst(asm) real_arch_inv) by(erule exE conjE)+ note M=this
-    show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (i m) (i n) < e" apply(rule_tac x=M in exI,rule,rule,rule,rule)
-    proof- case goal1 have "e/4>0" using `e>0` by auto note * = i[unfolded has_integral,rule_format,OF this]
-      from *[of m] guess gm apply-by(erule conjE exE)+ note gm=this[rule_format]
-      from *[of n] guess gn apply-by(erule conjE exE)+ note gn=this[rule_format]
-      from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b] guess p . note p=this
-      have lem2:"\<And>s1 s2 i1 i2. norm(s2 - s1) \<le> e/2 \<Longrightarrow> norm(s1 - i1) < e / 4 \<Longrightarrow> norm(s2 - i2) < e / 4 \<Longrightarrow>norm(i1 - i2) < e"
-      proof- case goal1 have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
-          using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
-          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:group_simps)
-        also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
-        finally show ?case .
-      qed
-      show ?case unfolding vector_dist_norm apply(rule lem2) defer
-        apply(rule gm(2)[OF conjI[OF p(1)]],rule_tac[2] gn(2)[OF conjI[OF p(1)]])
-        using conjunctD2[OF p(2)[unfolded fine_inter]] apply- apply assumption+ apply(rule order_trans)
-        apply(rule rsum_diff_bound[OF p(1), where e="2 / real M"])
-      proof show "2 / real M * content {a..b} \<le> e / 2" unfolding divide_inverse 
-          using M as by(auto simp add:field_simps)
-        fix x assume x:"x \<in> {a..b}"
-        have "norm (f x - g n x) + norm (f x - g m x) \<le> inverse (real n + 1) + inverse (real m + 1)" 
-            using g(1)[OF x, of n] g(1)[OF x, of m] by auto
-        also have "\<dots> \<le> inverse (real M) + inverse (real M)" apply(rule add_mono)
-          apply(rule_tac[!] le_imp_inverse_le) using goal1 M by auto
-        also have "\<dots> = 2 / real M" unfolding real_divide_def by auto
-        finally show "norm (g n x - g m x) \<le> 2 / real M"
-          using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
-          by(auto simp add:group_simps simp add:norm_minus_commute)
-      qed qed qed
-  from this[unfolded convergent_eq_cauchy[THEN sym]] guess s .. note s=this
-
-  show ?thesis unfolding integrable_on_def apply(rule_tac x=s in exI) unfolding has_integral
-  proof(rule,rule)  
-    case goal1 hence *:"e/3 > 0" by auto
-    from s[unfolded Lim_sequentially,rule_format,OF this] guess N1 .. note N1=this
-    from goal1 as have "e / 3 / content {a..b} > 0" by(auto simp add:field_simps)
-    from real_arch_invD[OF this] guess N2 apply-by(erule exE conjE)+ note N2=this
-    from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format]
-    have lem:"\<And>sf sg i. norm(sf - sg) \<le> e / 3 \<Longrightarrow> norm(i - s) < e / 3 \<Longrightarrow> norm(sg - i) < e / 3 \<Longrightarrow> norm(sf - s) < e"
-    proof- case goal1 have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
-        using norm_triangle_ineq[of "sf - sg" "sg - s"]
-        using norm_triangle_ineq[of "sg -  i" " i - s"] by(auto simp add:group_simps)
-      also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
-      finally show ?case .
-    qed
-    show ?case apply(rule_tac x=g' in exI) apply(rule,rule g')
-    proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> g' fine p" note * = g'(2)[OF this]
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e" apply-apply(rule lem[OF _ _ *])
-        apply(rule order_trans,rule rsum_diff_bound[OF p[THEN conjunct1]]) apply(rule,rule g,assumption)
-      proof- have "content {a..b} < e / 3 * (real N2)"
-          using N2 unfolding inverse_eq_divide using as by(auto simp add:field_simps)
-        hence "content {a..b} < e / 3 * (real (N1 + N2) + 1)"
-          apply-apply(rule less_le_trans,assumption) using `e>0` by auto 
-        thus "inverse (real (N1 + N2) + 1) * content {a..b} \<le> e / 3"
-          unfolding inverse_eq_divide by(auto simp add:field_simps)
-        show "norm (i (N1 + N2) - s) < e / 3" by(rule N1[rule_format,unfolded vector_dist_norm],auto)
-      qed qed qed qed
-
-subsection {* Negligible sets. *}
-
-definition "indicator s \<equiv> (\<lambda>x. if x \<in> s then 1 else (0::real))"
-
-lemma dest_vec1_indicator:
- "indicator s x = (if x \<in> s then 1 else 0)" unfolding indicator_def by auto
-
-lemma indicator_pos_le[intro]: "0 \<le> (indicator s x)" unfolding indicator_def by auto
-
-lemma indicator_le_1[intro]: "(indicator s x) \<le> 1" unfolding indicator_def by auto
-
-lemma dest_vec1_indicator_abs_le_1: "abs(indicator s x) \<le> 1"
-  unfolding indicator_def by auto
-
-definition "negligible (s::(real^'n) set) \<equiv> (\<forall>a b. ((indicator s) has_integral 0) {a..b})"
-
-lemma indicator_simps[simp]:"x\<in>s \<Longrightarrow> indicator s x = 1" "x\<notin>s \<Longrightarrow> indicator s x = 0"
-  unfolding indicator_def by auto
-
-subsection {* Negligibility of hyperplane. *}
-
-lemma vsum_nonzero_image_lemma: 
-  assumes "finite s" "g(a) = 0"
-  "\<forall>x\<in>s. \<forall>y\<in>s. f x = f y \<and> x \<noteq> y \<longrightarrow> g(f x) = 0"
-  shows "setsum g {f x |x. x \<in> s \<and> f x \<noteq> a} = setsum (g o f) s"
-  unfolding setsum_iterate[OF assms(1)] apply(subst setsum_iterate) defer
-  apply(rule iterate_nonzero_image_lemma) apply(rule assms monoidal_monoid)+
-  unfolding assms using neutral_add unfolding neutral_add using assms by auto 
-
-lemma interval_doublesplit: shows "{a..b} \<inter> {x . abs(x$k - c) \<le> (e::real)} =
-  {(\<chi> i. if i = k then max (a$k) (c - e) else a$i) .. (\<chi> i. if i = k then min (b$k) (c + e) else b$i)}"
-proof- have *:"\<And>x c e::real. abs(x - c) \<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 by(rule refl) qed
-
-lemma division_doublesplit: assumes "p division_of {a..b::real^'n}" 
-  shows "{l \<inter> {x. abs(x$k - c) \<le> e} |l. l \<in> p \<and> l \<inter> {x. abs(x$k - c) \<le> e} \<noteq> {}} division_of ({a..b} \<inter> {x. abs(x$k - c) \<le> e})"
-proof- have *:"\<And>x c. abs(x - c) \<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" and k=k,unfolded interval_split]
-  note division_split(2)[OF this, where c="c-e" and k=k] 
-  thus ?thesis apply(rule **) unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit
-    apply(rule set_ext) unfolding mem_Collect_eq apply rule apply(erule conjE exE)+ apply(rule_tac x=la in exI) defer
-    apply(erule conjE exE)+ apply(rule_tac x="l \<inter> {x. c + e \<ge> x $ k}" in exI) apply rule defer apply rule
-    apply(rule_tac x=l in exI) by blast+ qed
-
-lemma content_doublesplit: assumes "0 < e"
-  obtains d where "0 < d" "content({a..b} \<inter> {x. abs(x$k - c) \<le> d}) < e"
-proof(cases "content {a..b} = 0")
-  case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit
-    apply(rule le_less_trans[OF content_subset]) defer apply(subst True)
-    unfolding interval_doublesplit[THEN sym] using assms by auto 
-next case False def d \<equiv> "e / 3 / setprod (\<lambda>i. b$i - a$i) (UNIV - {k})"
-  note False[unfolded content_eq_0 not_ex not_le, rule_format]
-  hence prod0:"0 < setprod (\<lambda>i. b$i - a$i) (UNIV - {k})" apply-apply(rule setprod_pos) by smt
-  hence "d > 0" unfolding d_def using assms by(auto simp add:field_simps) thus ?thesis
-  proof(rule that[of d]) have *:"UNIV = insert k (UNIV - {k})" by auto
-    have **:"{a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} \<noteq> {} \<Longrightarrow> 
-      (\<Prod>i\<in>UNIV - {k}. interval_upperbound ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) $ i - interval_lowerbound ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) $ i)
-      = (\<Prod>i\<in>UNIV - {k}. b$i - a$i)" apply(rule setprod_cong,rule refl)
-      unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds by auto
-    show "content ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms)
-      unfolding if_not_P apply(subst *) apply(subst setprod_insert) unfolding **
-      unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds unfolding Cart_lambda_beta if_P[OF refl]
-    proof- have "(min (b $ k) (c + d) - max (a $ k) (c - d)) \<le> 2 * d" by auto
-      also have "... < e / (\<Prod>i\<in>UNIV - {k}. b $ i - a $ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps)
-      finally show "(min (b $ k) (c + d) - max (a $ k) (c - d)) * (\<Prod>i\<in>UNIV - {k}. b $ i - a $ i) < e"
-        unfolding pos_less_divide_eq[OF prod0] . qed auto qed qed
-
-lemma negligible_standard_hyperplane[intro]: "negligible {x. x$k = (c::real)}" 
-  unfolding negligible_def has_integral apply(rule,rule,rule,rule)
-proof- case goal1 from content_doublesplit[OF this,of a b k c] guess d . note d=this let ?i = "indicator {x. x$k = c}"
-  show ?case apply(rule_tac x="\<lambda>x. ball x d" in exI) apply(rule,rule gauge_ball,rule d)
-  proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> (\<lambda>x. ball x d) fine p"
-    have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x$k - c) \<le> d}) *\<^sub>R ?i x)"
-      apply(rule setsum_cong2) unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
-      apply(cases,rule disjI1,assumption,rule disjI2)
-    proof- fix x l assume as:"(x,l)\<in>p" "?i x \<noteq> 0" hence xk:"x$k = c" unfolding indicator_def apply-by(rule ccontr,auto)
-      show "content l = content (l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d})" apply(rule arg_cong[where f=content])
-        apply(rule set_ext,rule,rule) unfolding mem_Collect_eq
-      proof- fix y assume y:"y\<in>l" note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv]
-        note this[unfolded subset_eq mem_ball vector_dist_norm,rule_format,OF y] note le_less_trans[OF component_le_norm[of _ k] this]
-        thus "\<bar>y $ k - c\<bar> \<le> d" unfolding Cart_nth.diff xk by auto
-      qed auto qed
-    note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]]
-    show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e" unfolding diff_0_right * unfolding real_scaleR_def real_norm_def
-      apply(subst abs_of_nonneg) apply(rule setsum_nonneg,rule) unfolding split_paired_all split_conv
-      apply(rule mult_nonneg_nonneg) apply(drule p'(4)) apply(erule exE)+ apply(rule_tac b=b in back_subst)
-      prefer 2 apply(subst(asm) eq_commute) apply assumption
-      apply(subst interval_doublesplit) apply(rule content_pos_le) apply(rule indicator_pos_le)
-    proof- have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}))"
-        apply(rule setsum_mono) unfolding split_paired_all split_conv 
-        apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit intro!:content_pos_le)
-      also have "... < e" apply(subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]])
-      proof- case goal1 have "content ({u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<le> content {u..v}"
-          unfolding interval_doublesplit apply(rule content_subset) unfolding interval_doublesplit[THEN sym] by auto
-        thus ?case unfolding goal1 unfolding interval_doublesplit using content_pos_le by smt
-      next have *:"setsum content {l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} |l. l \<in> snd ` p \<and> l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} \<noteq> {}} \<ge> 0"
-          apply(rule setsum_nonneg,rule) unfolding mem_Collect_eq image_iff apply(erule exE bexE conjE)+ unfolding split_paired_all 
-        proof- fix x l a b assume as:"x = l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}" "(a, b) \<in> p" "l = snd (a, b)"
-          guess u v using p'(4)[OF as(2)] apply-by(erule exE)+ note * = this
-          show "content x \<ge> 0" unfolding as snd_conv * interval_doublesplit by(rule content_pos_le)
-        qed have **:"norm (1::real) \<le> 1" by auto note division_doublesplit[OF p'',unfolded interval_doublesplit]
-        note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym]]
-        note this[unfolded real_scaleR_def real_norm_def class_semiring.semiring_rules, of k c d] note le_less_trans[OF this d(2)]
-        from this[unfolded abs_of_nonneg[OF *]] show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d})) < e"
-          apply(subst vsum_nonzero_image_lemma[of "snd ` p" content "{}", unfolded o_def,THEN sym])
-          apply(rule finite_imageI p' content_empty)+ unfolding forall_in_division[OF p'']
-        proof(rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v
-          assume as:"{m..n} \<in> snd ` p" "{u..v} \<in> snd ` p" "{m..n} \<noteq> {u..v}"  "{m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} = {u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}"
-          have "({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<inter> ({u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<subseteq> {m..n} \<inter> {u..v}" by blast
-          note subset_interior[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]]
-          hence "interior ({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) = {}" unfolding as Int_absorb by auto
-          thus "content ({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) = 0" unfolding interval_doublesplit content_eq_0_interior[THEN sym] .
-        qed qed
-      finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) * ?i x) < e" .
-    qed qed qed
-
-subsection {* A technical lemma about "refinement" of division. *}
-
-lemma tagged_division_finer: fixes p::"((real^'n) \<times> ((real^'n) set)) set"
-  assumes "p tagged_division_of {a..b}" "gauge d"
-  obtains q where "q tagged_division_of {a..b}" "d fine q" "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
-proof-
-  let ?P = "\<lambda>p. p tagged_partial_division_of {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 {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
-    thus ?thesis apply-apply(rule that[of q]) unfolding tagged_division_ofD[OF assms(1)] by auto
-  } fix p::"((real^'n) \<times> ((real^'n) set)) set" assume as:"finite p"
-  show "?P p" apply(rule,rule) using as proof(induct p) 
-    case empty show ?case apply(rule_tac x="{}" in exI) unfolding fine_def by auto
-  next case (insert xk p) guess x k using surj_pair[of xk] apply- by(erule 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 apply-by(erule exE)+ note uv=this
-
-    have "finite {k. \<exists>x. (x, k) \<in> p}" 
-      apply(rule finite_subset[of _ "snd ` p"],rule) unfolding subset_eq image_iff mem_Collect_eq
-      apply(erule exE,rule_tac x="(xa,x)" in bexI) using p by auto
-    hence int:"interior {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],assumption)      
-      apply(rule p(5))  unfolding uv xk apply(rule insertI1,rule insertI2) apply assumption
-      using insert(2) unfolding uv xk by auto
-
-    show ?case proof(cases "{u..v} \<subseteq> d x")
-      case True thus ?thesis apply(rule_tac x="{(x,{u..v})} \<union> q1" in exI) apply rule
-        unfolding * uv apply(rule tagged_division_union,rule tagged_division_of_self)
-        apply(rule p[unfolded xk uv] insertI1)+  apply(rule q1,rule int) 
-        apply(rule,rule fine_union,subst fine_def) defer apply(rule q1)
-        unfolding Ball_def split_paired_All split_conv apply(rule,rule,rule,rule)
-        apply(erule insertE) defer apply(rule UnI2) apply(drule q1(3)[rule_format]) unfolding xk uv by auto
-    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,rule,rule,rule) apply(erule insertE)
-        apply(rule UnI2) defer apply(drule q1(3)[rule_format])using False unfolding xk uv by auto
-    qed qed qed
-
-subsection {* Hence the main theorem about negligible sets. *}
-
-lemma finite_product_dependent: assumes "finite s" "\<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 by auto qed auto
-
-lemma sum_sum_product: assumes "finite s" "\<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_Un_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 by auto
-    show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}" apply(rule finite_product_dependent) using insert by auto
-  qed(insert insert, auto) qed auto
-
-lemma has_integral_negligible: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "negligible s" "\<forall>x\<in>(t - s). f x = 0"
-  shows "(f has_integral 0) t"
-proof- presume P:"\<And>f::real^'n \<Rightarrow> 'a. \<And>a b. (\<forall>x. ~(x \<in> s) \<longrightarrow> f x = 0) \<Longrightarrow> (f has_integral 0) ({a..b})"
-  let ?f = "(\<lambda>x. if x \<in> t then f x else 0)"
-  show ?thesis apply(rule_tac f="?f" in has_integral_eq) apply(rule) unfolding if_P apply(rule refl)
-    apply(subst has_integral_alt) apply(cases,subst if_P,assumption) unfolding if_not_P
-  proof- assume "\<exists>a b. t = {a..b}" then guess a b apply-by(erule exE)+ note t = this
-    show "(?f has_integral 0) t" unfolding t apply(rule P) using assms(2) unfolding t by auto
-  next show "\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b} \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> t then ?f x else 0) has_integral z) {a..b} \<and> norm (z - 0) < e)"
-      apply(safe,rule_tac x=1 in exI,rule) apply(rule zero_less_one,safe) apply(rule_tac x=0 in exI)
-      apply(rule,rule P) using assms(2) by auto
-  qed
-next fix f::"real^'n \<Rightarrow> 'a" and a b::"real^'n" assume assm:"\<forall>x. x \<notin> s \<longrightarrow> f x = 0" 
-  show "(f has_integral 0) {a..b}" unfolding has_integral
-  proof(safe) case goal1
-    hence "\<And>n. e / 2 / ((real n+1) * (2 ^ n)) > 0" 
-      apply-apply(rule divide_pos_pos) defer apply(rule mult_pos_pos) by(auto simp add:field_simps)
-    note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b] note allI[OF this,of "\<lambda>x. x"] 
-    from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]]
-    show ?case apply(rule_tac x="\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x" in exI) 
-    proof safe show "gauge (\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x)" using d(1) unfolding gauge_def by auto
-      fix p assume as:"p tagged_division_of {a..b}" "(\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x) fine p" 
-      let ?goal = "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
-      { presume "p\<noteq>{} \<Longrightarrow> ?goal" thus ?goal apply(cases "p={}") using goal1 by auto  }
-      assume as':"p \<noteq> {}" from real_arch_simple[of "Sup((\<lambda>(x,k). norm(f x)) ` p)"] guess N ..
-      hence N:"\<forall>x\<in>(\<lambda>(x, k). norm (f x)) ` p. x \<le> real N" apply(subst(asm) Sup_finite_le_iff) using as as' by auto
-      have "\<forall>i. \<exists>q. q tagged_division_of {a..b} \<and> (d i) fine q \<and> (\<forall>(x, k)\<in>p. k \<subseteq> (d i) x \<longrightarrow> (x, k) \<in> q)"
-        apply(rule,rule tagged_division_finer[OF as(1) d(1)]) by auto
-      from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]]
-      have *:"\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> 0" apply(rule setsum_nonneg,safe) 
-        unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) apply(drule tagged_division_ofD(4)[OF q(1)]) by auto
-      have **:"\<And>f g s t. finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow> (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t"
-      proof- case goal1 thus ?case apply-apply(rule setsum_le_included[of s t g snd f]) prefer 4
-          apply safe apply(erule_tac x=x in ballE) apply(erule exE) apply(rule_tac x="(xa,x)" in bexI) by auto qed
-      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
-                     norm(setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x) (q i))) {0..N+1}"
-        unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
-        apply(rule order_trans,rule setsum_norm) defer apply(subst sum_sum_product) prefer 3 
-      proof(rule **,safe) show "finite {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i}" apply(rule finite_product_dependent) using q by auto
-        fix i a b assume as'':"(a,b) \<in> q i" show "0 \<le> (real i + 1) * (content b *\<^sub>R indicator s a)"
-          unfolding real_scaleR_def apply(rule mult_nonneg_nonneg) defer apply(rule mult_nonneg_nonneg)
-          using tagged_division_ofD(4)[OF q(1) as''] by auto
-      next fix i::nat show "finite (q i)" using q by auto
-      next fix x k assume xk:"(x,k) \<in> p" def n \<equiv> "nat \<lfloor>norm (f x)\<rfloor>"
-        have *:"norm (f x) \<in> (\<lambda>(x, k). norm (f x)) ` p" using xk by auto
-        have nfx:"real n \<le> norm(f x)" "norm(f x) \<le> real n + 1" unfolding n_def by auto
-        hence "n \<in> {0..N + 1}" using N[rule_format,OF *] by auto
-        moreover  note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv]
-        note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this] note this[unfolded n_def[symmetric]]
-        moreover have "norm (content k *\<^sub>R f x) \<le> (real n + 1) * (content k * indicator s x)"
-        proof(cases "x\<in>s") case False thus ?thesis using assm by auto
-        next case True have *:"content k \<ge> 0" using tagged_division_ofD(4)[OF as(1) xk] by auto
-          moreover have "content k * norm (f x) \<le> content k * (real n + 1)" apply(rule mult_mono) using nfx * by auto
-          ultimately show ?thesis unfolding abs_mult using nfx True by(auto simp add:field_simps)
-        qed ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le> (real y + 1) * (content k *\<^sub>R indicator s x)"
-          apply(rule_tac x=n in exI,safe) apply(rule_tac x=n in exI,rule_tac x="(x,k)" in exI,safe) by auto
-      qed(insert as, auto)
-      also have "... \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {0..N+1}" apply(rule setsum_mono) 
-      proof- case goal1 thus ?case apply(subst mult_commute, subst pos_le_divide_eq[THEN sym])
-          using d(2)[rule_format,of "q i" i] using q[rule_format] by(auto simp add:field_simps)
-      qed also have "... < e * inverse 2 * 2" unfolding real_divide_def setsum_right_distrib[THEN sym]
-        apply(rule mult_strict_left_mono) unfolding power_inverse atLeastLessThanSuc_atLeastAtMost[THEN sym]
-        apply(subst sumr_geometric) using goal1 by auto
-      finally show "?goal" by auto qed qed qed
-
-lemma has_integral_spike: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "negligible s" "(\<forall>x\<in>(t - s). g x = f x)" "(f has_integral y) t"
-  shows "(g has_integral y) t"
-proof- { fix a b::"real^'n" and f g ::"real^'n \<Rightarrow> 'a" and y::'a
-    assume as:"\<forall>x \<in> {a..b} - s. g x = f x" "(f has_integral y) {a..b}"
-    have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) {a..b}" apply(rule has_integral_add[OF as(2)])
-      apply(rule has_integral_negligible[OF assms(1)]) using as by auto
-    hence "(g has_integral y) {a..b}" by auto } note * = this
-  show ?thesis apply(subst has_integral_alt) using assms(2-) apply-apply(rule cond_cases,safe)
-    apply(rule *, assumption+) apply(subst(asm) has_integral_alt) unfolding if_not_P
-    apply(erule_tac x=e in allE,safe,rule_tac x=B in exI,safe) apply(erule_tac x=a in allE,erule_tac x=b in allE,safe)
-    apply(rule_tac x=z in exI,safe) apply(rule *[where fa2="\<lambda>x. if x\<in>t then f x else 0"]) by auto qed
-
-lemma has_integral_spike_eq:
-  assumes "negligible s" "\<forall>x\<in>(t - s). g x = f x"
-  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
-  apply rule apply(rule_tac[!] has_integral_spike[OF assms(1)]) using assms(2) by auto
-
-lemma integrable_spike: assumes "negligible s" "\<forall>x\<in>(t - s). g x = f x" "f integrable_on t"
-  shows "g integrable_on  t"
-  using assms unfolding integrable_on_def apply-apply(erule exE)
-  apply(rule,rule has_integral_spike) by fastsimp+
-
-lemma integral_spike: assumes "negligible s" "\<forall>x\<in>(t - s). g x = f x"
-  shows "integral t f = integral t g"
-  unfolding integral_def using has_integral_spike_eq[OF assms] by auto
-
-subsection {* Some other trivialities about negligible sets. *}
-
-lemma negligible_subset[intro]: assumes "negligible s" "t \<subseteq> s" shows "negligible t" unfolding negligible_def 
-proof(safe) case goal1 show ?case using assms(1)[unfolded negligible_def,rule_format,of a b]
-    apply-apply(rule has_integral_spike[OF assms(1)]) defer apply assumption
-    using assms(2) unfolding indicator_def by auto qed
-
-lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible(s - t)" using assms by auto
-
-lemma negligible_inter: assumes "negligible s \<or> negligible t" shows "negligible(s \<inter> t)" using assms by auto
-
-lemma negligible_union: assumes "negligible s" "negligible t" shows "negligible (s \<union> t)" unfolding negligible_def 
-proof safe case goal1 note assm = assms[unfolded negligible_def,rule_format,of a b]
-  thus ?case apply(subst has_integral_spike_eq[OF assms(2)])
-    defer apply assumption unfolding indicator_def by auto qed
-
-lemma negligible_union_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> (negligible s \<and> negligible t)"
-  using negligible_union by auto
-
-lemma negligible_sing[intro]: "negligible {a::real^'n}" 
-proof- guess x using UNIV_witness[where 'a='n] ..
-  show ?thesis using negligible_standard_hyperplane[of x "a$x"] by auto qed
-
-lemma negligible_insert[simp]: "negligible(insert a s) \<longleftrightarrow> negligible s"
-  apply(subst insert_is_Un) unfolding negligible_union_eq by auto
-
-lemma negligible_empty[intro]: "negligible {}" by auto
-
-lemma negligible_finite[intro]: assumes "finite s" shows "negligible s"
-  using assms apply(induct s) by auto
-
-lemma negligible_unions[intro]: assumes "finite s" "\<forall>t\<in>s. negligible t" shows "negligible(\<Union>s)"
-  using assms by(induct,auto) 
-
-lemma negligible:  "negligible s \<longleftrightarrow> (\<forall>t::(real^'n) set. (indicator s has_integral 0) t)"
-  apply safe defer apply(subst negligible_def)
-proof- fix t::"(real^'n) set" assume as:"negligible s"
-  have *:"(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)" by(rule ext,auto)  
-  show "(indicator s has_integral 0) t" apply(subst has_integral_alt)
-    apply(cases,subst if_P,assumption) unfolding if_not_P apply(safe,rule as[unfolded negligible_def,rule_format])
-    apply(rule_tac x=1 in exI) apply(safe,rule zero_less_one) apply(rule_tac x=0 in exI)
-    using negligible_subset[OF as,of "s \<inter> t"] unfolding negligible_def indicator_def unfolding * by auto qed auto
-
-subsection {* Finite case of the spike theorem is quite commonly needed. *}
-
-lemma has_integral_spike_finite: assumes "finite s" "\<forall>x\<in>t-s. g x = f x" 
-  "(f has_integral y) t" shows "(g has_integral y) t"
-  apply(rule has_integral_spike) using assms by auto
-
-lemma has_integral_spike_finite_eq: assumes "finite s" "\<forall>x\<in>t-s. g x = f x"
-  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
-  apply rule apply(rule_tac[!] has_integral_spike_finite) using assms by auto
-
-lemma integrable_spike_finite:
-  assumes "finite s" "\<forall>x\<in>t-s. g x = f x" "f integrable_on t" shows "g integrable_on  t"
-  using assms unfolding integrable_on_def apply safe apply(rule_tac x=y in exI)
-  apply(rule has_integral_spike_finite) by auto
-
-subsection {* In particular, the boundary of an interval is negligible. *}
-
-lemma negligible_frontier_interval: "negligible({a..b} - {a<..<b})"
-proof- let ?A = "\<Union>((\<lambda>k. {x. x$k = a$k} \<union> {x. x$k = b$k}) ` UNIV)"
-  have "{a..b} - {a<..<b} \<subseteq> ?A" apply rule unfolding Diff_iff mem_interval not_all
-    apply(erule conjE exE)+ apply(rule_tac X="{x. x $ xa = a $ xa} \<union> {x. x $ xa = b $ xa}" in UnionI)
-    apply(erule_tac[!] x=xa in allE) by auto
-  thus ?thesis apply-apply(rule negligible_subset[of ?A]) apply(rule negligible_unions[OF finite_imageI]) by auto qed
-
-lemma has_integral_spike_interior:
-  assumes "\<forall>x\<in>{a<..<b}. g x = f x" "(f has_integral y) ({a..b})" shows "(g has_integral y) ({a..b})"
-  apply(rule has_integral_spike[OF negligible_frontier_interval _ assms(2)]) using assms(1) by auto
-
-lemma has_integral_spike_interior_eq:
-  assumes "\<forall>x\<in>{a<..<b}. g x = f x" shows "((f has_integral y) ({a..b}) \<longleftrightarrow> (g has_integral y) ({a..b}))"
-  apply rule apply(rule_tac[!] has_integral_spike_interior) using assms by auto
-
-lemma integrable_spike_interior: assumes "\<forall>x\<in>{a<..<b}. g x = f x" "f integrable_on {a..b}" shows "g integrable_on {a..b}"
-  using  assms unfolding integrable_on_def using has_integral_spike_interior[OF assms(1)] by auto
-
-subsection {* Integrability of continuous functions. *}
-
-lemma neutral_and[simp]: "neutral op \<and> = True"
-  unfolding neutral_def apply(rule some_equality) by auto
-
-lemma monoidal_and[intro]: "monoidal op \<and>" unfolding monoidal_def by auto
-
-lemma iterate_and[simp]: assumes "finite s" shows "(iterate op \<and>) s p \<longleftrightarrow> (\<forall>x\<in>s. p x)" using assms
-apply induct unfolding iterate_insert[OF monoidal_and] by auto
-
-lemma operative_division_and: assumes "operative op \<and> P" "d division_of {a..b}"
-  shows "(\<forall>i\<in>d. P i) \<longleftrightarrow> P {a..b}"
-  using operative_division[OF monoidal_and assms] division_of_finite[OF assms(2)] by auto
-
-lemma operative_approximable: assumes "0 \<le> e" fixes f::"real^'n \<Rightarrow> 'a::banach"
-  shows "operative op \<and> (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::real^'n)) \<le> e) \<and> g integrable_on i)" unfolding operative_def neutral_and
-proof safe fix a b::"real^'n" { assume "content {a..b} = 0"
-    thus "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" 
-      apply(rule_tac x=f in exI) using assms by(auto intro!:integrable_on_null) }
-  { fix c k g assume as:"\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}"
-    show "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. x $ k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. x $ k \<le> c}"
-      "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $ k}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. c \<le> x $ k}"
-      apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2)] by auto }
-  fix c k g1 g2 assume as:"\<forall>x\<in>{a..b} \<inter> {x. x $ k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on {a..b} \<inter> {x. x $ k \<le> c}"
-                          "\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $ k}. norm (f x - g2 x) \<le> e" "g2 integrable_on {a..b} \<inter> {x. c \<le> x $ k}"
-  let ?g = "\<lambda>x. if x$k = c then f x else if x$k \<le> c then g1 x else g2 x"
-  show "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" apply(rule_tac x="?g" in exI)
-  proof safe case goal1 thus ?case apply- apply(cases "x$k=c", case_tac "x$k < c") using as assms by auto
-  next case goal2 presume "?g integrable_on {a..b} \<inter> {x. x $ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $ k \<ge> c}"
-    then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this]
-    show ?case unfolding integrable_on_def by auto
-  next show "?g integrable_on {a..b} \<inter> {x. x $ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $ k \<ge> c}"
-      apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using as(2,4) by auto qed qed
-
-lemma approximable_on_division: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  assumes "0 \<le> e" "d division_of {a..b}" "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
-  obtains g where "\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}"
-proof- note * = operative_division[OF monoidal_and operative_approximable[OF assms(1)] assms(2)]
-  note this[unfolded iterate_and[OF division_of_finite[OF assms(2)]]] from assms(3)[unfolded this[of f]]
-  guess g .. thus thesis apply-apply(rule that[of g]) by auto qed
-
-lemma integrable_continuous: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}"
-proof(rule integrable_uniform_limit,safe) fix e::real assume e:"0 < e"
-  from compact_uniformly_continuous[OF assms compact_interval,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d ..
-  note d=conjunctD2[OF this,rule_format]
-  from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this
-  note p' = tagged_division_ofD[OF p(1)]
-  have *:"\<forall>i\<in>snd ` p. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
-  proof(safe,unfold snd_conv) fix x l assume as:"(x,l) \<in> p" 
-    from p'(4)[OF this] guess a b apply-by(erule exE)+ note l=this
-    show "\<exists>g. (\<forall>x\<in>l. norm (f x - g x) \<le> e) \<and> g integrable_on l" apply(rule_tac x="\<lambda>y. f x" in exI)
-    proof safe show "(\<lambda>y. f x) integrable_on l" unfolding integrable_on_def l by(rule,rule has_integral_const)
-      fix y assume y:"y\<in>l" note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this]
-      note d(2)[OF _ _ this[unfolded mem_ball]]
-      thus "norm (f y - f x) \<le> e" using y p'(2-3)[OF as] unfolding vector_dist_norm l norm_minus_commute by fastsimp qed qed
-  from e have "0 \<le> e" by auto from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g .
-  thus "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" by auto qed 
-
-subsection {* Specialization of additivity to one dimension. *}
-
-lemma operative_1_lt: assumes "monoidal opp"
-  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real^1} = neutral opp) \<and>
-                (\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
-  unfolding operative_def content_eq_0_1 forall_1 vector_le_def vector_less_def
-proof safe fix a b c::"real^1" assume as:"\<forall>a b c. f {a..b} = opp (f ({a..b} \<inter> {x. x $ 1 \<le> c})) (f ({a..b} \<inter> {x. c \<le> x $ 1}))" "a $ 1 < c $ 1" "c $ 1 < b $ 1"
-    from this(2-) have "{a..b} \<inter> {x. x $ 1 \<le> c $ 1} = {a..c}" "{a..b} \<inter> {x. x $ 1 \<ge> c $ 1} = {c..b}" by auto
-    thus "opp (f {a..c}) (f {c..b}) = f {a..b}" unfolding as(1)[rule_format,of a b "c$1"] by auto
-next fix a b::"real^1" and c::real
-  assume as:"\<forall>a b. b $ 1 \<le> a $ 1 \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a $ 1 < c $ 1 \<and> c $ 1 < b $ 1 \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}"
-  show "f {a..b} = opp (f ({a..b} \<inter> {x. x $ 1 \<le> c})) (f ({a..b} \<inter> {x. c \<le> x $ 1}))"
-  proof(cases "c \<in> {a$1 .. b$1}")
-    case False hence "c<a$1 \<or> c>b$1" by auto
-    thus ?thesis apply-apply(erule disjE)
-    proof- assume "c<a$1" hence *:"{a..b} \<inter> {x. x $ 1 \<le> c} = {1..0}"  "{a..b} \<inter> {x. c \<le> x $ 1} = {a..b}" by auto
-      show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
-    next   assume "b$1<c" hence *:"{a..b} \<inter> {x. x $ 1 \<le> c} = {a..b}"  "{a..b} \<inter> {x. c \<le> x $ 1} = {1..0}" by auto
-      show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
-    qed
-  next case True hence *:"min (b $ 1) c = c" "max (a $ 1) c = c" by auto
-    show ?thesis unfolding interval_split num1_eq_iff if_True * vec_def[THEN sym]
-    proof(cases "c = a$1 \<or> c = b$1")
-      case False thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})"
-        apply-apply(subst as(2)[rule_format]) using True by auto
-    next case True thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})" apply-
-      proof(erule disjE) assume "c=a$1" hence *:"a = vec1 c" unfolding Cart_eq by auto 
-        hence "f {a..vec1 c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
-        thus ?thesis using assms unfolding * by auto
-      next assume "c=b$1" hence *:"b = vec1 c" unfolding Cart_eq by auto 
-        hence "f {vec1 c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
-        thus ?thesis using assms unfolding * by auto qed qed qed qed
-
-lemma operative_1_le: assumes "monoidal opp"
-  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real^1} = neutral opp) \<and>
-                (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
-unfolding operative_1_lt[OF assms]
-proof safe fix a b c::"real^1" assume as:"\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a < c" "c < b"
-  show "opp (f {a..c}) (f {c..b}) = f {a..b}" apply(rule as(1)[rule_format]) using as(2-) unfolding vector_le_def vector_less_def by auto
-next fix a b c ::"real^1"
-  assume "\<forall>a b. b \<le> a \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a \<le> c" "c \<le> b"
-  note as = this[rule_format]
-  show "opp (f {a..c}) (f {c..b}) = f {a..b}"
-  proof(cases "c = a \<or> c = b")
-    case False thus ?thesis apply-apply(subst as(2)) using as(3-) unfolding vector_le_def vector_less_def Cart_eq by(auto simp del:dest_vec1_eq)
-    next case True thus ?thesis apply-
-      proof(erule disjE) assume *:"c=a" hence "f {a..c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
-        thus ?thesis using assms unfolding * by auto
-      next               assume *:"c=b" hence "f {c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
-        thus ?thesis using assms unfolding * by auto qed qed qed 
-
-subsection {* Special case of additivity we need for the FCT. *}
-
-lemma additive_tagged_division_1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
-  assumes "dest_vec1 a \<le> dest_vec1 b" "p tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a"
-proof- let ?f = "(\<lambda>k::(real^1) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
-  have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty_1
-    by(auto simp add:not_less interval_bound_1 vector_less_def)
-  have **:"{a..b} \<noteq> {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)]
-  note * = this[unfolded if_not_P[OF **] interval_bound_1[OF assms(1)],THEN sym ]
-  show ?thesis unfolding * apply(subst setsum_iterate[THEN sym]) defer
-    apply(rule setsum_cong2) unfolding split_paired_all split_conv using assms(2) by auto qed
-
-subsection {* A useful lemma allowing us to factor out the content size. *}
-
-lemma has_integral_factor_content:
-  "(f has_integral i) {a..b} \<longleftrightarrow> (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p
-    \<longrightarrow> norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a..b}))"
-proof(cases "content {a..b} = 0")
-  case True show ?thesis unfolding has_integral_null_eq[OF True] apply safe
-    apply(rule,rule,rule gauge_trivial,safe) unfolding setsum_content_null[OF True] True defer 
-    apply(erule_tac x=1 in allE,safe) defer apply(rule fine_division_exists[of _ a b],assumption)
-    apply(erule_tac x=p in allE) unfolding setsum_content_null[OF True] by auto
-next case False note F = this[unfolded content_lt_nz[THEN sym]]
-  let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)"
-  show ?thesis apply(subst has_integral)
-  proof safe fix e::real assume e:"e>0"
-    { assume "\<forall>e>0. ?P e op <" thus "?P (e * content {a..b}) op \<le>" apply(erule_tac x="e * content {a..b}" in allE)
-        apply(erule impE) defer apply(erule exE,rule_tac x=d in exI)
-        using F e by(auto simp add:field_simps intro:mult_pos_pos) }
-    {  assume "\<forall>e>0. ?P (e * content {a..b}) op \<le>" thus "?P e op <" apply(erule_tac x="e / 2 / content {a..b}" in allE)
-        apply(erule impE) defer apply(erule exE,rule_tac x=d in exI)
-        using F e by(auto simp add:field_simps intro:mult_pos_pos) } qed qed
-
-subsection {* Fundamental theorem of calculus. *}
-
-lemma fundamental_theorem_of_calculus: fixes f::"real^1 \<Rightarrow> 'a::banach"
-  assumes "a \<le> b"  "\<forall>x\<in>{a..b}. ((f o vec1) has_vector_derivative f'(vec1 x)) (at x within {a..b})"
-  shows "(f' has_integral (f(vec1 b) - f(vec1 a))) ({vec1 a..vec1 b})"
-unfolding has_integral_factor_content
-proof safe fix e::real assume e:"e>0" have ab:"dest_vec1 (vec1 a) \<le> dest_vec1 (vec1 b)" using assms(1) by auto
-  note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt]
-  have *:"\<And>P Q. \<forall>x\<in>{a..b}. P x \<and> (\<forall>e>0. \<exists>d>0. Q x e d) \<Longrightarrow> \<forall>x. \<exists>(d::real)>0. x\<in>{a..b} \<longrightarrow> Q x e d" using e by blast
-  note this[OF assm,unfolded gauge_existence_lemma] from choice[OF this,unfolded Ball_def[symmetric]]
-  guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
-  show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
-                 norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \<le> e * content {vec1 a..vec1 b})"
-    apply(rule_tac x="\<lambda>x. ball x (d (dest_vec1 x))" in exI,safe)
-    apply(rule gauge_ball_dependent,rule,rule d(1))
-  proof- fix p assume as:"p tagged_division_of {vec1 a..vec1 b}" "(\<lambda>x. ball x (d (dest_vec1 x))) fine p"
-    show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \<le> e * content {vec1 a..vec1 b}" 
-      unfolding content_1[OF ab] additive_tagged_division_1[OF ab as(1),of f,THEN sym]
-      unfolding vector_minus_component[THEN sym] additive_tagged_division_1[OF ab as(1),of "\<lambda>x. x",THEN sym]
-      apply(subst dest_vec1_setsum) unfolding setsum_right_distrib defer unfolding setsum_subtractf[THEN sym] 
-    proof(rule setsum_norm_le,safe) fix x k assume "(x,k)\<in>p"
-      note xk = tagged_division_ofD(2-4)[OF as(1) this] from this(3) guess u v apply-by(erule exE)+ note k=this
-      have *:"dest_vec1 u \<le> dest_vec1 v" using xk unfolding k by auto
-      have ball:"\<forall>xa\<in>k. xa \<in> ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
-      have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \<le> norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)"
-        apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
-        unfolding scaleR.diff_left by(auto simp add:group_simps)
-      also have "... \<le> e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)"
-        apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4
-        apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1])
-        using ball[rule_format,of u] ball[rule_format,of v] 
-        using xk(1-2) unfolding k subset_eq by(auto simp add:vector_dist_norm norm_real)
-      also have "... \<le> e * dest_vec1 (interval_upperbound k - interval_lowerbound k)"
-        unfolding k interval_bound_1[OF *] using xk(1) unfolding k by(auto simp add:vector_dist_norm norm_real field_simps)
-      finally show "norm (content k *\<^sub>R f' x - (f (interval_upperbound k) - f (interval_lowerbound k))) \<le>
-        e * dest_vec1 (interval_upperbound k - interval_lowerbound k)" unfolding k interval_bound_1[OF *] content_1[OF *] .
-    qed(insert as, auto) qed qed
-
-subsection {* Attempt a systematic general set of "offset" results for components. *}
-
-lemma gauge_modify:
-  assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
-  shows "gauge (\<lambda>x y. d (f x) (f y))"
-  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) unfolding mem_def Collect_def by auto
-
-subsection {* Only need trivial subintervals if the interval itself is trivial. *}
-
-lemma division_of_nontrivial: fixes s::"(real^'n) set set"
-  assumes "s division_of {a..b}" "content({a..b}) \<noteq> 0"
-  shows "{k. k \<in> s \<and> content k \<noteq> 0} division_of {a..b}" using assms(1) apply-
-proof(induct "card s" arbitrary:s rule:nat_less_induct)
-  fix s::"(real^'n) set set" assume assm:"s division_of {a..b}"
-    "\<forall>m<card s. \<forall>x. m = card x \<longrightarrow> x division_of {a..b} \<longrightarrow> {k \<in> x. content k \<noteq> 0} division_of {a..b}" 
-  note s = division_ofD[OF assm(1)] let ?thesis = "{k \<in> s. content k \<noteq> 0} division_of {a..b}"
-  { presume *:"{k \<in> s. content k \<noteq> 0} \<noteq> s \<Longrightarrow> ?thesis"
-    show ?thesis apply cases defer apply(rule *,assumption) using assm(1) by auto }
-  assume noteq:"{k \<in> s. content k \<noteq> 0} \<noteq> s"
-  then obtain k where k:"k\<in>s" "content k = 0" by auto
-  from s(4)[OF k(1)] guess c d apply-by(erule exE)+ note k=k this
-  from k have "card s > 0" unfolding card_gt_0_iff using assm(1) by auto
-  hence card:"card (s - {k}) < card s" using assm(1) k(1) apply(subst card_Diff_singleton_if) by auto
-  have *:"closed (\<Union>(s - {k}))" apply(rule closed_Union) defer apply rule apply(drule DiffD1,drule s(4))
-    apply safe apply(rule closed_interval) using assm(1) by auto
-  have "k \<subseteq> \<Union>(s - {k})" apply safe apply(rule *[unfolded closed_limpt,rule_format]) unfolding islimpt_approachable
-  proof safe fix x and e::real assume as:"x\<in>k" "e>0"
-    from k(2)[unfolded k content_eq_0] guess i .. 
-    hence i:"c$i = d$i" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by smt
-    hence xi:"x$i = d$i" using as unfolding k mem_interval by smt
-    def y \<equiv> "(\<chi> j. if j = i then if c$i \<le> (a$i + b$i) / 2 then c$i + min e (b$i - c$i) / 2 else c$i - min e (c$i - a$i) / 2 else x$j)"
-    show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> dist x' x < e" apply(rule_tac x=y in bexI) 
-    proof have "d \<in> {c..d}" using s(3)[OF k(1)] unfolding k interval_eq_empty mem_interval by(fastsimp simp add: not_less)
-      hence "d \<in> {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN spec[where x=i]]
-      hence xyi:"y$i \<noteq> x$i" unfolding y_def unfolding i xi Cart_lambda_beta if_P[OF refl]
-        apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2) using assms(2)[unfolded content_eq_0] by smt+ 
-      thus "y \<noteq> x" unfolding Cart_eq by auto
-      have *:"UNIV = insert i (UNIV - {i})" by auto
-      have "norm (y - x) < e + setsum (\<lambda>i. 0) (UNIV::'n set)" apply(rule le_less_trans[OF norm_le_l1])
-        apply(subst *,subst setsum_insert) prefer 3 apply(rule add_less_le_mono)
-      proof- show "\<bar>(y - x) $ i\<bar> < e" unfolding y_def Cart_lambda_beta vector_minus_component if_P[OF refl]
-          apply(cases) apply(subst if_P,assumption) unfolding if_not_P unfolding i xi using di as(2) by auto
-        show "(\<Sum>i\<in>UNIV - {i}. \<bar>(y - x) $ i\<bar>) \<le> (\<Sum>i\<in>UNIV. 0)" unfolding y_def by auto 
-      qed auto thus "dist y x < e" unfolding vector_dist_norm by auto
-      have "y\<notin>k" unfolding k mem_interval apply rule apply(erule_tac x=i in allE) using xyi unfolding k i xi by auto
-      moreover have "y \<in> \<Union>s" unfolding s mem_interval
-      proof note simps = y_def Cart_lambda_beta if_not_P
-        fix j::'n show "a $ j \<le> y $ j \<and> y $ j \<le> b $ j" 
-        proof(cases "j = i") case False have "x \<in> {a..b}" using s(2)[OF k(1)] as(1) by auto
-          thus ?thesis unfolding simps if_not_P[OF False] unfolding mem_interval by auto
-        next case True note T = this show ?thesis
-          proof(cases "c $ i \<le> (a $ i + b $ i) / 2")
-            case True show ?thesis unfolding simps if_P[OF T] if_P[OF True] unfolding i
-              using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) 
-          next case False thus ?thesis unfolding simps if_P[OF T] if_not_P[OF False] unfolding i
-              using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps)
-          qed qed qed
-      ultimately show "y \<in> \<Union>(s - {k})" by auto
-    qed qed hence "\<Union>(s - {k}) = {a..b}" unfolding s(6)[THEN sym] by auto
-  hence  "{ka \<in> s - {k}. content ka \<noteq> 0} division_of {a..b}" apply-apply(rule assm(2)[rule_format,OF card refl])
-    apply(rule division_ofI) defer apply(rule_tac[1-4] s) using assm(1) by auto
-  moreover have "{ka \<in> s - {k}. content ka \<noteq> 0} = {k \<in> s. content k \<noteq> 0}" using k by auto ultimately show ?thesis by auto qed
-
-subsection {* Integrabibility on subintervals. *}
-
-lemma operative_integrable: fixes f::"real^'n \<Rightarrow> 'a::banach" shows
-  "operative op \<and> (\<lambda>i. f integrable_on i)"
-  unfolding operative_def neutral_and apply safe apply(subst integrable_on_def)
-  unfolding has_integral_null_eq apply(rule,rule refl) apply(rule,assumption)+
-  unfolding integrable_on_def by(auto intro: has_integral_split)
-
-lemma integrable_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach" 
-  assumes "f integrable_on {a..b}" "{c..d} \<subseteq> {a..b}" shows "f integrable_on {c..d}" 
-  apply(cases "{c..d} = {}") defer apply(rule partial_division_extend_1[OF assms(2)],assumption)
-  using operative_division_and[OF operative_integrable,THEN sym,of _ _ _ f] assms(1) by auto
-
-subsection {* Combining adjacent intervals in 1 dimension. *}
-
-lemma has_integral_combine: assumes "(a::real^1) \<le> c" "c \<le> b"
-  "(f has_integral i) {a..c}" "(f has_integral (j::'a::banach)) {c..b}"
-  shows "(f has_integral (i + j)) {a..b}"
-proof- note operative_integral[of f, unfolded operative_1_le[OF monoidal_lifted[OF monoidal_monoid]]]
-  note conjunctD2[OF this,rule_format] note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]]
-  hence "f integrable_on {a..b}" apply- apply(rule ccontr) apply(subst(asm) if_P) defer
-    apply(subst(asm) if_P) using assms(3-) by auto
-  with * show ?thesis apply-apply(subst(asm) if_P) defer apply(subst(asm) if_P) defer apply(subst(asm) if_P)
-    unfolding lifted.simps using assms(3-) by(auto simp add: integrable_on_def integral_unique) qed
-
-lemma integral_combine: fixes f::"real^1 \<Rightarrow> 'a::banach"
-  assumes "a \<le> c" "c \<le> b" "f integrable_on ({a..b})"
-  shows "integral {a..c} f + integral {c..b} f = integral({a..b}) f"
-  apply(rule integral_unique[THEN sym]) apply(rule has_integral_combine[OF assms(1-2)])
-  apply(rule_tac[!] integrable_integral integrable_subinterval[OF assms(3)])+ using assms(1-2) by auto
-
-lemma integrable_combine: fixes f::"real^1 \<Rightarrow> 'a::banach"
-  assumes "a \<le> c" "c \<le> b" "f integrable_on {a..c}" "f integrable_on {c..b}"
-  shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by(fastsimp intro!:has_integral_combine)
-
-subsection {* Reduce integrability to "local" integrability. *}
-
-lemma integrable_on_little_subintervals: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  assumes "\<forall>x\<in>{a..b}. \<exists>d>0. \<forall>u v. x \<in> {u..v} \<and> {u..v} \<subseteq> ball x d \<and> {u..v} \<subseteq> {a..b} \<longrightarrow> f integrable_on {u..v}"
-  shows "f integrable_on {a..b}"
-proof- have "\<forall>x. \<exists>d. x\<in>{a..b} \<longrightarrow> d>0 \<and> (\<forall>u v. x \<in> {u..v} \<and> {u..v} \<subseteq> ball x d \<and> {u..v} \<subseteq> {a..b} \<longrightarrow> f integrable_on {u..v})"
-    using assms by auto note this[unfolded gauge_existence_lemma] from choice[OF this] guess d .. note d=this[rule_format]
-  guess p apply(rule fine_division_exists[OF gauge_ball_dependent,of d a b]) using d by auto note p=this(1-2)
-  note division_of_tagged_division[OF this(1)] note * = operative_division_and[OF operative_integrable,OF this,THEN sym,of f]
-  show ?thesis unfolding * apply safe unfolding snd_conv
-  proof- fix x k assume "(x,k) \<in> p" note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this]
-    thus "f integrable_on k" apply safe apply(rule d[THEN conjunct2,rule_format,of x]) by auto qed qed
-
-subsection {* Second FCT or existence of antiderivative. *}
-
-lemma integrable_const[intro]:"(\<lambda>x. c) integrable_on {a..b}"
-  unfolding integrable_on_def by(rule,rule has_integral_const)
-
-lemma integral_has_vector_derivative: fixes f::"real \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a..b} f" "x \<in> {a..b}"
-  shows "((\<lambda>u. integral {vec a..vec u} (f o dest_vec1)) has_vector_derivative f(x)) (at x within {a..b})"
-  unfolding has_vector_derivative_def has_derivative_within_alt
-apply safe apply(rule scaleR.bounded_linear_left)
-proof- fix e::real assume e:"e>0"
-  note compact_uniformly_continuous[OF assms(1) compact_real_interval,unfolded uniformly_continuous_on_def]
-  from this[rule_format,OF e] guess d apply-by(erule conjE exE)+ note d=this[rule_format]
-  let ?I = "\<lambda>a b. integral {vec1 a..vec1 b} (f \<circ> dest_vec1)"
-  show "\<exists>d>0. \<forall>y\<in>{a..b}. norm (y - x) < d \<longrightarrow> norm (?I a y - ?I a x - (y - x) *\<^sub>R f x) \<le> e * norm (y - x)"
-  proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x")
-      case False have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous)
-        apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
-      hence *:"?I a y - ?I a x = ?I x y" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
-        using False unfolding not_less using assms(2) goal1 by auto
-      have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto
-      show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
-        defer apply(rule has_integral_sub) apply(rule integrable_integral)
-        apply(rule integrable_subinterval,rule integrable_continuous) apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+
-      proof- show "{vec1 x..vec1 y} \<subseteq> {vec1 a..vec1 b}" using goal1 assms(2) by auto
-        have *:"y - x = norm(y - x)" using False by auto
-        show "((\<lambda>xa. f x) has_integral (y - x) *\<^sub>R f x) {vec1 x..vec1 y}" apply(subst *) unfolding ** by auto
-        show "\<forall>xa\<in>{vec1 x..vec1 y}. norm (f (dest_vec1 xa) - f x) \<le> e" apply safe apply(rule less_imp_le)
-          apply(rule d(2)[unfolded vector_dist_norm]) using assms(2) using goal1 by auto
-      qed(insert e,auto)
-    next case True have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous)
-        apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
-      hence *:"?I a x - ?I a y = ?I y x" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
-        using True using assms(2) goal1 by auto
-      have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto
-      have ***:"\<And>fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto 
-      show ?thesis apply(subst ***) unfolding norm_minus_cancel **
-        apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
-        defer apply(rule has_integral_sub) apply(subst minus_minus[THEN sym]) unfolding minus_minus
-        apply(rule integrable_integral) apply(rule integrable_subinterval,rule integrable_continuous)
-        apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+
-      proof- show "{vec1 y..vec1 x} \<subseteq> {vec1 a..vec1 b}" using goal1 assms(2) by auto
-        have *:"x - y = norm(y - x)" using True by auto
-        show "((\<lambda>xa. f x) has_integral (x - y) *\<^sub>R f x) {vec1 y..vec1 x}" apply(subst *) unfolding ** by auto
-        show "\<forall>xa\<in>{vec1 y..vec1 x}. norm (f (dest_vec1 xa) - f x) \<le> e" apply safe apply(rule less_imp_le)
-          apply(rule d(2)[unfolded vector_dist_norm]) using assms(2) using goal1 by auto
-      qed(insert e,auto) qed qed qed
-
-lemma integral_has_vector_derivative': fixes f::"real^1 \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a..b} f" "x \<in> {a..b}"
-  shows "((\<lambda>u. (integral {a..vec u} f)) has_vector_derivative f x) (at (x$1) within {a$1..b$1})"
-  using integral_has_vector_derivative[OF continuous_on_o_vec1[OF assms(1)], of "x$1"]
-  unfolding o_def vec1_dest_vec1 using assms(2) by auto
-
-lemma antiderivative_continuous: assumes "continuous_on {a..b::real} f"
-  obtains g where "\<forall>x\<in> {a..b}. (g has_vector_derivative (f(x)::_::banach)) (at x within {a..b})"
-  apply(rule that,rule) using integral_has_vector_derivative[OF assms] by auto
-
-subsection {* Combined fundamental theorem of calculus. *}
-
-lemma antiderivative_integral_continuous: fixes f::"real \<Rightarrow> 'a::banach" assumes "continuous_on {a..b} f"
-  obtains g where "\<forall>u\<in>{a..b}. \<forall>v \<in> {a..b}. u \<le> v \<longrightarrow> ((f o dest_vec1) has_integral (g v - g u)) {vec u..vec v}"
-proof- from antiderivative_continuous[OF assms] guess g . note g=this
-  show ?thesis apply(rule that[of g])
-  proof safe case goal1 have "\<forall>x\<in>{u..v}. (g has_vector_derivative f x) (at x within {u..v})"
-      apply(rule,rule has_vector_derivative_within_subset) apply(rule g[rule_format]) using goal1(1-2) by auto
-    thus ?case using fundamental_theorem_of_calculus[OF goal1(3),of "g o dest_vec1" "f o dest_vec1"]
-      unfolding o_def vec1_dest_vec1 by auto qed qed
-
-subsection {* General "twiddling" for interval-to-interval function image. *}
-
-lemma has_integral_twiddle:
-  assumes "0 < r" "\<forall>x. h(g x) = x" "\<forall>x. g(h x) = x" "\<forall>x. continuous (at x) g"
-  "\<forall>u v. \<exists>w z. g ` {u..v} = {w..z}"
-  "\<forall>u v. \<exists>w z. h ` {u..v} = {w..z}"
-  "\<forall>u v. content(g ` {u..v}) = r * content {u..v}"
-  "(f has_integral i) {a..b}"
-  shows "((\<lambda>x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` {a..b})"
-proof- { presume *:"{a..b} \<noteq> {} \<Longrightarrow> ?thesis"
-    show ?thesis apply cases defer apply(rule *,assumption)
-    proof- case goal1 thus ?thesis unfolding goal1 assms(8)[unfolded goal1 has_integral_empty_eq] by auto qed }
-  assume "{a..b} \<noteq> {}" from assms(6)[rule_format,of a b] guess w z apply-by(erule exE)+ note wz=this
-  have inj:"inj g" "inj h" unfolding inj_on_def apply safe apply(rule_tac[!] ccontr)
-    using assms(2) apply(erule_tac x=x in allE) using assms(2) apply(erule_tac x=y in allE) defer
-    using assms(3) apply(erule_tac x=x in allE) using assms(3) apply(erule_tac x=y in allE) by auto
-  show ?thesis unfolding has_integral_def has_integral_compact_interval_def apply(subst if_P) apply(rule,rule,rule wz)
-  proof safe fix e::real assume e:"e>0" hence "e * r > 0" using assms(1) by(rule mult_pos_pos)
-    from assms(8)[unfolded has_integral,rule_format,OF this] guess d apply-by(erule exE conjE)+ note d=this[rule_format]
-    def d' \<equiv> "\<lambda>x y. d (g x) (g y)" have d':"\<And>x. d' x = {y. g y \<in> (d (g x))}" unfolding d'_def by(auto simp add:mem_def)
-    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)"
-    proof(rule_tac x=d' in exI,safe) show "gauge d'" using d(1) unfolding gauge_def d' using continuous_open_preimage_univ[OF assms(4)] by auto
-      fix p assume as:"p tagged_division_of h ` {a..b}" "d' fine p" note p = tagged_division_ofD[OF as(1)] 
-      have "(\<lambda>(x, k). (g x, g ` k)) ` p tagged_division_of {a..b} \<and> d fine (\<lambda>(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of 
-      proof safe show "finite ((\<lambda>(x, k). (g x, g ` k)) ` p)" using as by auto
-        show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p" using as(2) unfolding fine_def d' by auto
-        fix x k assume xk[intro]:"(x,k) \<in> p" show "g x \<in> g ` k" using p(2)[OF xk] by auto
-        show "\<exists>u v. g ` k = {u..v}" using p(4)[OF xk] using assms(5-6) by auto
-        { fix y assume "y \<in> k" thus "g y \<in> {a..b}" "g y \<in> {a..b}" using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"]
-            using assms(2)[rule_format,of y] unfolding inj_image_mem_iff[OF inj(2)] by auto }
-        fix x' k' assume xk':"(x',k') \<in> p" fix z assume "z \<in> interior (g ` k)" "z \<in> interior (g ` k')"
-        hence *:"interior (g ` k) \<inter> interior (g ` k') \<noteq> {}" by auto
-        have same:"(x, k) = (x', k')" apply-apply(rule ccontr,drule p(5)[OF xk xk'])
-        proof- assume as:"interior k \<inter> interior k' = {}" from nonempty_witness[OF *] guess z .
-          hence "z \<in> g ` (interior k \<inter> interior k')" using interior_image_subset[OF assms(4) inj(1)]
-            unfolding image_Int[OF inj(1)] by auto thus False using as by blast
-        qed thus "g x = g x'" by auto
-        { fix z assume "z \<in> k"  thus  "g z \<in> g ` k'" using same by auto }
-        { fix z assume "z \<in> k'" thus  "g z \<in> g ` k"  using same by auto }
-      next fix x assume "x \<in> {a..b}" hence "h x \<in>  \<Union>{k. \<exists>x. (x, k) \<in> p}" using p(6) by auto
-        then guess X unfolding Union_iff .. note X=this from this(1) guess y unfolding mem_Collect_eq ..
-        thus "x \<in> \<Union>{k. \<exists>x. (x, k) \<in> (\<lambda>(x, k). (g x, g ` k)) ` p}" apply-
-          apply(rule_tac X="g ` X" in UnionI) defer apply(rule_tac x="h x" in image_eqI)
-          using X(2) assms(3)[rule_format,of x] by auto
-      qed note ** = d(2)[OF this] have *:"inj_on (\<lambda>(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastsimp
-       have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding group_simps add_left_cancel
-        unfolding setsum_reindex[OF *] apply(subst scaleR_right.setsum) defer apply(rule setsum_cong2) unfolding o_def split_paired_all split_conv
-        apply(drule p(4)) apply safe unfolding assms(7)[rule_format] using p by auto
-      also have "... = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") unfolding scaleR.diff_right scaleR.scaleR_left[THEN sym]
-        unfolding real_scaleR_def using assms(1) by auto finally have *:"?l = ?r" .
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using ** unfolding * unfolding norm_scaleR
-        using assms(1) by(auto simp add:field_simps) qed qed qed
-
-subsection {* Special case of a basic affine transformation. *}
-
-lemma interval_image_affinity_interval: shows "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::real^'n) + c) ` {a..b} = {u..v}"
-  unfolding image_affinity_interval by auto
-
-lemmas Cart_simps = Cart_nth.add Cart_nth.minus Cart_nth.zero Cart_nth.diff Cart_nth.scaleR real_scaleR_def Cart_lambda_beta
-   Cart_eq vector_le_def vector_less_def
-
-lemma setprod_cong2: assumes "\<And>x. x \<in> A \<Longrightarrow> f x = g x" shows "setprod f A = setprod g A"
-  apply(rule setprod_cong) using assms by auto
-
-lemma content_image_affinity_interval: 
- "content((\<lambda>x::real^'n. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ CARD('n) * content {a..b}" (is "?l = ?r")
-proof- { presume *:"{a..b}\<noteq>{} \<Longrightarrow> ?thesis" show ?thesis apply(cases,rule *,assumption)
-      unfolding not_not using content_empty by auto }
-  assume as:"{a..b}\<noteq>{}" show ?thesis proof(cases "m \<ge> 0")
-    case True show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_P[OF True]
-      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') 
-      defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym]
-      apply(rule setprod_cong2) using True as unfolding interval_ne_empty Cart_simps not_le  
-      by(auto simp add:field_simps intro:mult_left_mono)
-  next case False show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_not_P[OF False]
-      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') 
-      defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym]
-      apply(rule setprod_cong2) using False as unfolding interval_ne_empty Cart_simps not_le 
-      by(auto simp add:field_simps mult_le_cancel_left_neg) qed qed
-
-lemma has_integral_affinity: assumes "(f has_integral i) {a..b::real^'n}" "m \<noteq> 0"
-  shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ CARD('n::finite))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})"
-  apply(rule has_integral_twiddle,safe) unfolding Cart_eq Cart_simps apply(rule zero_less_power)
-  defer apply(insert assms(2), simp add:field_simps) apply(insert assms(2), simp add:field_simps)
-  apply(rule continuous_intros)+ apply(rule interval_image_affinity_interval)+ apply(rule content_image_affinity_interval) using assms by auto
-
-lemma integrable_affinity: assumes "f integrable_on {a..b}" "m \<noteq> 0"
-  shows "(\<lambda>x. f(m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` {a..b})"
-  using assms unfolding integrable_on_def apply safe apply(drule has_integral_affinity) by auto
-
-subsection {* Special case of stretching coordinate axes separately. *}
-
-lemma image_stretch_interval:
-  "(\<lambda>x. \<chi> k. m k * x$k) ` {a..b::real^'n} =
-  (if {a..b} = {} then {} else {(\<chi> k. min (m(k) * a$k) (m(k) * b$k)) ..  (\<chi> k. max (m(k) * a$k) (m(k) * b$k))})" (is "?l = ?r")
-proof(cases "{a..b}={}") case True thus ?thesis unfolding True by auto
-next have *:"\<And>P Q. (\<forall>i. P i) \<and> (\<forall>i. Q i) \<longleftrightarrow> (\<forall>i. P i \<and> Q i)" by auto
-  case False note ab = this[unfolded interval_ne_empty]
-  show ?thesis apply-apply(rule set_ext)
-  proof- fix x::"real^'n" have **:"\<And>P Q. (\<forall>i. P i = Q i) \<Longrightarrow> (\<forall>i. P i) = (\<forall>i. Q i)" by auto
-    show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" unfolding if_not_P[OF False] 
-      unfolding image_iff mem_interval Bex_def Cart_simps Cart_eq *
-      unfolding lambda_skolem[THEN sym,of "\<lambda> i xa. (a $ i \<le> xa \<and> xa \<le> b $ i) \<and> x $ i = m i * xa"]
-    proof(rule **,rule) fix i::'n show "(\<exists>xa. (a $ i \<le> xa \<and> xa \<le> b $ i) \<and> x $ i = m i * xa) =
-        (min (m i * a $ i) (m i * b $ i) \<le> x $ i \<and> x $ i \<le> max (m i * a $ i) (m i * b $ i))"
-      proof(cases "m i = 0") case True thus ?thesis using ab by auto
-      next case False hence "0 < m i \<or> 0 > m i" by auto thus ?thesis apply-
-        proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $ i) (m i * b $ i) = m i * a $ i"
-            "max (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab unfolding min_def max_def by auto
-          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI)
-            using as by(auto simp add:field_simps)
-        next assume as:"0 > m i" hence *:"max (m i * a $ i) (m i * b $ i) = m i * a $ i"
-            "min (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab as unfolding min_def max_def 
-            by(auto simp add:field_simps mult_le_cancel_left_neg intro:real_le_antisym)
-          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI)
-            using as by(auto simp add:field_simps) qed qed qed qed qed 
-
-lemma interval_image_stretch_interval: "\<exists>u v. (\<lambda>x. \<chi> k. m k * x$k) ` {a..b::real^'n} = {u..v}"
-  unfolding image_stretch_interval by auto 
-
-lemma content_image_stretch_interval:
-  "content((\<lambda>x::real^'n. \<chi> k. m k * x$k) ` {a..b}) = abs(setprod m UNIV) * content({a..b})"
-proof(cases "{a..b} = {}") case True thus ?thesis
-    unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto
-next case False hence "(\<lambda>x. \<chi> k. m k * x $ k) ` {a..b} \<noteq> {}" by auto
-  thus ?thesis using False unfolding content_def image_stretch_interval apply- unfolding interval_bounds' if_not_P
-    unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding Cart_lambda_beta
-  proof- fix i::'n have "(m i < 0 \<or> m i > 0) \<or> m i = 0" by auto
-    thus "max (m i * a $ i) (m i * b $ i) - min (m i * a $ i) (m i * b $ i) = \<bar>m i\<bar> * (b $ i - a $ i)"
-      apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] 
-      by(auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) qed qed
-
-lemma has_integral_stretch: assumes "(f has_integral i) {a..b}" "\<forall>k. ~(m k = 0)"
-  shows "((\<lambda>x. f(\<chi> k. m k * x$k)) has_integral
-             ((1/(abs(setprod m UNIV))) *\<^sub>R i)) ((\<lambda>x. \<chi> k. 1/(m k) * x$k) ` {a..b})"
-  apply(rule has_integral_twiddle) unfolding zero_less_abs_iff content_image_stretch_interval
-  unfolding image_stretch_interval empty_as_interval Cart_eq using assms
-proof- show "\<forall>x. continuous (at x) (\<lambda>x. \<chi> k. m k * x $ k)"
-   apply(rule,rule linear_continuous_at) unfolding linear_linear
-   unfolding linear_def Cart_simps Cart_eq by(auto simp add:field_simps) qed auto
-
-lemma integrable_stretch: 
-  assumes "f integrable_on {a..b}" "\<forall>k. ~(m k = 0)"
-  shows "(\<lambda>x. f(\<chi> k. m k * x$k)) integrable_on ((\<lambda>x. \<chi> k. 1/(m k) * x$k) ` {a..b})"
-  using assms unfolding integrable_on_def apply-apply(erule exE) apply(drule has_integral_stretch) by auto
-
-subsection {* even more special cases. *}
-
-lemma uminus_interval_vector[simp]:"uminus ` {a..b} = {-b .. -a::real^'n}"
-  apply(rule set_ext,rule) defer unfolding image_iff
-  apply(rule_tac x="-x" in bexI) by(auto simp add:vector_le_def minus_le_iff le_minus_iff)
-
-lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) {a..b}"
-  shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
-  using has_integral_affinity[OF assms, of "-1" 0] by auto
-
-lemma has_integral_reflect[simp]: "((\<lambda>x. f(-x)) has_integral i) {-b..-a} \<longleftrightarrow> (f has_integral i) ({a..b})"
-  apply rule apply(drule_tac[!] has_integral_reflect_lemma) by auto
-
-lemma integrable_reflect[simp]: "(\<lambda>x. f(-x)) integrable_on {-b..-a} \<longleftrightarrow> f integrable_on {a..b}"
-  unfolding integrable_on_def by auto
-
-lemma integral_reflect[simp]: "integral {-b..-a} (\<lambda>x. f(-x)) = integral ({a..b}) f"
-  unfolding integral_def by auto
-
-subsection {* Stronger form of FCT; quite a tedious proof. *}
-
-(** move this **)
-declare norm_triangle_ineq4[intro] 
-
-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" "p tagged_division_of {vec1 a..vec1 b}"
-  shows "setsum (\<lambda>(x,k). f (dest_vec1 (interval_upperbound k)) - f(dest_vec1 (interval_lowerbound k))) p = f b - f a"
-  using additive_tagged_division_1[OF _ assms(2), of "f o dest_vec1"]
-  unfolding o_def vec1_dest_vec1 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"
-  unfolding split_def by(rule refl)
-
-lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
-  apply(subst(asm)(2) norm_minus_cancel[THEN sym])
-  apply(drule norm_triangle_le) by(auto simp add:group_simps)
-
-lemma fundamental_theorem_of_calculus_interior:
-  assumes"a \<le> b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_vector_derivative f'(x)) (at x)"
-  shows "((f' o dest_vec1) has_integral (f b - f a)) {vec a..vec b}"
-proof- { presume *:"a < b \<Longrightarrow> ?thesis" 
-    show ?thesis proof(cases,rule *,assumption)
-      assume "\<not> a < b" hence "a = b" using assms(1) by auto
-      hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" apply(auto simp add: Cart_simps) by smt
-      show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0_1 using * `a=b` by auto
-    qed } assume ab:"a < b"
-  let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
-                   norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f' \<circ> dest_vec1) x) - (f b - f a)) \<le> e * content {vec1 a..vec1 b})"
-  { presume "\<And>e. e>0 \<Longrightarrow> ?P e" thus ?thesis unfolding has_integral_factor_content by auto }
-  fix e::real assume e:"e>0"
-  note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib]
-  note conjunctD2[OF this] note bounded=this(1) and this(2)
-  from this(2) have "\<forall>x\<in>{a<..<b}. \<exists>d>0. \<forall>y. norm (y - x) < d \<longrightarrow> norm (f y - f x - (y - x) *\<^sub>R f' x) \<le> e/2 * norm (y - x)"
-    apply-apply safe apply(erule_tac x=x in ballE,erule_tac x="e/2" in allE) using e by auto note this[unfolded bgauge_existence_lemma]
-  from choice[OF this] guess d .. note conjunctD2[OF this[rule_format]] note d = this[rule_format]
-  have "bounded (f ` {a..b})" apply(rule compact_imp_bounded compact_continuous_image)+ using compact_real_interval assms by auto
-  from this[unfolded bounded_pos] guess B .. note B = this[rule_format]
-
-  have "\<exists>da. 0 < da \<and> (\<forall>c. a \<le> c \<and> {a..c} \<subseteq> {a..b} \<and> {a..c} \<subseteq> ball a da
-    \<longrightarrow> norm(content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b - a)) / 4)"
-  proof- have "a\<in>{a..b}" using ab by auto
-    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
-    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps)
-    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
-    have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' a) \<le> (e * (b - a)) / 8"
-    proof(cases "f' a = 0") case True
-      thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) 
-    next case False thus ?thesis 
-        apply(rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI)
-        using ab e by(auto simp add:field_simps)
-    qed then guess l .. note l = conjunctD2[OF this]
-    show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+)
-    proof- fix c assume as:"a \<le> c" "{a..c} \<subseteq> {a..b}" "{a..c} \<subseteq> ball a (min k l)" 
-      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval]
-      have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \<le> norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by(rule norm_triangle_ineq4)
-      also have "... \<le> e * (b - a) / 8 + e * (b - a) / 8" 
-      proof(rule add_mono) case goal1 have "\<bar>c - a\<bar> \<le> \<bar>l\<bar>" using as' by auto
-        thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto
-      next case goal2 show ?case apply(rule less_imp_le) apply(cases "a = c") defer
-          apply(rule k(2)[unfolded vector_dist_norm]) using as' e ab by(auto simp add:field_simps)
-      qed finally show "norm (content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto
-    qed qed then guess da .. note da=conjunctD2[OF this,rule_format]
-
-  have "\<exists>db>0. \<forall>c\<le>b. {c..b} \<subseteq> {a..b} \<and> {c..b} \<subseteq> ball b db \<longrightarrow> norm(content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b - a)) / 4"
-  proof- have "b\<in>{a..b}" using ab by auto
-    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
-    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps)
-    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
-    have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' b) \<le> (e * (b - a)) / 8"
-    proof(cases "f' b = 0") case True
-      thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) 
-    next case False thus ?thesis 
-        apply(rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI)
-        using ab e by(auto simp add:field_simps)
-    qed then guess l .. note l = conjunctD2[OF this]
-    show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+)
-    proof- fix c assume as:"c \<le> b" "{c..b} \<subseteq> {a..b}" "{c..b} \<subseteq> ball b (min k l)" 
-      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_interval]
-      have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \<le> norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by(rule norm_triangle_ineq4)
-      also have "... \<le> e * (b - a) / 8 + e * (b - a) / 8" 
-      proof(rule add_mono) case goal1 have "\<bar>c - b\<bar> \<le> \<bar>l\<bar>" using as' by auto
-        thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto
-      next case goal2 show ?case apply(rule less_imp_le) apply(cases "b = c") defer apply(subst norm_minus_commute)
-          apply(rule k(2)[unfolded vector_dist_norm]) using as' e ab by(auto simp add:field_simps)
-      qed finally show "norm (content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto
-    qed qed then guess db .. note db=conjunctD2[OF this,rule_format]
-
-  let ?d = "(\<lambda>x. ball x (if x=vec1 a then da else if x=vec b then db else d (dest_vec1 x)))"
-  show "?P e" apply(rule_tac x="?d" in exI)
-  proof safe case goal1 show ?case apply(rule gauge_ball_dependent) using ab db(1) da(1) d(1) by auto
-  next case goal2 note as=this let ?A = "{t. fst t \<in> {vec1 a, vec1 b}}" note p = tagged_division_ofD[OF goal2(1)]
-    have pA:"p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"  using goal2 by auto
-    note * = additive_tagged_division_1'[OF assms(1) goal2(1), THEN sym]
-    have **:"\<And>n1 s1 n2 s2::real. n2 \<le> s2 / 2 \<Longrightarrow> n1 - s1 \<le> s2 / 2 \<Longrightarrow> n1 + n2 \<le> s1 + s2" by arith
-    show ?case unfolding content_1'[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[THEN sym] split_minus
-      unfolding setsum_right_distrib apply(subst(2) pA,subst pA) unfolding setsum_Un_disjoint[OF pA(2-)]
-    proof(rule norm_triangle_le,rule **) 
-      case goal1 show ?case apply(rule order_trans,rule setsum_norm_le) apply(rule pA) defer apply(subst divide.setsum)
-      proof(rule order_refl,safe,unfold not_le o_def split_conv fst_conv,rule ccontr) fix x k assume as:"(x,k) \<in> p"
-          "e * (dest_vec1 (interval_upperbound k) - dest_vec1 (interval_lowerbound k)) / 2
-          < norm (content k *\<^sub>R f' (dest_vec1 x) - (f (dest_vec1 (interval_upperbound k)) - f (dest_vec1 (interval_lowerbound k))))"
-        from p(4)[OF this(1)] guess u v apply-by(erule exE)+ note k=this
-        hence "\<forall>i. u$i \<le> v$i" and uv:"{u,v}\<subseteq>{u..v}" using p(2)[OF as(1)] by auto note this(1) this(1)[unfolded forall_1]
-        note result = as(2)[unfolded k interval_bounds[OF this(1)] content_1[OF this(2)]]
-
-        assume as':"x \<noteq> vec1 a" "x \<noteq> vec1 b" hence "x$1 \<in> {a<..<b}" using p(2-3)[OF as(1)] by(auto simp add:Cart_simps) note  * = d(2)[OF this] 
-        have "norm ((v$1 - u$1) *\<^sub>R f' (x$1) - (f (v$1) - f (u$1))) =
-          norm ((f (u$1) - f (x$1) - (u$1 - x$1) *\<^sub>R f' (x$1)) - (f (v$1) - f (x$1) - (v$1 - x$1) *\<^sub>R f' (x$1)))" 
-          apply(rule arg_cong[of _ _ norm]) unfolding scaleR_left.diff by auto 
-        also have "... \<le> e / 2 * norm (u$1 - x$1) + e / 2 * norm (v$1 - x$1)" apply(rule norm_triangle_le_sub)
-          apply(rule add_mono) apply(rule_tac[!] *) using fineD[OF goal2(2) as(1)] as' unfolding k subset_eq
-          apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) using uv by(auto simp add:dist_real)
-        also have "... \<le> e / 2 * norm (v$1 - u$1)" using p(2)[OF as(1)] unfolding k by(auto simp add:field_simps)
-        finally have "e * (dest_vec1 v - dest_vec1 u) / 2 < e * (dest_vec1 v - dest_vec1 u) / 2"
-          apply- apply(rule less_le_trans[OF result]) using uv by auto thus False by auto qed
-
-    next have *:"\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2) / 2 \<Longrightarrow> x - s1 \<le> s2 / 2" by auto
-      case goal2 show ?case apply(rule *) apply(rule setsum_nonneg) apply(rule,unfold split_paired_all split_conv)
-        defer unfolding setsum_Un_disjoint[OF pA(2-),THEN sym] pA(1)[THEN sym] unfolding setsum_right_distrib[THEN sym] 
-        apply(subst additive_tagged_division_1[OF _ as(1)]) unfolding vec1_dest_vec1 apply(rule assms)
-      proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {vec1 a, vec1 b}}" note xk=IntD1[OF this]
-        from p(4)[OF this] guess u v apply-by(erule exE)+ note uv=this
-        with p(2)[OF xk] have "{u..v} \<noteq> {}" by auto
-        thus "0 \<le> e * ((interval_upperbound k)$1 - (interval_lowerbound k)$1)"
-          unfolding uv using e by(auto simp add:field_simps)
-      next have *:"\<And>s f t e. setsum f s = setsum f t \<Longrightarrow> norm(setsum f t) \<le> e \<Longrightarrow> norm(setsum f s) \<le> e" by auto
-        show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R (f' \<circ> dest_vec1) x -
-          (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) \<le> e * (b - a) / 2" 
-          apply(rule *[where t="p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content(snd t) \<noteq> 0}"])
-          apply(rule setsum_mono_zero_right[OF pA(2)]) defer apply(rule) unfolding split_paired_all split_conv o_def
-        proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {vec1 a, vec1 b}} - p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content (snd t) \<noteq> 0}"
-          hence xk:"(x,k)\<in>p" "content k = 0" by auto from p(4)[OF xk(1)] guess u v apply-by(erule exE)+ note uv=this
-          have "k\<noteq>{}" using p(2)[OF xk(1)] by auto hence *:"u = v" using xk unfolding uv content_eq_0_1 interval_eq_empty by auto
-          thus "content k *\<^sub>R (f' (x$1)) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1)) = 0" using xk unfolding uv by auto
-        next have *:"p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content(snd t) \<noteq> 0} = 
-            {t. t\<in>p \<and> fst t = vec1 a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = vec1 b \<and> content(snd t) \<noteq> 0}" by blast
-          have **:"\<And>s f. \<And>e::real. (\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y) \<Longrightarrow> (\<forall>x. x \<in> s \<longrightarrow> norm(f x) \<le> e) \<Longrightarrow> e>0 \<Longrightarrow> norm(setsum f s) \<le> e"
-          proof(case_tac "s={}") case goal2 then obtain x where "x\<in>s" by auto hence *:"s = {x}" using goal2(1) by auto
-            thus ?case using `x\<in>s` goal2(2) by auto
-          qed auto
-          case goal2 show ?case apply(subst *, subst setsum_Un_disjoint) prefer 4 apply(rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) 
-            apply(rule norm_triangle_le,rule add_mono) apply(rule_tac[1-2] **)
-          proof- let ?B = "\<lambda>x. {t \<in> p. fst t = vec1 x \<and> content (snd t) \<noteq> 0}"
-            have pa:"\<And>k. (vec1 a, k) \<in> p \<Longrightarrow> \<exists>v. k = {vec1 a .. v} \<and> vec1 a \<le> v" 
-            proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this
-              have *:"u \<le> v" using p(2)[OF goal1] unfolding uv by auto
-              have u:"u = vec1 a" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
-                have "u \<ge> vec1 a" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "u\<noteq>vec1 a" ultimately
-                have "u > vec1 a" unfolding Cart_simps by auto
-                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps)
-              qed thus ?case apply(rule_tac x=v in exI) unfolding uv using * by auto
-            qed
-            have pb:"\<And>k. (vec1 b, k) \<in> p \<Longrightarrow> \<exists>v. k = {v .. vec1 b} \<and> vec1 b \<ge> v" 
-            proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this
-              have *:"u \<le> v" using p(2)[OF goal1] unfolding uv by auto
-              have u:"v = vec1 b" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
-                have "v \<le> vec1 b" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "v\<noteq>vec1 b" ultimately
-                have "v < vec1 b" unfolding Cart_simps by auto
-                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps)
-              qed thus ?case apply(rule_tac x=u in exI) unfolding uv using * by auto
-            qed
-
-            show "\<forall>x y. x \<in> ?B a \<and> y \<in> ?B a \<longrightarrow> x = y" apply(rule,rule,rule,unfold split_paired_all)
-              unfolding mem_Collect_eq fst_conv snd_conv apply safe
-            proof- fix x k k' assume k:"(vec1 a, k) \<in> p" "(vec1 a, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
-              guess v using pa[OF k(1)] .. note v = conjunctD2[OF this]
-              guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (min (v$1) (v'$1))"
-              have "{vec1 a <..< ?v} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter]
-              moreover have "vec1 ((a + ?v$1)/2) \<in> {vec1 a <..< ?v}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps)
-              ultimately have "vec1 ((a + ?v$1)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
-              hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto
-              { assume "x\<in>k" thus "x\<in>k'" unfolding * . } { assume "x\<in>k'" thus "x\<in>k" unfolding * . }
-            qed 
-            show "\<forall>x y. x \<in> ?B b \<and> y \<in> ?B b \<longrightarrow> x = y" apply(rule,rule,rule,unfold split_paired_all)
-              unfolding mem_Collect_eq fst_conv snd_conv apply safe
-            proof- fix x k k' assume k:"(vec1 b, k) \<in> p" "(vec1 b, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
-              guess v using pb[OF k(1)] .. note v = conjunctD2[OF this]
-              guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (max (v$1) (v'$1))"
-              have "{?v <..< vec1 b} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter]
-              moreover have "vec1 ((b + ?v$1)/2) \<in> {?v <..< vec1 b}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps)
-              ultimately have "vec1 ((b + ?v$1)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
-              hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto
-              { assume "x\<in>k" thus "x\<in>k'" unfolding * . } { assume "x\<in>k'" thus "x\<in>k" unfolding * . }
-            qed
-
-            let ?a = a and ?b = b (* a is something else while proofing the next theorem. *)
-            show "\<forall>x. x \<in> ?B a \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x)
-              \<le> e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1
-            proof- case goal1 guess v using pa[OF goal1(1)] .. note v = conjunctD2[OF this]
-              have "vec1 ?a\<in>{vec1 ?a..v}" using v(2) by auto hence "dest_vec1 v \<le> ?b" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
-              moreover have "{?a..dest_vec1 v} \<subseteq> ball ?a da" using fineD[OF as(2) goal1(1)]
-                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE)
-                by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately
-              show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply-
-                apply(rule da(2)[of "v$1",unfolded vec1_dest_vec1])
-                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto
-            qed
-            show "\<forall>x. x \<in> ?B b \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x)
-              \<le> e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1
-            proof- case goal1 guess v using pb[OF goal1(1)] .. note v = conjunctD2[OF this]
-              have "vec1 ?b\<in>{v..vec1 ?b}" using v(2) by auto hence "dest_vec1 v \<ge> ?a" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
-              moreover have "{dest_vec1 v..?b} \<subseteq> ball ?b db" using fineD[OF as(2) goal1(1)]
-                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE) using ab
-                by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately
-              show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply-
-                apply(rule db(2)[of "v$1",unfolded vec1_dest_vec1])
-                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto
-            qed
-          qed(insert p(1) ab e, auto simp add:field_simps) qed auto qed qed qed qed
-
-subsection {* Stronger form with finite number of exceptional points. *}
-
-lemma fundamental_theorem_of_calculus_interior_strong: fixes f::"real \<Rightarrow> 'a::banach"
-  assumes"finite s" "a \<le> b" "continuous_on {a..b} f"
-  "\<forall>x\<in>{a<..<b} - s. (f has_vector_derivative f'(x)) (at x)"
-  shows "((f' o dest_vec1) has_integral (f b - f a)) {vec a..vec b}" using assms apply- 
-proof(induct "card s" arbitrary:s a b)
-  case 0 show ?case apply(rule fundamental_theorem_of_calculus_interior) using 0 by auto
-next case (Suc n) from this(2) guess c s' apply-apply(subst(asm) eq_commute) unfolding card_Suc_eq
-    apply(subst(asm)(2) eq_commute) by(erule exE conjE)+ note cs = this[rule_format]
-  show ?case proof(cases "c\<in>{a<..<b}")
-    case False thus ?thesis apply- apply(rule Suc(1)[OF cs(3) _ Suc(4,5)]) apply safe defer
-      apply(rule Suc(6)[rule_format]) using Suc(3) unfolding cs by auto
-  next have *:"f b - f a = (f c - f a) + (f b - f c)" by auto
-    case True hence "vec1 a \<le> vec1 c" "vec1 c \<le> vec1 b" by auto
-    thus ?thesis apply(subst *) apply(rule has_integral_combine) apply assumption+
-      apply(rule_tac[!] Suc(1)[OF cs(3)]) using Suc(3) unfolding cs
-    proof- show "continuous_on {a..c} f" "continuous_on {c..b} f"
-        apply(rule_tac[!] continuous_on_subset[OF Suc(5)]) using True by auto
-      let ?P = "\<lambda>i j. \<forall>x\<in>{i<..<j} - s'. (f has_vector_derivative f' x) (at x)"
-      show "?P a c" "?P c b" apply safe apply(rule_tac[!] Suc(6)[rule_format]) using True unfolding cs by auto
-    qed auto qed qed
-
-lemma fundamental_theorem_of_calculus_strong: fixes f::"real \<Rightarrow> 'a::banach"
-  assumes "finite s" "a \<le> b" "continuous_on {a..b} f"
-  "\<forall>x\<in>{a..b} - s. (f has_vector_derivative f'(x)) (at x)"
-  shows "((f' o dest_vec1) has_integral (f(b) - f(a))) {vec1 a..vec1 b}"
-  apply(rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f'])
-  using assms(4) by auto
-
-end
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,5 @@
 theory Multivariate_Analysis
-imports Determinants Integration_MV
+imports Determinants Integration Real_Integration
 begin
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Real_Integration.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,72 @@
+header{*Integration on real intervals*}
+
+theory Real_Integration
+imports Integration
+begin
+
+text{*We follow John Harrison in formalizing the Gauge integral.*}
+
+definition Integral :: "real set \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> bool" where
+  "Integral s f k = (f o dest_vec1 has_integral k) (vec1 ` s)"
+
+lemmas integral_unfold = Integral_def split_conv o_def vec1_interval
+
+lemma Integral_unique:
+    "[| Integral{a..b} f k1; Integral{a..b} f k2 |] ==> k1 = k2"
+  unfolding integral_unfold apply(rule has_integral_unique) by assumption+
+
+lemma Integral_zero [simp]: "Integral{a..a} f 0"
+  unfolding integral_unfold by auto
+
+lemma Integral_eq_diff_bounds: assumes "a \<le> b" shows "Integral{a..b} (%x. 1) (b - a)"
+  unfolding integral_unfold using has_integral_const[of "1::real" "vec1 a" "vec1 b"]
+  unfolding content_1'[OF assms] by auto
+
+lemma Integral_mult_const: assumes "a \<le> b" shows "Integral{a..b} (%x. c)  (c*(b - a))"
+  unfolding integral_unfold using has_integral_const[of "c::real" "vec1 a" "vec1 b"]
+  unfolding content_1'[OF assms] by(auto simp add:field_simps)
+
+lemma Integral_mult: assumes "Integral{a..b} f k" shows "Integral{a..b} (%x. c * f x) (c * k)"
+  using assms unfolding integral_unfold  apply(drule_tac has_integral_cmul[where c=c]) by auto
+
+lemma Integral_add:
+  assumes "Integral {a..b} f x1" "Integral {b..c} f x2"  "a \<le> b" and "b \<le> c"
+  shows "Integral {a..c} f (x1 + x2)"
+  using assms unfolding integral_unfold apply-
+  apply(rule has_integral_combine[of "vec1 a" "vec1 b" "vec1 c"]) by  auto
+
+lemma FTC1: assumes "a \<le> b" "\<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x)"
+  shows "Integral{a..b} f' (f(b) - f(a))"
+proof-note fundamental_theorem_of_calculus[OF assms(1), of"f o dest_vec1" "f' o dest_vec1"]
+  note * = this[unfolded o_def vec1_dest_vec1]
+  have **:"\<And>x. (\<lambda>xa\<Colon>real. xa * f' x) =  op * (f' x)" apply(rule ext) by(auto simp add:field_simps)
+  show ?thesis unfolding integral_unfold apply(rule *)
+    using assms(2) unfolding DERIV_conv_has_derivative has_vector_derivative_def
+    apply safe apply(rule has_derivative_at_within) by(auto simp add:**) qed
+
+lemma Integral_subst: "[| Integral{a..b} f k1; k2=k1 |] ==> Integral{a..b} f k2"
+by simp
+
+subsection {* Additivity Theorem of Gauge Integral *}
+
+text{* Bartle/Sherbert: Theorem 10.1.5 p. 278 *}
+lemma Integral_add_fun: "[| Integral{a..b} f k1; Integral{a..b} g k2 |] ==> Integral{a..b} (%x. f x + g x) (k1 + k2)"
+  unfolding integral_unfold apply(rule has_integral_add) by assumption+
+
+lemma norm_vec1'[simp]:"norm (vec1 x) = norm x"
+  using norm_vector_1[of "vec1 x"] by auto
+
+lemma Integral_le: assumes "a \<le> b" "\<forall>x. a \<le> x & x \<le> b --> f(x) \<le> g(x)" "Integral{a..b} f k1" "Integral{a..b} g k2" shows "k1 \<le> k2"
+proof- note assms(3-4)[unfolded integral_unfold] note has_integral_vec1[OF this(1)] has_integral_vec1[OF this(2)]
+  note has_integral_component_le[OF this,of 1] thus ?thesis using assms(2) by auto qed
+
+lemma monotonic_anti_derivative:
+  fixes f g :: "real => real" shows
+     "[| a \<le> b; \<forall>c. a \<le> c & c \<le> b --> f' c \<le> g' c;
+         \<forall>x. DERIV f x :> f' x; \<forall>x. DERIV g x :> g' x |]
+      ==> f b - f a \<le> g b - g a"
+apply (rule Integral_le, assumption)
+apply (auto intro: FTC1)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mutabelle/ROOT.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -0,0 +1,3 @@
+
+use_thy "MutabelleExtra";
+
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -10,9 +10,11 @@
 val take_random : int -> 'a list -> 'a list
 
 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
-type mtd = string * (theory -> term -> outcome)
+type timing = (string * int) list
 
-type mutant_subentry = term * (string * outcome) list
+type mtd = string * (theory -> term -> outcome * timing)
+
+type mutant_subentry = term * (string * (outcome * timing)) list
 type detailed_entry = string * bool * term * mutant_subentry list
 
 type subentry = string * int * int * int * int * int * int
@@ -52,7 +54,7 @@
 
 (* quickcheck options *)
 (*val quickcheck_generator = "SML"*)
-val iterations = 100
+val iterations = 1
 val size = 5
 
 exception RANDOM;
@@ -75,12 +77,13 @@
   else l + Real.floor (rmod (random ()) (real (h - l + 1)));
 
 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
-type mtd = string * (theory -> term -> outcome)
+type timing = (string * int) list
 
-type mutant_subentry = term * (string * outcome) list
+type mtd = string * (theory -> term -> outcome * timing)
+
+type mutant_subentry = term * (string * (outcome * timing)) list
 type detailed_entry = string * bool * term * mutant_subentry list
 
-
 type subentry = string * int * int * int * int * int * int
 type entry = string * bool * subentry list
 type report = entry list
@@ -94,11 +97,11 @@
 fun invoke_quickcheck quickcheck_generator thy t =
   TimeLimit.timeLimit (Time.fromSeconds (! Auto_Counterexample.time_limit))
       (fn _ =>
-          case Quickcheck.test_term (ProofContext.init thy) false (SOME quickcheck_generator)
+          case Quickcheck.timed_test_term (ProofContext.init thy) false (SOME quickcheck_generator)
                                     size iterations (preprocess thy [] t) of
-            NONE => NoCex
-          | SOME _ => GenuineCex) ()
-  handle TimeLimit.TimeOut => Timeout
+            (NONE, time_report) => (NoCex, time_report)
+          | (SOME _, time_report) => (GenuineCex, time_report)) ()
+  handle TimeLimit.TimeOut => (Timeout, [("timelimit", !Auto_Counterexample.time_limit)])
 
 fun quickcheck_mtd quickcheck_generator =
   ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck quickcheck_generator)
@@ -189,9 +192,9 @@
 
 val forbidden =
  [(* (@{const_name "power"}, "'a"), *)
-  (@{const_name HOL.induct_equal}, "'a"),
-  (@{const_name HOL.induct_implies}, "'a"),
-  (@{const_name HOL.induct_conj}, "'a"),
+  (*(@{const_name induct_equal}, "'a"),
+  (@{const_name induct_implies}, "'a"),
+  (@{const_name induct_conj}, "'a"),*)
   (@{const_name "undefined"}, "'a"),
   (@{const_name "default"}, "'a"),
   (@{const_name "dummy_pattern"}, "'a::{}") (*,
@@ -245,17 +248,26 @@
       Library.nth xs j :: take_random (n - 1) (nth_drop j xs)
     end
 
+fun cpu_time description f =
+  let
+    val start = start_timing ()
+    val result = Exn.capture f ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, (description, time)) end
+
 fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
   let
     val _ = priority ("Invoking " ^ mtd_name)
-    val res = case try (invoke_mtd thy) t of
-                SOME res => res
-              | NONE => (priority ("**** PROBLEMS WITH " ^
-                                 Syntax.string_of_term_global thy t); Error)
+    val ((res, timing), time) = cpu_time "total time"
+      (fn () => case try (invoke_mtd thy) t of
+          SOME (res, timing) => (res, timing)
+        | NONE => (priority ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
+           (Error , [])))
     val _ = priority (" Done")
-  in res end
+  in (res, time :: timing) end
 
 (* theory -> term list -> mtd -> subentry *)
+(*
 fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) =
   let
      val res = map (safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants
@@ -266,7 +278,7 @@
 
 fun create_entry thy thm exec mutants mtds =
   (Thm.get_name thm, exec, map (test_mutants_using_one_method thy mutants) mtds)
-
+*)
 fun create_detailed_entry thy thm exec mutants mtds =
   let
     fun create_mutant_subentry mutant = (mutant,
@@ -322,15 +334,22 @@
   | string_of_outcome Timeout = "Timeout"
   | string_of_outcome Error = "Error"
 
-fun string_of_mutant_subentry thy (t, results) =
+fun string_of_mutant_subentry thy thm_name (t, results) =
   "mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^
-  space_implode "; " (map (fn (mtd_name, outcome) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
+  space_implode "; "
+    (map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
   "\n"
 
+fun string_of_mutant_subentry' thy thm_name (t, results) =
+  "mutant of " ^ thm_name ^ ":" ^
+    cat_lines (map (fn (mtd_name, (outcome, timing)) =>
+      mtd_name ^ ": " ^ string_of_outcome outcome ^ "; " ^
+      space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing)) results)
+
 fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
    thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
    Syntax.string_of_term_global thy t ^ "\n" ^
-   cat_lines (map (string_of_mutant_subentry thy) mutant_subentries) ^ "\n"
+   cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n"
 
 (* subentry -> string *)
 fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno,
--- a/src/HOL/NSA/Hyperreal.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/NSA/Hyperreal.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -7,7 +7,7 @@
 *)
 
 theory Hyperreal
-imports Ln Deriv Taylor Integration HLog
+imports Ln Deriv Taylor HLog
 begin
 
 end
--- a/src/HOL/Nitpick.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Nitpick.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -36,7 +36,6 @@
            and bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
            and bisim_iterator_max :: bisim_iterator
            and Quot :: "'a \<Rightarrow> 'b"
-           and quot_normal :: "'a \<Rightarrow> 'a"
            and Tha :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
 
 datatype ('a, 'b) pair_box = PairBox 'a 'b
@@ -237,11 +236,10 @@
 setup {* Nitpick_Isar.setup *}
 
 hide (open) const unknown is_unknown undefined_fast_The undefined_fast_Eps bisim 
-    bisim_iterator_max Quot quot_normal Tha PairBox FunBox Word refl' wf'
-    wf_wfrec wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd
-    int_lcm Frac Abs_Frac Rep_Frac zero_frac one_frac num denom norm_frac frac
-    plus_frac times_frac uminus_frac number_of_frac inverse_frac less_eq_frac
-    of_frac
+    bisim_iterator_max Quot Tha PairBox FunBox Word refl' wf' wf_wfrec wf_wfrec'
+    wfrec' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac
+    Abs_Frac Rep_Frac zero_frac one_frac num denom norm_frac frac plus_frac
+    times_frac uminus_frac number_of_frac inverse_frac less_eq_frac of_frac
 hide (open) type bisim_iterator pair_box fun_box unsigned_bit signed_bit word
 hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def
     wf'_def wf_wfrec'_def wfrec'_def card'_def setsum'_def fold_graph'_def
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -13,7 +13,7 @@
 
 chapter {* 3. First Steps *}
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1]
+nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 15 s]
 
 subsection {* 3.1. Propositional Logic *}
 
@@ -70,7 +70,7 @@
 oops
 
 lemma "P Suc"
-nitpick [card = 1-6]
+nitpick
 oops
 
 lemma "P (op +\<Colon>nat\<Rightarrow>nat\<Rightarrow>nat)"
@@ -149,7 +149,7 @@
 "\<lbrakk>even' m; even' n\<rbrakk> \<Longrightarrow> even' (m + n)"
 
 lemma "\<exists>n \<in> {0, 2, 4, 6, 8}. \<not> even' n"
-nitpick [card nat = 10, unary_ints, verbose, show_consts] (* FIXME: should be genuine *)
+nitpick [card nat = 10, unary_ints, verbose, show_consts]
 oops
 
 lemma "even' (n - 2) \<Longrightarrow> even' n"
@@ -210,7 +210,7 @@
 lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>1 \<sigma> t = t"
 nitpick [verbose]
 nitpick [eval = "subst\<^isub>1 \<sigma> t"]
-nitpick [dont_box]
+(* nitpick [dont_box] *)
 oops
 
 primrec subst\<^isub>2 where
@@ -220,7 +220,7 @@
 "subst\<^isub>2 \<sigma> (App t u) = App (subst\<^isub>2 \<sigma> t) (subst\<^isub>2 \<sigma> u)"
 
 lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>2 \<sigma> t = t"
-nitpick
+nitpick [card = 1\<midarrow>6]
 sorry
 
 subsection {* 3.11. Scope Monotonicity *}
@@ -243,7 +243,7 @@
 "n \<in> reach \<Longrightarrow> n + 2 \<in> reach"
 
 lemma "n \<in> reach \<Longrightarrow> 2 dvd n"
-nitpick
+nitpick [unary_ints]
 apply (induct set: reach)
   apply auto
  nitpick
@@ -252,7 +252,7 @@
 oops
 
 lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0"
-nitpick
+nitpick [unary_ints]
 apply (induct set: reach)
   apply auto
  nitpick
@@ -289,12 +289,12 @@
           if b \<in> labels t then labels t else (labels t - {a}) \<union> {b}
         else
           if b \<in> labels t then (labels t - {b}) \<union> {a} else labels t)"
-nitpick
+(* nitpick *)
 proof (induct t)
   case Leaf thus ?case by simp
 next
   case (Branch t u) thus ?case
-  nitpick [non_std "'a bin_tree", show_consts]
+  nitpick [non_std, show_all]
   by auto
 qed
 
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -110,12 +110,12 @@
 
 lemma "\<exists>one \<in> {1}. \<exists>two \<in> {2}.
        f5 (\<lambda>a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x"
-nitpick [expect = potential] (* unfortunate *)
+nitpick [expect = genuine]
 oops
 
 lemma "\<exists>two \<in> {2}. \<exists>one \<in> {1}.
        f5 (\<lambda>a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x"
-nitpick [expect = potential] (* unfortunate *)
+nitpick [expect = genuine]
 oops
 
 lemma "\<forall>a. g a = a
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -143,11 +143,11 @@
 by (rule Rep_Sum_inverse)
 
 lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
-(* nitpick [expect = none] FIXME *)
+nitpick [expect = none]
 by (rule Zero_nat_def_raw)
 
 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Rep_Nat n))"
-(* nitpick [expect = none] FIXME *)
+nitpick [expect = none]
 by (rule Suc_def)
 
 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
--- a/src/HOL/Orderings.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Orderings.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -5,7 +5,7 @@
 header {* Abstract orderings *}
 
 theory Orderings
-imports Algebras
+imports HOL
 uses
   "~~/src/Provers/order.ML"
   "~~/src/Provers/quasi.ML"  (* FIXME unused? *)
--- a/src/HOL/Prolog/Func.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Prolog/Func.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -5,7 +5,7 @@
 header {* Untyped functional language, with call by value semantics *}
 
 theory Func
-imports HOHH Algebras
+imports HOHH
 begin
 
 typedecl tm
--- a/src/HOL/Quotient.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Quotient.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -2,6 +2,8 @@
     Author:     Cezary Kaliszyk and Christian Urban
 *)
 
+header {* Definition of Quotient Types *}
+
 theory Quotient
 imports Plain ATP_Linkup
 uses
@@ -80,7 +82,7 @@
   shows "((op =) OOO R) = R"
   by (auto simp add: expand_fun_eq)
 
-section {* Respects predicate *}
+subsection {* Respects predicate *}
 
 definition
   Respects
@@ -92,7 +94,7 @@
   unfolding mem_def Respects_def
   by simp
 
-section {* Function map and function relation *}
+subsection {* Function map and function relation *}
 
 definition
   fun_map (infixr "--->" 55)
@@ -124,7 +126,7 @@
   using a by auto
 
 
-section {* Quotient Predicate *}
+subsection {* Quotient Predicate *}
 
 definition
   "Quotient E Abs Rep \<equiv>
@@ -285,7 +287,7 @@
   shows "R2 (f x) (g y)"
   using a by simp
 
-section {* lemmas for regularisation of ball and bex *}
+subsection {* lemmas for regularisation of ball and bex *}
 
 lemma ball_reg_eqv:
   fixes P :: "'a \<Rightarrow> bool"
@@ -387,7 +389,7 @@
   shows "(\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y)"
   using assms by auto
 
-section {* Bounded abstraction *}
+subsection {* Bounded abstraction *}
 
 definition
   Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
@@ -465,7 +467,7 @@
   using a unfolding Quotient_def Bex_def in_respects fun_map_def id_apply
   by metis
 
-section {* @{text Bex1_rel} quantifier *}
+subsection {* @{text Bex1_rel} quantifier *}
 
 definition
   Bex1_rel :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
@@ -569,7 +571,7 @@
   apply auto
   done
 
-section {* Various respects and preserve lemmas *}
+subsection {* Various respects and preserve lemmas *}
 
 lemma quot_rel_rsp:
   assumes a: "Quotient R Abs Rep"
@@ -706,7 +708,7 @@
 
 end
 
-section {* ML setup *}
+subsection {* ML setup *}
 
 text {* Auxiliary data for the quotient package *}
 
@@ -762,7 +764,7 @@
 text {* Tactics for proving the lifted theorems *}
 use "~~/src/HOL/Tools/Quotient/quotient_tacs.ML"
 
-section {* Methods / Interface *}
+subsection {* Methods / Interface *}
 
 method_setup lifting =
   {* Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt thms))) *}
--- a/src/HOL/Rational.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Rational.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -707,13 +707,14 @@
 subsubsection {* Rationals are an Archimedean field *}
 
 lemma rat_floor_lemma:
-  assumes "0 < b"
   shows "of_int (a div b) \<le> Fract a b \<and> Fract a b < of_int (a div b + 1)"
 proof -
   have "Fract a b = of_int (a div b) + Fract (a mod b) b"
-    using `0 < b` by (simp add: of_int_rat)
+    by (cases "b = 0", simp, simp add: of_int_rat)
   moreover have "0 \<le> Fract (a mod b) b \<and> Fract (a mod b) b < 1"
-    using `0 < b` by (simp add: zero_le_Fract_iff Fract_less_one_iff)
+    unfolding Fract_of_int_quotient
+    by (rule linorder_cases [of b 0])
+       (simp add: divide_nonpos_neg, simp, simp add: divide_nonneg_pos)
   ultimately show ?thesis by simp
 qed
 
@@ -723,15 +724,14 @@
   show "\<exists>z. r \<le> of_int z"
   proof (induct r)
     case (Fract a b)
-    then have "Fract a b \<le> of_int (a div b + 1)"
-      using rat_floor_lemma [of b a] by simp
+    have "Fract a b \<le> of_int (a div b + 1)"
+      using rat_floor_lemma [of a b] by simp
     then show "\<exists>z. Fract a b \<le> of_int z" ..
   qed
 qed
 
-lemma floor_Fract:
-  assumes "0 < b" shows "floor (Fract a b) = a div b"
-  using rat_floor_lemma [OF `0 < b`, of a]
+lemma floor_Fract: "floor (Fract a b) = a div b"
+  using rat_floor_lemma [of a b]
   by (simp add: floor_unique)
 
 
--- a/src/HOL/Rings.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Rings.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -13,19 +13,6 @@
 imports Groups
 begin
 
-text {*
-  The theory of partially ordered rings is taken from the books:
-  \begin{itemize}
-  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
-  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
-  \end{itemize}
-  Most of the used notions can also be looked up in 
-  \begin{itemize}
-  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
-  \item \emph{Algebra I} by van der Waerden, Springer.
-  \end{itemize}
-*}
-
 class semiring = ab_semigroup_add + semigroup_mult +
   assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c"
   assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c"
@@ -506,6 +493,19 @@
   assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
   assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
 
+text {*
+  The theory of partially ordered rings is taken from the books:
+  \begin{itemize}
+  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
+  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
+  \end{itemize}
+  Most of the used notions can also be looked up in 
+  \begin{itemize}
+  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
+  \item \emph{Algebra I} by van der Waerden, Springer.
+  \end{itemize}
+*}
+
 class ordered_semiring = mult_mono + semiring_0 + ordered_ab_semigroup_add 
 begin
 
--- a/src/HOL/SEQ.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/SEQ.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -435,7 +435,7 @@
 
 lemma LIMSEQ_diff_approach_zero2:
   fixes L :: "'a::real_normed_vector"
-  shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L";
+  shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L"
 by (drule (1) LIMSEQ_diff, simp)
 
 text{*A sequence tends to zero iff its abs does*}
@@ -1047,6 +1047,17 @@
   shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
 unfolding Cauchy_def dist_norm ..
 
+lemma Cauchy_iff2:
+     "Cauchy X =
+      (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))"
+apply (simp add: Cauchy_iff, auto)
+apply (drule reals_Archimedean, safe)
+apply (drule_tac x = n in spec, auto)
+apply (rule_tac x = M in exI, auto)
+apply (drule_tac x = m in spec, simp)
+apply (drule_tac x = na in spec, auto)
+done
+
 lemma CauchyI:
   fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
   shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
--- a/src/HOL/TLA/Action.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/TLA/Action.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -16,8 +16,8 @@
   'a trfun = "(state * state) => 'a"
   action   = "bool trfun"
 
-instance
-  "*" :: (world, world) world ..
+arities
+  "*" :: (world, world) world
 
 consts
   (** abstract syntax **)
--- a/src/HOL/TLA/Init.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/TLA/Init.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -12,7 +12,7 @@
 begin
 
 typedecl behavior
-instance behavior :: world ..
+arities behavior :: world
 
 types
   temporal = "behavior form"
--- a/src/HOL/TLA/Intensional.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/TLA/Intensional.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -10,8 +10,8 @@
 imports Main
 begin
 
-axclass
-  world < type
+classes world
+classrel world < type
 
 (** abstract syntax **)
 
@@ -171,7 +171,7 @@
   "_liftMem"    :: "[lift, lift] => lift"                ("(_/ \<in> _)" [50, 51] 50)
   "_liftNotMem" :: "[lift, lift] => lift"                ("(_/ \<notin> _)" [50, 51] 50)
 
-axioms
+defs
   Valid_def:   "|- A    ==  ALL w. w |= A"
 
   unl_con:     "LIFT #c w  ==  c"
--- a/src/HOL/TLA/Stfun.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/TLA/Stfun.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -11,7 +11,7 @@
 
 typedecl state
 
-instance state :: world ..
+arities state :: world
 
 types
   'a stfun = "state => 'a"
--- a/src/HOL/Tools/Datatype/datatype_prop.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -22,6 +22,8 @@
     (string * sort) list -> theory -> term list list
   val make_splits : string list -> descr list ->
     (string * sort) list -> theory -> (term * term) list
+  val make_case_combs : string list -> descr list ->
+    (string * sort) list -> theory -> string -> term list
   val make_weak_case_congs : string list -> descr list ->
     (string * sort) list -> theory -> term list
   val make_case_congs : string list -> descr list ->
--- a/src/HOL/Tools/Function/function.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Function/function.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -97,6 +97,7 @@
           |> addsmps (conceal_partial o Binding.qualify false "partial")
                "psimps" conceal_partial psimp_attribs psimps
           ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps
+          ||> fold_option (Spec_Rules.add Spec_Rules.Equational o (pair fs)) trsimps
           ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
                  [Attrib.internal (K (Rule_Cases.case_names cnames)),
                   Attrib.internal (K (Rule_Cases.consumes 1)),
@@ -126,7 +127,7 @@
 val add_function =
   gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
 val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
-
+                                                
 fun gen_termination_proof prep_term raw_term_opt lthy =
   let
     val term_opt = Option.map (prep_term lthy) raw_term_opt
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -1004,9 +1004,11 @@
   handle Option.Option => raise SYNTAX ("Kodkod.read_next_problems",
                                         "expected number after \"PROBLEM\"")
 
-(* Path.T -> (int * raw_bound list) list * int list *)
+(* Path.T -> bool * ((int * raw_bound list) list * int list) *)
 fun read_output_file path =
-  read_next_problems (Substring.full (File.read path), [], []) |>> rev ||> rev
+  (false, read_next_problems (Substring.full (File.read path), [], [])
+          |>> rev ||> rev)
+  handle IO.Io _ => (true, ([], [])) | OS.SysErr _ => (true, ([], []))
 
 (* The fudge term below is to account for Kodkodi's slow start-up time, which
    is partly due to the JVM and partly due to the ML "bash" function. *)
@@ -1046,7 +1048,7 @@
         (* unit -> unit *)
         fun remove_temporary_files () =
           if overlord then ()
-          else List.app File.rm [in_path, out_path, err_path]
+          else List.app (K () o try File.rm) [in_path, out_path, err_path]
       in
         let
           val ms =
@@ -1076,11 +1078,13 @@
                       " < " ^ File.shell_path in_path ^
                       " > " ^ File.shell_path out_path ^
                       " 2> " ^ File.shell_path err_path)
-              val (ps, nontriv_js) = read_output_file out_path
-                                     |>> map (apfst reindex) ||> map reindex
+              val (io_error, (ps, nontriv_js)) =
+                read_output_file out_path
+                ||> apfst (map (apfst reindex)) ||> apsnd (map reindex)
               val js = triv_js @ nontriv_js
               val first_error =
                 File.fold_lines (fn line => fn "" => line | s => s) err_path ""
+                handle IO.Io _ => "" | OS.SysErr _ => ""
             in
               if null ps then
                 if code = 2 then
@@ -1092,6 +1096,8 @@
                 else if first_error <> "" then
                   Error (first_error |> perhaps (try (unsuffix "."))
                                      |> perhaps (try (unprefix "Error: ")), js)
+                else if io_error then
+                  Error ("I/O error", js)
                 else if code <> 0 then
                   Error ("Unknown error", js)
                 else
@@ -1102,7 +1108,8 @@
         in remove_temporary_files (); outcome end
         handle Exn.Interrupt =>
                let
-                 val nontriv_js = map reindex (snd (read_output_file out_path))
+                 val nontriv_js =
+                   read_output_file out_path |> snd |> snd |> map reindex
                in
                  (remove_temporary_files ();
                   Interrupted (SOME (triv_js @ nontriv_js)))
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -54,6 +54,7 @@
   val numeral_prefix : string
   val ubfp_prefix : string
   val lbfp_prefix : string
+  val quot_normal_prefix : string
   val skolem_prefix : string
   val special_prefix : string
   val uncurry_prefix : string
@@ -173,7 +174,7 @@
   val inverse_axioms_for_rep_fun : theory -> styp -> term list
   val optimized_typedef_axioms : theory -> string * typ list -> term list
   val optimized_quot_type_axioms :
-    theory -> (typ option * bool) list -> string * typ list -> term list
+    Proof.context -> (typ option * bool) list -> string * typ list -> term list
   val def_of_const : theory -> const_table -> styp -> term option
   val fixpoint_kind_of_const :
     theory -> const_table -> string * typ -> fixpoint_kind
@@ -268,6 +269,7 @@
 val step_prefix = nitpick_prefix ^ "step" ^ name_sep
 val ubfp_prefix = nitpick_prefix ^ "ubfp" ^ name_sep
 val lbfp_prefix = nitpick_prefix ^ "lbfp" ^ name_sep
+val quot_normal_prefix = nitpick_prefix ^ "qn" ^ name_sep
 val skolem_prefix = nitpick_prefix ^ "sk"
 val special_prefix = nitpick_prefix ^ "sp"
 val uncurry_prefix = nitpick_prefix ^ "unc"
@@ -277,6 +279,9 @@
 
 (* int -> string *)
 fun sel_prefix_for j = sel_prefix ^ string_of_int j ^ name_sep
+(* Proof.context -> typ -> string *)
+fun quot_normal_name_for_type ctxt T =
+  quot_normal_prefix ^ unyxml (Syntax.string_of_typ ctxt T)
 
 (* string -> string * string *)
 val strip_first_name_sep =
@@ -559,14 +564,15 @@
 (* theory -> typ -> typ -> typ -> typ *)
 fun instantiate_type thy T1 T1' T2 =
   Same.commit (Envir.subst_type_same
-                   (Sign.typ_match thy (Logic.varifyT T1, T1') Vartab.empty))
-              (Logic.varifyT T2)
+                   (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
   handle Type.TYPE_MATCH =>
          raise TYPE ("Nitpick_HOL.instantiate_type", [T1, T1'], [])
+fun varify_and_instantiate_type thy T1 T1' T2 =
+  instantiate_type thy (Logic.varifyT T1) T1' (Logic.varifyT T2)
 
 (* theory -> typ -> typ -> styp *)
 fun repair_constr_type thy body_T' T =
-  instantiate_type thy (body_type T) body_T' T
+  varify_and_instantiate_type thy (body_type T) body_T' T
 
 (* string -> (string * string) list -> theory -> theory *)
 fun register_frac_type frac_s ersaetze thy =
@@ -889,7 +895,8 @@
              [(@{const_name Quot}, rep_type_for_quot_type thy T --> T)]
            else case typedef_info thy s of
              SOME {abs_type, rep_type, Abs_name, ...} =>
-             [(Abs_name, instantiate_type thy abs_type T rep_type --> T)]
+             [(Abs_name,
+               varify_and_instantiate_type thy abs_type T rep_type --> T)]
            | NONE =>
              if T = @{typ ind} then
                [dest_Const @{const Zero_Rep}, dest_Const @{const Suc_Rep}]
@@ -1385,7 +1392,7 @@
     else case typedef_info thy abs_s of
       SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
       let
-        val rep_T = instantiate_type thy abs_type abs_T rep_type
+        val rep_T = varify_and_instantiate_type thy abs_type abs_T rep_type
         val rep_t = Const (Rep_name, abs_T --> rep_T)
         val set_t = Const (set_name, rep_T --> bool_T)
         val set_t' =
@@ -1400,8 +1407,10 @@
       end
     | NONE => []
   end
-fun optimized_quot_type_axioms thy stds abs_z =
+(* Proof.context -> string * typ list -> term list *)
+fun optimized_quot_type_axioms ctxt stds abs_z =
   let
+    val thy = ProofContext.theory_of ctxt
     val abs_T = Type abs_z
     val rep_T = rep_type_for_quot_type thy abs_T
     val equiv_rel = equiv_relation_for_quot_type thy abs_T
@@ -1410,7 +1419,7 @@
     val y_var = Var (("y", 0), rep_T)
     val x = (@{const_name Quot}, rep_T --> abs_T)
     val sel_a_t = select_nth_constr_arg thy stds x a_var 0 rep_T
-    val normal_t = Const (@{const_name quot_normal}, rep_T --> rep_T)
+    val normal_t = Const (quot_normal_name_for_type ctxt abs_T, rep_T --> rep_T)
     val normal_x = normal_t $ x_var
     val normal_y = normal_t $ y_var
     val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
@@ -1639,7 +1648,7 @@
                 in
                   (Abs (Name.uu, rep_T,
                         Const (@{const_name Quot}, rep_T --> abs_T)
-                               $ (Const (@{const_name quot_normal},
+                               $ (Const (quot_normal_name_for_type ctxt abs_T,
                                          rep_T --> rep_T) $ Bound 0)), ts)
                 end
               else if is_quot_rep_fun ctxt x then
@@ -1819,8 +1828,7 @@
                              termination_tacs
            in Synchronized.change cached_wf_props (cons (prop, wf)); wf end
        end)
-    handle List.Empty => false
-         | NO_TRIPLE () => false
+    handle List.Empty => false | NO_TRIPLE () => false
 
 (* The type constraint below is a workaround for a Poly/ML crash. *)
 
@@ -2231,6 +2239,10 @@
        else if String.isPrefix step_prefix s then
          (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
           format_type default_format default_format T)
+       else if String.isPrefix quot_normal_prefix s then
+         let val t = Const (nitpick_prefix ^ "normalize quotient type", T) in
+           (t, format_term_type thy def_table formats t)
+         end
        else if String.isPrefix skolem_prefix s then
          let
            val ss = the (AList.lookup (op =) (!skolems) s)
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -1595,12 +1595,7 @@
           KK.Atom (offset_of_type ofs nat_T)
         else
           fold kk_join (map to_integer [u1, u2]) (KK.Rel nat_subtract_rel)
-      | Op2 (Apply, _, R, u1, u2) =>
-        if is_Cst Unrep u2 andalso is_set_type (type_of u1) andalso
-           is_FreeName u1 then
-          false_atom
-        else
-          to_apply R u1 u2
+      | Op2 (Apply, _, R, u1, u2) => to_apply R u1 u2
       | Op2 (Lambda, _, R as Opt (Atom (1, j0)), u1, u2) =>
         to_guard [u1, u2] R (KK.Atom j0)
       | Op2 (Lambda, _, Func (_, Formula Neut), u1, u2) =>
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -95,7 +95,6 @@
   val nickname_of : nut -> string
   val is_skolem_name : nut -> bool
   val is_eval_name : nut -> bool
-  val is_FreeName : nut -> bool
   val is_Cst : cst -> nut -> bool
   val fold_nut : (nut -> 'a -> 'a) -> nut -> 'a -> 'a
   val map_nut : (nut -> nut) -> nut -> nut
@@ -369,8 +368,6 @@
 fun is_eval_name u =
   String.isPrefix eval_prefix (nickname_of u)
   handle NUT ("Nitpick_Nut.nickname_of", _) => false
-fun is_FreeName (FreeName _) = true
-  | is_FreeName _ = false
 (* cst -> nut -> bool *)
 fun is_Cst cst (Cst (cst', _, _)) = (cst = cst')
   | is_Cst _ _ = false
@@ -794,9 +791,9 @@
   end
 
 (* A nut is said to be constructive if whenever it evaluates to unknown in our
-   three-valued logic, it would evaluate to a unrepresentable value ("unrep")
+   three-valued logic, it would evaluate to a unrepresentable value ("Unrep")
    according to the HOL semantics. For example, "Suc n" is constructive if "n"
-   is representable or "Unrep", because unknown implies unrep. *)
+   is representable or "Unrep", because unknown implies "Unrep". *)
 (* nut -> bool *)
 fun is_constructive u =
   is_Cst Unrep u orelse
@@ -819,6 +816,16 @@
 fun unknown_boolean T R =
   Cst (case R of Formula Pos => False | Formula Neg => True | _ => Unknown,
        T, R)
+(* nut -> bool *)
+fun is_fully_representable_set u =
+  not (is_opt_rep (rep_of u)) andalso
+  case u of
+    FreeName _ => true
+  | Op1 (SingletonSet, _, _, _) => true
+  | Op2 (oper, _, _, u1, u2) =>
+    member (op =) [Union, SetDifference, Intersect] oper andalso
+    forall is_fully_representable_set [u1, u2]
+  | _ => false
 
 (* op1 -> typ -> rep -> nut -> nut *)
 fun s_op1 oper T R u1 =
@@ -860,7 +867,7 @@
         if is_constructive u1 then
           Cst (Unrep, T, R)
         else if is_boolean_type T then
-          if is_FreeName u1 then Cst (False, T, Formula Neut)
+          if is_fully_representable_set u1 then Cst (False, T, Formula Neut)
           else unknown_boolean T R
         else case u1 of
           Op2 (Apply, _, _, ConstName (@{const_name List.append}, _, _), _) =>
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -293,15 +293,15 @@
         $ do_term new_Ts old_Ts polar t2
       | Const (s as @{const_name The}, T) => do_description_operator s T
       | Const (s as @{const_name Eps}, T) => do_description_operator s T
-      | Const (@{const_name quot_normal}, Type ("fun", [T, _])) =>
-        let val T' = box_type hol_ctxt InFunLHS T in
-          Const (@{const_name quot_normal}, T' --> T')
-        end
       | Const (s as @{const_name Tha}, T) => do_description_operator s T
       | Const (x as (s, T)) =>
         Const (s, if s = @{const_name converse} orelse
                      s = @{const_name trancl} then
                     box_relational_operator_type T
+                  else if String.isPrefix quot_normal_prefix s then
+                    let val T' = box_type hol_ctxt InFunLHS (domain_type T) in
+                      T' --> T'
+                    end
                   else if is_built_in_const thy stds fast_descrs x orelse
                           s = @{const_name Sigma} then
                     T
@@ -1022,8 +1022,9 @@
 
 (* hol_context -> term -> (term list * term list) * (bool * bool) *)
 fun axioms_for_term
-        (hol_ctxt as {thy, max_bisim_depth, stds, user_axioms, fast_descrs,
-                      evals, def_table, nondef_table, user_nondefs, ...}) t =
+        (hol_ctxt as {thy, ctxt, max_bisim_depth, stds, user_axioms,
+                      fast_descrs, evals, def_table, nondef_table, user_nondefs,
+                      ...}) t =
   let
     type accumulator = styp list * (term list * term list)
     (* (term list * term list -> term list)
@@ -1090,7 +1091,8 @@
                else
                  accum |> fold (add_nondef_axiom depth)
                                (nondef_props_for_const thy false nondef_table x)
-                       |> is_funky_typedef thy (range_type T)
+                       |> (is_funky_typedef thy (range_type T) orelse
+                           range_type T = nat_T)
                           ? fold (add_maybe_def_axiom depth)
                                  (nondef_props_for_const thy true
                                                     (extra_table def_table s) x)
@@ -1100,7 +1102,8 @@
                else
                  accum |> fold (add_nondef_axiom depth)
                                (nondef_props_for_const thy false nondef_table x)
-                       |> is_funky_typedef thy (range_type T)
+                       |> (is_funky_typedef thy (range_type T) orelse
+                           range_type T = nat_T)
                           ? fold (add_maybe_def_axiom depth)
                                  (nondef_props_for_const thy true
                                                     (extra_table def_table s) x)
@@ -1134,7 +1137,8 @@
         #> (if is_pure_typedef thy T then
               fold (add_maybe_def_axiom depth) (optimized_typedef_axioms thy z)
             else if is_quot_type thy T then
-              fold (add_def_axiom depth) (optimized_quot_type_axioms thy stds z)
+              fold (add_def_axiom depth)
+                   (optimized_quot_type_axioms ctxt stds z)
             else if max_bisim_depth >= 0 andalso is_codatatype thy T then
               fold (add_maybe_def_axiom depth)
                    (codatatype_bisim_axioms hol_ctxt T)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -8,13 +8,13 @@
 sig
   val setup : theory -> theory
   val preprocess : Predicate_Compile_Aux.options -> string -> theory -> theory
+  val present_graph : bool Unsynchronized.ref
 end;
 
 structure Predicate_Compile (*: PREDICATE_COMPILE*) =
 struct
 
-(* options *)
-val fail_safe_mode = true
+val present_graph = Unsynchronized.ref false
 
 open Predicate_Compile_Aux;
 
@@ -33,10 +33,12 @@
            commas (map (Display.string_of_thm_global thy) intros)) intross)))
   else ()
       
-fun print_specs thy specs =
-  map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
-    ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
-
+fun print_specs options thy specs =
+  if show_intermediate_results options then
+    map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
+      ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
+    |> space_implode "\n" |> tracing
+  else ()
 fun overload_const thy s = the_default s (Option.map fst (AxClass.inst_of_param thy s))
 
 fun map_specs f specs =
@@ -47,15 +49,21 @@
     val _ = print_step options "Compiling predicates to flat introrules..."
     val specs = map (apsnd (map
       (fn th => if is_equationlike th then Predicate_Compile_Data.normalize_equation thy' th else th))) specs
-    val (intross1, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess specs thy')
+    val (intross1, thy'') =
+      apfst flat (fold_map (Predicate_Compile_Pred.preprocess options) specs thy')
     val _ = print_intross options thy'' "Flattened introduction rules: " intross1
     val _ = print_step options "Replacing functions in introrules..."
     val intross2 =
-      if fail_safe_mode then
-        case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
-          SOME intross => intross
-        | NONE => let val _ = warning "Function replacement failed!" in intross1 end
-      else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
+      if function_flattening options then
+        if fail_safe_function_flattening options then
+          case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
+            SOME intross => intross
+          | NONE =>
+            (if show_caught_failures options then tracing "Function replacement failed!" else ();
+            intross1)
+        else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
+      else
+        intross1
     val _ = print_intross options thy'' "Introduction rules with replaced functions: " intross2
     val _ = print_step options "Introducing new constants for abstractions at higher-order argument positions..."
     val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'')
@@ -70,35 +78,45 @@
   end
 
 fun preprocess_strong_conn_constnames options gr ts thy =
-  let
-    fun get_specs ts = map_filter (fn t =>
-      TermGraph.get_node gr t |>
-      (fn ths => if null ths then NONE else SOME (fst (dest_Const t), ths)))
-      ts
-    val _ = print_step options ("Preprocessing scc of " ^
-      commas (map (Syntax.string_of_term_global thy) ts))
-    val (prednames, funnames) = List.partition (fn t => body_type (fastype_of t) = @{typ bool}) ts
-    (* untangle recursion by defining predicates for all functions *)
-    val _ = print_step options
-      ("Compiling functions (" ^ commas (map (Syntax.string_of_term_global thy) funnames) ^
-        ") to predicates...")
-    val (fun_pred_specs, thy') =
-      if not (null funnames) then Predicate_Compile_Fun.define_predicates
-      (get_specs funnames) thy else ([], thy)
-    val _ = print_specs thy' fun_pred_specs
-    val specs = (get_specs prednames) @ fun_pred_specs
-    val (intross3, thy''') = process_specification options specs thy'
-    val _ = print_intross options thy''' "Introduction rules with new constants: " intross3
-    val intross4 = map_specs (maps remove_pointless_clauses) intross3
-    val _ = print_intross options thy''' "After removing pointless clauses: " intross4
-    val intross5 = map (fn (s, ths) => (overload_const thy''' s, map (AxClass.overload thy''') ths)) intross4
-    val intross6 = map_specs (map (expand_tuples thy''')) intross5
-    val _ = print_intross options thy''' "introduction rules before registering: " intross6
-    val _ = print_step options "Registering introduction rules..."
-    val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
-  in
-    thy''''
-  end;
+  if forall (fn (Const (c, _)) => Predicate_Compile_Core.is_registered thy c) ts then
+    thy
+  else
+    let
+      fun get_specs ts = map_filter (fn t =>
+        TermGraph.get_node gr t |>
+        (fn ths => if null ths then NONE else SOME (fst (dest_Const t), ths)))
+        ts
+      val _ = print_step options ("Preprocessing scc of " ^
+        commas (map (Syntax.string_of_term_global thy) ts))
+      val (prednames, funnames) = List.partition (fn t => body_type (fastype_of t) = @{typ bool}) ts
+      (* untangle recursion by defining predicates for all functions *)
+      val _ = print_step options
+        ("Compiling functions (" ^ commas (map (Syntax.string_of_term_global thy) funnames) ^
+          ") to predicates...")
+      val (fun_pred_specs, thy') =
+        (if function_flattening options andalso (not (null funnames)) then
+          if fail_safe_function_flattening options then
+            case try (Predicate_Compile_Fun.define_predicates (get_specs funnames)) thy of
+              SOME (intross, thy) => (intross, thy)
+            | NONE => ([], thy)
+          else Predicate_Compile_Fun.define_predicates (get_specs funnames) thy
+        else ([], thy))
+        (*||> Theory.checkpoint*)
+      val _ = print_specs options thy' fun_pred_specs
+      val specs = (get_specs prednames) @ fun_pred_specs
+      val (intross3, thy''') = process_specification options specs thy'
+      val _ = print_intross options thy''' "Introduction rules with new constants: " intross3
+      val intross4 = map_specs (maps remove_pointless_clauses) intross3
+      val _ = print_intross options thy''' "After removing pointless clauses: " intross4
+      val intross5 =
+        map (fn (s, ths) => (overload_const thy''' s, map (AxClass.overload thy''') ths)) intross4
+      val intross6 = map_specs (map (expand_tuples thy''')) intross5
+      val _ = print_intross options thy''' "introduction rules before registering: " intross6
+      val _ = print_step options "Registering introduction rules..."
+      val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
+    in
+      thy''''
+    end;
 
 fun preprocess options t thy =
   let
@@ -106,6 +124,7 @@
     val gr = Output.cond_timeit (!Quickcheck.timing) "preprocess-obtain graph"
           (fn () => Predicate_Compile_Data.obtain_specification_graph options thy t
           |> (fn gr => TermGraph.subgraph (member (op =) (TermGraph.all_succs gr [t])) gr))
+    val _ = if !present_graph then Predicate_Compile_Data.present_graph gr else ()
   in
     Output.cond_timeit (!Quickcheck.timing) "preprocess-process"
       (fn () => (fold_rev (preprocess_strong_conn_constnames options gr)
@@ -128,7 +147,12 @@
       show_modes = chk "show_modes",
       show_mode_inference = chk "show_mode_inference",
       show_compilation = chk "show_compilation",
+      show_caught_failures = false,
       skip_proof = chk "skip_proof",
+      function_flattening = not (chk "no_function_flattening"),
+      fail_safe_function_flattening = false,
+      no_topmost_reordering = false,
+      no_higher_order_predicate = [],
       inductify = chk "inductify",
       compilation = compilation
     }
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -29,6 +29,7 @@
 fun find_indices f xs =
   map_filter (fn (i, true) => SOME i | (i, false) => NONE) (map_index (apsnd f) xs)
 
+fun assert check = if check then () else error "Assertion failed!"
 (* mode *)
 
 datatype mode = Bool | Input | Output | Pair of mode * mode | Fun of mode * mode
@@ -57,21 +58,47 @@
 fun dest_tuple_mode (Pair (mode, mode')) = mode :: dest_tuple_mode mode'
   | dest_tuple_mode _ = []
 
-fun all_modes_of_typ (T as Type ("fun", _)) = 
+
+fun all_modes_of_typ' (T as Type ("fun", _)) = 
+  let
+    val (S, U) = strip_type T
+  in
+    if U = HOLogic.boolT then
+      fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2)
+        (map all_modes_of_typ' S) [Bool]
+    else
+      [Input, Output]
+  end
+  | all_modes_of_typ' (Type ("*", [T1, T2])) = 
+    map_product (curry Pair) (all_modes_of_typ' T1) (all_modes_of_typ' T2)
+  | all_modes_of_typ' _ = [Input, Output]
+
+fun all_modes_of_typ (T as Type ("fun", _)) =
   let
     val (S, U) = strip_type T
   in
     if U = HOLogic.boolT then
       fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2)
-        (map all_modes_of_typ S) [Bool]
+        (map all_modes_of_typ' S) [Bool]
     else
       [Input, Output]
   end
-  | all_modes_of_typ (Type ("*", [T1, T2])) = 
-    map_product (curry Pair) (all_modes_of_typ T1) (all_modes_of_typ T2)
   | all_modes_of_typ (Type ("bool", [])) = [Bool]
-  | all_modes_of_typ _ = [Input, Output]
+  | all_modes_of_typ T = all_modes_of_typ' T
 
+fun all_smodes_of_typ (T as Type ("fun", _)) =
+  let
+    val (S, U) = strip_type T
+    fun all_smodes (Type ("*", [T1, T2])) = 
+      map_product (curry Pair) (all_smodes T1) (all_smodes T2)
+      | all_smodes _ = [Input, Output]
+  in
+    if U = HOLogic.boolT then
+      fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_smodes S) [Bool]
+    else
+      error "all_smodes_of_typ: invalid type for predicate"
+  end
+(*
 fun extract_params arg =
   case fastype_of arg of
     (T as Type ("fun", _)) =>
@@ -86,7 +113,7 @@
         extract_params t1 @ extract_params t2
       end
   | _ => []
-
+*)
 fun ho_arg_modes_of mode =
   let
     fun ho_arg_mode (m as Fun _) =  [m]
@@ -144,9 +171,10 @@
         in
           (comb_option HOLogic.mk_prod (i1, i2), comb_option HOLogic.mk_prod (o1, o2))
         end
-      | split_arg_mode' Input t = (SOME t, NONE)
-      | split_arg_mode' Output t = (NONE,  SOME t)
-      | split_arg_mode' _ _ = error "split_map_mode: mode and term do not match"
+      | split_arg_mode' m t =
+        if eq_mode (m, Input) then (SOME t, NONE)
+        else if eq_mode (m, Output) then (NONE,  SOME t)
+        else error "split_map_mode: mode and term do not match"
   in
     (pairself (map_filter I) o split_list) (map2 split_arg_mode' (strip_fun_mode mode) ts)
   end
@@ -269,7 +297,6 @@
   let
     val T = (Sign.the_const_type thy constname)
   in body_type T = @{typ "bool"} end;
-  
 
 fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
   | is_predT _ = false
@@ -373,7 +400,60 @@
   in
     Logic.list_implies (maps f premises, head)
   end
+
+
+(* split theorems of case expressions *)
+
+(*
+fun has_split_rule_cname @{const_name "nat_case"} = true
+  | has_split_rule_cname @{const_name "list_case"} = true
+  | has_split_rule_cname _ = false
   
+fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true 
+  | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true 
+  | has_split_rule_term thy _ = false
+
+fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
+  | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
+  | has_split_rule_term' thy c = has_split_rule_term thy c
+
+*)
+fun prepare_split_thm ctxt split_thm =
+    (split_thm RS @{thm iffD2})
+    |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
+      @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
+
+fun find_split_thm thy (Const (name, typ)) =
+  let
+    fun split_name str =
+      case first_field "." str
+        of (SOME (field, rest)) => field :: split_name rest
+         | NONE => [str]
+    val splitted_name = split_name name
+  in
+    if length splitted_name > 0 andalso
+       String.isSuffix "_case" (List.last splitted_name)
+    then
+      (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
+      |> space_implode "."
+      |> PureThy.get_thm thy
+      |> SOME
+      handle ERROR msg => NONE
+    else NONE
+  end
+  | find_split_thm _ _ = NONE
+
+
+(* TODO: split rules for let and if are different *)
+fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
+  | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
+  | find_split_thm' thy c = find_split_thm thy c
+
+fun has_split_thm thy t = is_some (find_split_thm thy t)
+
+fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
+
+
 (* lifting term operations to theorems *)
 
 fun map_term thy f th =
@@ -388,7 +468,16 @@
 
 (* Different options for compiler *)
 
-datatype compilation = Pred | Random | Depth_Limited | DSeq | Annotated | Random_DSeq
+datatype compilation = Pred | Random | Depth_Limited | DSeq | Annotated
+  | Pos_Random_DSeq | Neg_Random_DSeq
+
+
+fun negative_compilation_of Pos_Random_DSeq = Neg_Random_DSeq
+  | negative_compilation_of Neg_Random_DSeq = Pos_Random_DSeq
+  | negative_compilation_of c = c
+  
+fun compilation_for_polarity false Pos_Random_DSeq = Neg_Random_DSeq
+  | compilation_for_polarity _ c = c
 
 fun string_of_compilation c = case c of
     Pred => ""
@@ -396,8 +485,9 @@
   | Depth_Limited => "depth limited"
   | DSeq => "dseq"
   | Annotated => "annotated"
-  | Random_DSeq => "random dseq"
-
+  | Pos_Random_DSeq => "pos_random dseq"
+  | Neg_Random_DSeq => "neg_random_dseq"
+  
 (*datatype compilation_options =
   Pred | Random of int | Depth_Limited of int | DSeq of int | Annotated*)
 
@@ -411,8 +501,12 @@
   show_mode_inference : bool,
   show_modes : bool,
   show_compilation : bool,
+  show_caught_failures : bool,
   skip_proof : bool,
-
+  no_topmost_reordering : bool,
+  function_flattening : bool,
+  fail_safe_function_flattening : bool,
+  no_higher_order_predicate : string list,
   inductify : bool,
   compilation : compilation
 };
@@ -428,8 +522,15 @@
 fun show_modes (Options opt) = #show_modes opt
 fun show_mode_inference (Options opt) = #show_mode_inference opt
 fun show_compilation (Options opt) = #show_compilation opt
+fun show_caught_failures (Options opt) = #show_caught_failures opt
+
 fun skip_proof (Options opt) = #skip_proof opt
 
+fun function_flattening (Options opt) = #function_flattening opt
+fun fail_safe_function_flattening (Options opt) = #fail_safe_function_flattening opt
+fun no_topmost_reordering (Options opt) = #no_topmost_reordering opt
+fun no_higher_order_predicate (Options opt) = #no_higher_order_predicate opt
+
 fun is_inductify (Options opt) = #inductify opt
 
 fun compilation (Options opt) = #compilation opt
@@ -444,18 +545,22 @@
   show_modes = false,
   show_mode_inference = false,
   show_compilation = false,
+  show_caught_failures = false,
   skip_proof = true,
-  
+  no_topmost_reordering = false,
+  function_flattening = false,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
   inductify = false,
   compilation = Pred
 }
 
 val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes",
-  "show_mode_inference", "show_compilation", "skip_proof", "inductify"]
+  "show_mode_inference", "show_compilation", "skip_proof", "inductify", "no_function_flattening"]
 
 val compilation_names = [("pred", Pred),
   (*("random", Random), ("depth_limited", Depth_Limited), ("annotated", Annotated),*)
-  ("dseq", DSeq), ("random_dseq", Random_DSeq)]
+  ("dseq", DSeq), ("random_dseq", Pos_Random_DSeq)]
 
 fun print_step options s =
   if show_steps options then tracing s else ()
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -16,7 +16,7 @@
   val register_intros : string * thm list -> theory -> theory
   val is_registered : theory -> string -> bool
   val function_name_of : Predicate_Compile_Aux.compilation -> theory
-    -> string -> Predicate_Compile_Aux.mode -> string
+    -> string -> bool * Predicate_Compile_Aux.mode -> string
   val predfun_intro_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
   val predfun_elim_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
   val all_preds_of : theory -> string list
@@ -153,7 +153,7 @@
   | mode_of (Term m) = m
   | mode_of (Mode_App (d1, d2)) =
     (case mode_of d1 of Fun (m, m') =>
-        (if m = mode_of d2 then m' else error "mode_of")
+        (if eq_mode (m, mode_of d2) then m' else error "mode_of")
       | _ => error "mode_of2")
   | mode_of (Mode_Pair (d1, d2)) =
     Pair (mode_of d1, mode_of d2)
@@ -182,7 +182,7 @@
 
 type moded_clause = term list * (indprem * mode_derivation) list
 
-type 'a pred_mode_table = (string * (mode * 'a) list) list
+type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
 
 (* book-keeping *)
 
@@ -257,8 +257,9 @@
       ^ "functions defined for predicate " ^ quote name)
   | SOME fun_names => fun_names
 
-fun function_name_of compilation thy name mode =
-  case AList.lookup (op =) (function_names_of compilation thy name) mode of
+fun function_name_of compilation thy name (pol, mode) =
+  case AList.lookup eq_mode
+    (function_names_of (compilation_for_polarity pol compilation) thy name) mode of
     NONE => error ("No " ^ string_of_compilation compilation
       ^ "function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name)
   | SOME function_name => function_name
@@ -296,12 +297,12 @@
   if show_modes options then
     tracing ("Inferred modes:\n" ^
       cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
-        string_of_mode ms)) modes))
+        (fn (p, m) => string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
   else ()
 
 fun print_pred_mode_table string_of_entry thy pred_mode_table =
   let
-    fun print_mode pred (mode, entry) =  "mode : " ^ string_of_mode mode
+    fun print_mode pred ((pol, mode), entry) =  "mode : " ^ string_of_mode mode
       ^ string_of_entry pred mode entry
     fun print_pred (pred, modes) =
       "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
@@ -364,7 +365,7 @@
     SOME (s, ms) => (case AList.lookup (op =) modes s of
       SOME modes =>
         let
-          val modes' = modes
+          val modes' = map snd modes
         in
           if not (eq_set eq_mode (ms, modes')) then
             error ("expected modes were not inferred:\n"
@@ -381,7 +382,7 @@
       SOME inferred_ms =>
         let
           val preds_without_modes = map fst (filter (null o snd) (modes @ extra_modes))
-          val modes' = inferred_ms
+          val modes' = map snd inferred_ms
         in
           if not (eq_set eq_mode (ms, modes')) then
             error ("expected modes were not inferred:\n"
@@ -880,8 +881,6 @@
       Random_Sequence_CompFuns.mk_random_dseqT T) $ random
   end;
 
-
-
 (* for external use with interactive mode *)
 val pred_compfuns = PredicateCompFuns.compfuns
 val randompred_compfuns = Random_Sequence_CompFuns.compfuns;
@@ -898,6 +897,8 @@
 
 (** mode analysis **)
 
+(* options for mode analysis  are: #use_random, #reorder_premises *)
+
 fun is_constrt thy =
   let
     val cnstrs = flat (maps
@@ -935,7 +936,7 @@
     in merge (map (fn ks => i::ks) is) is end
   else [[]];
 
-fun print_failed_mode options thy modes p m rs is =
+fun print_failed_mode options thy modes p (pol, m) rs is =
   if show_mode_inference options then
     let
       val _ = tracing ("Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
@@ -943,7 +944,7 @@
     in () end
   else ()
 
-fun error_of p m is =
+fun error_of p (pol, m) is =
   ("  Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
         p ^ " violates mode " ^ string_of_mode m)
 
@@ -992,7 +993,7 @@
 fun is_invertible_function thy (Const (f, _)) = is_constr thy f
   | is_invertible_function thy _ = false
 
-fun non_invertible_subterms thy (Free _) = []
+fun non_invertible_subterms thy (t as Free _) = []
   | non_invertible_subterms thy t = 
   case (strip_comb t) of (f, args) =>
     if is_invertible_function thy f then
@@ -1029,6 +1030,9 @@
   forall
     (fn t => is_eqT (fastype_of t) andalso forall (member (op =) vs) (term_vs t))
       (non_invertible_subterms thy t)
+  andalso
+    (forall (is_eqT o snd)
+      (inter (fn ((f', _), f) => f = f') vs (Term.add_frees t [])))
 
 fun vars_of_destructable_term thy (Free (x, _)) = [x]
   | vars_of_destructable_term thy t =
@@ -1042,18 +1046,52 @@
 
 fun missing_vars vs t = subtract (op =) vs (term_vs t)
 
-fun derivations_of thy modes vs t Input = 
-    [(Term Input, missing_vars vs t)]
-  | derivations_of thy modes vs t Output =
-    if is_possible_output thy vs t then [(Term Output, [])] else []
-  | derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
+fun output_terms (Const ("Pair", _) $ t1 $ t2, Mode_Pair (d1, d2)) =
+    output_terms (t1, d1)  @ output_terms (t2, d2)
+  | output_terms (t1 $ t2, Mode_App (d1, d2)) =
+    output_terms (t1, d1)  @ output_terms (t2, d2)
+  | output_terms (t, Term Output) = [t]
+  | output_terms _ = []
+
+fun lookup_mode modes (Const (s, T)) =
+   (case (AList.lookup (op =) modes s) of
+      SOME ms => SOME (map (fn m => (Context m, [])) ms)
+    | NONE => NONE)
+  | lookup_mode modes (Free (x, _)) =
+    (case (AList.lookup (op =) modes x) of
+      SOME ms => SOME (map (fn m => (Context m , [])) ms)
+    | NONE => NONE)
+
+fun derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
     map_product
       (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
         (derivations_of thy modes vs t1 m1) (derivations_of thy modes vs t2 m2)
+  | derivations_of thy modes vs t (m as Fun _) =
+    (*let
+      val (p, args) = strip_comb t
+    in
+      (case lookup_mode modes p of
+        SOME ms => map_filter (fn (Context m, []) => let
+          val ms = strip_fun_mode m
+          val (argms, restms) = chop (length args) ms
+          val m' = fold_rev (curry Fun) restms Bool
+        in
+          if forall (fn m => eq_mode (Input, m)) argms andalso eq_mode (m', mode) then
+            SOME (fold (curry Mode_App) (map Term argms) (Context m), missing_vars vs t)
+          else NONE
+        end) ms
+      | NONE => (if is_all_input mode then [(Context mode, [])] else []))
+    end*)
+    (case try (all_derivations_of thy modes vs) t  of
+      SOME derivs =>
+        filter (fn (d, mvars) => eq_mode (mode_of d, m) andalso null (output_terms (t, d))) derivs
+    | NONE => (if is_all_input m then [(Context m, [])] else []))
   | derivations_of thy modes vs t m =
-    (case try (all_derivations_of thy modes vs) t of
-      SOME derivs => filter (fn (d, mvars) => mode_of d = m) derivs
-    | NONE => (if is_all_input m then [(Context m, [])] else []))
+    if eq_mode (m, Input) then
+      [(Term Input, missing_vars vs t)]
+    else if eq_mode (m, Output) then
+      (if is_possible_output thy vs t then [(Term Output, [])] else [])
+    else []
 and all_derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) =
   let
     val derivs1 = all_derivations_of thy modes vs t1
@@ -1073,14 +1111,8 @@
           (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of thy modes vs t2 m')
         | _ => error "Something went wrong") derivs1
   end
-  | all_derivations_of thy modes vs (Const (s, T)) =
-    (case (AList.lookup (op =) modes s) of
-      SOME ms => map (fn m => (Context m, [])) ms
-    | NONE => error ("No mode for constant " ^ s))
-  | all_derivations_of _ modes vs (Free (x, _)) =
-    (case (AList.lookup (op =) modes x) of
-      SOME ms => map (fn m => (Context m , [])) ms
-    | NONE => error ("No mode for parameter variable " ^ x))
+  | all_derivations_of thy modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
+  | all_derivations_of thy modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
   | all_derivations_of _ modes vs _ = error "all_derivations_of"
 
 fun rev_option_ord ord (NONE, NONE) = EQUAL
@@ -1097,7 +1129,7 @@
     SOME (s, _) =>
       (case AList.lookup (op =) modes s of
         SOME ms =>
-          (case AList.lookup (op =) ms (head_mode_of deriv) of
+          (case AList.lookup (op =) (map (fn ((p, m), r) => (m, r)) ms) (head_mode_of deriv) of
             SOME r => r
           | NONE => false)
       | NONE => false)
@@ -1146,51 +1178,56 @@
   tracing ("modes: " ^ (commas (map (fn (s, ms) => s ^ ": " ^
     commas (map (fn (m, r) => string_of_mode m ^ (if r then " random " else " not ")) ms)) modes)))
 
-fun select_mode_prem' thy modes vs ps =
+fun select_mode_prem mode_analysis_options thy pol (modes, (pos_modes, neg_modes)) vs ps =
   let
-    val modes' = map (fn (s, ms) => (s, map fst ms)) modes
+    fun choose_mode_of_prem (Prem t) = partial_hd
+        (sort (deriv_ord2 thy modes t) (all_derivations_of thy pos_modes vs t))
+      | choose_mode_of_prem (Sidecond t) = SOME (Context Bool, missing_vars vs t)
+      | choose_mode_of_prem (Negprem t) = partial_hd
+          (sort (deriv_ord2 thy modes t) (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
+             (all_derivations_of thy neg_modes vs t)))
+      | choose_mode_of_prem p = error ("choose_mode_of_prem: " ^ string_of_prem thy p)
   in
-    partial_hd (sort (premise_ord thy modes) (ps ~~ map
-    (fn Prem t =>
-      partial_hd
-        (sort (deriv_ord2 thy modes t) (all_derivations_of thy modes' vs t))
-     | Sidecond t => SOME (Context Bool, missing_vars vs t)
-     | Negprem t =>
-         partial_hd
-          (sort (deriv_ord2 thy modes t) (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
-             (all_derivations_of thy modes' vs t)))
-     | p => error (string_of_prem thy p))
-    ps))
+    if #reorder_premises mode_analysis_options then
+      partial_hd (sort (premise_ord thy modes) (ps ~~ map choose_mode_of_prem ps))
+    else
+      SOME (hd ps, choose_mode_of_prem (hd ps))
   end
 
-fun check_mode_clause' use_random thy param_vs modes mode (ts, ps) =
+fun check_mode_clause' mode_analysis_options thy param_vs (modes :
+  (string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, ps) =
   let
     val vTs = distinct (op =) (fold Term.add_frees (map term_of_prem ps) (fold Term.add_frees ts []))
-    val modes' = modes @ (param_vs ~~ map (fn x => [(x, false)]) (ho_arg_modes_of mode))
-    val (in_ts, out_ts) = split_mode mode ts    
+    val modes' = modes @ (param_vs ~~ map (fn x => [((true, x), false), ((false, x), false)]) (ho_arg_modes_of mode))
+    fun retrieve_modes_of_pol pol = map (fn (s, ms) =>
+      (s, map_filter (fn ((p, m), r) => if p = pol then SOME m else NONE | _ => NONE) ms))
+    val (pos_modes', neg_modes') =
+      if #infer_pos_and_neg_modes mode_analysis_options then
+        (retrieve_modes_of_pol pol modes', retrieve_modes_of_pol (not pol) modes')
+      else
+        let
+          val modes = map (fn (s, ms) => (s, map (fn ((p, m), r) => m) ms)) modes'
+        in (modes, modes) end
+    val (in_ts, out_ts) = split_mode mode ts
     val in_vs = maps (vars_of_destructable_term thy) in_ts
     val out_vs = terms_vs out_ts
+    fun known_vs_after p vs = (case p of
+        Prem t => union (op =) vs (term_vs t)
+      | Sidecond t => union (op =) vs (term_vs t)
+      | Negprem t => union (op =) vs (term_vs t)
+      | _ => error "I do not know")
     fun check_mode_prems acc_ps rnd vs [] = SOME (acc_ps, vs, rnd)
       | check_mode_prems acc_ps rnd vs ps =
-        (case select_mode_prem' thy modes' vs ps of
-          SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd (*TODO: uses random? *)
-            (case p of
-                Prem t => union (op =) vs (term_vs t)
-              | Sidecond t => vs
-              | Negprem t => union (op =) vs (term_vs t)
-              | _ => error "I do not know")
-            (filter_out (equal p) ps)
+        (case
+          (select_mode_prem mode_analysis_options thy pol (modes', (pos_modes', neg_modes')) vs ps) of
+          SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd
+            (known_vs_after p vs) (filter_out (equal p) ps)
         | SOME (p, SOME (deriv, missing_vars)) =>
-          if use_random then
+          if #use_random mode_analysis_options andalso pol then
             check_mode_prems ((p, deriv) :: (map
-              (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output)) missing_vars)
-                @ acc_ps) true
-            (case p of
-                Prem t => union (op =) vs (term_vs t)
-              | Sidecond t => union (op =) vs (term_vs t)
-              | Negprem t => union (op =) vs (term_vs t)
-              | _ => error "I do not know")
-            (filter_out (equal p) ps)
+              (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
+                (distinct (op =) missing_vars))
+                @ acc_ps) true (known_vs_after p vs) (filter_out (equal p) ps)
           else NONE
         | SOME (p, NONE) => NONE
         | NONE => NONE)
@@ -1201,11 +1238,11 @@
       if forall (is_constructable thy vs) (in_ts @ out_ts) then
         SOME (ts, rev acc_ps, rnd)
       else
-        if use_random then
+        if #use_random mode_analysis_options andalso pol then
           let
-            val generators = map
+             val generators = map
               (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
-                (subtract (op =) vs (terms_vs out_ts))
+                (subtract (op =) vs (terms_vs (in_ts @ out_ts)))
           in
             SOME (ts, rev (generators @ acc_ps), true)
           end
@@ -1215,66 +1252,120 @@
 
 datatype result = Success of bool | Error of string
 
-fun check_modes_pred' use_random options thy param_vs clauses modes (p, ms) =
+fun check_modes_pred' mode_analysis_options options thy param_vs clauses modes (p, (ms : ((bool * mode) * bool) list)) =
   let
     fun split xs =
       let
         fun split' [] (ys, zs) = (rev ys, rev zs)
           | split' ((m, Error z) :: xs) (ys, zs) = split' xs (ys, z :: zs)
-          | split' ((m, Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
+          | split' (((m : bool * mode), Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
        in
          split' xs ([], [])
        end
     val rs = these (AList.lookup (op =) clauses p)
     fun check_mode m =
       let
-        val res = map (check_mode_clause' use_random thy param_vs modes m) rs
+        val res = Output.cond_timeit false "work part of check_mode for one mode" (fn _ => 
+          map (check_mode_clause' mode_analysis_options thy param_vs modes m) rs)
       in
+        Output.cond_timeit false "aux part of check_mode for one mode" (fn _ => 
         case find_indices is_none res of
           [] => Success (exists (fn SOME (_, _, true) => true | _ => false) res)
-        | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is))
+        | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is)))
       end
-    val res = map (fn (m, _) => (m, check_mode m)) ms
+    val _ = if show_mode_inference options then
+        tracing ("checking " ^ string_of_int (length ms) ^ " modes ...")
+      else ()
+    val res = Output.cond_timeit false "check_mode" (fn _ => map (fn (m, _) => (m, check_mode m)) ms)
     val (ms', errors) = split res
   in
-    ((p, ms'), errors)
+    ((p, (ms' : ((bool * mode) * bool) list)), errors)
   end;
 
-fun get_modes_pred' use_random thy param_vs clauses modes (p, ms) =
+fun get_modes_pred' mode_analysis_options thy param_vs clauses modes (p, ms) =
   let
     val rs = these (AList.lookup (op =) clauses p)
   in
     (p, map (fn (m, rnd) =>
-      (m, map ((fn (ts, ps, rnd) => (ts, ps)) o the o check_mode_clause' use_random thy param_vs modes m) rs)) ms)
+      (m, map
+        ((fn (ts, ps, rnd) => (ts, ps)) o the o
+          check_mode_clause' mode_analysis_options thy param_vs modes m) rs)) ms)
   end;
 
-fun fixp f x =
+fun fixp f (x : (string * ((bool * mode) * bool) list) list) =
   let val y = f x
   in if x = y then x else fixp f y end;
 
-fun fixp_with_state f (x, state) =
+fun fixp_with_state f (x : (string * ((bool * mode) * bool) list) list, state) =
   let
     val (y, state') = f (x, state)
   in
     if x = y then (y, state') else fixp_with_state f (y, state')
   end
 
-fun infer_modes use_random options preds extra_modes param_vs clauses thy =
+fun string_of_ext_mode ((pol, mode), rnd) =
+  string_of_mode mode ^ "(" ^ (if pol then "pos" else "neg") ^ ", "
+  ^ (if rnd then "rnd" else "nornd") ^ ")"
+
+fun print_extra_modes options modes =
+  if show_mode_inference options then
+    tracing ("Modes of inferred predicates: " ^
+      cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_ext_mode ms)) modes))
+  else ()
+
+fun infer_modes mode_analysis_options options compilation preds all_modes param_vs clauses thy =
   let
-    val all_modes = map (fn (s, T) => (s, map (rpair false) (all_modes_of_typ T))) preds
-    fun needs_random s m = (m, member (op =) (#needs_random (the_pred_data thy s)) m)
-    val extra_modes = map (fn (s, ms) => (s, map (needs_random s) ms)) extra_modes
-    val (modes, errors) =
-      fixp_with_state (fn (modes, errors) =>
+    val collect_errors = false
+    fun appair f (x1, x2) (y1, y2) = (f x1 y1, f x2 y2)
+    fun needs_random s (false, m) = ((false, m), false)
+      | needs_random s (true, m) = ((true, m), member (op =) (#needs_random (the_pred_data thy s)) m)
+    fun add_polarity_and_random_bit s b ms = map (fn m => needs_random s (b, m)) ms
+    val prednames = map fst preds
+    (* extramodes contains all modes of all constants, should we only use the necessary ones
+       - what is the impact on performance? *)
+    val extra_modes =
+      if #infer_pos_and_neg_modes mode_analysis_options then
         let
-          val res = map
-            (check_modes_pred' use_random options thy param_vs clauses (modes @ extra_modes)) modes
-        in (map fst res, errors @ maps snd res) end)
-          (all_modes, [])
+          val pos_extra_modes =
+            all_modes_of compilation thy |> filter_out (fn (name, _) => member (op =) prednames name)
+          val neg_extra_modes =
+            all_modes_of (negative_compilation_of compilation) thy
+            |> filter_out (fn (name, _) => member (op =) prednames name)
+        in
+          map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)
+                @ add_polarity_and_random_bit s false (the (AList.lookup (op =) neg_extra_modes s))))
+            pos_extra_modes
+        end
+      else
+        map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)))
+          (all_modes_of compilation thy |> filter_out (fn (name, _) => member (op =) prednames name))
+    val _ = print_extra_modes options extra_modes
+    val start_modes =
+      if #infer_pos_and_neg_modes mode_analysis_options then
+        map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms @
+          (map (fn m => ((false, m), false)) ms))) all_modes
+      else
+        map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms)) all_modes
+    fun iteration modes = map
+      (check_modes_pred' mode_analysis_options options thy param_vs clauses (modes @ extra_modes))
+        modes
+    val ((modes : (string * ((bool * mode) * bool) list) list), errors) =
+      Output.cond_timeit false "Fixpount computation of mode analysis" (fn () =>
+      if collect_errors then
+        fixp_with_state (fn (modes, errors) =>
+          let
+            val (modes', new_errors) = split_list (iteration modes)
+          in (modes', errors @ flat new_errors) end) (start_modes, [])
+        else
+          (fixp (fn modes => map fst (iteration modes)) start_modes, []))
+    val moded_clauses = map (get_modes_pred' mode_analysis_options thy param_vs clauses
+      (modes @ extra_modes)) modes
     val thy' = fold (fn (s, ms) => if member (op =) (map fst preds) s then
-      set_needs_random s (map fst (filter (fn (_, rnd) => rnd = true) ms)) else I) modes thy
+      set_needs_random s (map_filter (fn ((true, m), true) => SOME m | _ => NONE) ms) else I)
+      modes thy
+
   in
-    ((map (get_modes_pred' use_random thy param_vs clauses (modes @ extra_modes)) modes, errors), thy')
+    ((moded_clauses, errors), thy')
   end;
 
 (* term construction *)
@@ -1414,14 +1505,25 @@
        (v', mk_bot compfuns U')]))
   end;
 
-fun compile_expr compilation_modifiers compfuns thy (t, deriv) additional_arguments =
+fun string_of_tderiv thy (t, deriv) = 
+  (case (t, deriv) of
+    (t1 $ t2, Mode_App (deriv1, deriv2)) =>
+      string_of_tderiv thy (t1, deriv1) ^ " $ " ^ string_of_tderiv thy (t2, deriv2)
+  | (Const ("Pair", _) $ t1 $ t2, Mode_Pair (deriv1, deriv2)) =>
+    "(" ^ string_of_tderiv thy (t1, deriv1) ^ ", " ^ string_of_tderiv thy (t2, deriv2) ^ ")"
+  | (t, Term Input) => Syntax.string_of_term_global thy t ^ "[Input]"
+  | (t, Term Output) => Syntax.string_of_term_global thy t ^ "[Output]"
+  | (t, Context m) => Syntax.string_of_term_global thy t ^ "[" ^ string_of_mode m ^ "]")
+
+fun compile_expr compilation_modifiers compfuns thy pol (t, deriv) additional_arguments =
   let
     fun expr_of (t, deriv) =
       (case (t, deriv) of
         (t, Term Input) => SOME t
       | (t, Term Output) => NONE
       | (Const (name, T), Context mode) =>
-        SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy name mode,
+        SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy name
+          (pol, mode),
           Comp_Mod.funT_of compilation_modifiers mode T))
       | (Free (s, T), Context m) =>
         SOME (Free (s, Comp_Mod.funT_of compilation_modifiers m T))
@@ -1446,7 +1548,7 @@
   end
 
 fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments
-  mode inp (ts, moded_ps) =
+  (pol, mode) inp (ts, moded_ps) =
   let
     val iss = ho_arg_modes_of mode
     val compile_match = compile_match compilation_modifiers compfuns
@@ -1479,7 +1581,7 @@
                  let
                    val u =
                      compile_expr compilation_modifiers compfuns thy
-                       (t, deriv) additional_arguments'
+                       pol (t, deriv) additional_arguments'
                    val (_, out_ts''') = split_mode mode (snd (strip_comb t))
                    val rest = compile_prems out_ts''' vs' names'' ps
                  in
@@ -1489,7 +1591,7 @@
                  let
                    val u = mk_not compfuns
                      (compile_expr compilation_modifiers compfuns thy
-                       (t, deriv) additional_arguments')
+                       (not pol) (t, deriv) additional_arguments')
                    val (_, out_ts''') = split_mode mode (snd (strip_comb t))
                    val rest = compile_prems out_ts''' vs' names'' ps
                  in
@@ -1506,6 +1608,7 @@
              | Generator (v, T) =>
                  let
                    val u = mk_random T
+                   
                    val rest = compile_prems [Free (v, T)]  vs' names'' ps;
                  in
                    (u, rest)
@@ -1519,7 +1622,7 @@
     mk_bind compfuns (mk_single compfuns inp, prem_t)
   end
 
-fun compile_pred compilation_modifiers thy all_vs param_vs s T mode moded_cls =
+fun compile_pred compilation_modifiers thy all_vs param_vs s T (pol, mode) moded_cls =
   let
     (* TODO: val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers
       (all_vs @ param_vs)
@@ -1547,14 +1650,14 @@
       (fn t as Free (x, _) => if member (op =) param_vs x then NONE else SOME t | t => SOME t)) in_ts
     val cl_ts =
       map (compile_clause compilation_modifiers compfuns
-        thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts')) moded_cls;
+        thy all_vs param_vs additional_arguments (pol, mode) (HOLogic.mk_tuple in_ts')) moded_cls;
     val compilation = Comp_Mod.wrap_compilation compilation_modifiers compfuns
       s T mode additional_arguments
       (if null cl_ts then
         mk_bot compfuns (HOLogic.mk_tupleT outTs)
       else foldr1 (mk_sup compfuns) cl_ts)
     val fun_const =
-      Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy s mode, funT)
+      Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy s (pol, mode), funT)
   in
     HOLogic.mk_Trueprop
       (HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation))
@@ -1583,13 +1686,20 @@
   | strip_split_abs (Abs (_, _, t)) = strip_split_abs t
   | strip_split_abs t = t
 
-fun mk_args is_eval (Pair (m1, m2), Type ("*", [T1, T2])) names =
-    let
-      val (t1, names') = mk_args is_eval (m1, T1) names
-      val (t2, names'') = mk_args is_eval (m2, T2) names'
-    in
-      (HOLogic.mk_prod (t1, t2), names'')
-    end
+fun mk_args is_eval (m as Pair (m1, m2), T as Type ("*", [T1, T2])) names =
+    if eq_mode (m, Input) orelse eq_mode (m, Output) then
+      let
+        val x = Name.variant names "x"
+      in
+        (Free (x, T), x :: names)
+      end
+    else
+      let
+        val (t1, names') = mk_args is_eval (m1, T1) names
+        val (t2, names'') = mk_args is_eval (m2, T2) names'
+      in
+        (HOLogic.mk_prod (t1, t2), names'')
+      end
   | mk_args is_eval ((m as Fun _), T) names =
     let
       val funT = funT_of PredicateCompFuns.compfuns m T
@@ -1828,11 +1938,11 @@
 
 (* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
 
-fun prove_sidecond thy modes t =
+fun prove_sidecond thy t =
   let
     fun preds_of t nameTs = case strip_comb t of 
       (f as Const (name, T), args) =>
-        if AList.defined (op =) modes name then (name, T) :: nameTs
+        if is_registered thy name then (name, T) :: nameTs
           else fold preds_of args nameTs
       | _ => nameTs
     val preds = preds_of t []
@@ -1847,7 +1957,7 @@
     (* need better control here! *)
   end
 
-fun prove_clause options thy nargs modes mode (_, clauses) (ts, moded_ps) =
+fun prove_clause options thy nargs mode (_, clauses) (ts, moded_ps) =
   let
     val (in_ts, clause_out_ts) = split_mode mode ts;
     fun prove_prems out_ts [] =
@@ -1911,7 +2021,7 @@
           | Sidecond t =>
            rtac @{thm if_predI} 1
            THEN print_tac' options "before sidecond:"
-           THEN prove_sidecond thy modes t
+           THEN prove_sidecond thy t
            THEN print_tac' options "after sidecond:"
            THEN prove_prems [] ps)
       in (prove_match options thy out_ts)
@@ -1929,7 +2039,7 @@
   | select_sup _ 1 = [rtac @{thm supI1}]
   | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
 
-fun prove_one_direction options thy clauses preds modes pred mode moded_clauses =
+fun prove_one_direction options thy clauses preds pred mode moded_clauses =
   let
     val T = the (AList.lookup (op =) preds pred)
     val nargs = length (binder_types T)
@@ -1942,7 +2052,7 @@
     THEN (EVERY (map
            (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
              (1 upto (length moded_clauses))))
-    THEN (EVERY (map2 (prove_clause options thy nargs modes mode) clauses moded_clauses))
+    THEN (EVERY (map2 (prove_clause options thy nargs mode) clauses moded_clauses))
     THEN print_tac' options "proved one direction"
   end;
 
@@ -2026,10 +2136,10 @@
 (* FIXME: what is this for? *)
 (* replace defined by has_mode thy pred *)
 (* TODO: rewrite function *)
-fun prove_sidecond2 thy modes t = let
+fun prove_sidecond2 thy t = let
   fun preds_of t nameTs = case strip_comb t of 
     (f as Const (name, T), args) =>
-      if AList.defined (op =) modes name then (name, T) :: nameTs
+      if is_registered thy name then (name, T) :: nameTs
         else fold preds_of args nameTs
     | _ => nameTs
   val preds = preds_of t []
@@ -2044,7 +2154,7 @@
    THEN print_tac "after sidecond2 simplification"
    end
   
-fun prove_clause2 thy modes pred mode (ts, ps) i =
+fun prove_clause2 thy pred mode (ts, ps) i =
   let
     val pred_intro_rule = nth (intros_of thy pred) (i - 1)
     val (in_ts, clause_out_ts) = split_mode mode ts;
@@ -2102,7 +2212,7 @@
         | Sidecond t =>
           etac @{thm bindE} 1
           THEN etac @{thm if_predE} 1
-          THEN prove_sidecond2 thy modes t 
+          THEN prove_sidecond2 thy t
           THEN prove_prems2 [] ps)
       in print_tac "before prove_match2:"
          THEN prove_match2 thy out_ts
@@ -2119,11 +2229,11 @@
     THEN prems_tac
   end;
  
-fun prove_other_direction options thy modes pred mode moded_clauses =
+fun prove_other_direction options thy pred mode moded_clauses =
   let
     fun prove_clause clause i =
       (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
-      THEN (prove_clause2 thy modes pred mode clause i)
+      THEN (prove_clause2 thy pred mode clause i)
   in
     (DETERM (TRY (rtac @{thm unit.induct} 1)))
      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
@@ -2136,7 +2246,7 @@
 
 (** proof procedure **)
 
-fun prove_pred options thy clauses preds modes pred mode (moded_clauses, compiled_term) =
+fun prove_pred options thy clauses preds pred (pol, mode) (moded_clauses, compiled_term) =
   let
     val ctxt = ProofContext.init thy
     val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
@@ -2146,9 +2256,9 @@
         (fn _ =>
         rtac @{thm pred_iffI} 1
         THEN print_tac' options "after pred_iffI"
-        THEN prove_one_direction options thy clauses preds modes pred mode moded_clauses
+        THEN prove_one_direction options thy clauses preds pred mode moded_clauses
         THEN print_tac' options "proved one direction"
-        THEN prove_other_direction options thy modes pred mode moded_clauses
+        THEN prove_other_direction options thy pred mode moded_clauses
         THEN print_tac' options "proved other direction")
       else (fn _ => Skip_Proof.cheat_tac thy))
   end;
@@ -2173,11 +2283,11 @@
   map_preds_modes (fn pred => compile_pred comp_modifiers thy all_vs param_vs pred
       (the (AList.lookup (op =) preds pred))) moded_clauses
 
-fun prove options thy clauses preds modes moded_clauses compiled_terms =
-  map_preds_modes (prove_pred options thy clauses preds modes)
+fun prove options thy clauses preds moded_clauses compiled_terms =
+  map_preds_modes (prove_pred options thy clauses preds)
     (join_preds_modes moded_clauses compiled_terms)
 
-fun prove_by_skip options thy _ _ _ _ compiled_terms =
+fun prove_by_skip options thy _ _ _ compiled_terms =
   map_preds_modes
     (fn pred => fn mode => fn t => Drule.export_without_context (Skip_Proof.make_thm thy t))
     compiled_terms
@@ -2204,9 +2314,13 @@
     val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs]
       (ProofContext.init thy)
     val preds = map dest_Const preds
-    val extra_modes =
-      all_modes_of compilation thy |> filter_out (fn (name, _) => member (op =) prednames name)
     val all_vs = terms_vs intrs
+    val all_modes = 
+      map (fn (s, T) =>
+        (s,
+            (if member (op =) (no_higher_order_predicate options) s then
+               (all_smodes_of_typ T)
+            else (all_modes_of_typ T)))) preds
     val params =
       case intrs of
         [] =>
@@ -2219,8 +2333,12 @@
           in
             map2 (curry Free) param_names paramTs
           end
-      | (intr :: _) => maps extract_params
-          (snd (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))))
+      | (intr :: _) =>
+        let
+          val (p, args) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)) 
+        in
+          ho_args_of (hd (the (AList.lookup (op =) all_modes (fst (dest_Const p))))) args
+        end
     val param_vs = map (fst o dest_Free) params
     fun add_clause intr clauses =
       let
@@ -2232,7 +2350,7 @@
       end;
     val clauses = fold add_clause intrs []
   in
-    (preds, all_vs, param_vs, extra_modes, clauses)
+    (preds, all_vs, param_vs, all_modes, clauses)
   end;
 
 (* sanity check of introduction rules *)
@@ -2294,7 +2412,7 @@
             val arg_names = Name.variant_list []
               (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
             val args = map2 (curry Free) arg_names Ts
-            val predfun = Const (function_name_of Pred thy predname full_mode,
+            val predfun = Const (function_name_of Pred thy predname (true, full_mode),
               Ts ---> PredicateCompFuns.mk_predT @{typ unit})
             val rhs = @{term Predicate.holds} $ (list_comb (predfun, args))
             val eq_term = HOLogic.mk_Trueprop
@@ -2317,20 +2435,20 @@
 
 datatype steps = Steps of
   {
-  define_functions : options -> (string * typ) list -> string * mode list -> theory -> theory,
-  infer_modes : options -> (string * typ) list -> (string * mode list) list
+  define_functions : options -> (string * typ) list -> string * (bool * mode) list -> theory -> theory,
+  (*infer_modes : options -> (string * typ) list -> (string * mode list) list
     -> string list -> (string * (term list * indprem list) list) list
-    -> theory -> ((moded_clause list pred_mode_table * string list) * theory),
-  prove : options -> theory -> (string * (term list * indprem list) list) list
-    -> (string * typ) list -> (string * mode list) list
+    -> theory -> ((moded_clause list pred_mode_table * string list) * theory),*)
+  prove : options -> theory -> (string * (term list * indprem list) list) list -> (string * typ) list
     -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
   add_code_equations : theory -> (string * typ) list
     -> (string * thm list) list -> (string * thm list) list,
   comp_modifiers : Comp_Mod.comp_modifiers,
+  use_random : bool,
   qname : bstring
   }
 
-fun add_equations_of steps options prednames thy =
+fun add_equations_of steps mode_analysis_options options prednames thy =
   let
     fun dest_steps (Steps s) = s
     val _ = print_step options
@@ -2338,14 +2456,20 @@
       (*val _ = check_intros_elim_match thy prednames*)
       (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
     val compilation = Comp_Mod.compilation (#comp_modifiers (dest_steps steps))
-    val (preds, all_vs, param_vs, extra_modes, clauses) =
+    val _ =
+      if show_intermediate_results options then
+        tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
+      else ()
+    val (preds, all_vs, param_vs, all_modes, clauses) =
       prepare_intrs options compilation thy prednames (maps (intros_of thy) prednames)
     val _ = print_step options "Infering modes..."
     val ((moded_clauses, errors), thy') =
-      #infer_modes (dest_steps steps) options preds extra_modes param_vs clauses thy
+      (*Output.cond_timeit true "Infering modes"
+      (fn _ =>*) infer_modes mode_analysis_options
+        options compilation preds all_modes param_vs clauses thy
     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
     val _ = check_expected_modes preds options modes
-    val _ = check_proposed_modes preds options modes extra_modes errors
+    (*val _ = check_proposed_modes preds options modes (fst extra_modes) errors*)
     val _ = print_modes options thy' modes
     val _ = print_step options "Defining executable functions..."
     val thy'' = fold (#define_functions (dest_steps steps) options preds) modes thy'
@@ -2355,8 +2479,8 @@
       compile_preds (#comp_modifiers (dest_steps steps)) thy'' all_vs param_vs preds moded_clauses
     val _ = print_compiled_terms options thy'' compiled_terms
     val _ = print_step options "Proving equations..."
-    val result_thms = #prove (dest_steps steps) options thy'' clauses preds (extra_modes @ modes)
-      moded_clauses compiled_terms
+    val result_thms =
+      #prove (dest_steps steps) options thy'' clauses preds moded_clauses compiled_terms
     val result_thms' = #add_code_equations (dest_steps steps) thy'' preds
       (maps_modes result_thms)
     val qname = #qname (dest_steps steps)
@@ -2398,7 +2522,14 @@
     val thy'' = fold_rev
       (fn preds => fn thy =>
         if not (forall (defined thy) preds) then
-          add_equations_of steps options preds thy
+          let
+            val mode_analysis_options = {use_random = #use_random (dest_steps steps),
+              reorder_premises =
+                not (no_topmost_reordering options andalso not (null (inter (op =) preds names))),
+              infer_pos_and_neg_modes = #use_random (dest_steps steps)}
+          in
+            add_equations_of steps mode_analysis_options options preds thy
+          end
         else thy)
       scc thy' |> Theory.checkpoint
   in thy'' end
@@ -2468,11 +2599,15 @@
   }
 
 val add_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes false,
-  define_functions = create_definitions,
+  (Steps {
+  define_functions =
+    fn options => fn preds => fn (s, modes) =>
+      create_definitions
+      options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
   prove = prove,
   add_code_equations = add_code_equations,
   comp_modifiers = predicate_comp_modifiers,
+  use_random = false,
   qname = "equation"})
 
 val annotated_comp_modifiers = Comp_Mod.Comp_Modifiers
@@ -2499,9 +2634,9 @@
   transform_additional_arguments = K I : (indprem -> term list -> term list)
   }
 
-val random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
+val pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
   {
-  compilation = Random_DSeq,
+  compilation = Pos_Random_DSeq,
   function_name_prefix = "random_dseq_",
   compfuns = Random_Sequence_CompFuns.compfuns,
   additional_arguments = K [],
@@ -2510,6 +2645,17 @@
   transform_additional_arguments = K I : (indprem -> term list -> term list)
   }
 
+val neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
+  {
+  compilation = Neg_Random_DSeq,
+  function_name_prefix = "random_dseq_neg_",
+  compfuns = Random_Sequence_CompFuns.compfuns,
+  additional_arguments = K [],
+  wrap_compilation = K (K (K (K (K I))))
+   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+  transform_additional_arguments = K I : (indprem -> term list -> term list)
+  }
+
 (*
 val add_depth_limited_equations = gen_add_equations
   (Steps {infer_modes = infer_modes,
@@ -2521,11 +2667,15 @@
   qname = "depth_limited_equation"})
 *)
 val add_annotated_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes false,
-  define_functions = define_functions annotated_comp_modifiers PredicateCompFuns.compfuns,
+  (Steps {
+  define_functions =
+    fn options => fn preds => fn (s, modes) =>
+      define_functions annotated_comp_modifiers PredicateCompFuns.compfuns options preds
+      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
   prove = prove_by_skip,
   add_code_equations = K (K I),
   comp_modifiers = annotated_comp_modifiers,
+  use_random = false,
   qname = "annotated_equation"})
 (*
 val add_quickcheck_equations = gen_add_equations
@@ -2538,19 +2688,33 @@
   qname = "random_equation"})
 *)
 val add_dseq_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes false,
-  define_functions = define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns,
+  (Steps {
+  define_functions =
+  fn options => fn preds => fn (s, modes) =>
+    define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns
+    options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
   prove = prove_by_skip,
   add_code_equations = K (K I),
   comp_modifiers = dseq_comp_modifiers,
+  use_random = false,
   qname = "dseq_equation"})
 
 val add_random_dseq_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes true,
-  define_functions = define_functions random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns,
+  (Steps {
+  define_functions =
+    fn options => fn preds => fn (s, modes) =>
+    let
+      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
+      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
+    in define_functions pos_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
+      options preds (s, pos_modes)
+      #> define_functions neg_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
+      options preds (s, neg_modes)
+    end,
   prove = prove_by_skip,
   add_code_equations = K (K I),
-  comp_modifiers = random_dseq_comp_modifiers,
+  comp_modifiers = pos_random_dseq_comp_modifiers,
+  use_random = true,
   qname = "random_dseq_equation"})
 
 
@@ -2700,8 +2864,8 @@
           | Depth_Limited => depth_limited_comp_modifiers
           | Annotated => annotated_comp_modifiers*)
           | DSeq => dseq_comp_modifiers
-          | Random_DSeq => random_dseq_comp_modifiers
-        val t_pred = compile_expr comp_modifiers compfuns thy (body, deriv) additional_arguments;
+          | Random_DSeq => pos_random_dseq_comp_modifiers
+        val t_pred = compile_expr comp_modifiers compfuns thy true (body, deriv) additional_arguments;
         val T_pred = dest_predT compfuns (fastype_of t_pred)
         val arrange = split_lambda (HOLogic.mk_tuple outargs) output_tuple
       in
@@ -2717,7 +2881,7 @@
       case compilation of
         Random => RandomPredCompFuns.compfuns
       | DSeq => DSequence_CompFuns.compfuns
-      | Random_DSeq => Random_Sequence_CompFuns.compfuns
+      | Pos_Random_DSeq => Random_Sequence_CompFuns.compfuns
       | _ => PredicateCompFuns.compfuns
     val t = analyze_compr thy compfuns param_user_modes options t_compr;
     val T = dest_predT compfuns (fastype_of t);
@@ -2729,7 +2893,7 @@
           (Code_Eval.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
             (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
             |> Random_Engine.run))
-      | Random_DSeq =>
+      | Pos_Random_DSeq =>
           let
             val [nrandom, size, depth] = arguments
           in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -6,11 +6,17 @@
 
 signature PREDICATE_COMPILE_DATA =
 sig
-  type specification_table;
-  (*val make_const_spec_table : Predicate_Compile_Aux.options -> theory -> specification_table*)
-  val get_specification : theory -> term -> thm list
+  val ignore_consts : string list -> theory -> theory
+  val keep_functions : string list -> theory -> theory
+  val keep_function : theory -> string -> bool
+  val processed_specs : theory -> string -> (string * thm list) list option
+  val store_processed_specs : (string * (string * thm list) list) -> theory -> theory
+  
+  val get_specification : Predicate_Compile_Aux.options -> theory -> term -> thm list
   val obtain_specification_graph :
     Predicate_Compile_Aux.options -> theory -> term -> thm list TermGraph.T
+    
+  val present_graph : thm list TermGraph.T -> unit
   val normalize_equation : theory -> thm -> thm
 end;
 
@@ -22,20 +28,39 @@
 structure Data = Theory_Data
 (
   type T =
-    {const_spec_table : thm list Symtab.table};
+    {ignore_consts : unit Symtab.table,
+     keep_functions : unit Symtab.table,
+     processed_specs : ((string * thm list) list) Symtab.table};
   val empty =
-    {const_spec_table = Symtab.empty};
+    {ignore_consts = Symtab.empty,
+     keep_functions = Symtab.empty,
+     processed_specs =  Symtab.empty};
   val extend = I;
   fun merge
-    ({const_spec_table = const_spec_table1},
-     {const_spec_table = const_spec_table2}) =
-     {const_spec_table = Symtab.merge (K true) (const_spec_table1, const_spec_table2)}
+    ({ignore_consts = c1, keep_functions = k1, processed_specs = s1},
+     {ignore_consts = c2, keep_functions = k2, processed_specs = s2}) =
+     {ignore_consts = Symtab.merge (K true) (c1, c2),
+      keep_functions = Symtab.merge (K true) (k1, k2),
+      processed_specs = Symtab.merge (K true) (s1, s2)}
 );
 
-fun mk_data c = {const_spec_table = c}
-fun map_data f {const_spec_table = c} = mk_data (f c)
+
+
+fun mk_data (c, k, s) = {ignore_consts = c, keep_functions = k, processed_specs = s}
+fun map_data f {ignore_consts = c, keep_functions = k, processed_specs = s} = mk_data (f (c, k, s))
+
+fun ignore_consts cs = Data.map (map_data (apfst3 (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
+
+fun keep_functions cs = Data.map (map_data (apsnd3 (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
 
-type specification_table = thm list Symtab.table
+fun keep_function thy = Symtab.defined (#keep_functions (Data.get thy))
+
+fun processed_specs thy = Symtab.lookup (#processed_specs (Data.get thy))
+
+fun store_processed_specs (constname, specs) =
+  Data.map (map_data (aptrd3 (Symtab.update_new (constname, specs))))
+(* *)
+
 
 fun defining_term_of_introrule_term t =
   let
@@ -120,17 +145,11 @@
     val t' = Pattern.rewrite_term thy rewr [] t
     val tac = Skip_Proof.cheat_tac thy
     val th'' = Goal.prove ctxt (Term.add_free_names t' []) [] t' (fn _ => tac)
-    val th''' = LocalDefs.unfold ctxt [@{thm split_conv}] th''
+    val th''' = LocalDefs.unfold ctxt [@{thm split_conv}, @{thm fst_conv}, @{thm snd_conv}] th''
   in
     th'''
   end;
 
-fun normalize_equation thy th =
-  mk_meta_equation th
-  |> Predicate_Compile_Set.unfold_set_notation
-  |> full_fun_cong_expand
-  |> split_all_pairs thy
-  |> tap check_equation_format
 
 fun inline_equations thy th =
   let
@@ -143,69 +162,58 @@
   in
     th'
   end
-(*
-fun store_thm_in_table options ignore thy th=
-  let
-    val th = th
-      |> inline_equations options thy
-      |> Predicate_Compile_Set.unfold_set_notation
-      |> AxClass.unoverload thy
-    val (const, th) =
-      if is_equationlike th then
-        let
-          val eq = normalize_equation thy th
-        in
-          (defining_const_of_equation eq, eq)
-        end
-      else if is_introlike th then (defining_const_of_introrule th, th)
-      else error "store_thm: unexpected definition format"
-  in
-    if ignore const then I else Symtab.cons_list (const, th)
-  end
+
+fun normalize_equation thy th =
+  mk_meta_equation th
+  |> full_fun_cong_expand
+  |> split_all_pairs thy
+  |> tap check_equation_format
+  |> inline_equations thy
 
-fun make_const_spec_table options thy =
+fun normalize_intros thy th =
+  split_all_pairs thy th
+  |> inline_equations thy
+
+fun normalize thy th =
+  if is_equationlike th then
+    normalize_equation thy th
+  else
+    normalize_intros thy th
+
+fun get_specification options thy t =
   let
-    fun store ignore f =
-      fold (store_thm_in_table options ignore thy)
-        (map (Thm.transfer thy) (f ))
-    val table = Symtab.empty
-      |> store (K false) Predicate_Compile_Alternative_Defs.get
-    val ignore = Symtab.defined table
-  in
-    table
-    |> store ignore (fn ctxt => maps
-      else [])
-        
-    |> store ignore Nitpick_Simps.get
-    |> store ignore Nitpick_Intros.get
-  end
-
-fun get_specification table constname =
-  case Symtab.lookup table constname of
-    SOME thms => thms                  
-  | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
-*)
-
-fun get_specification thy t =
-  Output.cond_timeit true "get_specification" (fn () =>
-  let
+    (*val (c, T) = dest_Const t
+    val t = Const (AxClass.unoverload_const thy (c, T), T)*)
+    val _ = if show_steps options then
+        tracing ("getting specification of " ^ Syntax.string_of_term_global thy t ^
+          " with type " ^ Syntax.string_of_typ_global thy (fastype_of t))
+      else ()
     val ctxt = ProofContext.init thy
     fun filtering th =
       if is_equationlike th andalso
-        defining_const_of_equation (normalize_equation thy th) = (fst (dest_Const t)) then
+        defining_const_of_equation (normalize_equation thy th) = fst (dest_Const t) then
         SOME (normalize_equation thy th)
       else
         if is_introlike th andalso defining_term_of_introrule th = t then
           SOME th
         else
           NONE
-  in
-    case map_filter filtering (map (Thm.transfer thy) (Predicate_Compile_Alternative_Defs.get ctxt))
+    val spec = case map_filter filtering (map (normalize thy o Thm.transfer thy) (Predicate_Compile_Alternative_Defs.get ctxt))
      of [] => (case Spec_Rules.retrieve ctxt t
-       of [] => rev (map_filter filtering (map (Thm.transfer thy) (Nitpick_Intros.get ctxt)))
+       of [] => (case rev ( 
+         (map_filter filtering (map (normalize_intros thy o Thm.transfer thy)
+           (Nitpick_Intros.get ctxt))))
+         of [] => error ("No specification for " ^ (Syntax.string_of_term_global thy t))
+          | ths => ths)
        | ((_, (_, ths)) :: _) => map (normalize_equation thy o Thm.transfer thy) ths)
      | ths => rev ths
-  end)
+    val _ =
+      if show_intermediate_results options then
+        Output.tracing (commas (map (Display.string_of_thm_global thy) spec))
+      else ()
+  in
+    spec
+  end
 
 val logic_operator_names =
   [@{const_name "=="}, 
@@ -216,7 +224,8 @@
    @{const_name "op -->"},
    @{const_name "All"},
    @{const_name "Ex"}, 
-   @{const_name "op &"}]
+   @{const_name "op &"},
+   @{const_name "op |"}]
 
 fun special_cases (c, T) = member (op =) [
   @{const_name Product_Type.Unity},
@@ -233,7 +242,11 @@
   @{const_name Int.Bit1},
   @{const_name Int.Pls},
   @{const_name Int.zero_int_inst.zero_int},
-  @{const_name List.filter}] c
+  @{const_name List.filter},
+  @{const_name HOL.If},
+  @{const_name Groups.minus}
+  ] c
+
 
 fun print_specification options thy constname specs = 
   if show_intermediate_results options then
@@ -254,19 +267,43 @@
       |> filter_out has_code_pred_intros
       |> filter_out case_consts
       |> filter_out special_cases
+      |> filter_out (fn (c, _) => Symtab.defined (#ignore_consts (Data.get thy)) c)
+      |> map (fn (c, _) => (c, Sign.the_const_constraint thy c))
       |> map Const
       (*
       |> filter is_defining_constname*)
     fun extend t =
       let
-        val specs = get_specification thy t
-          |> map (inline_equations thy)
+        val specs = get_specification options thy t
           (*|> Predicate_Compile_Set.unfold_set_notation*)
         (*val _ = print_specification options thy constname specs*)
       in (specs, defiants_of specs) end;
   in
     TermGraph.extend extend t TermGraph.empty
   end;
-  
+
+
+fun present_graph gr =
+  let
+    fun eq_cname (Const (c1, _), Const (c2, _)) = (c1 = c2)
+    fun string_of_const (Const (c, _)) = c
+      | string_of_const _ = error "string_of_const: unexpected term"
+    val constss = TermGraph.strong_conn gr;
+    val mapping = Termtab.empty |> fold (fn consts => fold (fn const =>
+      Termtab.update (const, consts)) consts) constss;
+    fun succs consts = consts
+      |> maps (TermGraph.imm_succs gr)
+      |> subtract eq_cname consts
+      |> map (the o Termtab.lookup mapping)
+      |> distinct (eq_list eq_cname);
+    val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
+    
+    fun namify consts = map string_of_const consts
+      |> commas;
+    val prgr = map (fn (consts, constss) =>
+      { name = namify consts, ID = namify consts, dir = "", unfold = true,
+        path = "", parents = map namify constss }) conn;
+  in Present.display_graph prgr end;
+
 
 end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -9,6 +9,8 @@
   val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory
   val rewrite_intro : theory -> thm -> thm list
   val pred_of_function : theory -> string -> string option
+  
+  val add_function_predicate_translation : (term * term) -> theory -> theory
 end;
 
 structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN =
@@ -16,19 +18,36 @@
 
 open Predicate_Compile_Aux;
 
-(* Table from constant name (string) to term of inductive predicate *)
-structure Pred_Compile_Preproc = Theory_Data
+(* Table from function to inductive predicate *)
+structure Fun_Pred = Theory_Data
 (
-  type T = string Symtab.table;
-  val empty = Symtab.empty;
+  type T = (term * term) Item_Net.T;
+  val empty = Item_Net.init (op aconv o pairself fst) (single o fst);
   val extend = I;
-  fun merge data : T = Symtab.merge (op =) data;   (* FIXME handle Symtab.DUP ?? *)
+  val merge = Item_Net.merge;
 )
 
-fun pred_of_function thy name = Symtab.lookup (Pred_Compile_Preproc.get thy) name
+fun lookup thy net t =
+  case Item_Net.retrieve net t of
+    [] => NONE
+  | [(f, p)] =>
+    let
+      val subst = Pattern.match thy (f, t) (Vartab.empty, Vartab.empty)
+    in
+      SOME (Envir.subst_term subst p)
+    end
+  | _ => error ("Multiple matches possible for lookup of " ^ Syntax.string_of_term_global thy t)
 
-fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy) 
+fun pred_of_function thy name =
+  case Item_Net.retrieve (Fun_Pred.get thy) (Const (name, Term.dummyT)) of
+    [] => NONE
+  | [(f, p)] => SOME (fst (dest_Const p))
+  | _ => error ("Multiple matches possible for lookup of constant " ^ name)
 
+fun defined_const thy name = is_some (pred_of_function thy name)
+
+fun add_function_predicate_translation (f, p) =
+  Fun_Pred.map (Item_Net.update (f, p))
 
 fun transform_ho_typ (T as Type ("fun", _)) =
   let
@@ -63,27 +82,6 @@
       (Free (Long_Name.base_name name ^ "P", pred_type T))
   end
 
-fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
-  | mk_param thy lookup_pred t =
-  if Predicate_Compile_Aux.is_predT (fastype_of t) then
-    t
-  else
-    let
-      val (vs, body) = strip_abs t
-      val names = Term.add_free_names body []
-      val vs_names = Name.variant_list names (map fst vs)
-      val vs' = map2 (curry Free) vs_names (map snd vs)
-      val body' = subst_bounds (rev vs', body)
-      val (f, args) = strip_comb body'
-      val resname = Name.variant (vs_names @ names) "res"
-      val resvar = Free (resname, body_type (fastype_of body'))
-      (*val P = case try lookup_pred f of SOME P => P | NONE => error "mk_param"
-      val pred_body = list_comb (P, args @ [resvar])
-      *)
-      val pred_body = HOLogic.mk_eq (body', resvar)
-      val param = fold_rev lambda (vs' @ [resvar]) pred_body
-    in param end
-    
 (* creates the list of premises for every intro rule *)
 (* theory -> term -> (string list, term list list) *)
 
@@ -92,22 +90,6 @@
   val (func, args) = strip_comb lhs
 in ((func, args), rhs) end;
 
-fun string_of_typ T = Syntax.string_of_typ_global @{theory} T
-
-fun string_of_term t =
-  case t of
-    Const (c, T) => "Const (" ^ c ^ ", " ^ string_of_typ T ^ ")"
-  | Free (c, T) => "Free (" ^ c ^ ", " ^ string_of_typ T ^ ")"
-  | Var ((c, i), T) => "Var ((" ^ c ^ ", " ^ string_of_int i ^ "), " ^ string_of_typ T ^ ")"
-  | Bound i => "Bound " ^ string_of_int i
-  | Abs (x, T, t) => "Abs (" ^ x ^ ", " ^ string_of_typ T ^ ", " ^ string_of_term t ^ ")"
-  | t1 $ t2 => "(" ^ string_of_term t1 ^ ") $ (" ^ string_of_term t2 ^ ")"
-  
-fun ind_package_get_nparams thy name =
-  case try (Inductive.the_inductive (ProofContext.init thy)) name of
-    SOME (_, result) => length (Inductive.params_of (#raw_induct result))
-  | NONE => error ("No such predicate: " ^ quote name) 
-
 (* TODO: does not work with higher order functions yet *)
 fun mk_rewr_eq (func, pred) =
   let
@@ -122,49 +104,6 @@
       (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res]))
   end;
 
-fun has_split_rule_cname @{const_name "nat_case"} = true
-  | has_split_rule_cname @{const_name "list_case"} = true
-  | has_split_rule_cname _ = false
-  
-fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true 
-  | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true 
-  | has_split_rule_term thy _ = false
-
-fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
-  | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
-  | has_split_rule_term' thy c = has_split_rule_term thy c
-  
-fun prepare_split_thm ctxt split_thm =
-    (split_thm RS @{thm iffD2})
-    |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
-      @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
-
-fun find_split_thm thy (Const (name, typ)) =
-  let
-    fun split_name str =
-      case first_field "." str
-        of (SOME (field, rest)) => field :: split_name rest
-         | NONE => [str]
-    val splitted_name = split_name name
-  in
-    if length splitted_name > 0 andalso
-       String.isSuffix "_case" (List.last splitted_name)
-    then
-      (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
-      |> space_implode "."
-      |> PureThy.get_thm thy
-      |> SOME
-      handle ERROR msg => NONE
-    else NONE
-  end
-  | find_split_thm _ _ = NONE
-
-fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
-  | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
-  | find_split_thm' thy c = find_split_thm thy c
-
-fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
-
 fun folds_map f xs y =
   let
     fun folds_map' acc [] y = [(rev acc, y)]
@@ -174,23 +113,91 @@
       folds_map' [] xs y
     end;
 
-fun mk_prems thy (lookup_pred, get_nparams) t (names, prems) =
+fun keep_functions thy t =
+  case try dest_Const (fst (strip_comb t)) of
+    SOME (c, _) => Predicate_Compile_Data.keep_function thy c
+  | _ => false
+
+fun mk_prems thy lookup_pred t (names, prems) =
   let
     fun mk_prems' (t as Const (name, T)) (names, prems) =
-      if is_constr thy name orelse (is_none (try lookup_pred t)) then
+      (if is_constr thy name orelse (is_none (lookup_pred t)) then
         [(t, (names, prems))]
-      else [(lookup_pred t, (names, prems))]
+      else
+       (*(if is_none (try lookup_pred t) then
+          [(Abs ("uu", fastype_of t, HOLogic.mk_eq (t, Bound 0)), (names, prems))]
+        else*) [(the (lookup_pred t), (names, prems))])
     | mk_prems' (t as Free (f, T)) (names, prems) = 
-      [(lookup_pred t, (names, prems))]
+      (case lookup_pred t of
+        SOME t' => [(t', (names, prems))]
+      | NONE => [(t, (names, prems))])
     | mk_prems' (t as Abs _) (names, prems) =
       if Predicate_Compile_Aux.is_predT (fastype_of t) then
-      [(t, (names, prems))] else error "mk_prems': Abs "
-      (* mk_param *)
+        ([(Envir.eta_contract t, (names, prems))])
+      else
+        let
+          val (vars, body) = strip_abs t
+          val _ = assert (fastype_of body = body_type (fastype_of body))
+          val absnames = Name.variant_list names (map fst vars)
+          val frees = map2 (curry Free) absnames (map snd vars)
+          val body' = subst_bounds (rev frees, body)
+          val resname = Name.variant (absnames @ names) "res"
+          val resvar = Free (resname, fastype_of body)
+          val t = mk_prems' body' ([], [])
+            |> map (fn (res, (inner_names, inner_prems)) =>
+              let
+                fun mk_exists (x, T) t = HOLogic.mk_exists (x, T, t)
+                val vTs = 
+                  fold Term.add_frees inner_prems []
+                  |> filter (fn (x, T) => member (op =) inner_names x)
+                val t = 
+                  fold mk_exists vTs
+                  (foldr1 HOLogic.mk_conj (HOLogic.mk_eq (resvar, res) ::
+                    map HOLogic.dest_Trueprop inner_prems))
+              in
+                t
+              end)
+              |> foldr1 HOLogic.mk_disj
+              |> fold lambda (resvar :: rev frees)
+        in
+          [(t, (names, prems))]
+        end
     | mk_prems' t (names, prems) =
-      if Predicate_Compile_Aux.is_constrt thy t then
+      if Predicate_Compile_Aux.is_constrt thy t orelse keep_functions thy t then
         [(t, (names, prems))]
       else
-        if has_split_rule_term' thy (fst (strip_comb t)) then
+        case (fst (strip_comb t)) of
+          Const (@{const_name "If"}, _) =>
+            (let
+              val (_, [B, x, y]) = strip_comb t
+            in
+              (mk_prems' x (names, prems)
+              |> map (fn (res, (names, prems)) => (res, (names, (HOLogic.mk_Trueprop B) :: prems))))
+              @ (mk_prems' y (names, prems)
+              |> map (fn (res, (names, prems)) =>
+                (res, (names, (HOLogic.mk_Trueprop (HOLogic.mk_not B)) :: prems))))
+            end)
+        | Const (@{const_name "Let"}, _) => 
+            (let
+              val (_, [f, g]) = strip_comb t
+            in
+              mk_prems' f (names, prems)
+              |> maps (fn (res, (names, prems)) =>
+                mk_prems' (betapply (g, res)) (names, prems))
+            end)
+        | Const (@{const_name "split"}, _) => 
+            (let
+              val (_, [g, res]) = strip_comb t
+              val [res1, res2] = Name.variant_list names ["res1", "res2"]
+              val (T1, T2) = HOLogic.dest_prodT (fastype_of res)
+              val (resv1, resv2) = (Free (res1, T1), Free (res2, T2))
+            in
+              mk_prems' (betapplys (g, [resv1, resv2]))
+              (res1 :: res2 :: names,
+              HOLogic.mk_Trueprop (HOLogic.mk_eq (res, HOLogic.mk_prod (resv1, resv2))) :: prems)
+            end)
+        | _ =>
+        if has_split_thm thy (fst (strip_comb t)) then
           let
             val (f, args) = strip_comb t
             val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
@@ -208,8 +215,15 @@
                 val vars = map Free (var_names ~~ (map snd vTs))
                 val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
                 val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+                val (lhss : term list, rhss) =
+                  split_list (map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems')
               in
-                mk_prems' inner_t (var_names @ names, prems' @ prems)
+                folds_map mk_prems' lhss (var_names @ names, prems)
+                |> map (fn (ress, (names, prems)) =>
+                  let
+                    val prems' = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (ress ~~ rhss)
+                  in (names, prems' @ prems) end)
+                |> maps (mk_prems' inner_t)
               end
           in
             maps mk_prems_of_assm assms
@@ -219,53 +233,77 @@
             val (f, args) = strip_comb t
             (* TODO: special procedure for higher-order functions: split arguments in
               simple types and function types *)
+            val args = map (Pattern.eta_long []) args
             val resname = Name.variant names "res"
             val resvar = Free (resname, body_type (fastype_of t))
+            val _ = assert (fastype_of t = body_type (fastype_of t))
             val names' = resname :: names
             fun mk_prems'' (t as Const (c, _)) =
-              if is_constr thy c orelse (is_none (try lookup_pred t)) then
+              if is_constr thy c orelse (is_none (lookup_pred t)) then
+                let
+                  val _ = ()(*tracing ("not translating function " ^ Syntax.string_of_term_global thy t)*)
+                in
                 folds_map mk_prems' args (names', prems) |>
                 map
                   (fn (argvs, (names'', prems')) =>
                   let
                     val prem = HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
                   in (names'', prem :: prems') end)
+                end
               else
                 let
-                  val pred = lookup_pred t
-                  val nparams = get_nparams pred
-                  val (params, args) = chop nparams args
-                  val params' = map (mk_param thy lookup_pred) params
+                  (* lookup_pred is falsch für polymorphe Argumente und bool. *)
+                  val pred = the (lookup_pred t)
+                  val Ts = binder_types (fastype_of pred)
                 in
                   folds_map mk_prems' args (names', prems)
                   |> map (fn (argvs, (names'', prems')) =>
                     let
-                      val prem = HOLogic.mk_Trueprop (list_comb (pred, params' @ argvs @ [resvar]))
+                      fun lift_arg T t =
+                        if (fastype_of t) = T then t
+                        else
+                          let
+                            val _ = assert (T =
+                              (binder_types (fastype_of t) @ [@{typ bool}] ---> @{typ bool}))
+                            fun mk_if T (b, t, e) =
+                              Const (@{const_name If}, @{typ bool} --> T --> T --> T) $ b $ t $ e
+                            val Ts = binder_types (fastype_of t)
+                            val t = 
+                            list_abs (map (pair "x") Ts @ [("b", @{typ bool})],
+                              mk_if @{typ bool} (list_comb (t, map Bound (length Ts downto 1)),
+                              HOLogic.mk_eq (@{term True}, Bound 0),
+                              HOLogic.mk_eq (@{term False}, Bound 0)))
+                          in
+                            t
+                          end
+                      (*val _ = tracing ("Ts: " ^ commas (map (Syntax.string_of_typ_global thy) Ts))
+                      val _ = map2 check_arity Ts (map fastype_of (argvs @ [resvar]))*)
+                      val argvs' = map2 lift_arg (fst (split_last Ts)) argvs
+                      val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs' @ [resvar]))
                     in (names'', prem :: prems') end)
                 end
             | mk_prems'' (t as Free (_, _)) =
-                let
-                  (* higher order argument call *)
-                  val pred = lookup_pred t
-                in
-                  folds_map mk_prems' args (resname :: names, prems)
-                  |> map (fn (argvs, (names', prems')) =>
-                     let
-                       val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs @ [resvar]))
-                     in (names', prem :: prems') end)
-                end
+              folds_map mk_prems' args (names', prems) |>
+                map
+                  (fn (argvs, (names'', prems')) =>
+                  let
+                    val prem = 
+                      case lookup_pred t of
+                        NONE => HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
+                      | SOME p => HOLogic.mk_Trueprop (list_comb (p, argvs @ [resvar]))
+                  in (names'', prem :: prems') end)
             | mk_prems'' t =
               error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
           in
             map (pair resvar) (mk_prems'' f)
           end
   in
-    mk_prems' t (names, prems)
+    mk_prems' (Pattern.eta_long [] t) (names, prems)
   end;
 
 (* assumption: mutual recursive predicates all have the same parameters. *)  
 fun define_predicates specs thy =
-  if forall (fn (const, _) => member (op =) (Symtab.keys (Pred_Compile_Preproc.get thy)) const) specs then
+  if forall (fn (const, _) => defined_const thy const) specs then
     ([], thy)
   else
   let
@@ -275,36 +313,20 @@
       (* create prednames *)
     val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list
     val argss' = map (map transform_ho_arg) argss
-    val pnames = map dest_Free (distinct (op =) (maps (filter (is_funtype o fastype_of)) argss'))
+    (* TODO: higher order arguments also occur in tuples! *)
+    val ho_argss = distinct (op =) (maps (filter (is_funtype o fastype_of)) argss)
+    val params = distinct (op =) (maps (filter (is_funtype o fastype_of)) argss')
+    val pnames = map dest_Free params
     val preds = map pred_of funs
     val prednames = map (fst o dest_Free) preds
     val funnames = map (fst o dest_Const) funs
     val fun_pred_names = (funnames ~~ prednames)  
       (* mapping from term (Free or Const) to term *)
-    fun lookup_pred (Const (name, T)) =
-      (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
-          SOME c => Const (c, pred_type T)
-        | NONE =>
-          (case AList.lookup op = fun_pred_names name of
-            SOME f => Free (f, pred_type T)
-          | NONE => Const (name, T)))
-      | lookup_pred (Free (name, T)) =
-        if member op = (map fst pnames) name then
-          Free (name, transform_ho_typ T)
-        else
-          Free (name, T)
-      | lookup_pred t =
-         error ("lookup function is not defined for " ^ Syntax.string_of_term_global thy t)
-     
-        (* mapping from term (predicate term, not function term!) to int *)
-    fun get_nparams (Const (name, _)) =
-      the_default 0 (try (ind_package_get_nparams thy) name)
-    | get_nparams (Free (name, _)) =
-        (if member op = prednames name then
-          length pnames
-        else 0)
-    | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
-  
+    fun map_Free f = Free o f o dest_Free
+    val net = fold Item_Net.update
+      ((funs ~~ preds) @ (ho_argss ~~ params))
+        (Fun_Pred.get thy)
+    fun lookup_pred t = lookup thy net t
     (* create intro rules *)
   
     fun mk_intros ((func, pred), (args, rhs)) =
@@ -314,14 +336,15 @@
       else
         let
           val names = Term.add_free_names rhs []
-        in mk_prems thy (lookup_pred, get_nparams) rhs (names, [])
+        in mk_prems thy lookup_pred rhs (names, [])
           |> map (fn (resultt, (names', prems)) =>
             Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
         end
     fun mk_rewr_thm (func, pred) = @{thm refl}
   in
-    case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of
-      NONE => ([], thy) 
+    case (*try *)SOME (maps mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss))) of
+      NONE =>
+        let val _ = tracing "error occured!" in ([], thy) end
     | SOME intr_ts =>
         if is_some (try (map (cterm_of thy)) intr_ts) then
           let
@@ -333,53 +356,59 @@
                   no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
                 (map (fn (s, T) =>
                   ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds)))
-                pnames
+                []
                 (map (fn x => (Attrib.empty_binding, x)) intr_ts)
                 []
               ||> Sign.restore_naming thy
             val prednames = map (fst o dest_Const) (#preds ind_result)
             (* val rewr_thms = map mk_rewr_eq ((distinct (op =) funs) ~~ (#preds ind_result)) *)
             (* add constants to my table *)
+            
             val specs = map (fn predname => (predname, filter (Predicate_Compile_Aux.is_intro predname)
               (#intrs ind_result))) prednames
+            (*
             val thy'' = Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy'
+            *)
+            
+            val thy'' = Fun_Pred.map
+              (fold Item_Net.update (map (apfst Logic.varify)
+                (distinct (op =) funs ~~ (#preds ind_result)))) thy'
+            (*val _ = print_specs thy'' specs*)
           in
             (specs, thy'')
           end
         else
           let
-            val _ = tracing "Introduction rules of function_predicate are not welltyped"
+            val _ = Output.tracing (
+            "Introduction rules of function_predicate are not welltyped: " ^
+              commas (map (Syntax.string_of_term_global thy) intr_ts))
           in ([], thy) end
   end
 
 fun rewrite_intro thy intro =
   let
-    fun lookup_pred (Const (name, T)) =
+    (*val _ = tracing ("Rewriting intro with registered mapping for: " ^
+      commas (Symtab.keys (Pred_Compile_Preproc.get thy)))*)
+    (*fun lookup_pred (Const (name, T)) =
       (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
-        SOME c => Const (c, pred_type T)
-      | NONE => error ("Function " ^ name ^ " is not inductified"))
-    | lookup_pred (Free (name, T)) = Free (name, T)
-    | lookup_pred _ = error "lookup function is not defined!"
-
-    fun get_nparams (Const (name, _)) =
-      the_default 0 (try (ind_package_get_nparams thy) name)
-    | get_nparams (Free _) = 0
-    | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
-    
+        SOME c => SOME (Const (c, pred_type T))
+      | NONE => NONE)
+    | lookup_pred _ = NONE
+    *)
+    fun lookup_pred t = lookup thy (Fun_Pred.get thy) t
     val intro_t = (Logic.unvarify o prop_of) intro
     val (prems, concl) = Logic.strip_horn intro_t
     val frees = map fst (Term.add_frees intro_t [])
     fun rewrite prem names =
       let
+        (*val _ = tracing ("Rewriting premise " ^ Syntax.string_of_term_global thy prem ^ "...")*)
         val t = (HOLogic.dest_Trueprop prem)
         val (lit, mk_lit) = case try HOLogic.dest_not t of
             SOME t => (t, HOLogic.mk_not)
           | NONE => (t, I)
-        val (P, args) = (strip_comb lit) 
+        val (P, args) = (strip_comb lit)
       in
-        folds_map (
-          fn t => if (is_funtype (fastype_of t)) then (fn x => [(t, x)])
-            else mk_prems thy (lookup_pred, get_nparams) t) args (names, [])
+        folds_map (mk_prems thy lookup_pred) args (names, [])
         |> map (fn (resargs, (names', prems')) =>
           let
             val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -7,23 +7,85 @@
 signature PREDICATE_COMPILE_PRED =
 sig
   (* preprocesses an equation to a set of intro rules; defines new constants *)
-  (*
-  val preprocess_pred_equation : thm -> theory -> thm list * theory
-  *)
-  val preprocess : string -> theory -> (thm list list * theory) 
-  (* output is the term list of clauses of an unknown predicate *)
-  val preprocess_term : term -> theory -> (term list * theory)
-  
-  (*val rewrite : thm -> thm*)
-  
+  val preprocess : Predicate_Compile_Aux.options -> (string * thm list) -> theory
+    -> ((string * thm list) list * theory) 
+  val flat_higher_order_arguments : ((string * thm list) list * theory)
+    -> ((string * thm list) list * ((string * thm list) list * theory))
 end;
 
-(* : PREDICATE_COMPILE_PREPROC_PRED *)  (* FIXME *)
-structure Predicate_Compile_Pred =
+
+structure Predicate_Compile_Pred : PREDICATE_COMPILE_PRED =
 struct
 
 open Predicate_Compile_Aux
 
+
+fun datatype_names_of_case_name thy case_name =
+  map (#1 o #2) (#descr (the (Datatype_Data.info_of_case thy case_name)))
+
+fun make_case_rewrites new_type_names descr sorts thy =
+  let
+    val case_combs = Datatype_Prop.make_case_combs new_type_names descr sorts thy "f";
+    fun make comb =
+      let
+        val Type ("fun", [T, T']) = fastype_of comb;
+        val (Const (case_name, _), fs) = strip_comb comb
+        val used = Term.add_tfree_names comb []
+        val U = TFree (Name.variant used "'t", HOLogic.typeS)
+        val x = Free ("x", T)
+        val f = Free ("f", T' --> U)
+        fun apply_f f' =
+          let
+            val Ts = binder_types (fastype_of f')
+            val bs = map Bound ((length Ts - 1) downto 0)
+          in
+            fold (curry absdummy) (rev Ts) (f $ (list_comb (f', bs)))
+          end
+        val fs' = map apply_f fs
+        val case_c' = Const (case_name, (map fastype_of fs') @ [T] ---> U)
+      in
+        HOLogic.mk_eq (f $ (comb $ x), list_comb (case_c', fs') $ x)
+      end
+  in
+    map make case_combs
+  end
+
+fun case_rewrites thy Tcon =
+  let
+    val info = Datatype.the_info thy Tcon
+    val descr = #descr info
+    val sorts = #sorts info
+    val typ_names = the_default [Tcon] (#alt_names info)
+  in
+    map (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop)
+      (make_case_rewrites typ_names [descr] sorts thy)
+  end
+
+fun instantiated_case_rewrites thy Tcon =
+  let
+    val rew_ths = case_rewrites thy Tcon
+    val ctxt = ProofContext.init thy
+    fun instantiate th =
+    let
+      val f = (fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))))
+      val Type ("fun", [uninst_T, uninst_T']) = fastype_of f
+      val ([tname, tname', uname, yname], ctxt') = Variable.add_fixes ["'t", "'t'", "'u", "y"] ctxt
+      val T = TFree (tname, HOLogic.typeS)
+      val T' = TFree (tname', HOLogic.typeS)
+      val U = TFree (uname, HOLogic.typeS)
+      val y = Free (yname, U)
+      val f' = absdummy (U --> T', Bound 0 $ y)
+      val th' = Thm.certify_instantiate
+        ([(dest_TVar uninst_T, U --> T'), (dest_TVar uninst_T', T')],
+         [((fst (dest_Var f), (U --> T') --> T'), f')]) th
+      val [th'] = Variable.export ctxt' ctxt [th']
+   in
+     th'
+   end
+ in
+   map instantiate rew_ths
+ end
+
 fun is_compound ((Const ("Not", _)) $ t) =
     error "is_compound: Negation should not occur; preprocessing is defect"
   | is_compound ((Const ("Ex", _)) $ _) = true
@@ -35,6 +97,7 @@
 fun flatten constname atom (defs, thy) =
   if is_compound atom then
     let
+      val atom = Envir.beta_norm (Pattern.eta_long [] atom)
       val constname = Name.variant (map (Long_Name.base_name o fst) defs)
         ((Long_Name.base_name constname) ^ "_aux")
       val full_constname = Sign.full_bname thy constname
@@ -50,7 +113,82 @@
       (lhs, ((full_constname, [definition]) :: defs, thy'))
     end
   else
-    (atom, (defs, thy))
+    (case (fst (strip_comb atom)) of
+      (Const (@{const_name If}, _)) => let
+          val if_beta = @{lemma "(if c then x else y) z = (if c then x z else y z)" by simp}
+          val atom' = MetaSimplifier.rewrite_term thy
+            (map (fn th => th RS @{thm eq_reflection}) [@{thm if_bool_eq_disj}, if_beta]) [] atom
+          val _ = assert (not (atom = atom'))
+        in
+          flatten constname atom' (defs, thy)
+        end
+    | _ =>  
+      if (has_split_thm thy (fst (strip_comb atom))) then
+        let
+          val (f, args) = strip_comb atom
+          val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
+          (* TODO: contextify things - this line is to unvarify the split_thm *)
+          (*val ((_, [isplit_thm]), _) = Variable.import true [split_thm] (ProofContext.init thy)*)
+          val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
+          val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl) 
+          val Tcons = datatype_names_of_case_name thy (fst (dest_Const f))
+          val ths = maps (instantiated_case_rewrites thy) Tcons
+          val atom = MetaSimplifier.rewrite_term thy
+            (map (fn th => th RS @{thm eq_reflection}) ths) [] atom
+          val (f, args) = strip_comb atom
+          val subst = Pattern.match thy (split_t, atom) (Vartab.empty, Vartab.empty)
+          val (_, split_args) = strip_comb split_t
+          val match = split_args ~~ args
+          
+          (*
+          fun mk_prems_of_assm assm =
+            let
+              val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
+              val names = [] (* TODO *)
+              val var_names = Name.variant_list names (map fst vTs)
+              val vars = map Free (var_names ~~ (map snd vTs))
+              val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
+              val (HOLogic.dest_eq (HOLogic.dest_Trueprop prem))
+              val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+            in
+              (*mk_prems' inner_t (var_names @ names, prems' @ prems)*) error "asda"
+            end
+          *)
+          val names = Term.add_free_names atom []
+          val frees = map Free (Term.add_frees atom [])
+          val constname = Name.variant (map (Long_Name.base_name o fst) defs)
+            ((Long_Name.base_name constname) ^ "_aux")
+          val full_constname = Sign.full_bname thy constname
+          val constT = map fastype_of frees ---> HOLogic.boolT
+          val lhs = list_comb (Const (full_constname, constT), frees)
+          fun new_def assm =
+            let
+              val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
+              val var_names = Name.variant_list names (map fst vTs)
+              val vars = map Free (var_names ~~ (map snd vTs))
+              val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
+              fun mk_subst prem =
+                let
+                  val (Free (x, T), r) = HOLogic.dest_eq (HOLogic.dest_Trueprop prem)
+                in
+                  ((x, T), r)
+                end
+              val subst = map mk_subst prems'
+              val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+              val def = Logic.mk_equals (lhs, inner_t)
+            in
+              Envir.expand_term_frees subst def
+            end
+         val new_defs = map new_def assms
+         val (definition, thy') = thy
+          |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+          |> PureThy.add_axioms (map_index
+              (fn (i, t) => ((Binding.name (constname ^ "_def" ^ string_of_int i), t), [])) new_defs)
+        in
+          (lhs, ((full_constname, definition) :: defs, thy'))
+        end
+      else
+        (atom, (defs, thy)))
 
 fun flatten_intros constname intros thy =
   let
@@ -107,30 +245,60 @@
 
 val rewrite =
   Simplifier.simplify (HOL_basic_ss addsimps [@{thm Ball_def}, @{thm Bex_def}])
-  #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}]) 
+  #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
   #> Conv.fconv_rule nnf_conv 
   #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm ex_disj_distrib}])
 
-val rewrite_intros =
-(*  Simplifier.simplify (HOL_basic_ss addsimps @{thms HOL.simp_thms(9)}) *)
-  Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm not_not}])
-  
-fun preprocess (constname, specs) thy =
+fun split_conjs thy t =
+  let 
+    fun split_conjunctions (Const (@{const_name "op &"}, _) $ t1 $ t2) =
+    (split_conjunctions t1) @ (split_conjunctions t2)
+    | split_conjunctions t = [t]
+  in
+    map HOLogic.mk_Trueprop (split_conjunctions (HOLogic.dest_Trueprop t))
+  end
+
+fun rewrite_intros thy =
+  Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
+  #> Simplifier.full_simplify (HOL_basic_ss addsimps @{thms bool_simps} addsimps @{thms nnf_simps})
+  #> map_term thy (maps_premises (split_conjs thy))
+
+fun print_specs options thy msg ths =
+  if show_intermediate_results options then
+    (tracing (msg); tracing (commas (map (Display.string_of_thm_global thy) ths)))
+  else
+    ()
+(*
+fun split_cases thy th =
   let
-    val ctxt = ProofContext.init thy
+    
+  in
+    map_term thy th
+  end
+*)
+fun preprocess options (constname, specs) thy =
+(*  case Predicate_Compile_Data.processed_specs thy constname of
+    SOME specss => (specss, thy)
+  | NONE =>*)
+    let
+      val ctxt = ProofContext.init thy
       val intros =
-      if forall is_pred_equation specs then 
-        introrulify thy (map rewrite specs)
-      else if forall (is_intro constname) specs then
-        map rewrite_intros specs
-      else
-        error ("unexpected specification for constant " ^ quote constname ^ ":\n"
-          ^ commas (map (quote o Display.string_of_thm_global thy) specs))
-    val (intros', (local_defs, thy')) = flatten_intros constname intros thy
-    val (intross, thy'') = fold_map preprocess local_defs thy'
-  in
-    ((constname, intros') :: flat intross,thy'')
-  end;
+        if forall is_pred_equation specs then 
+          map (map_term thy (maps_premises (split_conjs thy))) (introrulify thy (map rewrite specs))
+        else if forall (is_intro constname) specs then
+          map (rewrite_intros thy) specs
+        else
+          error ("unexpected specification for constant " ^ quote constname ^ ":\n"
+            ^ commas (map (quote o Display.string_of_thm_global thy) specs))
+      val _ = print_specs options thy "normalized intros" intros
+      (*val intros = maps (split_cases thy) intros*)
+      val (intros', (local_defs, thy')) = flatten_intros constname intros thy
+      val (intross, thy'') = fold_map (preprocess options) local_defs thy'
+      val full_spec = (constname, intros') :: flat intross
+      (*val thy''' = Predicate_Compile_Data.store_processed_specs (constname, full_spec) thy''*)
+    in
+      (full_spec, thy'')
+    end;
 
 fun preprocess_term t thy = error "preprocess_pred_term: to implement" 
 
@@ -166,7 +334,8 @@
           else
             (arg, (new_defs, thy))
         
-        val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy)
+        val (args', (new_defs', thy')) = fold_map replace_abs_arg
+          (map Envir.beta_eta_contract args) (new_defs, thy)
       in
         (list_comb (pred, args'), (new_defs', thy'))
       end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -6,10 +6,18 @@
 
 signature PREDICATE_COMPILE_QUICKCHECK =
 sig
-  val quickcheck : Proof.context -> term -> int -> term list option
+  (*val quickcheck : Proof.context -> term -> int -> term list option*)
   val test_ref :
     ((unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option) Unsynchronized.ref
   val tracing : bool Unsynchronized.ref;
+  val quickcheck_compile_term : bool -> bool -> Proof.context -> term -> int -> term list option
+(*  val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *)
+  val quiet : bool Unsynchronized.ref;
+  val nrandom : int Unsynchronized.ref;
+  val depth : int Unsynchronized.ref;
+  val debug : bool Unsynchronized.ref;
+  val function_flattening : bool Unsynchronized.ref;
+  val no_higher_order_predicate : string list Unsynchronized.ref;
 end;
 
 structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
@@ -24,21 +32,106 @@
 
 val target = "Quickcheck"
 
+val quiet = Unsynchronized.ref false;
+
+val nrandom = Unsynchronized.ref 2;
+
+val depth = Unsynchronized.ref 8;
+
+val debug = Unsynchronized.ref false;
+val function_flattening = Unsynchronized.ref true;
+
+
+val no_higher_order_predicate = Unsynchronized.ref [];
+
 val options = Options {
   expected_modes = NONE,
   proposed_modes = NONE,
   proposed_names = [],
+  show_steps = false,
+  show_intermediate_results = false,
+  show_proof_trace = false,
+  show_modes = false,
+  show_mode_inference = false,
+  show_compilation = false,
+  show_caught_failures = false,
+  skip_proof = false,
+  compilation = Random,
+  inductify = true,
+  function_flattening = true,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
+  no_topmost_reordering = true
+}
+
+val debug_options = Options {
+  expected_modes = NONE,
+  proposed_modes = NONE,
+  proposed_names = [],
   show_steps = true,
   show_intermediate_results = true,
   show_proof_trace = false,
-  show_modes = false,
-  show_mode_inference = false,
+  show_modes = true,
+  show_mode_inference = true,
   show_compilation = false,
+  show_caught_failures = true,
   skip_proof = false,
   compilation = Random,
-  inductify = false
+  inductify = true,
+  function_flattening = true,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
+  no_topmost_reordering = true
 }
 
+
+fun set_function_flattening b
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f, 
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re}) =
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = b,
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re})
+
+fun set_fail_safe_function_flattening b
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f, 
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re}) =
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f,
+    fail_safe_function_flattening = b, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re})
+
+fun set_no_higher_order_predicate ss
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f, 
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re}) =
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f,
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss, no_topmost_reordering = re})
+
+
+fun get_options () = 
+  set_no_higher_order_predicate (!no_higher_order_predicate)
+    (set_function_flattening (!function_flattening)
+      (if !debug then debug_options else options))
+
 fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
 val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns)
 val mk_randompredT = #mk_predT (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
@@ -63,13 +156,15 @@
 
 fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
 
-fun quickcheck ctxt t =
+fun cpu_time description f =
   let
-    (*val () =
-      if !tracing then
-        tracing ("Starting quickcheck with " ^ (Syntax.string_of_term ctxt t))
-      else
-        ()*)
+    val start = start_timing ()
+    val result = Exn.capture f ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, (description, time)) end
+
+fun compile_term options ctxt t =
+  let
     val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
     val thy = (ProofContext.theory_of ctxt') 
     val (vs, t') = strip_abs t
@@ -82,44 +177,73 @@
     val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
     val const = Const (full_constname, constT)
     val t = Logic.list_implies
-      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),                               
+      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
        HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
     val tac = fn _ => Skip_Proof.cheat_tac thy1
     val intro = Goal.prove (ProofContext.init thy1) (map fst vs') [] t tac
-    (*val _ = tracing (Display.string_of_thm ctxt' intro)*)
-    val thy2 = (*Output.cond_timeit (!Quickcheck.timing) "predicate intros"
-      (fn () => *)(Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1)
-    val thy3 = (*Output.cond_timeit (!Quickcheck.timing) "predicate preprocessing"
-        (fn () =>*) (Predicate_Compile.preprocess options const thy2)
-    val thy4 = Output.cond_timeit (!Quickcheck.timing) "random_dseq compilation"
+    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
+    val (thy3, preproc_time) =  cpu_time "predicate preprocessing"
+        (fn () => Predicate_Compile.preprocess options const thy2)
+    val (thy4, core_comp_time) = cpu_time "random_dseq core compilation"
         (fn () => Predicate_Compile_Core.add_random_dseq_equations options [full_constname] thy3)
-    (*val depth_limited_modes = Predicate_Compile_Core.modes_of Depth_Limited thy'' full_constname*)
-    val modes = Predicate_Compile_Core.modes_of Random_DSeq thy4 full_constname
+    val _ = Predicate_Compile_Core.print_all_modes Pos_Random_DSeq thy4
+    val modes = Predicate_Compile_Core.modes_of Pos_Random_DSeq thy4 full_constname
     val output_mode = fold_rev (curry Fun) (map (K Output) (binder_types constT)) Bool
     val prog =
       if member eq_mode modes output_mode then
         let
-          val name = Predicate_Compile_Core.function_name_of Random_DSeq thy4 full_constname output_mode
+          val name = Predicate_Compile_Core.function_name_of Pos_Random_DSeq thy4
+            full_constname (true, output_mode)
           val T = (mk_randompredT (HOLogic.mk_tupleT (map snd vs')))
         in
           Const (name, T)
         end
-      (*else if member (op =) depth_limited_modes ([], []) then
-        let
-          val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], [])
-          val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs')))
-        in lift_pred (Const (name, T) $ Bound 0) end*)
-      else error "Predicate Compile Quickcheck failed"
+      else error ("Predicate Compile Quickcheck failed: " ^ commas (map string_of_mode modes))
     val qc_term = mk_bind (prog,
       mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
       (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
     val compilation =
-      Code_Eval.eval NONE ("Predicate_Compile_Quickcheck.test_ref", test_ref)
+      Code_Eval.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
         (fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc)
         thy4 qc_term []
   in
-    (fn size =>
-      Option.map fst (DSequence.yield (compilation size size |> Random_Engine.run) size true))
+    (fn size => fn nrandom => fn depth =>
+      Option.map fst (DSequence.yield (compilation nrandom size |> Random_Engine.run) depth true))
+  end
+
+fun try_upto quiet f i =
+  let
+    fun try' j =
+      if j <= i then
+        let
+          val _ = priority ("Executing depth " ^ string_of_int j)
+        in
+          case f j handle Match => (if quiet then ()
+             else warning "Exception Match raised during quickcheck"; NONE)
+          of NONE => try' (j + 1) | SOME q => SOME q
+        end
+      else
+        NONE
+  in
+    try' 0
+  end
+
+(* quickcheck interface functions *)
+
+fun compile_term' options ctxt t =
+  let
+    val c = compile_term options ctxt t
+  in
+    (fn size => try_upto (!quiet) (c size (!nrandom)) (!depth))
+  end
+
+fun quickcheck_compile_term function_flattening fail_safe_function_flattening ctxt t =
+  let
+     val options =
+       set_fail_safe_function_flattening fail_safe_function_flattening
+         (set_function_flattening function_flattening (get_options ()))
+  in
+    compile_term' options ctxt t
   end
 
 end;
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -37,7 +37,9 @@
   val equiv_rules_get: Proof.context -> thm list
   val equiv_rules_add: attribute
   val rsp_rules_get: Proof.context -> thm list
+  val rsp_rules_add: attribute
   val prs_rules_get: Proof.context -> thm list
+  val prs_rules_add: attribute
   val id_simps_get: Proof.context -> thm list
   val quotient_rules_get: Proof.context -> thm list
   val quotient_rules_add: attribute
@@ -241,6 +243,7 @@
    val description = "Respectfulness theorems.")
 
 val rsp_rules_get = RspRules.get
+val rsp_rules_add = RspRules.add
 
 (* preservation theorems *)
 structure PrsRules = Named_Thms
@@ -248,6 +251,7 @@
    val description = "Preservation theorems.")
 
 val prs_rules_get = PrsRules.get
+val prs_rules_add = PrsRules.add
 
 (* id simplification theorems *)
 structure IdSimps = Named_Thms
--- a/src/HOL/Typerep.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/Typerep.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -70,7 +70,8 @@
 
 add_typerep @{type_name fun}
 #> Typedef.interpretation ensure_typerep
-#> Code.type_interpretation (ensure_typerep o fst)
+#> Code.datatype_interpretation (ensure_typerep o fst)
+#> Code.abstype_interpretation (ensure_typerep o fst)
 
 end
 *}
--- a/src/HOL/W0/README.html	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<HTML>
-
-<HEAD>
-  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
-  <TITLE>HOL/W0/README</TITLE>
-</HEAD>
-
-<BODY>
-
-<H1>Type Inference for MiniML (without <tt>let</tt>)</H1>
-
-This theory defines the type inference rules and the type inference algorithm
-<em>W</em> for simply-typed lambda terms due to Milner. It proves the
-soundness and completeness of <em>W</em> w.r.t. to the rules. An optimized
-version <em>I</em> is shown to implement <em>W</em>.
-
-<P>
-
-A report describing the theory is found here:<br>
-<A HREF = "http://www4.informatik.tu-muenchen.de/~nipkow/pubs/tphol96.html">
-Formal Verification of Algorithm W: The Monomorphic Case</A>.
-
-<P>
-
-<B>NOTE:</B> This theory has been superseded by a more recent development
-which formalizes type inference for a language including <tt>let</tt>. For
-details click <A HREF="../MiniML/index.html">here</A>.
-</BODY>
-</HTML>
--- a/src/HOL/W0/ROOT.ML	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-use_thys ["W0"];
--- a/src/HOL/W0/W0.thy	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,925 +0,0 @@
-(*  Title:      HOL/W0/W0.thy
-    ID:         $Id$
-    Author:     Dieter Nazareth, Tobias Nipkow, Thomas Stauner, Markus Wenzel
-*)
-
-theory W0
-imports Main
-begin
-
-section {* Universal error monad *}
-
-datatype 'a maybe = Ok 'a | Fail
-
-definition
-  bind :: "'a maybe \<Rightarrow> ('a \<Rightarrow> 'b maybe) \<Rightarrow> 'b maybe"    (infixl "\<bind>" 60) where
-  "m \<bind> f = (case m of Ok r \<Rightarrow> f r | Fail \<Rightarrow> Fail)"
-
-syntax
-  "_bind" :: "patterns \<Rightarrow> 'a maybe \<Rightarrow> 'b \<Rightarrow> 'c"    ("(_ := _;//_)" 0)
-translations
-  "P := E; F" == "E \<bind> (\<lambda>P. F)"
-
-lemma bind_Ok [simp]: "(Ok s) \<bind> f = (f s)"
-  by (simp add: bind_def)
-
-lemma bind_Fail [simp]: "Fail \<bind> f = Fail"
-  by (simp add: bind_def)
-
-lemma split_bind:
-    "P (res \<bind> f) = ((res = Fail \<longrightarrow> P Fail) \<and> (\<forall>s. res = Ok s \<longrightarrow> P (f s)))"
-  by (induct res) simp_all
-
-lemma split_bind_asm:
-  "P (res \<bind> f) = (\<not> (res = Fail \<and> \<not> P Fail \<or> (\<exists>s. res = Ok s \<and> \<not> P (f s))))"
-  by (simp split: split_bind)
-
-lemmas bind_splits = split_bind split_bind_asm
-
-lemma bind_eq_Fail [simp]:
-  "((m \<bind> f) = Fail) = ((m = Fail) \<or> (\<exists>p. m = Ok p \<and> f p = Fail))"
-  by (simp split: split_bind)
-
-lemma rotate_Ok: "(y = Ok x) = (Ok x = y)"
-  by (rule eq_sym_conv)
-
-
-section {* MiniML-types and type substitutions *}
-
-axclass type_struct \<subseteq> type
-  -- {* new class for structures containing type variables *}
-
-datatype "typ" = TVar nat | TFun "typ" "typ"    (infixr "->" 70)
-  -- {* type expressions *}
-
-types subst = "nat => typ"
-  -- {* type variable substitution *}
-
-instance "typ" :: type_struct ..
-instance list :: (type_struct) type_struct ..
-instance "fun" :: (type, type_struct) type_struct ..
-
-
-subsection {* Substitutions *}
-
-consts
-  app_subst :: "subst \<Rightarrow> 'a::type_struct \<Rightarrow> 'a::type_struct"    ("$")
-  -- {* extension of substitution to type structures *}
-primrec (app_subst_typ)
-  app_subst_TVar: "$s (TVar n) = s n"
-  app_subst_Fun: "$s (t1 -> t2) = $s t1 -> $s t2"
-
-defs (overloaded)
-  app_subst_list: "$s \<equiv> map ($s)"
-
-consts
-  free_tv :: "'a::type_struct \<Rightarrow> nat set"
-  -- {* @{text "free_tv s"}: the type variables occuring freely in the type structure @{text s} *}
-
-primrec (free_tv_typ)
-  "free_tv (TVar m) = {m}"
-  "free_tv (t1 -> t2) = free_tv t1 \<union> free_tv t2"
-
-primrec (free_tv_list)
-  "free_tv [] = {}"
-  "free_tv (x # xs) = free_tv x \<union> free_tv xs"
-
-definition
-  dom :: "subst \<Rightarrow> nat set" where
-  "dom s = {n. s n \<noteq> TVar n}"
-  -- {* domain of a substitution *}
-
-definition
-  cod :: "subst \<Rightarrow> nat set" where
-  "cod s = (\<Union>m \<in> dom s. free_tv (s m))"
-  -- {* codomain of a substitutions: the introduced variables *}
-
-defs (overloaded)
-  free_tv_subst: "free_tv s \<equiv> dom s \<union> cod s"
-
-text {*
-  @{text "new_tv s n"} checks whether @{text n} is a new type variable
-  wrt.\ a type structure @{text s}, i.e.\ whether @{text n} is greater
-  than any type variable occuring in the type structure.
-*}
-
-definition
-  new_tv :: "nat \<Rightarrow> 'a::type_struct \<Rightarrow> bool" where
-  "new_tv n ts = (\<forall>m. m \<in> free_tv ts \<longrightarrow> m < n)"
-
-
-subsubsection {* Identity substitution *}
-
-definition
-  id_subst :: subst where
-  "id_subst = (\<lambda>n. TVar n)"
-
-lemma app_subst_id_te [simp]:
-  "$id_subst = (\<lambda>t::typ. t)"
-  -- {* application of @{text id_subst} does not change type expression *}
-proof
-  fix t :: "typ"
-  show "$id_subst t = t"
-    by (induct t) (simp_all add: id_subst_def)
-qed
-
-lemma app_subst_id_tel [simp]: "$id_subst = (\<lambda>ts::typ list. ts)"
-  -- {* application of @{text id_subst} does not change list of type expressions *}
-proof
-  fix ts :: "typ list"
-  show "$id_subst ts = ts"
-    by (induct ts) (simp_all add: app_subst_list)
-qed
-
-lemma o_id_subst [simp]: "$s o id_subst = s"
-  by (rule ext) (simp add: id_subst_def)
-
-lemma dom_id_subst [simp]: "dom id_subst = {}"
-  by (simp add: dom_def id_subst_def)
-
-lemma cod_id_subst [simp]: "cod id_subst = {}"
-  by (simp add: cod_def)
-
-lemma free_tv_id_subst [simp]: "free_tv id_subst = {}"
-  by (simp add: free_tv_subst)
-
-
-lemma cod_app_subst [simp]:
-  assumes free: "v \<in> free_tv (s n)"
-    and neq: "v \<noteq> n"
-  shows "v \<in> cod s"
-proof -
-  have "s n \<noteq> TVar n"
-  proof
-    assume "s n = TVar n"
-    with free have "v = n" by simp
-    with neq show False ..
-  qed
-  with free show ?thesis
-    by (auto simp add: dom_def cod_def)
-qed
-
-lemma subst_comp_te: "$g ($f t :: typ) = $(\<lambda>x. $g (f x)) t"
-  -- {* composition of substitutions *}
-  by (induct t) simp_all
-
-lemma subst_comp_tel: "$g ($f ts :: typ list) = $(\<lambda>x. $g (f x)) ts"
-  by (induct ts) (simp_all add: app_subst_list subst_comp_te)
-
-
-lemma app_subst_Nil [simp]: "$s [] = []"
-  by (simp add: app_subst_list)
-
-lemma app_subst_Cons [simp]: "$s (t # ts) = ($s t) # ($s ts)"
-  by (simp add: app_subst_list)
-
-lemma new_tv_TVar [simp]: "new_tv n (TVar m) = (m < n)"
-  by (simp add: new_tv_def)
-
-lemma new_tv_Fun [simp]:
-  "new_tv n (t1 -> t2) = (new_tv n t1 \<and> new_tv n t2)"
-  by (auto simp add: new_tv_def)
-
-lemma new_tv_Nil [simp]: "new_tv n []"
-  by (simp add: new_tv_def)
-
-lemma new_tv_Cons [simp]: "new_tv n (t # ts) = (new_tv n t \<and> new_tv n ts)"
-  by (auto simp add: new_tv_def)
-
-lemma new_tv_id_subst [simp]: "new_tv n id_subst"
-  by (simp add: id_subst_def new_tv_def free_tv_subst dom_def cod_def)
-
-lemma new_tv_subst:
-  "new_tv n s =
-    ((\<forall>m. n \<le> m \<longrightarrow> s m = TVar m) \<and>
-     (\<forall>l. l < n \<longrightarrow> new_tv n (s l)))"
-  apply (unfold new_tv_def)
-  apply (tactic "safe_tac HOL_cs")
-  -- {* @{text \<Longrightarrow>} *}
-    apply (tactic {* fast_tac (HOL_cs addDs [@{thm leD}] addss (@{simpset}
-      addsimps [thm "free_tv_subst", thm "dom_def"])) 1 *})
-   apply (subgoal_tac "m \<in> cod s \<or> s l = TVar l")
-    apply (tactic "safe_tac HOL_cs")
-     apply (tactic {* fast_tac (HOL_cs addDs [UnI2] addss (@{simpset}
-       addsimps [thm "free_tv_subst"])) 1 *})
-    apply (drule_tac P = "\<lambda>x. m \<in> free_tv x" in subst, assumption)
-    apply simp
-    apply (unfold free_tv_subst cod_def dom_def)
-    apply clarsimp
-  apply safe
-  apply metis
-  apply (metis linorder_not_less)+
-  done
-
-lemma new_tv_list: "new_tv n x = (\<forall>y \<in> set x. new_tv n y)"
-  by (induct x) simp_all
-
-lemma subst_te_new_tv [simp]:
-  "new_tv n (t::typ) \<Longrightarrow> $(\<lambda>x. if x = n then t' else s x) t = $s t"
-  -- {* substitution affects only variables occurring freely *}
-  by (induct t) simp_all
-
-lemma subst_tel_new_tv [simp]:
-  "new_tv n (ts::typ list) \<Longrightarrow> $(\<lambda>x. if x = n then t else s x) ts = $s ts"
-  by (induct ts) simp_all
-
-lemma new_tv_le: "n \<le> m \<Longrightarrow> new_tv n (t::typ) \<Longrightarrow> new_tv m t"
-  -- {* all greater variables are also new *}
-proof (induct t)
-  case (TVar n)
-  then show ?case by (auto intro: less_le_trans)
-next
-  case TFun
-  then show ?case by simp
-qed
-
-lemma [simp]: "new_tv n t \<Longrightarrow> new_tv (Suc n) (t::typ)"
-  by (rule lessI [THEN less_imp_le [THEN new_tv_le]])
-
-lemma new_tv_list_le:
-  assumes "n \<le> m"
-  shows "new_tv n (ts::typ list) \<Longrightarrow> new_tv m ts"
-proof (induct ts)
-  case Nil
-  then show ?case by simp
-next
-  case Cons
-  with `n \<le> m` show ?case by (auto intro: new_tv_le)
-qed
-
-lemma [simp]: "new_tv n ts \<Longrightarrow> new_tv (Suc n) (ts::typ list)"
-  by (rule lessI [THEN less_imp_le [THEN new_tv_list_le]])
-
-lemma new_tv_subst_le: "n \<le> m \<Longrightarrow> new_tv n (s::subst) \<Longrightarrow> new_tv m s"
-  apply (simp add: new_tv_subst)
-  apply clarify
-  apply (rule_tac P = "l < n" and Q = "n <= l" in disjE)
-    apply clarify
-    apply (simp_all add: new_tv_le)
-  done
-
-lemma [simp]: "new_tv n s \<Longrightarrow> new_tv (Suc n) (s::subst)"
-  by (rule lessI [THEN less_imp_le [THEN new_tv_subst_le]])
-
-lemma new_tv_subst_var:
-    "n < m \<Longrightarrow> new_tv m (s::subst) \<Longrightarrow> new_tv m (s n)"
-  -- {* @{text new_tv} property remains if a substitution is applied *}
-  by (simp add: new_tv_subst)
-
-lemma new_tv_subst_te [simp]:
-    "new_tv n s \<Longrightarrow> new_tv n (t::typ) \<Longrightarrow> new_tv n ($s t)"
-  by (induct t) (auto simp add: new_tv_subst)
-
-lemma new_tv_subst_tel [simp]:
-    "new_tv n s \<Longrightarrow> new_tv n (ts::typ list) \<Longrightarrow> new_tv n ($s ts)"
-  by (induct ts) (fastsimp simp add: new_tv_subst)+
-
-lemma new_tv_Suc_list: "new_tv n ts --> new_tv (Suc n) (TVar n # ts)"
-  -- {* auxilliary lemma *}
-  by (simp add: new_tv_list)
-
-lemma new_tv_subst_comp_1 [simp]:
-    "new_tv n (s::subst) \<Longrightarrow> new_tv n r \<Longrightarrow> new_tv n ($r o s)"
-  -- {* composition of substitutions preserves @{text new_tv} proposition *}
-  by (simp add: new_tv_subst)
-
-lemma new_tv_subst_comp_2 [simp]:
-    "new_tv n (s::subst) \<Longrightarrow> new_tv n r \<Longrightarrow> new_tv n (\<lambda>v. $r (s v))"
-  by (simp add: new_tv_subst)
-
-lemma new_tv_not_free_tv [simp]: "new_tv n ts \<Longrightarrow> n \<notin> free_tv ts"
-  -- {* new type variables do not occur freely in a type structure *}
-  by (auto simp add: new_tv_def)
-
-lemma ftv_mem_sub_ftv_list [simp]:
-    "(t::typ) \<in> set ts \<Longrightarrow> free_tv t \<subseteq> free_tv ts"
-  by (induct ts) auto
-
-text {*
-  If two substitutions yield the same result if applied to a type
-  structure the substitutions coincide on the free type variables
-  occurring in the type structure.
-*}
-
-lemma eq_subst_te_eq_free:
-    "$s1 (t::typ) = $s2 t \<Longrightarrow> n \<in> free_tv t \<Longrightarrow> s1 n = s2 n"
-  by (induct t) auto
-
-lemma eq_free_eq_subst_te:
-    "(\<forall>n. n \<in> free_tv t --> s1 n = s2 n) \<Longrightarrow> $s1 (t::typ) = $s2 t"
-  by (induct t) auto
-
-lemma eq_subst_tel_eq_free:
-    "$s1 (ts::typ list) = $s2 ts \<Longrightarrow> n \<in> free_tv ts \<Longrightarrow> s1 n = s2 n"
-  by (induct ts) (auto intro: eq_subst_te_eq_free)
-
-lemma eq_free_eq_subst_tel:
-    "(\<forall>n. n \<in> free_tv ts --> s1 n = s2 n) \<Longrightarrow> $s1 (ts::typ list) = $s2 ts"
-  by (induct ts) (auto intro: eq_free_eq_subst_te)
-
-text {*
-  \medskip Some useful lemmas.
-*}
-
-lemma codD: "v \<in> cod s \<Longrightarrow> v \<in> free_tv s"
-  by (simp add: free_tv_subst)
-
-lemma not_free_impl_id: "x \<notin> free_tv s \<Longrightarrow> s x = TVar x"
-  by (simp add: free_tv_subst dom_def)
-
-lemma free_tv_le_new_tv: "new_tv n t \<Longrightarrow> m \<in> free_tv t \<Longrightarrow> m < n"
-  by (unfold new_tv_def) fast
-
-lemma free_tv_subst_var: "free_tv (s (v::nat)) \<le> insert v (cod s)"
-  by (cases "v \<in> dom s") (auto simp add: cod_def dom_def)
-
-lemma free_tv_app_subst_te: "free_tv ($s (t::typ)) \<subseteq> cod s \<union> free_tv t"
-  by (induct t) (auto simp add: free_tv_subst_var)
-
-lemma free_tv_app_subst_tel: "free_tv ($s (ts::typ list)) \<subseteq> cod s \<union> free_tv ts"
-  apply (induct ts)
-   apply simp
-  apply (cut_tac free_tv_app_subst_te)
-  apply fastsimp
-  done
-
-lemma free_tv_comp_subst:
-    "free_tv (\<lambda>u::nat. $s1 (s2 u) :: typ) \<subseteq> free_tv s1 \<union> free_tv s2"
-  apply (unfold free_tv_subst dom_def)
-  apply (auto dest!: free_tv_subst_var [THEN subsetD] free_tv_app_subst_te [THEN subsetD]
-    simp add: cod_def dom_def simp del: bex_simps)
-  done
-
-
-subsection {* Most general unifiers *}
-
-consts
-  mgu :: "typ \<Rightarrow> typ \<Rightarrow> subst maybe"
-axioms
-  mgu_eq [simp]: "mgu t1 t2 = Ok u \<Longrightarrow> $u t1 = $u t2"
-  mgu_mg [simp]: "mgu t1 t2 = Ok u \<Longrightarrow> $s t1 = $s t2 \<Longrightarrow> \<exists>r. s = $r o u"
-  mgu_Ok: "$s t1 = $s t2 \<Longrightarrow> \<exists>u. mgu t1 t2 = Ok u"
-  mgu_free [simp]: "mgu t1 t2 = Ok u \<Longrightarrow> free_tv u \<subseteq> free_tv t1 \<union> free_tv t2"
-
-lemma mgu_new: "mgu t1 t2 = Ok u \<Longrightarrow> new_tv n t1 \<Longrightarrow> new_tv n t2 \<Longrightarrow> new_tv n u"
-  -- {* @{text mgu} does not introduce new type variables *}
-  by (unfold new_tv_def) (blast dest: mgu_free)
-
-
-section {* Mini-ML with type inference rules *}
-
-datatype
-  expr = Var nat | Abs expr | App expr expr
-
-
-text {* Type inference rules. *}
-
-inductive
-  has_type :: "typ list \<Rightarrow> expr \<Rightarrow> typ \<Rightarrow> bool"  ("((_) |-/ (_) :: (_))" [60, 0, 60] 60)
-  where
-    Var: "n < length a \<Longrightarrow> a |- Var n :: a ! n"
-  | Abs: "t1#a |- e :: t2 \<Longrightarrow> a |- Abs e :: t1 -> t2"
-  | App: "a |- e1 :: t2 -> t1 \<Longrightarrow> a |- e2 :: t2
-      \<Longrightarrow> a |- App e1 e2 :: t1"
-
-
-text {* Type assigment is closed wrt.\ substitution. *}
-
-lemma has_type_subst_closed: "a |- e :: t ==> $s a |- e :: $s t"
-proof (induct set: has_type)
-  case (Var n a)
-  then have "n < length (map ($ s) a)" by simp
-  then have "map ($ s) a |- Var n :: map ($ s) a ! n"
-    by (rule has_type.Var)
-  also have "map ($ s) a ! n = $ s (a ! n)"
-    by (rule nth_map) (rule Var)
-  also have "map ($ s) a = $ s a"
-    by (simp only: app_subst_list)
-  finally show ?case .
-next
-  case (Abs t1 a e t2)
-  then have "$ s t1 # map ($ s) a |- e :: $ s t2"
-    by (simp add: app_subst_list)
-  then have "map ($ s) a |- Abs e :: $ s t1 -> $ s t2"
-    by (rule has_type.Abs)
-  then show ?case
-    by (simp add: app_subst_list)
-next
-  case App
-  then show ?case by (simp add: has_type.App)
-qed
-
-
-section {* Correctness and completeness of the type inference algorithm W *}
-
-consts
-  "\<W>" :: "expr \<Rightarrow> typ list \<Rightarrow> nat \<Rightarrow> (subst \<times> typ \<times> nat) maybe"
-primrec
-  "\<W> (Var i) a n =
-    (if i < length a then Ok (id_subst, a ! i, n) else Fail)"
-  "\<W> (Abs e) a n =
-    ((s, t, m) := \<W> e (TVar n # a) (Suc n);
-     Ok (s, (s n) -> t, m))"
-  "\<W> (App e1 e2) a n =
-    ((s1, t1, m1) := \<W> e1 a n;
-     (s2, t2, m2) := \<W> e2 ($s1 a) m1;
-     u := mgu ($ s2 t1) (t2 -> TVar m2);
-     Ok ($u o $s2 o s1, $u (TVar m2), Suc m2))"
-
-theorem W_correct: "Ok (s, t, m) = \<W> e a n ==> $s a |- e :: t"
-proof (induct e arbitrary: a s t m n)
-  case (Var i)
-  from `Ok (s, t, m) = \<W> (Var i) a n`
-  show "$s a |- Var i :: t" by (simp add: has_type.Var split: if_splits)
-next
-  case (Abs e)
-  from `Ok (s, t, m) = \<W> (Abs e) a n`
-  obtain t' where "t = s n -> t'"
-      and "Ok (s, t', m) = \<W> e (TVar n # a) (Suc n)"
-    by (auto split: bind_splits)
-  with Abs.hyps show "$s a |- Abs e :: t"
-    by (force intro: has_type.Abs)
-next
-  case (App e1 e2)
-  from `Ok (s, t, m) = \<W> (App e1 e2) a n`
-  obtain s1 t1 n1 s2 t2 n2 u where
-          s: "s = $u o $s2 o s1"
-      and t: "t = u n2"
-      and mgu_ok: "mgu ($s2 t1) (t2 -> TVar n2) = Ok u"
-      and W1_ok: "Ok (s1, t1, n1) = \<W> e1 a n"
-      and W2_ok: "Ok (s2, t2, n2) = \<W> e2 ($s1 a) n1"
-    by (auto split: bind_splits simp: that)
-  show "$s a |- App e1 e2 :: t"
-  proof (rule has_type.App)
-    from s have s': "$u ($s2 ($s1 a)) = $s a"
-      by (simp add: subst_comp_tel o_def)
-    show "$s a |- e1 :: $u t2 -> t"
-    proof -
-      from W1_ok have "$s1 a |- e1 :: t1" by (rule App.hyps(1))
-      then have "$u ($s2 ($s1 a)) |- e1 :: $u ($s2 t1)"
-        by (intro has_type_subst_closed)
-      with s' t mgu_ok show ?thesis by simp
-    qed
-    show "$s a |- e2 :: $u t2"
-    proof -
-      from W2_ok have "$s2 ($s1 a) |- e2 :: t2" by (rule App.hyps(2))
-      then have "$u ($s2 ($s1 a)) |- e2 :: $u t2"
-        by (rule has_type_subst_closed)
-      with s' show ?thesis by simp
-    qed
-  qed
-qed
-
-
-inductive_cases has_type_casesE:
-  "s |- Var n :: t"
-  "s |- Abs e :: t"
-  "s |- App e1 e2 ::t"
-
-
-lemmas [simp] = Suc_le_lessD
-  and [simp del] = less_imp_le ex_simps all_simps
-
-lemma W_var_ge [simp]: "!!a n s t m. \<W> e a n = Ok (s, t, m) \<Longrightarrow> n \<le> m"
-  -- {* the resulting type variable is always greater or equal than the given one *}
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply clarsimp
-   txt {* case @{text "Abs e"} *}
-   apply (simp split add: split_bind)
-   apply (fast dest: Suc_leD)
-  txt {* case @{text "App e1 e2"} *}
-  apply (simp (no_asm) split add: split_bind)
-  apply (intro strip)
-  apply (rename_tac s t na sa ta nb sb)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = n in allE)
-  apply (erule_tac x = "$s a" in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = t in allE)
-  apply (erule_tac x = na in allE)
-  apply (erule_tac x = na in allE)
-  apply (simp add: eq_sym_conv)
-  done
-
-lemma W_var_geD: "Ok (s, t, m) = \<W> e a n \<Longrightarrow> n \<le> m"
-  by (simp add: eq_sym_conv)
-
-lemma new_tv_W: "!!n a s t m.
-  new_tv n a \<Longrightarrow> \<W> e a n = Ok (s, t, m) \<Longrightarrow> new_tv m s & new_tv m t"
-  -- {* resulting type variable is new *}
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply clarsimp
-    apply (force elim: list_ball_nth simp add: id_subst_def new_tv_list new_tv_subst)
-   txt {* case @{text "Abs e"} *}
-   apply (simp (no_asm) add: new_tv_subst new_tv_Suc_list split add: split_bind)
-   apply (intro strip)
-   apply (erule_tac x = "Suc n" in allE)
-   apply (erule_tac x = "TVar n # a" in allE)
-   apply (fastsimp simp add: new_tv_subst new_tv_Suc_list)
-  txt {* case @{text "App e1 e2"} *}
-  apply (simp (no_asm) split add: split_bind)
-  apply (intro strip)
-  apply (rename_tac s t na sa ta nb sb)
-  apply (erule_tac x = n in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = t in allE)
-  apply (erule_tac x = na in allE)
-  apply (erule_tac x = na in allE)
-  apply (simp add: eq_sym_conv)
-  apply (erule_tac x = "$s a" in allE)
-  apply (erule_tac x = sa in allE)
-  apply (erule_tac x = ta in allE)
-  apply (erule_tac x = nb in allE)
-  apply (simp add: o_def rotate_Ok)
-  apply (rule conjI)
-   apply (rule new_tv_subst_comp_2)
-    apply (rule new_tv_subst_comp_2)
-     apply (rule lessI [THEN less_imp_le, THEN new_tv_subst_le])
-     apply (rule_tac n = na in new_tv_subst_le)
-      apply (simp add: rotate_Ok)
-     apply (simp (no_asm_simp))
-    apply (fast dest: W_var_geD intro: new_tv_list_le new_tv_subst_tel
-      lessI [THEN less_imp_le, THEN new_tv_subst_le])
-   apply (erule sym [THEN mgu_new])
-    apply (best dest: W_var_geD intro: new_tv_subst_te new_tv_list_le new_tv_subst_tel
-      lessI [THEN less_imp_le, THEN new_tv_le] lessI [THEN less_imp_le, THEN new_tv_subst_le]
-      new_tv_le)
-   apply (tactic {* fast_tac (HOL_cs addDs [thm "W_var_geD"]
-     addIs [thm "new_tv_list_le", thm "new_tv_subst_tel", thm "new_tv_le"]
-     addss @{simpset}) 1 *})
-  apply (rule lessI [THEN new_tv_subst_var])
-  apply (erule sym [THEN mgu_new])
-    apply (bestsimp intro!: lessI [THEN less_imp_le, THEN new_tv_le] new_tv_subst_te
-      dest!: W_var_geD intro: new_tv_list_le new_tv_subst_tel
-        lessI [THEN less_imp_le, THEN new_tv_subst_le] new_tv_le)
-  apply (tactic {* fast_tac (HOL_cs addDs [thm "W_var_geD"]
-    addIs [thm "new_tv_list_le", thm "new_tv_subst_tel", thm "new_tv_le"]
-    addss @{simpset}) 1 *})
-  done
-
-lemma free_tv_W: "!!n a s t m v. \<W> e a n = Ok (s, t, m) \<Longrightarrow>
-    (v \<in> free_tv s \<or> v \<in> free_tv t) \<Longrightarrow> v < n \<Longrightarrow> v \<in> free_tv a"
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply clarsimp
-    apply (tactic {* fast_tac (HOL_cs addIs [thm "nth_mem", subsetD, thm "ftv_mem_sub_ftv_list"]) 1 *})
-   txt {* case @{text "Abs e"} *}
-   apply (simp add: free_tv_subst split add: split_bind)
-   apply (intro strip)
-   apply (rename_tac s t n1 v)
-   apply (erule_tac x = "Suc n" in allE)
-   apply (erule_tac x = "TVar n # a" in allE)
-   apply (erule_tac x = s in allE)
-   apply (erule_tac x = t in allE)
-   apply (erule_tac x = n1 in allE)
-   apply (erule_tac x = v in allE)
-   apply (force elim!: allE intro: cod_app_subst)
-  txt {* case @{text "App e1 e2"} *}
-  apply (simp (no_asm) split add: split_bind)
-  apply (intro strip)
-  apply (rename_tac s t n1 s1 t1 n2 s3 v)
-  apply (erule_tac x = n in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = t in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = v in allE)
-  txt {* second case *}
-  apply (erule_tac x = "$ s a" in allE)
-  apply (erule_tac x = s1 in allE)
-  apply (erule_tac x = t1 in allE)
-  apply (erule_tac x = n2 in allE)
-  apply (erule_tac x = v in allE)
-  apply (tactic "safe_tac (empty_cs addSIs [conjI, impI] addSEs [conjE])")
-   apply (simp add: rotate_Ok o_def)
-   apply (drule W_var_geD)
-   apply (drule W_var_geD)
-   apply (frule less_le_trans, assumption)
-   apply (fastsimp dest: free_tv_comp_subst [THEN subsetD] sym [THEN mgu_free] codD
-     free_tv_app_subst_te [THEN subsetD] free_tv_app_subst_tel [THEN subsetD] subsetD elim: UnE)
-  apply simp
-  apply (drule sym [THEN W_var_geD])
-  apply (drule sym [THEN W_var_geD])
-  apply (frule less_le_trans, assumption)
-  apply (tactic {* fast_tac (HOL_cs addDs [thm "mgu_free", thm "codD",
-    thm "free_tv_subst_var" RS subsetD,
-    thm "free_tv_app_subst_te" RS subsetD,
-    thm "free_tv_app_subst_tel" RS subsetD, @{thm less_le_trans}, subsetD]
-    addSEs [UnE] addss (@{simpset} setSolver unsafe_solver)) 1 *})
-      -- {* builtin arithmetic in simpset messes things up *}
-  done
-
-text {*
-  \medskip Completeness of @{text \<W>} wrt.\ @{text has_type}.
-*}
-
-lemma W_complete_aux: "!!s' a t' n. $s' a |- e :: t' \<Longrightarrow> new_tv n a \<Longrightarrow>
-    (\<exists>s t. (\<exists>m. \<W> e a n = Ok (s, t, m)) \<and> (\<exists>r. $s' a = $r ($s a) \<and> t' = $r t))"
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply (intro strip)
-    apply (simp (no_asm) cong add: conj_cong)
-    apply (erule has_type_casesE)
-    apply (simp add: eq_sym_conv app_subst_list)
-    apply (rule_tac x = s' in exI)
-    apply simp
-   txt {* case @{text "Abs e"} *}
-   apply (intro strip)
-   apply (erule has_type_casesE)
-   apply (erule_tac x = "\<lambda>x. if x = n then t1 else (s' x)" in allE)
-   apply (erule_tac x = "TVar n # a" in allE)
-   apply (erule_tac x = t2 in allE)
-   apply (erule_tac x = "Suc n" in allE)
-   apply (fastsimp cong add: conj_cong split add: split_bind)
-  txt {* case @{text "App e1 e2"} *}
-  apply (intro strip)
-  apply (erule has_type_casesE)
-  apply (erule_tac x = s' in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = "t2 -> t'" in allE)
-  apply (erule_tac x = n in allE)
-  apply (tactic "safe_tac HOL_cs")
-  apply (erule_tac x = r in allE)
-  apply (erule_tac x = "$s a" in allE)
-  apply (erule_tac x = t2 in allE)
-  apply (erule_tac x = m in allE)
-  apply simp
-  apply (tactic "safe_tac HOL_cs")
-   apply (tactic {* fast_tac (HOL_cs addIs [sym RS thm "W_var_geD",
-     thm "new_tv_W" RS conjunct1, thm "new_tv_list_le", thm "new_tv_subst_tel"]) 1 *})
-  apply (subgoal_tac
-    "$(\<lambda>x. if x = ma then t' else (if x \<in> free_tv t - free_tv sa then r x
-      else ra x)) ($ sa t) =
-    $(\<lambda>x. if x = ma then t' else (if x \<in> free_tv t - free_tv sa then r x
-      else ra x)) (ta -> (TVar ma))")
-   apply (rule_tac [2] t = "$(\<lambda>x. if x = ma then t'
-     else (if x \<in> (free_tv t - free_tv sa) then r x else ra x)) ($sa t)" and
-     s = "($ ra ta) -> t'" in ssubst)
-    prefer 2
-    apply (simp add: subst_comp_te)
-    apply (rule eq_free_eq_subst_te)
-    apply (intro strip)
-    apply (subgoal_tac "na \<noteq> ma")
-     prefer 2
-     apply (fast dest: new_tv_W sym [THEN W_var_geD] new_tv_not_free_tv new_tv_le)
-    apply (case_tac "na \<in> free_tv sa")
-     txt {* @{text "na \<notin> free_tv sa"} *}
-     prefer 2
-     apply (frule not_free_impl_id)
-     apply simp
-    txt {* @{text "na \<in> free_tv sa"} *}
-    apply (drule_tac ts1 = "$s a" and r = "$ r ($ s a)" in subst_comp_tel [THEN [2] trans])
-    apply (drule_tac eq_subst_tel_eq_free)
-     apply (fast intro: free_tv_W free_tv_le_new_tv dest: new_tv_W)
-    apply simp
-    apply (case_tac "na \<in> dom sa")
-     prefer 2
-     txt {* @{text "na \<noteq> dom sa"} *}
-     apply (simp add: dom_def)
-    txt {* @{text "na \<in> dom sa"} *}
-    apply (rule eq_free_eq_subst_te)
-    apply (intro strip)
-    apply (subgoal_tac "nb \<noteq> ma")
-     prefer 2
-     apply (frule new_tv_W, assumption)
-     apply (erule conjE)
-     apply (drule new_tv_subst_tel)
-      apply (fast intro: new_tv_list_le dest: sym [THEN W_var_geD])
-     apply (fastsimp dest: new_tv_W new_tv_not_free_tv simp add: cod_def free_tv_subst)
-    apply (fastsimp simp add: cod_def free_tv_subst)
-   prefer 2
-   apply (simp (no_asm))
-   apply (rule eq_free_eq_subst_te)
-   apply (intro strip)
-   apply (subgoal_tac "na \<noteq> ma")
-    prefer 2
-    apply (frule new_tv_W, assumption)
-    apply (erule conjE)
-    apply (drule sym [THEN W_var_geD])
-    apply (fast dest: new_tv_list_le new_tv_subst_tel new_tv_W new_tv_not_free_tv)
-   apply (case_tac "na \<in> free_tv t - free_tv sa")
-    prefer 2
-    txt {* case @{text "na \<notin> free_tv t - free_tv sa"} *}
-    apply simp
-    defer
-    txt {* case @{text "na \<in> free_tv t - free_tv sa"} *}
-    apply simp
-    apply (drule_tac ts1 = "$s a" and r = "$ r ($ s a)" in subst_comp_tel [THEN [2] trans])
-    apply (drule eq_subst_tel_eq_free)
-     apply (fast intro: free_tv_W free_tv_le_new_tv dest: new_tv_W)
-    apply (simp add: free_tv_subst dom_def)
-   prefer 2 apply fast
-  apply (simp (no_asm_simp) split add: split_bind)
-  apply (tactic "safe_tac HOL_cs")
-   apply (drule mgu_Ok)
-   apply fastsimp
-  apply (drule mgu_mg, assumption)
-  apply (erule exE)
-  apply (rule_tac x = rb in exI)
-  apply (rule conjI)
-   prefer 2
-   apply (drule_tac x = ma in fun_cong)
-   apply (simp add: eq_sym_conv)
-  apply (simp (no_asm) add: o_def subst_comp_tel [symmetric])
-  apply (rule subst_comp_tel [symmetric, THEN [2] trans])
-  apply (simp add: o_def eq_sym_conv)
-  apply (rule eq_free_eq_subst_tel)
-  apply (tactic "safe_tac HOL_cs")
-  apply (subgoal_tac "ma \<noteq> na")
-   prefer 2
-   apply (frule new_tv_W, assumption)
-   apply (erule conjE)
-   apply (drule new_tv_subst_tel)
-    apply (fast intro: new_tv_list_le dest: sym [THEN W_var_geD])
-   apply (frule_tac n = m in new_tv_W, assumption)
-   apply (erule conjE)
-   apply (drule free_tv_app_subst_tel [THEN subsetD])
-   apply (auto dest: W_var_geD [OF sym] new_tv_list_le
-     codD new_tv_not_free_tv)
-  apply (case_tac "na \<in> free_tv t - free_tv sa")
-   prefer 2
-   txt {* case @{text "na \<notin> free_tv t - free_tv sa"} *}
-   apply simp
-   defer
-   txt {* case @{text "na \<in> free_tv t - free_tv sa"} *}
-   apply simp
-   apply (drule free_tv_app_subst_tel [THEN subsetD])
-   apply (fastsimp dest: codD subst_comp_tel [THEN [2] trans]
-     eq_subst_tel_eq_free simp add: free_tv_subst dom_def)
-  done
-
-lemma W_complete: "[] |- e :: t' ==>
-    \<exists>s t. (\<exists>m. \<W> e [] n = Ok (s, t, m)) \<and> (\<exists>r. t' = $r t)"
-  apply (cut_tac a = "[]" and s' = id_subst and e = e and t' = t' in W_complete_aux)
-    apply simp_all
-  done
-
-
-section {* Equivalence of W and I *}
-
-text {*
-  Recursive definition of type inference algorithm @{text \<I>} for
-  Mini-ML.
-*}
-
-consts
-  "\<I>" :: "expr \<Rightarrow> typ list \<Rightarrow> nat \<Rightarrow> subst \<Rightarrow> (subst \<times> typ \<times> nat) maybe"
-primrec
-  "\<I> (Var i) a n s = (if i < length a then Ok (s, a ! i, n) else Fail)"
-  "\<I> (Abs e) a n s = ((s, t, m) := \<I> e (TVar n # a) (Suc n) s;
-    Ok (s, TVar n -> t, m))"
-  "\<I> (App e1 e2) a n s =
-    ((s1, t1, m1) := \<I> e1 a n s;
-    (s2, t2, m2) := \<I> e2 a m1 s1;
-    u := mgu ($s2 t1) ($s2 t2 -> TVar m2);
-    Ok($u o s2, TVar m2, Suc m2))"
-
-text {* \medskip Correctness. *}
-
-lemma I_correct_wrt_W: "!!a m s s' t n.
-    new_tv m a \<and> new_tv m s \<Longrightarrow> \<I> e a m s = Ok (s', t, n) \<Longrightarrow>
-    \<exists>r. \<W> e ($s a) m = Ok (r, $s' t, n) \<and> s' = ($r o s)"
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply (simp add: app_subst_list split: split_if)
-   txt {* case @{text "Abs e"} *}
-   apply (tactic {* asm_full_simp_tac
-     (@{simpset} setloop (split_inside_tac [thm "split_bind"])) 1 *})
-   apply (intro strip)
-   apply (rule conjI)
-    apply (intro strip)
-    apply (erule allE)+
-    apply (erule impE)
-     prefer 2 apply (fastsimp simp add: new_tv_subst)
-    apply (tactic {* fast_tac (HOL_cs addIs [thm "new_tv_Suc_list" RS mp,
-      thm "new_tv_subst_le", @{thm less_imp_le}, @{thm lessI}]) 1 *})
-   apply (intro strip)
-   apply (erule allE)+
-   apply (erule impE)
-    prefer 2 apply (fastsimp simp add: new_tv_subst)
-   apply (tactic {* fast_tac (HOL_cs addIs [thm "new_tv_Suc_list" RS mp,
-     thm "new_tv_subst_le", @{thm less_imp_le}, @{thm lessI}]) 1 *})
-  txt {* case @{text "App e1 e2"} *}
-  apply (tactic {* simp_tac (@{simpset} setloop (split_inside_tac [thm "split_bind"])) 1 *})
-  apply (intro strip)
-  apply (rename_tac s1' t1 n1 s2' t2 n2 sa)
-  apply (rule conjI)
-   apply fastsimp
-  apply (intro strip)
-  apply (rename_tac s1 t1' n1')
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = m in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = s1' in allE)
-  apply (erule_tac x = t1 in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = s1' in allE)
-  apply (erule_tac x = s2' in allE)
-  apply (erule_tac x = t2 in allE)
-  apply (erule_tac x = n2 in allE)
-  apply (rule conjI)
-   apply (intro strip)
-   apply (rule notI)
-   apply simp
-   apply (erule impE)
-    apply (frule new_tv_subst_tel, assumption)
-    apply (drule_tac a = "$s a" in new_tv_W, assumption)
-    apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-   apply (fastsimp simp add: subst_comp_tel)
-  apply (intro strip)
-  apply (rename_tac s2 t2' n2')
-  apply (rule conjI)
-   apply (intro strip)
-   apply (rule notI)
-   apply simp
-   apply (erule impE)
-   apply (frule new_tv_subst_tel, assumption)
-   apply (drule_tac a = "$s a" in new_tv_W, assumption)
-    apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-   apply (fastsimp simp add: subst_comp_tel subst_comp_te)
-  apply (intro strip)
-  apply (erule (1) notE impE)
-  apply (erule (1) notE impE)
-  apply (erule exE)
-  apply (erule conjE)
-  apply (erule impE)
-   apply (frule new_tv_subst_tel, assumption)
-   apply (drule_tac a = "$s a" in new_tv_W, assumption)
-   apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-  apply (erule (1) notE impE)
-  apply (erule exE conjE)+
-  apply (simp (asm_lr) add: subst_comp_tel subst_comp_te o_def, (erule conjE)+, hypsubst)+
-  apply (subgoal_tac "new_tv n2 s \<and> new_tv n2 r \<and> new_tv n2 ra")
-   apply (simp add: new_tv_subst)
-  apply (frule new_tv_subst_tel, assumption)
-  apply (drule_tac a = "$s a" in new_tv_W, assumption)
-  apply (tactic "safe_tac HOL_cs")
-    apply (bestsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-   apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-  apply (drule_tac e = e1 in sym [THEN W_var_geD])
-  apply (drule new_tv_subst_tel, assumption)
-  apply (drule_tac ts = "$s a" in new_tv_list_le, assumption)
-  apply (drule new_tv_subst_tel, assumption)
-  apply (bestsimp dest: new_tv_W simp add: subst_comp_tel)
-  done
-
-lemma I_complete_wrt_W: "!!a m s.
-    new_tv m a \<and> new_tv m s \<Longrightarrow> \<I> e a m s = Fail \<Longrightarrow> \<W> e ($s a) m = Fail"
-  apply (atomize (full))
-  apply (induct e)
-    apply (simp add: app_subst_list)
-   apply (simp (no_asm))
-   apply (intro strip)
-   apply (subgoal_tac "TVar m # $s a = $s (TVar m # a)")
-    apply (tactic {* asm_simp_tac (HOL_ss addsimps
-      [thm "new_tv_Suc_list", @{thm lessI} RS @{thm less_imp_le} RS thm "new_tv_subst_le"]) 1 *})
-    apply (erule conjE)
-    apply (drule new_tv_not_free_tv [THEN not_free_impl_id])
-    apply (simp (no_asm_simp))
-  apply (simp (no_asm_simp))
-  apply (intro strip)
-  apply (erule exE)+
-  apply (erule conjE)+
-  apply (drule I_correct_wrt_W [COMP swap_prems_rl])
-   apply fast
-  apply (erule exE)
-  apply (erule conjE)
-  apply hypsubst
-  apply (simp (no_asm_simp))
-  apply (erule disjE)
-   apply (rule disjI1)
-   apply (simp (no_asm_use) add: o_def subst_comp_tel)
-   apply (erule allE, erule allE, erule allE, erule impE, erule_tac [2] impE,
-     erule_tac [2] asm_rl, erule_tac [2] asm_rl)
-   apply (rule conjI)
-    apply (fast intro: W_var_ge [THEN new_tv_list_le])
-   apply (rule new_tv_subst_comp_2)
-    apply (fast intro: W_var_ge [THEN new_tv_subst_le])
-   apply (fast intro!: new_tv_subst_tel intro: new_tv_W [THEN conjunct1])
-  apply (rule disjI2)
-  apply (erule exE)+
-  apply (erule conjE)
-  apply (drule I_correct_wrt_W [COMP swap_prems_rl])
-   apply (rule conjI)
-   apply (fast intro: W_var_ge [THEN new_tv_list_le])
-   apply (rule new_tv_subst_comp_1)
-   apply (fast intro: W_var_ge [THEN new_tv_subst_le])
-   apply (fast intro!: new_tv_subst_tel intro: new_tv_W [THEN conjunct1])
-  apply (erule exE)
-  apply (erule conjE)
-  apply hypsubst
-  apply (simp add: o_def subst_comp_te [symmetric] subst_comp_tel [symmetric])
-  done
-
-end
--- a/src/HOL/W0/document/root.tex	Tue Feb 23 17:33:03 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-
-\documentclass[11pt,a4paper]{article}
-\usepackage{isabelle,isabellesym}
-\usepackage{pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-
-\newcommand{\isasymbind}{\textsf{bind}}
-
-\begin{document}
-
-\title{Type inference for let-free MiniML}
-\author{Dieter Nazareth, Tobias Nipkow, Thomas Stauner, Markus Wenzel}
-\maketitle
-
-\tableofcontents
-
-\parindent 0pt\parskip 0.5ex
-\input{session}
-
-%\bibliographystyle{abbrv}
-%\bibliography{root}
-
-\end{document}
--- a/src/HOL/ex/Codegenerator_Candidates.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/Codegenerator_Candidates.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -8,6 +8,8 @@
   Complex_Main
   AssocList
   Binomial
+  "~~/src/HOL/Decision_Procs/Commutative_Ring_Complete"
+  Dlist
   Fset
   Enum
   List_Prefix
@@ -17,12 +19,11 @@
   Permutation
   "~~/src/HOL/Number_Theory/Primes"
   Product_ord
+  "~~/src/HOL/ex/Records"
   SetsAndFunctions
   Tree
   While_Combinator
   Word
-  "~~/src/HOL/Decision_Procs/Commutative_Ring_Complete"
-  "~~/src/HOL/ex/Records"
 begin
 
 inductive sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
--- a/src/HOL/ex/PER.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/PER.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ex/PER.thy
-    ID:         $Id$
     Author:     Oscar Slotosch and Markus Wenzel, TU Muenchen
 *)
 
@@ -30,12 +29,10 @@
   but not necessarily reflexive.
 *}
 
-consts
-  eqv :: "'a => 'a => bool"    (infixl "\<sim>" 50)
-
-axclass partial_equiv < type
-  partial_equiv_sym [elim?]: "x \<sim> y ==> y \<sim> x"
-  partial_equiv_trans [trans]: "x \<sim> y ==> y \<sim> z ==> x \<sim> z"
+class partial_equiv =
+  fixes eqv :: "'a => 'a => bool"    (infixl "\<sim>" 50)
+  assumes partial_equiv_sym [elim?]: "x \<sim> y ==> y \<sim> x"
+  assumes partial_equiv_trans [trans]: "x \<sim> y ==> y \<sim> z ==> x \<sim> z"
 
 text {*
   \medskip The domain of a partial equivalence relation is the set of
@@ -70,7 +67,10 @@
   structural one corresponding to the congruence property.
 *}
 
-defs (overloaded)
+instantiation "fun" :: (partial_equiv, partial_equiv) partial_equiv
+begin
+
+definition
   eqv_fun_def: "f \<sim> g == \<forall>x \<in> domain. \<forall>y \<in> domain. x \<sim> y --> f x \<sim> g y"
 
 lemma partial_equiv_funI [intro?]:
@@ -86,8 +86,7 @@
   spaces (in \emph{both} argument positions).
 *}
 
-instance "fun" :: (partial_equiv, partial_equiv) partial_equiv
-proof
+instance proof
   fix f g h :: "'a::partial_equiv => 'b::partial_equiv"
   assume fg: "f \<sim> g"
   show "g \<sim> f"
@@ -110,6 +109,8 @@
   qed
 qed
 
+end
+
 
 subsection {* Total equivalence *}
 
@@ -120,8 +121,8 @@
   symmetric.
 *}
 
-axclass equiv < partial_equiv
-  eqv_refl [intro]: "x \<sim> x"
+class equiv =
+  assumes eqv_refl [intro]: "x \<sim> x"
 
 text {*
   On total equivalences all elements are reflexive, and congruence
@@ -150,7 +151,7 @@
   \emph{equivalence classes} over elements of the base type @{typ 'a}.
 *}
 
-typedef 'a quot = "{{x. a \<sim> x}| a::'a. True}"
+typedef 'a quot = "{{x. a \<sim> x}| a::'a::partial_equiv. True}"
   by blast
 
 lemma quotI [intro]: "{x. a \<sim> x} \<in> quot"
--- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -2,6 +2,21 @@
 imports "../Predicate_Compile"
 begin
 
+section {* Common constants *}
+
+declare HOL.if_bool_eq_disj[code_pred_inline]
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
+
+section {* Pairs *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name split}] *}
+
+section {* Bounded quantifiers *}
+
+declare Ball_def[code_pred_inline]
+declare Bex_def[code_pred_inline]
+
 section {* Set operations *}
 
 declare Collect_def[code_pred_inline]
@@ -9,13 +24,37 @@
 
 declare eq_reflection[OF empty_def, code_pred_inline]
 declare insert_code[code_pred_def]
+
+declare subset_iff[code_pred_inline]
+
+declare Int_def[code_pred_inline]
 declare eq_reflection[OF Un_def, code_pred_inline]
 declare eq_reflection[OF UNION_def, code_pred_inline]
 
+lemma Diff[code_pred_inline]:
+  "(A - B) = (%x. A x \<and> \<not> B x)"
+by (auto simp add: mem_def)
 
+lemma set_equality[code_pred_inline]:
+  "(A = B) = ((\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x))"
+by (fastsimp simp add: mem_def)
+
+section {* Setup for Numerals *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
 
 section {* Alternative list definitions *}
+
+text {* size simps are not yet added to the Spec_Rules interface. So they are just added manually here! *}
  
+lemma [code_pred_def]:
+  "length [] = 0"
+  "length (x # xs) = Suc (length xs)"
+by auto
+
 subsection {* Alternative rules for set *}
 
 lemma set_ConsI1 [code_pred_intro]:
--- a/src/HOL/ex/Predicate_Compile_Quickcheck.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/Predicate_Compile_Quickcheck.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -3,11 +3,14 @@
 header {* A Prototype of Quickcheck based on the Predicate Compiler *}
 
 theory Predicate_Compile_Quickcheck
-imports "../Predicate_Compile"
+imports "../Predicate_Compile" Quickcheck Predicate_Compile_Alternative_Defs
 uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
 begin
 
-setup {* Quickcheck.add_generator ("predicate_compile", Predicate_Compile_Quickcheck.quickcheck) *}
+setup {* Quickcheck.add_generator ("predicate_compile_wo_ff", Predicate_Compile_Quickcheck.quickcheck_compile_term false true) *}
+setup {* Quickcheck.add_generator ("predicate_compile_ff_fs", Predicate_Compile_Quickcheck.quickcheck_compile_term true true) *}
+setup {* Quickcheck.add_generator ("predicate_compile_ff_nofs", Predicate_Compile_Quickcheck.quickcheck_compile_term true false) *}
+
 (*
 datatype alphabet = a | b
 
--- a/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1,45 +1,39 @@
 theory Predicate_Compile_Quickcheck_ex
 imports Predicate_Compile_Quickcheck
-  Predicate_Compile_Alternative_Defs
 begin
 
-ML {* Predicate_Compile_Alternative_Defs.get *}
-
 section {* Sets *}
-(*
+
 lemma "x \<in> {(1::nat)} ==> False"
-quickcheck[generator=predicate_compile, iterations=10]
+quickcheck[generator=predicate_compile_wo_ff, iterations=10]
 oops
-*)
-(* TODO: some error with doubled negation *)
-(*
+
 lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x \<noteq> Suc 0"
-quickcheck[generator=predicate_compile]
+quickcheck[generator=predicate_compile_wo_ff]
 oops
-*)
-(*
+
 lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x = Suc 0"
-quickcheck[generator=predicate_compile]
+quickcheck[generator=predicate_compile_wo_ff]
 oops
-*) 
+ 
 lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x <= Suc 0"
-(*quickcheck[generator=predicate_compile]*)
+quickcheck[generator=predicate_compile_wo_ff]
 oops
 
 section {* Numerals *}
-(*
+
 lemma
   "x \<in> {1, 2, (3::nat)} ==> x = 1 \<or> x = 2"
-quickcheck[generator=predicate_compile]
+quickcheck[generator=predicate_compile_wo_ff]
 oops
-*)
+
 lemma "x \<in> {1, 2, (3::nat)} ==> x < 3"
-(*quickcheck[generator=predicate_compile]*)
+quickcheck[generator=predicate_compile_wo_ff]
 oops
 
 lemma
   "x \<in> {1, 2} \<union> {3, 4} ==> x = (1::nat) \<or> x = (2::nat)"
-(*quickcheck[generator=predicate_compile]*)
+quickcheck[generator=predicate_compile_wo_ff]
 oops
 
 section {* Context Free Grammar *}
@@ -53,33 +47,15 @@
 | "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
 | "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
 | "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
-(*
-code_pred [random_dseq inductify] "S\<^isub>1p" .
-*)
-(*thm B\<^isub>1p.random_dseq_equation*)
-(*
-values [random_dseq 2, 2, 4] 10 "{x. S\<^isub>1p x}"
-values [random_dseq 1, 1, 5] 20 "{x. S\<^isub>1p x}"
 
-ML {* set ML_Context.trace *}
-*)
-ML {* set Toplevel.debug *}
-(*
-quickcheck[generator = predicate_compile, size = 10, iterations = 1]
-oops
-*)
-ML {* Spec_Rules.get *}
-ML {* Item_Net.retrieve *}
-local_setup {* Local_Theory.checkpoint *}
-ML {* Predicate_Compile_Data.get_specification @{theory} @{term "append"} *}
 lemma
-  "w \<in> S\<^isub>1p \<Longrightarrow> w = []"
-quickcheck[generator = predicate_compile, iterations=1]
+  "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
+quickcheck[generator = predicate_compile_ff_nofs, iterations=1]
 oops
 
 theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=15]
+"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=predicate_compile_ff_nofs, size=15]
 oops
 
 
@@ -90,7 +66,7 @@
 | "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
 | "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
 | "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
-
+(*
 code_pred [random_dseq inductify] S\<^isub>2 .
 thm S\<^isub>2.random_dseq_equation
 thm A\<^isub>2.random_dseq_equation
@@ -118,10 +94,10 @@
 "w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
 quickcheck[generator=predicate_compile, size = 10, iterations = 1]
 oops
-
+*)
 theorem S\<^isub>2_sound:
 "w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=15, iterations=1]
+quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10]
 oops
 
 inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
@@ -141,17 +117,17 @@
 
 lemma S\<^isub>3_sound:
 "w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=10, iterations=10]
+quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10]
 oops
 
 lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
-quickcheck[size=10, generator = predicate_compile]
+quickcheck[size=10, generator = predicate_compile_ff_fs]
 oops
 
 theorem S\<^isub>3_complete:
 "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
 (*quickcheck[generator=SML]*)
-quickcheck[generator=predicate_compile, size=10, iterations=100]
+quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100]
 oops
 
 
@@ -166,20 +142,23 @@
 
 theorem S\<^isub>4_sound:
 "w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator = predicate_compile, size=5, iterations=1]
+quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
 oops
 
 theorem S\<^isub>4_complete:
 "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-quickcheck[generator = predicate_compile, size=5, iterations=1]
+quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
 oops
 
-hide const b
+hide const a b
 
 subsection {* Lexicographic order *}
+(* TODO *)
+(*
 lemma
   "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
-
+oops
+*)
 subsection {* IMP *}
 
 types
@@ -208,7 +187,7 @@
 
 lemma
   "exec c s s' ==> exec (Seq c c) s s'"
-quickcheck[generator = predicate_compile, size=3, iterations=1]
+(*quickcheck[generator = predicate_compile_wo_ff, size=2, iterations=10]*)
 oops
 
 subsection {* Lambda *}
@@ -263,28 +242,9 @@
 
 lemma
   "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
-quickcheck[generator = predicate_compile, size = 7, iterations = 10]
+quickcheck[generator = predicate_compile_ff_fs, size = 7, iterations = 10]
 oops
 
-(*
-code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
-thm typing.equation
-
-code_pred (modes: i => i => bool,  i => o => bool as reduce') beta .
-thm beta.equation
-
-values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
-
-definition "reduce t = Predicate.the (reduce' t)"
-
-value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
-
-code_pred [random] typing .
-code_pred [random_dseq] typing .
-
-(*values [random] 1 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
-*)*)
-
 subsection {* JAD *}
 
 definition matrix :: "('a :: semiring_0) list list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
@@ -300,9 +260,17 @@
 lemma [code_pred_intro]:
   "matrix [] 0 m"
   "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
-sorry
+proof -
+  show "matrix [] 0 m" unfolding matrix_def by auto
+next
+  show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
+    unfolding matrix_def by auto
+qed
 
-code_pred [random_dseq inductify] matrix sorry
+code_pred [random_dseq inductify] matrix
+  apply (cases x)
+  unfolding matrix_def apply fastsimp
+  apply fastsimp done
 
 
 values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}"
@@ -344,10 +312,10 @@
 
 definition
   "length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)"
-
+(*
 definition
   "transpose M = [map (\<lambda> xs. xs ! i) (takeWhile (\<lambda> xs. i < length xs) M). i \<leftarrow> [0 ..< length (M ! 0)]]"
-
+*)
 definition
   "inflate upds = foldr (\<lambda> (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)"
 
@@ -356,15 +324,14 @@
 
 definition
   "jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (\<lambda> (i, x). v ! i * x)))"
-ML {* ML_Context.trace := false *}
 
 lemma "matrix (M::int list list) rs cs \<Longrightarrow> False"
-quickcheck[generator = predicate_compile, size = 6]
+quickcheck[generator = predicate_compile_ff_nofs, size = 6]
 oops
 
 lemma
   "\<lbrakk> matrix M rs cs ; length v = cs \<rbrakk> \<Longrightarrow> jad_mv v (jad M) = mv M v"
-(*quickcheck[generator = predicate_compile]*)
+quickcheck[generator = predicate_compile_wo_ff]
 oops
 
 end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -252,10 +252,12 @@
   "one_or_two = {Suc 0, (Suc (Suc 0))}"
 
 code_pred [inductify] one_or_two .
+
 code_pred [dseq] one_or_two .
-(*code_pred [random_dseq] one_or_two .*)
+code_pred [random_dseq] one_or_two .
+thm one_or_two.dseq_equation
 values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
-(*values [random_dseq 1,1,2] "{x. one_or_two x}"*)
+values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
 
 inductive one_or_two' :: "nat => bool"
 where
@@ -269,12 +271,12 @@
 
 definition one_or_two'':
   "one_or_two'' == {1, (2::nat)}"
-ML {* prop_of @{thm one_or_two''} *}
-(*code_pred [inductify] one_or_two'' .
+
+code_pred [inductify] one_or_two'' .
 thm one_or_two''.equation
 
 values "{x. one_or_two'' x}"
-*)
+
 subsection {* even predicate *}
 
 inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
@@ -779,6 +781,25 @@
 thm divmod_rel.equation
 value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
 
+subsection {* Transforming predicate logic into logic programs *}
+
+subsection {* Transforming functions into logic programs *}
+definition
+  "case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"
+
+code_pred [inductify] case_f .
+thm case_fP.equation
+thm case_fP.intros
+
+fun fold_map_idx where
+  "fold_map_idx f i y [] = (y, [])"
+| "fold_map_idx f i y (x # xs) =
+ (let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs
+ in (y'', x' # xs'))"
+
+text {* mode analysis explores thousand modes - this is infeasible at the moment... *}
+(*code_pred [inductify, show_steps] fold_map_idx .*)
+
 subsection {* Minimum *}
 
 definition Min
@@ -883,9 +904,16 @@
 
 
 values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
-
-
-code_pred [inductify] lenlex .
+thm lenlex_conv
+thm lex_conv
+declare list.size(3,4)[code_pred_def]
+(*code_pred [inductify, show_steps, show_intermediate_results] length .*)
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name Orderings.top_class.top}] *}
+code_pred [inductify] lex .
+thm lex.equation
+thm lex_def
+declare lenlex_conv[code_pred_def]
+code_pred [inductify, show_steps, show_intermediate_results] lenlex .
 thm lenlex.equation
 
 code_pred [random_dseq inductify] lenlex .
@@ -893,10 +921,10 @@
 
 values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
 thm lists.intros
-(*
+
 code_pred [inductify] lists .
-*)
-(*thm lists.equation*)
+
+thm lists.equation
 
 subsection {* AVL Tree *}
 
@@ -974,13 +1002,17 @@
   (o * o => bool) => i => bool,
   (i * o => bool) => i => bool) [inductify] Domain .
 thm Domain.equation
+thm Range_def
+
 code_pred (modes:
   (o * o => bool) => o => bool,
   (o * o => bool) => i => bool,
   (o * i => bool) => i => bool) [inductify] Range .
 thm Range.equation
+
 code_pred [inductify] Field .
 thm Field.equation
+
 (*thm refl_on_def
 code_pred [inductify] refl_on .
 thm refl_on.equation*)
@@ -992,9 +1024,10 @@
 thm trans.equation
 code_pred [inductify] single_valued .
 thm single_valued.equation
-code_pred [inductify] inv_image .
+thm inv_image_def
+(*code_pred [inductify] inv_image .
 thm inv_image.equation
-
+*)
 subsection {* Inverting list functions *}
 
 (*code_pred [inductify] length .
--- a/src/HOL/ex/Refute_Examples.thy	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/HOL/ex/Refute_Examples.thy	Tue Feb 23 17:55:00 2010 +0100
@@ -1417,29 +1417,20 @@
 
 (*****************************************************************************)
 
-subsubsection {* Axiomatic type classes and overloading *}
+subsubsection {* Type classes and overloading *}
 
 text {* A type class without axioms: *}
 
-axclass classA
+class classA
 
 lemma "P (x::'a::classA)"
   refute
 oops
 
-text {* The axiom of this type class does not contain any type variables: *}
-
-axclass classB
-  classB_ax: "P | ~ P"
-
-lemma "P (x::'a::classB)"
-  refute
-oops
-
 text {* An axiom with a type variable (denoting types which have at least two elements): *}
 
-axclass classC < type
-  classC_ax: "\<exists>x y. x \<noteq> y"
+class classC =
+  assumes classC_ax: "\<exists>x y. x \<noteq> y"
 
 lemma "P (x::'a::classC)"
   refute
@@ -1451,11 +1442,9 @@
 
 text {* A type class for which a constant is defined: *}
 
-consts
-  classD_const :: "'a \<Rightarrow> 'a"
-
-axclass classD < type
-  classD_ax: "classD_const (classD_const x) = classD_const x"
+class classD =
+  fixes classD_const :: "'a \<Rightarrow> 'a"
+  assumes classD_ax: "classD_const (classD_const x) = classD_const x"
 
 lemma "P (x::'a::classD)"
   refute
@@ -1463,16 +1452,12 @@
 
 text {* A type class with multiple superclasses: *}
 
-axclass classE < classC, classD
+class classE = classC + classD
 
 lemma "P (x::'a::classE)"
   refute
 oops
 
-lemma "P (x::'a::{classB, classE})"
-  refute
-oops
-
 text {* OFCLASS: *}
 
 lemma "OFCLASS('a::type, type_class)"
@@ -1485,12 +1470,6 @@
   apply intro_classes
 done
 
-lemma "OFCLASS('a, classB_class)"
-  refute  -- {* no countermodel exists *}
-  apply intro_classes
-  apply simp
-done
-
 lemma "OFCLASS('a::type, classC_class)"
   refute
 oops
--- a/src/Pure/Isar/class_target.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/Pure/Isar/class_target.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -56,11 +56,6 @@
   (*tactics*)
   val intro_classes_tac: thm list -> tactic
   val default_intro_tac: Proof.context -> thm list -> tactic
-
-  (*old axclass layer*)
-  val axclass_cmd: binding * xstring list
-    -> (Attrib.binding * string list) list
-    -> theory -> class * theory
 end;
 
 structure Class_Target : CLASS_TARGET =
@@ -629,24 +624,5 @@
   Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac))
     "apply some intro/elim rule"));
 
-
-(** old axclass command **)
-
-fun axclass_cmd (class, raw_superclasses) raw_specs thy =
-  let
-    val _ = Output.legacy_feature "command \"axclass\" deprecated; use \"class\" instead.";
-    val ctxt = ProofContext.init thy;
-    val superclasses = map (Sign.read_class thy) raw_superclasses;
-    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
-      raw_specs;
-    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
-          raw_specs)
-      |> snd
-      |> (map o map) fst;
-  in
-    AxClass.define_class (class, superclasses) []
-      (name_atts ~~ axiomss) thy
-  end;
-
 end;
 
--- a/src/Pure/Isar/code.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/Pure/Isar/code.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -49,10 +49,13 @@
   val add_signature_cmd: string * string -> theory -> theory
   val add_datatype: (string * typ) list -> theory -> theory
   val add_datatype_cmd: string list -> theory -> theory
+  val datatype_interpretation:
+    (string * ((string * sort) list * (string * typ list) list)
+      -> theory -> theory) -> theory -> theory
   val add_abstype: string * typ -> string * typ -> theory -> Proof.state
   val add_abstype_cmd: string -> string -> theory -> Proof.state
-  val type_interpretation:
-    (string * ((string * sort) list * (string * typ list) list)
+  val abstype_interpretation:
+    (string * ((string * sort) list * ((string * typ) * (string * thm)))
       -> theory -> theory) -> theory -> theory
   val add_eqn: thm -> theory -> theory
   val add_nbe_eqn: thm -> theory -> theory
@@ -63,8 +66,8 @@
   val del_eqns: string -> theory -> theory
   val add_case: thm -> theory -> theory
   val add_undefined: string -> theory -> theory
-  val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
-  val get_datatype_of_constr_or_abstr: theory -> string -> (string * bool) option
+  val get_type: theory -> string -> ((string * sort) list * (string * typ list) list)
+  val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option
   val is_constr: theory -> string -> bool
   val is_abstr: theory -> string -> bool
   val get_cert: theory -> ((thm * bool) list -> (thm * bool) list) -> string -> cert
@@ -163,21 +166,21 @@
   signatures: int Symtab.table * typ Symtab.table,
   functions: ((bool * fun_spec) * (serial * fun_spec) list) Symtab.table
     (*with explicit history*),
-  datatypes: ((serial * ((string * sort) list * typ_spec)) list) Symtab.table
+  types: ((serial * ((string * sort) list * typ_spec)) list) Symtab.table
     (*with explicit history*),
   cases: (int * (int * string list)) Symtab.table * unit Symtab.table
 };
 
-fun make_spec (history_concluded, ((signatures, functions), (datatypes, cases))) =
+fun make_spec (history_concluded, ((signatures, functions), (types, cases))) =
   Spec { history_concluded = history_concluded,
-    signatures = signatures, functions = functions, datatypes = datatypes, cases = cases };
+    signatures = signatures, functions = functions, types = types, cases = cases };
 fun map_spec f (Spec { history_concluded = history_concluded, signatures = signatures,
-  functions = functions, datatypes = datatypes, cases = cases }) =
-  make_spec (f (history_concluded, ((signatures, functions), (datatypes, cases))));
+  functions = functions, types = types, cases = cases }) =
+  make_spec (f (history_concluded, ((signatures, functions), (types, cases))));
 fun merge_spec (Spec { history_concluded = _, signatures = (tycos1, sigs1), functions = functions1,
-    datatypes = datatypes1, cases = (cases1, undefs1) },
+    types = types1, cases = (cases1, undefs1) },
   Spec { history_concluded = _, signatures = (tycos2, sigs2), functions = functions2,
-    datatypes = datatypes2, cases = (cases2, undefs2) }) =
+    types = types2, cases = (cases2, undefs2) }) =
   let
     val signatures = (Symtab.merge (op =) (tycos1, tycos2),
       Symtab.merge typ_equiv (sigs1, sigs2));
@@ -190,15 +193,15 @@
           then raw_history else filtered_history;
       in ((false, (snd o hd) history), history) end;
     val functions = Symtab.join (K merge_functions) (functions1, functions2);
-    val datatypes = Symtab.join (K (AList.merge (op =) (K true))) (datatypes1, datatypes2);
+    val types = Symtab.join (K (AList.merge (op =) (K true))) (types1, types2);
     val cases = (Symtab.merge (K true) (cases1, cases2),
       Symtab.merge (K true) (undefs1, undefs2));
-  in make_spec (false, ((signatures, functions), (datatypes, cases))) end;
+  in make_spec (false, ((signatures, functions), (types, cases))) end;
 
 fun history_concluded (Spec { history_concluded, ... }) = history_concluded;
 fun the_signatures (Spec { signatures, ... }) = signatures;
 fun the_functions (Spec { functions, ... }) = functions;
-fun the_datatypes (Spec { datatypes, ... }) = datatypes;
+fun the_types (Spec { types, ... }) = types;
 fun the_cases (Spec { cases, ... }) = cases;
 val map_history_concluded = map_spec o apfst;
 val map_signatures = map_spec o apsnd o apfst o apfst;
@@ -423,11 +426,11 @@
       $ Free ("x", ty_abs)), Free ("x", ty_abs));
   in (tyco, (vs ~~ sorts, ((fst abs_ty, ty), (rep, cert)))) end;    
 
-fun get_datatype_entry thy tyco = case these (Symtab.lookup ((the_datatypes o the_exec) thy) tyco)
+fun get_type_entry thy tyco = case these (Symtab.lookup ((the_types o the_exec) thy) tyco)
  of (_, entry) :: _ => SOME entry
   | _ => NONE;
 
-fun get_datatype_spec thy tyco = case get_datatype_entry thy tyco
+fun get_type_spec thy tyco = case get_type_entry thy tyco
  of SOME (vs, spec) => apfst (pair vs) (constructors_of spec)
   | NONE => arity_number thy tyco
       |> Name.invents Name.context Name.aT
@@ -435,23 +438,23 @@
       |> rpair []
       |> rpair false;
 
-fun get_abstype_spec thy tyco = case get_datatype_entry thy tyco
+fun get_abstype_spec thy tyco = case get_type_entry thy tyco
  of SOME (vs, Abstractor spec) => (vs, spec)
   | NONE => error ("Not an abstract type: " ^ tyco);
  
-fun get_datatype thy = fst o get_datatype_spec thy;
+fun get_type thy = fst o get_type_spec thy;
 
-fun get_datatype_of_constr_or_abstr thy c =
+fun get_type_of_constr_or_abstr thy c =
   case (snd o strip_type o const_typ thy) c
-   of Type (tyco, _) => let val ((vs, cos), abstract) = get_datatype_spec thy tyco
+   of Type (tyco, _) => let val ((vs, cos), abstract) = get_type_spec thy tyco
         in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end
     | _ => NONE;
 
-fun is_constr thy c = case get_datatype_of_constr_or_abstr thy c
+fun is_constr thy c = case get_type_of_constr_or_abstr thy c
  of SOME (_, false) => true
    | _ => false;
 
-fun is_abstr thy c = case get_datatype_of_constr_or_abstr thy c
+fun is_abstr thy c = case get_type_of_constr_or_abstr thy c
  of SOME (_, true) => true
    | _ => false;
 
@@ -952,7 +955,7 @@
       |> Symtab.dest
       |> (map o apsnd) (snd o fst)
       |> sort (string_ord o pairself fst);
-    val datatypes = the_datatypes exec
+    val datatypes = the_types exec
       |> Symtab.dest
       |> map (fn (tyco, (_, (vs, spec)) :: _) =>
           ((tyco, vs), constructors_of spec))
@@ -1088,24 +1091,21 @@
   (map_exec_purge o map_cases o apsnd) (Symtab.update (c, ())) thy;
 
 
-(* datatypes *)
+(* types *)
 
-structure Type_Interpretation =
-  Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
-
-fun register_datatype (tyco, vs_spec) thy =
+fun register_type (tyco, vs_spec) thy =
   let
     val (old_constrs, some_old_proj) =
-      case these (Symtab.lookup ((the_datatypes o the_exec) thy) tyco)
+      case these (Symtab.lookup ((the_types o the_exec) thy) tyco)
        of (_, (_, Constructors cos)) :: _ => (map fst cos, NONE)
         | (_, (_, Abstractor (_, (co, _)))) :: _ => ([], SOME co)
         | [] => ([], NONE)
     val outdated_funs = case some_old_proj
-     of NONE => []
+     of NONE => old_constrs
       | SOME old_proj => Symtab.fold
           (fn (c, ((_, spec), _)) => if member (op =) (the_list (associated_abstype spec)) tyco
             then insert (op =) c else I)
-            ((the_functions o the_exec) thy) [old_proj];
+            ((the_functions o the_exec) thy) (old_proj :: old_constrs);
     fun drop_outdated_cases cases = fold Symtab.delete_safe
       (Symtab.fold (fn (c, (_, (_, cos))) =>
         if exists (member (op =) old_constrs) cos
@@ -1116,13 +1116,15 @@
     |> map_exec_purge
         ((map_typs o Symtab.map_default (tyco, [])) (cons (serial (), vs_spec))
         #> (map_cases o apfst) drop_outdated_cases)
-    |> Type_Interpretation.data (tyco, serial ())
   end;
 
-fun type_interpretation f = Type_Interpretation.interpretation
-  (fn (tyco, _) => fn thy => f (tyco, get_datatype thy tyco) thy);
+fun unoverload_const_typ thy (c, ty) = (AxClass.unoverload_const thy (c, ty), ty);
 
-fun unoverload_const_typ thy (c, ty) = (AxClass.unoverload_const thy (c, ty), ty);
+structure Datatype_Interpretation =
+  Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
+
+fun datatype_interpretation f = Datatype_Interpretation.interpretation
+  (fn (tyco, _) => fn thy => f (tyco, get_type thy tyco) thy);
 
 fun add_datatype proto_constrs thy =
   let
@@ -1131,20 +1133,29 @@
   in
     thy
     |> fold (del_eqns o fst) constrs
-    |> register_datatype (tyco, (vs, Constructors cos))
+    |> register_type (tyco, (vs, Constructors cos))
+    |> Datatype_Interpretation.data (tyco, serial ())
   end;
 
 fun add_datatype_cmd raw_constrs thy =
   add_datatype (map (read_bare_const thy) raw_constrs) thy;
 
+structure Abstype_Interpretation =
+  Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
+
+fun abstype_interpretation f = Abstype_Interpretation.interpretation
+  (fn (tyco, _) => fn thy => f (tyco, get_abstype_spec thy tyco) thy);
+
 fun add_abstype proto_abs proto_rep thy =
   let
     val (abs, rep) = pairself (unoverload_const_typ thy) (proto_abs, proto_rep);
     val (tyco, (vs, (abs_ty as (abs, ty), (rep, cert_prop)))) = abstype_cert thy abs (fst rep);
     fun after_qed [[cert]] = ProofContext.theory
-      (register_datatype (tyco, (vs, Abstractor (abs_ty, (rep, cert))))
+      (del_eqns abs
+      #> register_type (tyco, (vs, Abstractor (abs_ty, (rep, cert))))
       #> change_fun_spec false rep ((K o Proj)
-        (map_types Logic.varifyT (mk_proj tyco vs ty abs rep), tyco)));
+        (map_types Logic.varifyT (mk_proj tyco vs ty abs rep), tyco))
+      #> Abstype_Interpretation.data (tyco, serial ()));
   in
     thy
     |> ProofContext.init
@@ -1188,7 +1199,7 @@
            (mk_attribute o code_target_attr))
       || Scan.succeed (mk_attribute add_warning_eqn);
   in
-    Type_Interpretation.init
+    Datatype_Interpretation.init
     #> Attrib.setup (Binding.name "code") (Scan.lift code_attribute_parser)
         "declare theorems for code generation"
   end));
--- a/src/Pure/Isar/isar_syn.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -99,13 +99,6 @@
   OuterSyntax.command "defaultsort" "declare default sort" K.thy_decl
     (P.sort >> (Toplevel.theory o Sign.add_defsort));
 
-val _ =
-  OuterSyntax.command "axclass" "define axiomatic type class" K.thy_decl
-    (P.binding -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
-        P.!!! (P.list1 P.xname)) []
-        -- Scan.repeat (SpecParse.thm_name ":" -- (P.prop >> single))
-      >> (fn (x, y) => Toplevel.theory (snd o Class.axclass_cmd x y)));
-
 
 (* types *)
 
--- a/src/Tools/Code/code_eval.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/Tools/Code/code_eval.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -130,7 +130,7 @@
     val thy = ProofContext.theory_of background;
     val tyco = Sign.intern_type thy raw_tyco;
     val constrs = map (Code.check_const thy) raw_constrs;
-    val constrs' = (map fst o snd o Code.get_datatype thy) tyco;
+    val constrs' = (map fst o snd o Code.get_type thy) tyco;
     val _ = if eq_set (op =) (constrs, constrs') then ()
       else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
     val is_first = is_first_occ background;
--- a/src/Tools/Code/code_thingol.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/Tools/Code/code_thingol.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -256,7 +256,7 @@
     | thyname :: _ => thyname;
   fun thyname_of_const thy c = case AxClass.class_of_param thy c
    of SOME class => thyname_of_class thy class
-    | NONE => (case Code.get_datatype_of_constr_or_abstr thy c
+    | NONE => (case Code.get_type_of_constr_or_abstr thy c
        of SOME (tyco, _) => Codegen.thyname_of_type thy tyco
         | NONE => Codegen.thyname_of_const thy c);
   fun purify_base "==>" = "follows"
@@ -543,7 +543,7 @@
   let
     val stmt_datatype =
       let
-        val (vs, cos) = Code.get_datatype thy tyco;
+        val (vs, cos) = Code.get_type thy tyco;
       in
         fold_map (translate_tyvar_sort thy algbr eqngr) vs
         ##>> fold_map (fn (c, tys) =>
@@ -569,7 +569,7 @@
         ##>> fold_map (translate_eqn thy algbr eqngr) eqns
         #>> (fn info => Fun (c, info))
       end;
-    val stmt_const = case Code.get_datatype_of_constr_or_abstr thy c
+    val stmt_const = case Code.get_type_of_constr_or_abstr thy c
      of SOME (tyco, _) => stmt_datatypecons tyco
       | NONE => (case AxClass.class_of_param thy c
          of SOME class => stmt_classparam class
--- a/src/Tools/quickcheck.ML	Tue Feb 23 17:33:03 2010 +0100
+++ b/src/Tools/quickcheck.ML	Tue Feb 23 17:55:00 2010 +0100
@@ -10,6 +10,8 @@
   val timing : bool Unsynchronized.ref
   val test_term: Proof.context -> bool -> string option -> int -> int -> term ->
     (string * term) list option
+  val timed_test_term: Proof.context -> bool -> string option -> int -> int -> term ->
+    ((string * term) list option * (string * int) list)
   val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory
   val setup: theory -> theory
   val quickcheck: (string * string) list -> int -> Proof.state -> (string * term) list option
@@ -97,13 +99,20 @@
     val frees = Term.add_frees t [];
   in (map fst frees, list_abs_free (frees, t)) end
 
-fun test_term ctxt quiet generator_name size i t =
+fun cpu_time description f =
+  let
+    val start = start_timing ()
+    val result = Exn.capture f ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, (description, time)) end
+
+fun timed_test_term ctxt quiet generator_name size i t =
   let
     val (names, t') = prep_test_term t;
-    val testers = (*cond_timeit (!timing) "quickcheck compilation"
-      (fn () => *)(case generator_name
+    val (testers, comp_time) = cpu_time "quickcheck compilation"
+      (fn () => (case generator_name
        of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t'
-        | SOME name => [mk_tester_select name ctxt t']);
+        | SOME name => [mk_tester_select name ctxt t']));
     fun iterate f 0 = NONE
       | iterate f j = case f () handle Match => (if quiet then ()
              else warning "Exception Match raised during quickcheck"; NONE)
@@ -117,13 +126,17 @@
       else (if quiet then () else priority ("Test data size: " ^ string_of_int k);
         case with_testers k testers
          of NONE => with_size (k + 1) | SOME q => SOME q);
+    val (result, exec_time) = cpu_time "quickcheck execution"
+      (fn () => case with_size 1
+        of NONE => NONE
+        | SOME ts => SOME (names ~~ ts))
   in
-    cond_timeit (!timing) "quickcheck execution"
-    (fn () => case with_size 1
-      of NONE => NONE
-      | SOME ts => SOME (names ~~ ts))
+    (result, [exec_time, comp_time])
   end;
 
+fun test_term ctxt quiet generator_name size i t =
+  fst (timed_test_term ctxt quiet generator_name size i t)
+
 fun monomorphic_term thy insts default_T = 
   let
     fun subst (T as TFree (v, S)) =