merged
authorwenzelm
Tue, 05 Nov 2019 22:56:06 +0100
changeset 71059 9b531e611d66
parent 71045 9858f391ed2d (diff)
parent 71058 6ca9e8377613 (current diff)
child 71060 295609359b58
child 71061 1d19e844fa4d
merged
--- a/src/HOL/Analysis/Cartesian_Space.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Cartesian_Space.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -583,7 +583,7 @@
     finally show ?thesis .
   qed
   then show ?thesis
-    by (simp add: dim_span)
+    by (simp)
 qed
 
 lemma column_rank_def:
@@ -1028,7 +1028,9 @@
 qed
 
 
-subsection \<open> We can find an orthogonal matrix taking any unit vector to any other\<close>
+subsection \<open>Finding an Orthogonal Matrix\<close>
+
+text \<open>We can find an orthogonal matrix taking any unit vector to any other.\<close>
 
 lemma  orthogonal_matrix_transpose [simp]:
      "orthogonal_matrix(transpose A) \<longleftrightarrow> orthogonal_matrix A"
@@ -1123,9 +1125,9 @@
 qed
 
 
-subsection  \<open>Linearity of scaling, and hence isometry, that preserves origin\<close>
+subsection  \<open>Scaling and isometry\<close>
 
-lemma  scaling_linear:
+proposition scaling_linear:
   fixes f :: "'a::real_inner \<Rightarrow> 'a::real_inner"
   assumes f0: "f 0 = 0"
     and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
@@ -1158,7 +1160,7 @@
   by (metis dist_0_norm)
 
 
-subsection  \<open>Can extend an isometry from unit sphere\<close>
+text \<open>Can extend an isometry from unit sphere:\<close>
 
 lemma  isometry_sphere_extend:
   fixes f:: "'a::real_inner \<Rightarrow> 'a"
@@ -1393,8 +1395,7 @@
     fix A B
     assume "P ((*v) A)" and "P ((*v) B)"
     then show "P ((*v) (A ** B))"
-      by (auto simp add: matrix_vector_mult_matrix_matrix_mult_compose matrix_vector_mul_linear
-          intro!: comp)
+      by (auto simp add: matrix_vector_mult_matrix_matrix_mult_compose intro!: comp)
   next
     fix A :: "real^'n^'n" and i
     assume "row i A = 0"
--- a/src/HOL/Analysis/Convex.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Convex.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -14,7 +14,7 @@
   "HOL-Library.Set_Algebras"
 begin
 
-subsection \<open>Convexity\<close>
+subsection \<open>Convex Sets\<close>
 
 definition\<^marker>\<open>tag important\<close> convex :: "'a::real_vector set \<Rightarrow> bool"
   where "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
@@ -340,7 +340,7 @@
   done
 
 
-subsection \<open>Functions that are convex on a set\<close>
+subsection \<open>Convex Functions on a Set\<close>
 
 definition\<^marker>\<open>tag important\<close> convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
   where "convex_on s f \<longleftrightarrow>
@@ -1865,7 +1865,7 @@
 qed
 
 
-subsection \<open>Affine dependence and consequential theorems\<close>
+subsection \<open>Affine Dependence\<close>
 
 text "Formalized by Lars Schewe."
 
@@ -2696,8 +2696,7 @@
 proof -
   obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
     using assms affine_hull_explicit[of S] by auto
-  then have "\<exists>v\<in>s. u v \<noteq> 0"
-    using sum_not_0[of "u" "s"] by auto
+  then have "\<exists>v\<in>s. u v \<noteq> 0" by auto
   then have "finite s \<and> s \<subseteq> S \<and> (\<exists>v\<in>s. u v \<noteq> 0 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0)"
     using s_u by auto
   then show ?thesis
@@ -3483,7 +3482,7 @@
     have "sum c s = 0"
       by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite s\<close> sum_negf)
     moreover have "\<not> (\<forall>v\<in>s. c v = 0)"
-      by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum_not_0 zero_neq_one)
+      by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum.neutral zero_neq_one)
     moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
       by (simp add: c_def if_smult sum_negf
              comm_monoid_add_class.sum.If_cases \<open>finite s\<close>)
--- a/src/HOL/Analysis/Determinants.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Determinants.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -2,7 +2,7 @@
     Author:     Amine Chaieb, University of Cambridge; proofs reworked by LCP
 *)
 
-section \<open>Traces, Determinant of square matrices and some properties\<close>
+section \<open>Traces and Determinants of Square Matrices\<close>
 
 theory Determinants
 imports
--- a/src/HOL/Analysis/Elementary_Topology.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Elementary_Topology.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -22,95 +22,6 @@
   using openI by auto
 
 
-subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Archimedean properties and useful consequences\<close>
-
-text\<open>Bernoulli's inequality\<close>
-proposition Bernoulli_inequality:
-  fixes x :: real
-  assumes "-1 \<le> x"
-    shows "1 + n * x \<le> (1 + x) ^ n"
-proof (induct n)
-  case 0
-  then show ?case by simp
-next
-  case (Suc n)
-  have "1 + Suc n * x \<le> 1 + (Suc n)*x + n * x^2"
-    by (simp add: algebra_simps)
-  also have "... = (1 + x) * (1 + n*x)"
-    by (auto simp: power2_eq_square algebra_simps  of_nat_Suc)
-  also have "... \<le> (1 + x) ^ Suc n"
-    using Suc.hyps assms mult_left_mono by fastforce
-  finally show ?case .
-qed
-
-corollary Bernoulli_inequality_even:
-  fixes x :: real
-  assumes "even n"
-    shows "1 + n * x \<le> (1 + x) ^ n"
-proof (cases "-1 \<le> x \<or> n=0")
-  case True
-  then show ?thesis
-    by (auto simp: Bernoulli_inequality)
-next
-  case False
-  then have "real n \<ge> 1"
-    by simp
-  with False have "n * x \<le> -1"
-    by (metis linear minus_zero mult.commute mult.left_neutral mult_left_mono_neg neg_le_iff_le order_trans zero_le_one)
-  then have "1 + n * x \<le> 0"
-    by auto
-  also have "... \<le> (1 + x) ^ n"
-    using assms
-    using zero_le_even_power by blast
-  finally show ?thesis .
-qed
-
-corollary real_arch_pow:
-  fixes x :: real
-  assumes x: "1 < x"
-  shows "\<exists>n. y < x^n"
-proof -
-  from x have x0: "x - 1 > 0"
-    by arith
-  from reals_Archimedean3[OF x0, rule_format, of y]
-  obtain n :: nat where n: "y < real n * (x - 1)" by metis
-  from x0 have x00: "x- 1 \<ge> -1" by arith
-  from Bernoulli_inequality[OF x00, of n] n
-  have "y < x^n" by auto
-  then show ?thesis by metis
-qed
-
-corollary real_arch_pow_inv:
-  fixes x y :: real
-  assumes y: "y > 0"
-    and x1: "x < 1"
-  shows "\<exists>n. x^n < y"
-proof (cases "x > 0")
-  case True
-  with x1 have ix: "1 < 1/x" by (simp add: field_simps)
-  from real_arch_pow[OF ix, of "1/y"]
-  obtain n where n: "1/y < (1/x)^n" by blast
-  then show ?thesis using y \<open>x > 0\<close>
-    by (auto simp add: field_simps)
-next
-  case False
-  with y x1 show ?thesis
-    by (metis less_le_trans not_less power_one_right)
-qed
-
-lemma forall_pos_mono:
-  "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
-    (\<And>n::nat. n \<noteq> 0 \<Longrightarrow> P (inverse (real n))) \<Longrightarrow> (\<And>e. 0 < e \<Longrightarrow> P e)"
-  by (metis real_arch_inverse)
-
-lemma forall_pos_mono_1:
-  "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
-    (\<And>n. P (inverse (real (Suc n)))) \<Longrightarrow> 0 < e \<Longrightarrow> P e"
-  apply (rule forall_pos_mono)
-  apply auto
-  apply (metis Suc_pred of_nat_Suc)
-  done
-
 subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Affine transformations of intervals\<close>
 
 lemma real_affinity_le: "0 < m \<Longrightarrow> m * x + c \<le> y \<longleftrightarrow> x \<le> inverse m * y + - (c / m)"
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -135,8 +135,7 @@
       by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff)
     have "CARD('a) ^ CARD('b) = card (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
       (is "_ = card ?A")
-      by (subst card_PiE) (auto simp: prod_constant)
-    
+      by (subst card_PiE) (auto)
     also have "?A = Pi UNIV (\<lambda>_. UNIV)" 
       by auto
     finally show "card \<dots> = CARD('a) ^ CARD('b)" ..
@@ -1075,7 +1074,7 @@
 
 lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
   apply (vector matrix_vector_mult_def mat_def)
-  apply (simp add: if_distrib if_distribR sum.delta' cong del: if_weak_cong)
+  apply (simp add: if_distrib if_distribR cong del: if_weak_cong)
   done
 
 lemma matrix_transpose_mul:
--- a/src/HOL/Analysis/Linear_Algebra.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Linear_Algebra.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -50,9 +50,6 @@
 lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d"
   by (rule independent_mono[OF independent_Basis])
 
-lemma sum_not_0: "sum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0"
-  by (rule ccontr) auto
-
 lemma subset_translation_eq [simp]:
     fixes a :: "'a::real_vector" shows "(+) a ` s \<subseteq> (+) a ` t \<longleftrightarrow> s \<subseteq> t"
   by auto
@@ -493,96 +490,6 @@
   by (rule adjoint_unique, simp add: adjoint_clauses [OF lf])
 
 
-subsection \<open>Archimedean properties and useful consequences\<close>
-
-text\<open>Bernoulli's inequality\<close>
-proposition Bernoulli_inequality:
-  fixes x :: real
-  assumes "-1 \<le> x"
-    shows "1 + n * x \<le> (1 + x) ^ n"
-proof (induct n)
-  case 0
-  then show ?case by simp
-next
-  case (Suc n)
-  have "1 + Suc n * x \<le> 1 + (Suc n)*x + n * x^2"
-    by (simp add: algebra_simps)
-  also have "... = (1 + x) * (1 + n*x)"
-    by (auto simp: power2_eq_square algebra_simps  of_nat_Suc)
-  also have "... \<le> (1 + x) ^ Suc n"
-    using Suc.hyps assms mult_left_mono by fastforce
-  finally show ?case .
-qed
-
-corollary Bernoulli_inequality_even:
-  fixes x :: real
-  assumes "even n"
-    shows "1 + n * x \<le> (1 + x) ^ n"
-proof (cases "-1 \<le> x \<or> n=0")
-  case True
-  then show ?thesis
-    by (auto simp: Bernoulli_inequality)
-next
-  case False
-  then have "real n \<ge> 1"
-    by simp
-  with False have "n * x \<le> -1"
-    by (metis linear minus_zero mult.commute mult.left_neutral mult_left_mono_neg neg_le_iff_le order_trans zero_le_one)
-  then have "1 + n * x \<le> 0"
-    by auto
-  also have "... \<le> (1 + x) ^ n"
-    using assms
-    using zero_le_even_power by blast
-  finally show ?thesis .
-qed
-
-corollary real_arch_pow:
-  fixes x :: real
-  assumes x: "1 < x"
-  shows "\<exists>n. y < x^n"
-proof -
-  from x have x0: "x - 1 > 0"
-    by arith
-  from reals_Archimedean3[OF x0, rule_format, of y]
-  obtain n :: nat where n: "y < real n * (x - 1)" by metis
-  from x0 have x00: "x- 1 \<ge> -1" by arith
-  from Bernoulli_inequality[OF x00, of n] n
-  have "y < x^n" by auto
-  then show ?thesis by metis
-qed
-
-corollary real_arch_pow_inv:
-  fixes x y :: real
-  assumes y: "y > 0"
-    and x1: "x < 1"
-  shows "\<exists>n. x^n < y"
-proof (cases "x > 0")
-  case True
-  with x1 have ix: "1 < 1/x" by (simp add: field_simps)
-  from real_arch_pow[OF ix, of "1/y"]
-  obtain n where n: "1/y < (1/x)^n" by blast
-  then show ?thesis using y \<open>x > 0\<close>
-    by (auto simp add: field_simps)
-next
-  case False
-  with y x1 show ?thesis
-    by (metis less_le_trans not_less power_one_right)
-qed
-
-lemma forall_pos_mono:
-  "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
-    (\<And>n::nat. n \<noteq> 0 \<Longrightarrow> P (inverse (real n))) \<Longrightarrow> (\<And>e. 0 < e \<Longrightarrow> P e)"
-  by (metis real_arch_inverse)
-
-lemma forall_pos_mono_1:
-  "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
-    (\<And>n. P (inverse (real (Suc n)))) \<Longrightarrow> 0 < e \<Longrightarrow> P e"
-  apply (rule forall_pos_mono)
-  apply auto
-  apply (metis Suc_pred of_nat_Suc)
-  done
-
-
 subsection\<^marker>\<open>tag unimportant\<close> \<open>Euclidean Spaces as Typeclass\<close>
 
 lemma independent_Basis: "independent Basis"
@@ -899,7 +806,7 @@
     by (simp add: span_span)
   from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB
   have iC: "independent C"
-    by (simp add: dim_span)
+    by (simp)
   from C fB have "card C \<le> dim V"
     by simp
   moreover have "dim V \<le> card C"
@@ -1441,8 +1348,7 @@
   show ?thesis
     apply (rule that[OF b(1)])
     apply (rule subspace_dim_equal)
-    by (auto simp: assms b dim_hyperplane dim_span subspace_hyperplane
-        subspace_span)
+    by (auto simp: assms b dim_hyperplane subspace_hyperplane)
 qed
 
 lemma dim_eq_hyperplane:
@@ -1451,7 +1357,7 @@
 by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane)
 
 
-subsection\<open> Orthogonal bases, Gram-Schmidt process, and related theorems\<close>
+subsection\<open> Orthogonal bases and Gram-Schmidt process\<close>
 
 lemma pairwise_orthogonal_independent:
   assumes "pairwise orthogonal S" and "0 \<notin> S"
@@ -1512,7 +1418,7 @@
        if "x \<in> S" for x
   proof -
     have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)"
-      by (simp add: \<open>finite S\<close> inner_commute sum.delta that)
+      by (simp add: \<open>finite S\<close> inner_commute that)
     also have "... =  (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))"
       apply (rule sum.cong [OF refl], simp)
       by (meson S orthogonal_def pairwise_def that)
@@ -1665,8 +1571,7 @@
   obtain B where "B \<subseteq> span S" and orthB: "pairwise orthogonal B"
              and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
              and "independent B" "card B = dim S" "span B = span S"
-    by (rule orthonormal_basis_subspace [of "span S", OF subspace_span])
-      (auto simp: dim_span)
+    by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) (auto)
   with assms obtain u where spanBT: "span B \<subseteq> span T" and "u \<notin> span B" "u \<in> span T"
     by auto
   obtain C where orthBC: "pairwise orthogonal (B \<union> C)" and spanBC: "span (B \<union> C) = span (B \<union> {u})"
@@ -1708,11 +1613,11 @@
   assumes "dim S < DIM('a)"
   obtains x where "x \<noteq> 0" "\<And>y. y \<in> span S \<Longrightarrow> orthogonal x y"
 proof -
-have "span S \<subset> UNIV"
+  have "span S \<subset> UNIV"
   by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane
       mem_Collect_eq top.extremum_strict top.not_eq_extremum)
   with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis
-    by (auto simp: span_UNIV)
+    by (auto)
 qed
 
 corollary\<^marker>\<open>tag unimportant\<close> orthogonal_to_vector_exists:
@@ -1787,7 +1692,7 @@
       using \<open>independent (S - {0})\<close> independent_imp_finite by blast
     show "card (S - {0}) = DIM('a)"
       using span_delete_0 [of S] S
-      by (simp add: \<open>independent (S - {0})\<close> indep_card_eq_dim_span dim_UNIV)
+      by (simp add: \<open>independent (S - {0})\<close> indep_card_eq_dim_span)
   qed (use S \<open>a \<noteq> 0\<close> in auto)
 qed
 
@@ -1838,7 +1743,7 @@
   then have 0: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0"
     by simp
   have "dim(A \<union> B) = dim (span (A \<union> B))"
-    by (simp add: dim_span)
+    by (simp)
   also have "span (A \<union> B) = ((\<lambda>(a, b). a + b) ` (span A \<times> span B))"
     by (auto simp add: span_Un image_def)
   also have "dim \<dots> = dim {x + y |x y. x \<in> span A \<and> y \<in> span B}"
@@ -1846,9 +1751,9 @@
   also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)"
     by (auto simp: dest: 0)
   also have "... = dim (span A) + dim (span B)"
-    by (rule dim_sums_Int) (auto simp: subspace_span)
+    by (rule dim_sums_Int) (auto)
   also have "... = dim A + dim B"
-    by (simp add: dim_span)
+    by (simp)
   finally show ?thesis .
 qed
 
@@ -1880,8 +1785,7 @@
         using \<open>y \<in> span A\<close> add.commute by blast
     qed
     show "span B \<subseteq> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
-      by (rule span_minimal)
-        (auto intro: * span_minimal simp: subspace_span)
+      by (rule span_minimal) (auto intro: * span_minimal)
   qed
   then show ?thesis
     by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq
@@ -1938,7 +1842,7 @@
         show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x
           using that B
           apply (simp add: field_split_simps)
-          by (metis \<open>linear h\<close> le_less_trans linear_diff mult.commute)
+          by (metis \<open>linear h\<close> le_less_trans linear_diff)
       qed
     qed
   qed
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -32,7 +32,7 @@
     by (auto intro!: real_le_rsqrt)
 qed
 
-subsection \<open>Continuity of the representation WRT an orthogonal basis\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuity of the representation WRT an orthogonal basis\<close>
 
 lemma orthogonal_Basis: "pairwise orthogonal Basis"
   by (simp add: inner_not_same_Basis orthogonal_def pairwise_def)
@@ -353,8 +353,8 @@
 
 subsection \<open>Boxes\<close>
 
-abbreviation One :: "'a::euclidean_space"
-  where "One \<equiv> \<Sum>Basis"
+abbreviation\<^marker>\<open>tag important\<close> One :: "'a::euclidean_space" where
+"One \<equiv> \<Sum>Basis"
 
 lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False
 proof -
@@ -366,14 +366,14 @@
   with independent_Basis show False by force
 qed
 
-corollary One_neq_0[iff]: "One \<noteq> 0"
+corollary\<^marker>\<open>tag unimportant\<close> One_neq_0[iff]: "One \<noteq> 0"
   by (metis One_non_0)
 
-corollary Zero_neq_One[iff]: "0 \<noteq> One"
+corollary\<^marker>\<open>tag unimportant\<close> Zero_neq_One[iff]: "0 \<noteq> One"
   by (metis One_non_0)
 
-definition\<^marker>\<open>tag important\<close> (in euclidean_space) eucl_less (infix "<e" 50)
-  where "eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)"
+definition\<^marker>\<open>tag important\<close> (in euclidean_space) eucl_less (infix "<e" 50) where 
+"eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)"
 
 definition\<^marker>\<open>tag important\<close> box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
 definition\<^marker>\<open>tag important\<close> "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}"
--- a/src/HOL/Decision_Procs/Approximation.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -73,554 +73,93 @@
 
 subsection "Implement approximation function"
 
-fun lift_bin :: "(float * float) option \<Rightarrow> (float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float * float) option) \<Rightarrow> (float * float) option" where
-"lift_bin (Some (l1, u1)) (Some (l2, u2)) f = f l1 u1 l2 u2" |
+fun lift_bin :: "(float interval) option \<Rightarrow> (float interval) option \<Rightarrow> (float interval \<Rightarrow> float interval \<Rightarrow> (float interval) option) \<Rightarrow> (float interval) option" where
+"lift_bin (Some ivl1) (Some ivl2) f = f ivl1 ivl2" |
 "lift_bin a b f = None"
 
-fun lift_bin' :: "(float * float) option \<Rightarrow> (float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float * float)) \<Rightarrow> (float * float) option" where
-"lift_bin' (Some (l1, u1)) (Some (l2, u2)) f = Some (f l1 u1 l2 u2)" |
+fun lift_bin' :: "(float interval) option \<Rightarrow> (float interval) option \<Rightarrow> (float interval \<Rightarrow> float interval \<Rightarrow> float interval) \<Rightarrow> (float interval) option" where
+"lift_bin' (Some ivl1) (Some ivl2) f = Some (f ivl1 ivl2)" |
 "lift_bin' a b f = None"
 
-fun lift_un :: "(float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> ((float option) * (float option))) \<Rightarrow> (float * float) option" where
-"lift_un (Some (l1, u1)) f = (case (f l1 u1) of (Some l, Some u) \<Rightarrow> Some (l, u)
-                                             | t \<Rightarrow> None)" |
+fun lift_un :: "float interval option \<Rightarrow> (float interval \<Rightarrow> float interval option) \<Rightarrow> float interval option" where
+"lift_un (Some ivl) f = f ivl" |
 "lift_un b f = None"
 
-fun lift_un' :: "(float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> (float * float)) \<Rightarrow> (float * float) option" where
-"lift_un' (Some (l1, u1)) f = Some (f l1 u1)" |
+lemma lift_un_eq:\<comment> \<open>TODO\<close> "lift_un x f = Option.bind x f"
+  by (cases x) auto
+
+fun lift_un' :: "(float interval) option \<Rightarrow> (float interval \<Rightarrow> (float interval)) \<Rightarrow> (float interval) option" where
+"lift_un' (Some ivl1) f = Some (f ivl1)" |
 "lift_un' b f = None"
 
-definition bounded_by :: "real list \<Rightarrow> (float \<times> float) option list \<Rightarrow> bool" where 
-  "bounded_by xs vs \<longleftrightarrow>
-  (\<forall> i < length vs. case vs ! i of None \<Rightarrow> True
-         | Some (l, u) \<Rightarrow> xs ! i \<in> { real_of_float l .. real_of_float u })"
-                                                                     
+definition bounded_by :: "real list \<Rightarrow> (float interval) option list \<Rightarrow> bool" where
+  "bounded_by xs vs \<longleftrightarrow> (\<forall> i < length vs. case vs ! i of None \<Rightarrow> True | Some ivl \<Rightarrow> xs ! i \<in>\<^sub>r ivl)"
+
 lemma bounded_byE:
   assumes "bounded_by xs vs"
   shows "\<And> i. i < length vs \<Longrightarrow> case vs ! i of None \<Rightarrow> True
-         | Some (l, u) \<Rightarrow> xs ! i \<in> { real_of_float l .. real_of_float u }"
-  using assms bounded_by_def by blast
+         | Some ivl \<Rightarrow> xs ! i \<in>\<^sub>r ivl"
+  using assms
+  by (auto simp: bounded_by_def)
 
 lemma bounded_by_update:
   assumes "bounded_by xs vs"
-    and bnd: "xs ! i \<in> { real_of_float l .. real_of_float u }"
-  shows "bounded_by xs (vs[i := Some (l,u)])"
-proof -
-  {
-    fix j
-    let ?vs = "vs[i := Some (l,u)]"
-    assume "j < length ?vs"
-    hence [simp]: "j < length vs" by simp
-    have "case ?vs ! j of None \<Rightarrow> True | Some (l, u) \<Rightarrow> xs ! j \<in> { real_of_float l .. real_of_float u }"
-    proof (cases "?vs ! j")
-      case (Some b)
-      thus ?thesis
-      proof (cases "i = j")
-        case True
-        thus ?thesis using \<open>?vs ! j = Some b\<close> and bnd by auto
-      next
-        case False
-        thus ?thesis using \<open>bounded_by xs vs\<close> unfolding bounded_by_def by auto
-      qed
-    qed auto
-  }
-  thus ?thesis unfolding bounded_by_def by auto
-qed
+    and bnd: "xs ! i \<in>\<^sub>r ivl"
+  shows "bounded_by xs (vs[i := Some ivl])"
+  using assms
+  by (cases "i < length vs") (auto simp: bounded_by_def nth_list_update split: option.splits)
 
 lemma bounded_by_None: "bounded_by xs (replicate (length xs) None)"
   unfolding bounded_by_def by auto
 
-fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option" where
-"approx' prec a bs          = (case (approx prec a bs) of Some (l, u) \<Rightarrow> Some (float_round_down prec l, float_round_up prec u) | None \<Rightarrow> None)" |
-"approx prec (Add a b) bs   =
-  lift_bin' (approx' prec a bs) (approx' prec b bs)
-    (\<lambda> l1 u1 l2 u2. (float_plus_down prec l1 l2, float_plus_up prec u1 u2))" |
-"approx prec (Minus a) bs   = lift_un' (approx' prec a bs) (\<lambda> l u. (-u, -l))" |
-"approx prec (Mult a b) bs  =
-  lift_bin' (approx' prec a bs) (approx' prec b bs) (bnds_mult prec)" |
-"approx prec (Inverse a) bs = lift_un (approx' prec a bs) (\<lambda> l u. if (0 < l \<or> u < 0) then (Some (float_divl prec 1 u), Some (float_divr prec 1 l)) else (None, None))" |
-"approx prec (Cos a) bs     = lift_un' (approx' prec a bs) (bnds_cos prec)" |
-"approx prec Pi bs          = Some (lb_pi prec, ub_pi prec)" |
-"approx prec (Min a b) bs   = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (min l1 l2, min u1 u2))" |
-"approx prec (Max a b) bs   = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (max l1 l2, max u1 u2))" |
-"approx prec (Abs a) bs     = lift_un' (approx' prec a bs) (\<lambda>l u. (if l < 0 \<and> 0 < u then 0 else min \<bar>l\<bar> \<bar>u\<bar>, max \<bar>l\<bar> \<bar>u\<bar>))" |
-"approx prec (Arctan a) bs  = lift_un' (approx' prec a bs) (\<lambda> l u. (lb_arctan prec l, ub_arctan prec u))" |
-"approx prec (Sqrt a) bs    = lift_un' (approx' prec a bs) (\<lambda> l u. (lb_sqrt prec l, ub_sqrt prec u))" |
-"approx prec (Exp a) bs     = lift_un' (approx' prec a bs) (\<lambda> l u. (lb_exp prec l, ub_exp prec u))" |
-"approx prec (Powr a b) bs  = lift_bin (approx' prec a bs) (approx' prec b bs) (bnds_powr prec)" |
-"approx prec (Ln a) bs      = lift_un (approx' prec a bs) (\<lambda> l u. (lb_ln prec l, ub_ln prec u))" |
-"approx prec (Power a n) bs = lift_un' (approx' prec a bs) (float_power_bnds prec n)" |
-"approx prec (Floor a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (floor_fl l, floor_fl u))" |
-"approx prec (Num f) bs     = Some (f, f)" |
-"approx prec (Var i) bs    = (if i < length bs then bs ! i else None)"
+fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float interval) option list \<Rightarrow> (float interval) option"
+  where
+    "approx' prec a bs          = (case (approx prec a bs) of Some ivl \<Rightarrow> Some (round_interval prec ivl) | None \<Rightarrow> None)" |
+    "approx prec (Add a b) bs   = lift_bin' (approx' prec a bs) (approx' prec b bs) (+)" |
+    "approx prec (Minus a) bs   = lift_un' (approx' prec a bs) uminus" |
+    "approx prec (Mult a b) bs  = lift_bin' (approx' prec a bs) (approx' prec b bs) (mult_float_interval prec)" |
+    "approx prec (Inverse a) bs = lift_un (approx' prec a bs) (inverse_float_interval prec)" |
+    "approx prec (Cos a) bs     = lift_un' (approx' prec a bs) (cos_float_interval prec)" |
+    "approx prec Pi bs          = Some (pi_float_interval prec)" |
+    "approx prec (Min a b) bs   = lift_bin' (approx' prec a bs) (approx' prec b bs) (min_interval)" |
+    "approx prec (Max a b) bs   = lift_bin' (approx' prec a bs) (approx' prec b bs) (max_interval)" |
+    "approx prec (Abs a) bs     = lift_un' (approx' prec a bs) (abs_interval)" |
+    "approx prec (Arctan a) bs  = lift_un' (approx' prec a bs) (arctan_float_interval prec)" |
+    "approx prec (Sqrt a) bs    = lift_un' (approx' prec a bs) (sqrt_float_interval prec)" |
+    "approx prec (Exp a) bs     = lift_un' (approx' prec a bs) (exp_float_interval prec)" |
+    "approx prec (Powr a b) bs  = lift_bin (approx' prec a bs) (approx' prec b bs) (powr_float_interval prec)" |
+    "approx prec (Ln a) bs      = lift_un (approx' prec a bs) (ln_float_interval prec)" |
+    "approx prec (Power a n) bs = lift_un' (approx' prec a bs) (power_float_interval prec n)" |
+    "approx prec (Floor a) bs = lift_un' (approx' prec a bs) (floor_float_interval)" |
+    "approx prec (Num f) bs     = Some (interval_of f)" |
+    "approx prec (Var i) bs    = (if i < length bs then bs ! i else None)"
 
 lemma approx_approx':
-  assumes Pa: "\<And>l u. Some (l, u) = approx prec a vs \<Longrightarrow>
-      l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-    and approx': "Some (l, u) = approx' prec a vs"
-  shows "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-proof -
-  obtain l' u' where S: "Some (l', u') = approx prec a vs"
-    using approx' unfolding approx'.simps by (cases "approx prec a vs") auto
-  have l': "l = float_round_down prec l'" and u': "u = float_round_up prec u'"
-    using approx' unfolding approx'.simps S[symmetric] by auto
-  show ?thesis unfolding l' u'
-    using order_trans[OF Pa[OF S, THEN conjunct2] float_round_up[of u']]
-    using order_trans[OF float_round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto
-qed
-
-lemma lift_bin_ex:
-  assumes lift_bin_Some: "Some (l, u) = lift_bin a b f"
-  shows "\<exists> l1 u1 l2 u2. Some (l1, u1) = a \<and> Some (l2, u2) = b"
-proof (cases a)
-  case None
-  hence "None = lift_bin a b f"
-    unfolding None lift_bin.simps ..
-  thus ?thesis
-    using lift_bin_Some by auto
-next
-  case (Some a')
-  show ?thesis
-  proof (cases b)
-    case None
-    hence "None = lift_bin a b f"
-      unfolding None lift_bin.simps ..
-    thus ?thesis using lift_bin_Some by auto
-  next
-    case (Some b')
-    obtain la ua where a': "a' = (la, ua)"
-      by (cases a') auto
-    obtain lb ub where b': "b' = (lb, ub)"
-      by (cases b') auto
-    thus ?thesis
-      unfolding \<open>a = Some a'\<close> \<open>b = Some b'\<close> a' b' by auto
-  qed
-qed
-
-lemma lift_bin_f:
-  assumes lift_bin_Some: "Some (l, u) = lift_bin (g a) (g b) f"
-    and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"
-    and Pb: "\<And>l u. Some (l, u) = g b \<Longrightarrow> P l u b"
-  shows "\<exists> l1 u1 l2 u2. P l1 u1 a \<and> P l2 u2 b \<and> Some (l, u) = f l1 u1 l2 u2"
-proof -
-  obtain l1 u1 l2 u2
-    where Sa: "Some (l1, u1) = g a"
-      and Sb: "Some (l2, u2) = g b"
-    using lift_bin_ex[OF assms(1)] by auto
-  have lu: "Some (l, u) = f l1 u1 l2 u2"
-    using lift_bin_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin.simps] by auto
-  thus ?thesis
-    using Pa[OF Sa] Pb[OF Sb] by auto
-qed
-
-lemma lift_bin:
-  assumes lift_bin_Some: "Some (l, u) = lift_bin (approx' prec a bs) (approx' prec b bs) f"
-    and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>
-      real_of_float l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real_of_float u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
-    and Pb: "\<And>l u. Some (l, u) = approx prec b bs \<Longrightarrow>
-      real_of_float l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real_of_float u"
-  shows "\<exists>l1 u1 l2 u2. (real_of_float l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real_of_float u1) \<and>
-                       (real_of_float l2 \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real_of_float u2) \<and>
-                       Some (l, u) = (f l1 u1 l2 u2)"
-proof -
-  { fix l u assume "Some (l, u) = approx' prec a bs"
-    with approx_approx'[of prec a bs, OF _ this] Pa
-    have "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u" by auto } note Pa = this
-  { fix l u assume "Some (l, u) = approx' prec b bs"
-    with approx_approx'[of prec b bs, OF _ this] Pb
-    have "l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u" by auto } note Pb = this
-
-  from lift_bin_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_bin_Some, OF Pa Pb]
-  show ?thesis by auto
-qed
-
-lemma lift_bin'_ex:
-  assumes lift_bin'_Some: "Some (l, u) = lift_bin' a b f"
-  shows "\<exists> l1 u1 l2 u2. Some (l1, u1) = a \<and> Some (l2, u2) = b"
-proof (cases a)
-  case None
-  hence "None = lift_bin' a b f"
-    unfolding None lift_bin'.simps ..
-  thus ?thesis
-    using lift_bin'_Some by auto
-next
-  case (Some a')
-  show ?thesis
-  proof (cases b)
-    case None
-    hence "None = lift_bin' a b f"
-      unfolding None lift_bin'.simps ..
-    thus ?thesis using lift_bin'_Some by auto
-  next
-    case (Some b')
-    obtain la ua where a': "a' = (la, ua)"
-      by (cases a') auto
-    obtain lb ub where b': "b' = (lb, ub)"
-      by (cases b') auto
-    thus ?thesis
-      unfolding \<open>a = Some a'\<close> \<open>b = Some b'\<close> a' b' by auto
-  qed
-qed
-
-lemma lift_bin'_f:
-  assumes lift_bin'_Some: "Some (l, u) = lift_bin' (g a) (g b) f"
-    and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"
-    and Pb: "\<And>l u. Some (l, u) = g b \<Longrightarrow> P l u b"
-  shows "\<exists> l1 u1 l2 u2. P l1 u1 a \<and> P l2 u2 b \<and> l = fst (f l1 u1 l2 u2) \<and> u = snd (f l1 u1 l2 u2)"
-proof -
-  obtain l1 u1 l2 u2
-    where Sa: "Some (l1, u1) = g a"
-      and Sb: "Some (l2, u2) = g b"
-    using lift_bin'_ex[OF assms(1)] by auto
-  have lu: "(l, u) = f l1 u1 l2 u2"
-    using lift_bin'_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin'.simps] by auto
-  have "l = fst (f l1 u1 l2 u2)" and "u = snd (f l1 u1 l2 u2)"
-    unfolding lu[symmetric] by auto
-  thus ?thesis
-    using Pa[OF Sa] Pb[OF Sb] by auto
-qed
-
-lemma lift_bin':
-  assumes lift_bin'_Some: "Some (l, u) = lift_bin' (approx' prec a bs) (approx' prec b bs) f"
-    and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>
-      l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
-    and Pb: "\<And>l u. Some (l, u) = approx prec b bs \<Longrightarrow>
-      l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u"
-  shows "\<exists>l1 u1 l2 u2. (l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u1) \<and>
-                       (l2 \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u2) \<and>
-                       l = fst (f l1 u1 l2 u2) \<and> u = snd (f l1 u1 l2 u2)"
+  assumes Pa: "\<And>ivl. approx prec a vs = Some ivl \<Longrightarrow> interpret_floatarith a xs \<in>\<^sub>r ivl"
+    and approx': "approx' prec a vs = Some ivl"
+  shows "interpret_floatarith a xs \<in>\<^sub>r ivl"
 proof -
-  { fix l u assume "Some (l, u) = approx' prec a bs"
-    with approx_approx'[of prec a bs, OF _ this] Pa
-    have "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u" by auto } note Pa = this
-  { fix l u assume "Some (l, u) = approx' prec b bs"
-    with approx_approx'[of prec b bs, OF _ this] Pb
-    have "l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u" by auto } note Pb = this
-
-  from lift_bin'_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_bin'_Some, OF Pa Pb]
-  show ?thesis by auto
-qed
-
-lemma lift_un'_ex:
-  assumes lift_un'_Some: "Some (l, u) = lift_un' a f"
-  shows "\<exists> l u. Some (l, u) = a"
-proof (cases a)
-  case None
-  hence "None = lift_un' a f"
-    unfolding None lift_un'.simps ..
-  thus ?thesis
-    using lift_un'_Some by auto
-next
-  case (Some a')
-  obtain la ua where a': "a' = (la, ua)"
-    by (cases a') auto
-  thus ?thesis
-    unfolding \<open>a = Some a'\<close> a' by auto
-qed
-
-lemma lift_un'_f:
-  assumes lift_un'_Some: "Some (l, u) = lift_un' (g a) f"
-    and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"
-  shows "\<exists> l1 u1. P l1 u1 a \<and> l = fst (f l1 u1) \<and> u = snd (f l1 u1)"
-proof -
-  obtain l1 u1 where Sa: "Some (l1, u1) = g a"
-    using lift_un'_ex[OF assms(1)] by auto
-  have lu: "(l, u) = f l1 u1"
-    using lift_un'_Some[unfolded Sa[symmetric] lift_un'.simps] by auto
-  have "l = fst (f l1 u1)" and "u = snd (f l1 u1)"
-    unfolding lu[symmetric] by auto
-  thus ?thesis
-    using Pa[OF Sa] by auto
-qed
-
-lemma lift_un':
-  assumes lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f"
-    and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>
-      l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-      (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
-  shows "\<exists>l1 u1. (l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u1) \<and>
-    l = fst (f l1 u1) \<and> u = snd (f l1 u1)"
-proof -
-  have Pa: "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-    if "Some (l, u) = approx' prec a bs" for l u
-    using approx_approx'[of prec a bs, OF _ that] Pa
-     by auto
-  from lift_un'_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_un'_Some, OF Pa]
-  show ?thesis by auto
-qed
-
-lemma lift_un'_bnds:
-  assumes bnds: "\<forall> (x::real) lx ux. (l, u) = f lx ux \<and> x \<in> { lx .. ux } \<longrightarrow> l \<le> f' x \<and> f' x \<le> u"
-    and lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f"
-    and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>
-      l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-  shows "real_of_float l \<le> f' (interpret_floatarith a xs) \<and> f' (interpret_floatarith a xs) \<le> real_of_float u"
-proof -
-  from lift_un'[OF lift_un'_Some Pa]
-  obtain l1 u1 where "l1 \<le> interpret_floatarith a xs"
-    and "interpret_floatarith a xs \<le> u1"
-    and "l = fst (f l1 u1)"
-    and "u = snd (f l1 u1)"
-    by blast
-  hence "(l, u) = f l1 u1" and "interpret_floatarith a xs \<in> {l1 .. u1}"
-    by auto
-  thus ?thesis
-    using bnds by auto
-qed
-
-lemma lift_un_ex:
-  assumes lift_un_Some: "Some (l, u) = lift_un a f"
-  shows "\<exists>l u. Some (l, u) = a"
-proof (cases a)
-  case None
-  hence "None = lift_un a f"
-    unfolding None lift_un.simps ..
-  thus ?thesis
-    using lift_un_Some by auto
-next
-  case (Some a')
-  obtain la ua where a': "a' = (la, ua)"
-    by (cases a') auto
-  thus ?thesis
-    unfolding \<open>a = Some a'\<close> a' by auto
-qed
-
-lemma lift_un_f:
-  assumes lift_un_Some: "Some (l, u) = lift_un (g a) f"
-    and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"
-  shows "\<exists> l1 u1. P l1 u1 a \<and> Some l = fst (f l1 u1) \<and> Some u = snd (f l1 u1)"
-proof -
-  obtain l1 u1 where Sa: "Some (l1, u1) = g a"
-    using lift_un_ex[OF assms(1)] by auto
-  have "fst (f l1 u1) \<noteq> None \<and> snd (f l1 u1) \<noteq> None"
-  proof (rule ccontr)
-    assume "\<not> (fst (f l1 u1) \<noteq> None \<and> snd (f l1 u1) \<noteq> None)"
-    hence or: "fst (f l1 u1) = None \<or> snd (f l1 u1) = None" by auto
-    hence "lift_un (g a) f = None"
-    proof (cases "fst (f l1 u1) = None")
-      case True
-      then obtain b where b: "f l1 u1 = (None, b)"
-        by (cases "f l1 u1") auto
-      thus ?thesis
-        unfolding Sa[symmetric] lift_un.simps b by auto
-    next
-      case False
-      hence "snd (f l1 u1) = None"
-        using or by auto
-      with False obtain b where b: "f l1 u1 = (Some b, None)"
-        by (cases "f l1 u1") auto
-      thus ?thesis
-        unfolding Sa[symmetric] lift_un.simps b by auto
-    qed
-    thus False
-      using lift_un_Some by auto
-  qed
-  then obtain a' b' where f: "f l1 u1 = (Some a', Some b')"
-    by (cases "f l1 u1") auto
-  from lift_un_Some[unfolded Sa[symmetric] lift_un.simps f]
-  have "Some l = fst (f l1 u1)" and "Some u = snd (f l1 u1)"
-    unfolding f by auto
-  thus ?thesis
-    unfolding Sa[symmetric] lift_un.simps using Pa[OF Sa] by auto
-qed
-
-lemma lift_un:
-  assumes lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f"
-    and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>
-        l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-      (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")
-  shows "\<exists>l1 u1. (l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u1) \<and>
-                  Some l = fst (f l1 u1) \<and> Some u = snd (f l1 u1)"
-proof -
-  have Pa: "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-    if "Some (l, u) = approx' prec a bs" for l u
-    using approx_approx'[of prec a bs, OF _ that] Pa by auto
-  from lift_un_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_un_Some, OF Pa]
-  show ?thesis by auto
-qed
-
-lemma lift_un_bnds:
-  assumes bnds: "\<forall>(x::real) lx ux. (Some l, Some u) = f lx ux \<and> x \<in> { lx .. ux } \<longrightarrow> l \<le> f' x \<and> f' x \<le> u"
-    and lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f"
-    and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>
-      l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"
-  shows "real_of_float l \<le> f' (interpret_floatarith a xs) \<and> f' (interpret_floatarith a xs) \<le> real_of_float u"
-proof -
-  from lift_un[OF lift_un_Some Pa]
-  obtain l1 u1 where "l1 \<le> interpret_floatarith a xs"
-    and "interpret_floatarith a xs \<le> u1"
-    and "Some l = fst (f l1 u1)"
-    and "Some u = snd (f l1 u1)"
-    by blast
-  hence "(Some l, Some u) = f l1 u1" and "interpret_floatarith a xs \<in> {l1 .. u1}"
-    by auto
-  thus ?thesis
-    using bnds by auto
+  obtain ivl' where S: "approx prec a vs = Some ivl'" and ivl_def: "ivl = round_interval prec ivl'"
+    using approx' unfolding approx'.simps by (cases "approx prec a vs") auto
+  show ?thesis
+    by (auto simp: ivl_def intro!: in_round_intervalI Pa[OF S])
 qed
 
 lemma approx:
   assumes "bounded_by xs vs"
-    and "Some (l, u) = approx prec arith vs" (is "_ = ?g arith")
-  shows "l \<le> interpret_floatarith arith xs \<and> interpret_floatarith arith xs \<le> u" (is "?P l u arith")
-  using \<open>Some (l, u) = approx prec arith vs\<close>
-proof (induct arith arbitrary: l u)
-  case (Add a b)
-  from lift_bin'[OF Add.prems[unfolded approx.simps]] Add.hyps
-  obtain l1 u1 l2 u2 where "l = float_plus_down prec l1 l2"
-    and "u = float_plus_up prec u1 u2" "l1 \<le> interpret_floatarith a xs"
-    and "interpret_floatarith a xs \<le> u1" "l2 \<le> interpret_floatarith b xs"
-    and "interpret_floatarith b xs \<le> u2"
-    unfolding fst_conv snd_conv by blast
-  thus ?case
-    unfolding interpret_floatarith.simps by (auto intro!: float_plus_up_le float_plus_down_le)
-next
-  case (Minus a)
-  from lift_un'[OF Minus.prems[unfolded approx.simps]] Minus.hyps
-  obtain l1 u1 where "l = -u1" "u = -l1"
-    and "l1 \<le> interpret_floatarith a xs" "interpret_floatarith a xs \<le> u1"
-    unfolding fst_conv snd_conv by blast
-  thus ?case
-    unfolding interpret_floatarith.simps using minus_float.rep_eq by auto
-next
-  case (Mult a b)
-  from lift_bin'[OF Mult.prems[unfolded approx.simps]] Mult.hyps
-  obtain l1 u1 l2 u2
-    where l: "l = fst (bnds_mult prec l1 u1 l2 u2)"
-    and u: "u = snd (bnds_mult prec l1 u1 l2 u2)"
-    and a: "l1 \<le> interpret_floatarith a xs" "interpret_floatarith a xs \<le> u1"
-    and b: "l2 \<le> interpret_floatarith b xs" "interpret_floatarith b xs \<le> u2" unfolding fst_conv snd_conv by blast
-  from l u have lu: "(l, u) = bnds_mult prec l1 u1 l2 u2" by simp
-  from bnds_mult[OF lu] a b show ?case by simp
-next
-  case (Inverse a)
-  from lift_un[OF Inverse.prems[unfolded approx.simps], unfolded if_distrib[of fst] if_distrib[of snd] fst_conv snd_conv] Inverse.hyps
-  obtain l1 u1 where l': "Some l = (if 0 < l1 \<or> u1 < 0 then Some (float_divl prec 1 u1) else None)"
-    and u': "Some u = (if 0 < l1 \<or> u1 < 0 then Some (float_divr prec 1 l1) else None)"
-    and l1: "l1 \<le> interpret_floatarith a xs"
-    and u1: "interpret_floatarith a xs \<le> u1"
-    by blast
-  have either: "0 < l1 \<or> u1 < 0"
-  proof (rule ccontr)
-    assume P: "\<not> (0 < l1 \<or> u1 < 0)"
-    show False
-      using l' unfolding if_not_P[OF P] by auto
-  qed
-  moreover have l1_le_u1: "real_of_float l1 \<le> real_of_float u1"
-    using l1 u1 by auto
-  ultimately have "real_of_float l1 \<noteq> 0" and "real_of_float u1 \<noteq> 0"
-    by auto
-
-  have inv: "inverse u1 \<le> inverse (interpret_floatarith a xs)
-           \<and> inverse (interpret_floatarith a xs) \<le> inverse l1"
-  proof (cases "0 < l1")
-    case True
-    hence "0 < real_of_float u1" and "0 < real_of_float l1" "0 < interpret_floatarith a xs"
-      using l1_le_u1 l1 by auto
-    show ?thesis
-      unfolding inverse_le_iff_le[OF \<open>0 < real_of_float u1\<close> \<open>0 < interpret_floatarith a xs\<close>]
-        inverse_le_iff_le[OF \<open>0 < interpret_floatarith a xs\<close> \<open>0 < real_of_float l1\<close>]
-      using l1 u1 by auto
-  next
-    case False
-    hence "u1 < 0"
-      using either by blast
-    hence "real_of_float u1 < 0" and "real_of_float l1 < 0" "interpret_floatarith a xs < 0"
-      using l1_le_u1 u1 by auto
-    show ?thesis
-      unfolding inverse_le_iff_le_neg[OF \<open>real_of_float u1 < 0\<close> \<open>interpret_floatarith a xs < 0\<close>]
-        inverse_le_iff_le_neg[OF \<open>interpret_floatarith a xs < 0\<close> \<open>real_of_float l1 < 0\<close>]
-      using l1 u1 by auto
-  qed
-
-  from l' have "l = float_divl prec 1 u1"
-    by (cases "0 < l1 \<or> u1 < 0") auto
-  hence "l \<le> inverse u1"
-    unfolding nonzero_inverse_eq_divide[OF \<open>real_of_float u1 \<noteq> 0\<close>]
-    using float_divl[of prec 1 u1] by auto
-  also have "\<dots> \<le> inverse (interpret_floatarith a xs)"
-    using inv by auto
-  finally have "l \<le> inverse (interpret_floatarith a xs)" .
-  moreover
-  from u' have "u = float_divr prec 1 l1"
-    by (cases "0 < l1 \<or> u1 < 0") auto
-  hence "inverse l1 \<le> u"
-    unfolding nonzero_inverse_eq_divide[OF \<open>real_of_float l1 \<noteq> 0\<close>]
-    using float_divr[of 1 l1 prec] by auto
-  hence "inverse (interpret_floatarith a xs) \<le> u"
-    by (rule order_trans[OF inv[THEN conjunct2]])
-  ultimately show ?case
-    unfolding interpret_floatarith.simps using l1 u1 by auto
-next
-  case (Abs x)
-  from lift_un'[OF Abs.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Abs.hyps
-  obtain l1 u1 where l': "l = (if l1 < 0 \<and> 0 < u1 then 0 else min \<bar>l1\<bar> \<bar>u1\<bar>)"
-    and u': "u = max \<bar>l1\<bar> \<bar>u1\<bar>"
-    and l1: "l1 \<le> interpret_floatarith x xs"
-    and u1: "interpret_floatarith x xs \<le> u1"
-    by blast
-  thus ?case
-    unfolding l' u'
-    by (cases "l1 < 0 \<and> 0 < u1") (auto simp add: real_of_float_min real_of_float_max)
-next
-  case (Min a b)
-  from lift_bin'[OF Min.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Min.hyps
-  obtain l1 u1 l2 u2 where l': "l = min l1 l2" and u': "u = min u1 u2"
-    and l1: "l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> u1"
-    and l1: "l2 \<le> interpret_floatarith b xs" and u1: "interpret_floatarith b xs \<le> u2"
-    by blast
-  thus ?case
-    unfolding l' u' by (auto simp add: real_of_float_min)
-next
-  case (Max a b)
-  from lift_bin'[OF Max.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Max.hyps
-  obtain l1 u1 l2 u2 where l': "l = max l1 l2" and u': "u = max u1 u2"
-    and l1: "l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> u1"
-    and l1: "l2 \<le> interpret_floatarith b xs" and u1: "interpret_floatarith b xs \<le> u2"
-    by blast
-  thus ?case
-    unfolding l' u' by (auto simp add: real_of_float_max)
-next
-  case (Cos a)
-  with lift_un'_bnds[OF bnds_cos] show ?case by auto
-next
-  case (Arctan a)
-  with lift_un'_bnds[OF bnds_arctan] show ?case by auto
-next
-  case Pi
-  with pi_boundaries show ?case by auto
-next
-  case (Sqrt a)
-  with lift_un'_bnds[OF bnds_sqrt] show ?case by auto
-next
-  case (Exp a)
-  with lift_un'_bnds[OF bnds_exp] show ?case by auto
-next
-  case (Powr a b)
-  from lift_bin[OF Powr.prems[unfolded approx.simps]] Powr.hyps
-    obtain l1 u1 l2 u2 where lu: "Some (l, u) = bnds_powr prec l1 u1 l2 u2"
-      and l1: "l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> u1"
-      and l2: "l2 \<le> interpret_floatarith b xs" and u2: "interpret_floatarith b xs \<le> u2"
-      by blast
-  from bnds_powr[OF lu] l1 u1 l2 u2
-    show ?case by simp
-next
-  case (Ln a)
-  with lift_un_bnds[OF bnds_ln] show ?case by auto
-next
-  case (Power a n)
-  with lift_un'_bnds[OF bnds_power] show ?case by auto
-next
-  case (Floor a)
-  from lift_un'[OF Floor.prems[unfolded approx.simps] Floor.hyps]
-  show ?case by (auto simp: floor_fl.rep_eq floor_mono)
-next
-  case (Num f)
-  thus ?case by auto
-next
-  case (Var n)
-  from this[symmetric] \<open>bounded_by xs vs\<close>[THEN bounded_byE, of n]
-  show ?case by (cases "n < length vs") auto
-qed
+    and "approx prec arith vs = Some ivl"
+  shows "interpret_floatarith arith xs \<in>\<^sub>r ivl"
+  using \<open>approx prec arith vs = Some ivl\<close>
+  using  \<open>bounded_by xs vs\<close>[THEN bounded_byE]
+  by (induct arith arbitrary: ivl)
+    (force split: option.splits if_splits
+      intro!: plus_in_float_intervalI mult_float_intervalI uminus_in_float_intervalI
+      inverse_float_interval_eqI abs_in_float_intervalI
+      min_intervalI max_intervalI
+      cos_float_intervalI
+      arctan_float_intervalI pi_float_interval sqrt_float_intervalI exp_float_intervalI
+      powr_float_interval_eqI ln_float_interval_eqI power_float_intervalI floor_float_intervalI
+      intro: in_round_intervalI)+
 
 datatype form = Bound floatarith floatarith floatarith form
               | Assign floatarith floatarith form
@@ -639,30 +178,32 @@
 "interpret_form (Conj f g) vs \<longleftrightarrow> interpret_form f vs \<and> interpret_form g vs" |
 "interpret_form (Disj f g) vs \<longleftrightarrow> interpret_form f vs \<or> interpret_form g vs"
 
-fun approx_form' and approx_form :: "nat \<Rightarrow> form \<Rightarrow> (float * float) option list \<Rightarrow> nat list \<Rightarrow> bool" where
-"approx_form' prec f 0 n l u bs ss = approx_form prec f (bs[n := Some (l, u)]) ss" |
-"approx_form' prec f (Suc s) n l u bs ss =
-  (let m = (l + u) * Float 1 (- 1)
-   in (if approx_form' prec f s n l m bs ss then approx_form' prec f s n m u bs ss else False))" |
+fun approx_form' and approx_form :: "nat \<Rightarrow> form \<Rightarrow> (float interval) option list \<Rightarrow> nat list \<Rightarrow> bool" where
+"approx_form' prec f 0 n ivl bs ss = approx_form prec f (bs[n := Some ivl]) ss" |
+"approx_form' prec f (Suc s) n ivl bs ss =
+  (let (ivl1, ivl2) = split_float_interval ivl
+   in (if approx_form' prec f s n ivl1 bs ss then approx_form' prec f s n ivl2 bs ss else False))" |
 "approx_form prec (Bound (Var n) a b f) bs ss =
    (case (approx prec a bs, approx prec b bs)
-   of (Some (l, _), Some (_, u)) \<Rightarrow> approx_form' prec f (ss ! n) n l u bs ss
+   of (Some ivl1, Some ivl2) \<Rightarrow> approx_form' prec f (ss ! n) n (sup ivl1 ivl2) bs ss
     | _ \<Rightarrow> False)" |
 "approx_form prec (Assign (Var n) a f) bs ss =
    (case (approx prec a bs)
-   of (Some (l, u)) \<Rightarrow> approx_form' prec f (ss ! n) n l u bs ss
+   of (Some ivl) \<Rightarrow> approx_form' prec f (ss ! n) n ivl bs ss
     | _ \<Rightarrow> False)" |
 "approx_form prec (Less a b) bs ss =
    (case (approx prec a bs, approx prec b bs)
-   of (Some (l, u), Some (l', u')) \<Rightarrow> float_plus_up prec u (-l') < 0
+   of (Some ivl, Some ivl') \<Rightarrow> float_plus_up prec (upper ivl) (-lower ivl') < 0
     | _ \<Rightarrow> False)" |
 "approx_form prec (LessEqual a b) bs ss =
    (case (approx prec a bs, approx prec b bs)
-   of (Some (l, u), Some (l', u')) \<Rightarrow> float_plus_up prec u (-l') \<le> 0
+   of (Some ivl, Some ivl') \<Rightarrow> float_plus_up prec (upper ivl) (-lower ivl') \<le> 0
     | _ \<Rightarrow> False)" |
 "approx_form prec (AtLeastAtMost x a b) bs ss =
    (case (approx prec x bs, approx prec a bs, approx prec b bs)
-   of (Some (lx, ux), Some (l, u), Some (l', u')) \<Rightarrow> float_plus_up prec u (-lx) \<le> 0 \<and> float_plus_up prec ux (-l') \<le> 0
+   of (Some ivlx, Some ivl, Some ivl') \<Rightarrow>
+      float_plus_up prec (upper ivl) (-lower ivlx) \<le> 0 \<and>
+      float_plus_up prec (upper ivlx) (-lower ivl') \<le> 0
     | _ \<Rightarrow> False)" |
 "approx_form prec (Conj a b) bs ss \<longleftrightarrow> approx_form prec a bs ss \<and> approx_form prec b bs ss" |
 "approx_form prec (Disj a b) bs ss \<longleftrightarrow> approx_form prec a bs ss \<or> approx_form prec b bs ss" |
@@ -671,32 +212,32 @@
 lemma lazy_conj: "(if A then B else False) = (A \<and> B)" by simp
 
 lemma approx_form_approx_form':
-  assumes "approx_form' prec f s n l u bs ss"
-    and "(x::real) \<in> { l .. u }"
-  obtains l' u' where "x \<in> { l' .. u' }"
-    and "approx_form prec f (bs[n := Some (l', u')]) ss"
-using assms proof (induct s arbitrary: l u)
+  assumes "approx_form' prec f s n ivl bs ss"
+    and "(x::real) \<in>\<^sub>r ivl"
+  obtains ivl' where "x \<in>\<^sub>r ivl'"
+    and "approx_form prec f (bs[n := Some ivl']) ss"
+using assms proof (induct s arbitrary: ivl)
   case 0
-  from this(1)[of l u] this(2,3)
+  from this(1)[of ivl] this(2,3)
   show thesis by auto
 next
   case (Suc s)
-
-  let ?m = "(l + u) * Float 1 (- 1)"
-  have "real_of_float l \<le> ?m" and "?m \<le> real_of_float u"
-    unfolding less_eq_float_def using Suc.prems by auto
-
-  with \<open>x \<in> { l .. u }\<close>
-  have "x \<in> { l .. ?m} \<or> x \<in> { ?m .. u }" by auto
-  thus thesis
-  proof (rule disjE)
-    assume *: "x \<in> { l .. ?m }"
-    with Suc.hyps[OF _ _ *] Suc.prems
-    show thesis by (simp add: Let_def lazy_conj)
+  then obtain ivl1 ivl2 where ivl_split: "split_float_interval ivl = (ivl1, ivl2)"
+    by (fastforce dest: split_float_intervalD)
+  from split_float_interval_realD[OF this \<open>x \<in>\<^sub>r ivl\<close>]
+  consider "x \<in>\<^sub>r ivl1" | "x \<in>\<^sub>r ivl2"
+    by atomize_elim
+  then show thesis
+  proof cases
+    case *: 1
+    from Suc.hyps[OF _ _ *] Suc.prems
+    show ?thesis
+      by (simp add: lazy_conj ivl_split split: prod.splits)
   next
-    assume *: "x \<in> { ?m .. u }"
-    with Suc.hyps[OF _ _ *] Suc.prems
-    show thesis by (simp add: Let_def lazy_conj)
+    case *: 2
+    from Suc.hyps[OF _ _ *] Suc.prems
+    show ?thesis
+      by (simp add: lazy_conj ivl_split split: prod.splits)
   qed
 qed
 
@@ -709,23 +250,24 @@
   then obtain n
     where x_eq: "x = Var n" by (cases x) auto
 
-  with Bound.prems obtain l u' l' u
-    where l_eq: "Some (l, u') = approx prec a vs"
-    and u_eq: "Some (l', u) = approx prec b vs"
-    and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"
+  with Bound.prems obtain ivl1 ivl2
+    where l_eq: "approx prec a vs = Some ivl1"
+    and u_eq: "approx prec b vs = Some ivl2"
+    and approx_form': "approx_form' prec f (ss ! n) n (sup ivl1 ivl2) vs ss"
     by (cases "approx prec a vs", simp) (cases "approx prec b vs", auto)
 
   have "interpret_form f xs"
     if "xs ! n \<in> { interpret_floatarith a xs .. interpret_floatarith b xs }"
   proof -
     from approx[OF Bound.prems(2) l_eq] and approx[OF Bound.prems(2) u_eq] that
-    have "xs ! n \<in> { l .. u}" by auto
+    have "xs ! n \<in>\<^sub>r (sup ivl1 ivl2)"
+      by (auto simp: set_of_eq sup_float_def max_def inf_float_def min_def)
 
     from approx_form_approx_form'[OF approx_form' this]
-    obtain lx ux where bnds: "xs ! n \<in> { lx .. ux }"
-      and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .
+    obtain ivlx where bnds: "xs ! n \<in>\<^sub>r ivlx"
+      and approx_form: "approx_form prec f (vs[n := Some ivlx]) ss" .
 
-    from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some (lx, ux)])"
+    from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some ivlx])"
       by (rule bounded_by_update)
     with Bound.hyps[OF approx_form] show ?thesis
       by blast
@@ -737,22 +279,22 @@
   then obtain n where x_eq: "x = Var n"
     by (cases x) auto
 
-  with Assign.prems obtain l u
-    where bnd_eq: "Some (l, u) = approx prec a vs"
+  with Assign.prems obtain ivl
+    where bnd_eq: "approx prec a vs = Some ivl"
     and x_eq: "x = Var n"
-    and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"
+    and approx_form': "approx_form' prec f (ss ! n) n ivl vs ss"
     by (cases "approx prec a vs") auto
 
   have "interpret_form f xs"
     if bnds: "xs ! n = interpret_floatarith a xs"
   proof -
     from approx[OF Assign.prems(2) bnd_eq] bnds
-    have "xs ! n \<in> { l .. u}" by auto
+    have "xs ! n \<in>\<^sub>r ivl" by auto
     from approx_form_approx_form'[OF approx_form' this]
-    obtain lx ux where bnds: "xs ! n \<in> { lx .. ux }"
-      and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .
+    obtain ivlx where bnds: "xs ! n \<in>\<^sub>r ivlx"
+      and approx_form: "approx_form prec f (vs[n := Some ivlx]) ss" .
 
-    from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some (lx, ux)])"
+    from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some ivlx])"
       by (rule bounded_by_update)
     with Assign.hyps[OF approx_form] show ?thesis
       by blast
@@ -761,37 +303,38 @@
     using interpret_form.simps x_eq and interpret_floatarith.simps by simp
 next
   case (Less a b)
-  then obtain l u l' u'
-    where l_eq: "Some (l, u) = approx prec a vs"
-      and u_eq: "Some (l', u') = approx prec b vs"
-      and inequality: "real_of_float (float_plus_up prec u (-l')) < 0"
+  then obtain ivl ivl'
+    where l_eq: "approx prec a vs = Some ivl"
+      and u_eq: "approx prec b vs = Some ivl'"
+      and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivl')) < 0"
     by (cases "approx prec a vs", auto, cases "approx prec b vs", auto)
   from le_less_trans[OF float_plus_up inequality]
     approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq]
-  show ?case by auto
+  show ?case by (auto simp: set_of_eq)
 next
   case (LessEqual a b)
-  then obtain l u l' u'
-    where l_eq: "Some (l, u) = approx prec a vs"
-      and u_eq: "Some (l', u') = approx prec b vs"
-      and inequality: "real_of_float (float_plus_up prec u (-l')) \<le> 0"
+  then obtain ivl ivl'
+    where l_eq: "approx prec a vs = Some ivl"
+      and u_eq: "approx prec b vs = Some ivl'"
+      and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivl')) \<le> 0"
     by (cases "approx prec a vs", auto, cases "approx prec b vs", auto)
   from order_trans[OF float_plus_up inequality]
     approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq]
-  show ?case by auto
+  show ?case by (auto simp: set_of_eq)
 next
   case (AtLeastAtMost x a b)
-  then obtain lx ux l u l' u'
-    where x_eq: "Some (lx, ux) = approx prec x vs"
-    and l_eq: "Some (l, u) = approx prec a vs"
-    and u_eq: "Some (l', u') = approx prec b vs"
-    and inequality: "real_of_float (float_plus_up prec u (-lx)) \<le> 0" "real_of_float (float_plus_up prec ux (-l')) \<le> 0"
+  then obtain ivlx ivl ivl'
+    where x_eq: "approx prec x vs = Some ivlx"
+      and l_eq: "approx prec a vs = Some ivl"
+      and u_eq: "approx prec b vs = Some ivl'"
+    and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivlx)) \<le> 0"
+      "real_of_float (float_plus_up prec (upper ivlx) (-lower ivl')) \<le> 0"
     by (cases "approx prec x vs", auto,
       cases "approx prec a vs", auto,
       cases "approx prec b vs", auto)
   from order_trans[OF float_plus_up inequality(1)] order_trans[OF float_plus_up inequality(2)]
     approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq]
-  show ?case by auto
+  show ?case by (auto simp: set_of_eq)
 qed auto
 
 lemma approx_form:
@@ -907,12 +450,12 @@
 
 declare approx.simps[simp del]
 
-fun isDERIV_approx :: "nat \<Rightarrow> nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> bool" where
+fun isDERIV_approx :: "nat \<Rightarrow> nat \<Rightarrow> floatarith \<Rightarrow> (float interval) option list \<Rightarrow> bool" where
 "isDERIV_approx prec x (Add a b) vs         = (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs)" |
 "isDERIV_approx prec x (Mult a b) vs        = (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs)" |
 "isDERIV_approx prec x (Minus a) vs         = isDERIV_approx prec x a vs" |
 "isDERIV_approx prec x (Inverse a) vs       =
-  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l \<or> u < 0 | None \<Rightarrow> False))" |
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl \<or> upper ivl < 0 | None \<Rightarrow> False))" |
 "isDERIV_approx prec x (Cos a) vs           = isDERIV_approx prec x a vs" |
 "isDERIV_approx prec x (Arctan a) vs        = isDERIV_approx prec x a vs" |
 "isDERIV_approx prec x (Min a b) vs         = False" |
@@ -920,15 +463,15 @@
 "isDERIV_approx prec x (Abs a) vs           = False" |
 "isDERIV_approx prec x Pi vs                = True" |
 "isDERIV_approx prec x (Sqrt a) vs          =
-  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl | None \<Rightarrow> False))" |
 "isDERIV_approx prec x (Exp a) vs           = isDERIV_approx prec x a vs" |
 "isDERIV_approx prec x (Powr a b) vs        =
-  (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |
+  (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl | None \<Rightarrow> False))" |
 "isDERIV_approx prec x (Ln a) vs            =
-  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl | None \<Rightarrow> False))" |
 "isDERIV_approx prec x (Power a 0) vs       = True" |
 "isDERIV_approx prec x (Floor a) vs         =
-  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> l > floor u \<and> u < ceiling l | None \<Rightarrow> False))" |
+  (isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> lower ivl > floor (upper ivl) \<and> upper ivl < ceiling (lower ivl) | None \<Rightarrow> False))" |
 "isDERIV_approx prec x (Power a (Suc n)) vs = isDERIV_approx prec x a vs" |
 "isDERIV_approx prec x (Num f) vs           = True" |
 "isDERIV_approx prec x (Var n) vs           = True"
@@ -940,88 +483,65 @@
   using isDERIV_approx
 proof (induct f)
   case (Inverse a)
-  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
-    and *: "0 < l \<or> u < 0"
+  then obtain ivl where approx_Some: "approx prec a vs = Some ivl"
+    and *: "0 < lower ivl \<or> upper ivl < 0"
     by (cases "approx prec a vs") auto
   with approx[OF \<open>bounded_by xs vs\<close> approx_Some]
-  have "interpret_floatarith a xs \<noteq> 0" by auto
+  have "interpret_floatarith a xs \<noteq> 0" by (auto simp: set_of_eq)
   thus ?case using Inverse by auto
 next
   case (Ln a)
-  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
-    and *: "0 < l"
+  then obtain ivl where approx_Some: "approx prec a vs = Some ivl"
+    and *: "0 < lower ivl"
     by (cases "approx prec a vs") auto
   with approx[OF \<open>bounded_by xs vs\<close> approx_Some]
-  have "0 < interpret_floatarith a xs" by auto
+  have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq)
   thus ?case using Ln by auto
 next
   case (Sqrt a)
-  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
-    and *: "0 < l"
+  then obtain ivl where approx_Some: "approx prec a vs = Some ivl"
+    and *: "0 < lower ivl"
     by (cases "approx prec a vs") auto
   with approx[OF \<open>bounded_by xs vs\<close> approx_Some]
-  have "0 < interpret_floatarith a xs" by auto
+  have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq)
   thus ?case using Sqrt by auto
 next
   case (Power a n)
   thus ?case by (cases n) auto
 next
   case (Powr a b)
-  from Powr obtain l1 u1 where a: "Some (l1, u1) = approx prec a vs" and pos: "0 < l1"
+  from Powr obtain ivl1 where a: "approx prec a vs = Some ivl1"
+    and pos: "0 < lower ivl1"
     by (cases "approx prec a vs") auto
   with approx[OF \<open>bounded_by xs vs\<close> a]
-    have "0 < interpret_floatarith a xs" by auto
+  have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq)
   with Powr show ?case by auto
 next
   case (Floor a)
-  then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"
-    and "real_of_int \<lfloor>real_of_float u\<rfloor> < real_of_float l" "real_of_float u < real_of_int \<lceil>real_of_float l\<rceil>"
+  then obtain ivl where approx_Some: "approx prec a vs = Some ivl"
+    and "real_of_int \<lfloor>real_of_float (upper ivl)\<rfloor> < real_of_float (lower ivl)"
+      "real_of_float (upper ivl) < real_of_int \<lceil>real_of_float (lower ivl)\<rceil>"
     and "isDERIV x a xs"
     by (cases "approx prec a vs") auto
   with approx[OF \<open>bounded_by xs vs\<close> approx_Some] le_floor_iff
   show ?case
-    by (force elim!: Ints_cases)
+    by (force elim!: Ints_cases simp: set_of_eq)
 qed auto
 
 lemma bounded_by_update_var:
   assumes "bounded_by xs vs"
-    and "vs ! i = Some (l, u)"
-    and bnd: "x \<in> { real_of_float l .. real_of_float u }"
+    and "vs ! i = Some ivl"
+    and bnd: "x \<in>\<^sub>r ivl"
   shows "bounded_by (xs[i := x]) vs"
-proof (cases "i < length xs")
-  case False
-  thus ?thesis
-    using \<open>bounded_by xs vs\<close> by auto
-next
-  case True
-  let ?xs = "xs[i := x]"
-  from True have "i < length ?xs" by auto
-  have "case vs ! j of None \<Rightarrow> True | Some (l, u) \<Rightarrow> ?xs ! j \<in> {real_of_float l .. real_of_float u}"
-    if "j < length vs" for j
-  proof (cases "vs ! j")
-    case None
-    then show ?thesis by simp
-  next
-    case (Some b)
-    thus ?thesis
-    proof (cases "i = j")
-      case True
-      thus ?thesis using \<open>vs ! i = Some (l, u)\<close> Some and bnd \<open>i < length ?xs\<close>
-        by auto
-    next
-      case False
-      thus ?thesis
-        using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>j < length vs\<close>] Some by auto
-    qed
-  qed
-  thus ?thesis
-    unfolding bounded_by_def by auto
-qed
+  using assms
+  using nth_list_update
+  by (cases "i < length xs")
+    (force simp: bounded_by_def  split: option.splits)+
 
 lemma isDERIV_approx':
   assumes "bounded_by xs vs"
-    and vs_x: "vs ! x = Some (l, u)"
-    and X_in: "X \<in> {real_of_float l .. real_of_float u}"
+    and vs_x: "vs ! x = Some ivl"
+    and X_in: "X \<in>\<^sub>r ivl"
     and approx: "isDERIV_approx prec x f vs"
   shows "isDERIV x f (xs[x := X])"
 proof -
@@ -1033,14 +553,14 @@
   assumes "n < length xs"
     and bnd: "bounded_by xs vs"
     and isD: "isDERIV_approx prec n f vs"
-    and app: "Some (l, u) = approx prec (DERIV_floatarith n f) vs" (is "_ = approx _ ?D _")
-  shows "\<exists>(x::real). l \<le> x \<and> x \<le> u \<and>
+    and app: "approx prec (DERIV_floatarith n f) vs = Some ivl" (is "approx _ ?D _ = _")
+  shows "\<exists>(x::real). x \<in>\<^sub>r ivl \<and>
              DERIV (\<lambda> x. interpret_floatarith f (xs[n := x])) (xs!n) :> x"
-         (is "\<exists> x. _ \<and> _ \<and> DERIV (?i f) _ :> _")
-proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI[OF _ conjI])
+         (is "\<exists> x. _ \<and> DERIV (?i f) _ :> _")
+proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI)
   let "?i f" = "\<lambda>x. interpret_floatarith f (xs[n := x])"
   from approx[OF bnd app]
-  show "l \<le> ?i ?D (xs!n)" and "?i ?D (xs!n) \<le> u"
+  show "?i ?D (xs!n) \<in>\<^sub>r ivl"
     using \<open>n < length xs\<close> by auto
   from DERIV_floatarith[OF \<open>n < length xs\<close>, of f "xs!n"] isDERIV_approx[OF bnd isD]
   show "DERIV (?i f) (xs!n) :> (?i ?D (xs!n))"
@@ -1048,11 +568,11 @@
 qed
 
 lemma lift_bin_aux:
-  assumes lift_bin_Some: "Some (l, u) = lift_bin a b f"
-  obtains l1 u1 l2 u2
-  where "a = Some (l1, u1)"
-    and "b = Some (l2, u2)"
-    and "f l1 u1 l2 u2 = Some (l, u)"
+  assumes lift_bin_Some: "lift_bin a b f = Some ivl"
+  obtains ivl1 ivl2
+  where "a = Some ivl1"
+    and "b = Some ivl2"
+    and "f ivl1 ivl2 = Some ivl"
   using assms by (cases a, simp, cases b, simp, auto)
 
 
@@ -1060,22 +580,22 @@
 "approx_tse prec n 0 c k f bs = approx prec f bs" |
 "approx_tse prec n (Suc s) c k f bs =
   (if isDERIV_approx prec n f bs then
-    lift_bin (approx prec f (bs[n := Some (c,c)]))
+    lift_bin (approx prec f (bs[n := Some (interval_of c)]))
              (approx_tse prec n s c (Suc k) (DERIV_floatarith n f) bs)
-             (\<lambda> l1 u1 l2 u2. approx prec
+             (\<lambda>ivl1 ivl2. approx prec
                  (Add (Var 0)
                       (Mult (Inverse (Num (Float (int k) 0)))
                                  (Mult (Add (Var (Suc (Suc 0))) (Minus (Num c)))
-                                       (Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), bs!n])
+                                       (Var (Suc 0))))) [Some ivl1, Some ivl2, bs!n])
   else approx prec f bs)"
 
 lemma bounded_by_Cons:
   assumes bnd: "bounded_by xs vs"
-    and x: "x \<in> { real_of_float l .. real_of_float u }"
-  shows "bounded_by (x#xs) ((Some (l, u))#vs)"
+    and x: "x \<in>\<^sub>r ivl"
+  shows "bounded_by (x#xs) ((Some ivl)#vs)"
 proof -
-  have "case ((Some (l,u))#vs) ! i of Some (l, u) \<Rightarrow> (x#xs)!i \<in> { real_of_float l .. real_of_float u } | None \<Rightarrow> True"
-    if *: "i < length ((Some (l, u))#vs)" for i
+  have "case ((Some ivl)#vs) ! i of Some ivl \<Rightarrow> (x#xs)!i \<in>\<^sub>r ivl | None \<Rightarrow> True"
+    if *: "i < length ((Some ivl)#vs)" for i
   proof (cases i)
     case 0
     with x show ?thesis by auto
@@ -1093,25 +613,25 @@
   assumes "bounded_by xs vs"
     and bnd_c: "bounded_by (xs[x := c]) vs"
     and "x < length vs" and "x < length xs"
-    and bnd_x: "vs ! x = Some (lx, ux)"
-    and ate: "Some (l, u) = approx_tse prec x s c k f vs"
-  shows "\<exists> n. (\<forall> m < n. \<forall> (z::real) \<in> {lx .. ux}.
+    and bnd_x: "vs ! x = Some ivlx"
+    and ate: "approx_tse prec x s c k f vs = Some ivl"
+  shows "\<exists> n. (\<forall> m < n. \<forall> (z::real) \<in> set_of (real_interval ivlx).
       DERIV (\<lambda> y. interpret_floatarith ((DERIV_floatarith x ^^ m) f) (xs[x := y])) z :>
             (interpret_floatarith ((DERIV_floatarith x ^^ (Suc m)) f) (xs[x := z])))
-   \<and> (\<forall> (t::real) \<in> {lx .. ux}.  (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) *
+   \<and> (\<forall> (t::real) \<in> set_of (real_interval ivlx).  (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) *
                   interpret_floatarith ((DERIV_floatarith x ^^ i) f) (xs[x := c]) *
                   (xs!x - c)^i) +
       inverse (real (\<Prod> j \<in> {k..<k+n}. j)) *
       interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := t]) *
-      (xs!x - c)^n \<in> {l .. u})" (is "\<exists> n. ?taylor f k l u n")
+      (xs!x - c)^n \<in>\<^sub>r ivl)" (is "\<exists> n. ?taylor f k ivl n")
   using ate
-proof (induct s arbitrary: k f l u)
+proof (induct s arbitrary: k f ivl)
   case 0
   {
-    fix t::real assume "t \<in> {lx .. ux}"
+    fix t::real assume "t \<in>\<^sub>r ivlx"
     note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> bnd_x this]
     from approx[OF this 0[unfolded approx_tse.simps]]
-    have "(interpret_floatarith f (xs[x := t])) \<in> {l .. u}"
+    have "(interpret_floatarith f (xs[x := t])) \<in>\<^sub>r ivl"
       by (auto simp add: algebra_simps)
   }
   thus ?case by (auto intro!: exI[of _ 0])
@@ -1122,32 +642,32 @@
     case False
     note ap = Suc.prems[unfolded approx_tse.simps if_not_P[OF False]]
     {
-      fix t::real assume "t \<in> {lx .. ux}"
+      fix t::real assume "t \<in>\<^sub>r ivlx"
       note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> bnd_x this]
       from approx[OF this ap]
-      have "(interpret_floatarith f (xs[x := t])) \<in> {l .. u}"
+      have "(interpret_floatarith f (xs[x := t])) \<in>\<^sub>r ivl"
         by (auto simp add: algebra_simps)
     }
     thus ?thesis by (auto intro!: exI[of _ 0])
   next
     case True
     with Suc.prems
-    obtain l1 u1 l2 u2
-      where a: "Some (l1, u1) = approx prec f (vs[x := Some (c,c)])"
-        and ate: "Some (l2, u2) = approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs"
-        and final: "Some (l, u) = approx prec
+    obtain ivl1 ivl2
+      where a: "approx prec f (vs[x := Some (interval_of c)]) = Some ivl1"
+        and ate: "approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs = Some ivl2"
+        and final: "approx prec
           (Add (Var 0)
                (Mult (Inverse (Num (Float (int k) 0)))
                      (Mult (Add (Var (Suc (Suc 0))) (Minus (Num c)))
-                           (Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), vs!x]"
+                           (Var (Suc 0))))) [Some ivl1, Some ivl2, vs!x] = Some ivl"
       by (auto elim!: lift_bin_aux)
 
     from bnd_c \<open>x < length xs\<close>
-    have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (c,c)])"
+    have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (interval_of c)])"
       by (auto intro!: bounded_by_update)
 
     from approx[OF this a]
-    have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \<in> { l1 .. u1 }"
+    have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \<in>\<^sub>r ivl1"
               (is "?f 0 (real_of_float c) \<in> _")
       by auto
 
@@ -1155,16 +675,16 @@
       for f :: "'a \<Rightarrow> 'a" and n :: nat and x :: 'a
       by (induct n) auto
     from Suc.hyps[OF ate, unfolded this] obtain n
-      where DERIV_hyp: "\<And>m z. \<lbrakk> m < n ; (z::real) \<in> { lx .. ux } \<rbrakk> \<Longrightarrow>
+      where DERIV_hyp: "\<And>m z. \<lbrakk> m < n ; (z::real) \<in>\<^sub>r ivlx \<rbrakk> \<Longrightarrow>
         DERIV (?f (Suc m)) z :> ?f (Suc (Suc m)) z"
-      and hyp: "\<forall>t \<in> {real_of_float lx .. real_of_float ux}.
+      and hyp: "\<forall>t \<in> set_of (real_interval ivlx).
         (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {Suc k..<Suc k + i}. j)) * ?f (Suc i) c * (xs!x - c)^i) +
-          inverse (real (\<Prod> j \<in> {Suc k..<Suc k + n}. j)) * ?f (Suc n) t * (xs!x - c)^n \<in> {l2 .. u2}"
+          inverse (real (\<Prod> j \<in> {Suc k..<Suc k + n}. j)) * ?f (Suc n) t * (xs!x - c)^n \<in>\<^sub>r ivl2"
           (is "\<forall> t \<in> _. ?X (Suc k) f n t \<in> _")
       by blast
 
     have DERIV: "DERIV (?f m) z :> ?f (Suc m) z"
-      if "m < Suc n" and bnd_z: "z \<in> { lx .. ux }" for m and z::real
+      if "m < Suc n" and bnd_z: "z \<in>\<^sub>r ivlx" for m and z::real
     proof (cases m)
       case 0
       with DERIV_floatarith[OF \<open>x < length xs\<close>
@@ -1187,16 +707,16 @@
     define C where "C = xs!x - c"
 
     {
-      fix t::real assume t: "t \<in> {lx .. ux}"
+      fix t::real assume t: "t \<in>\<^sub>r ivlx"
       hence "bounded_by [xs!x] [vs!x]"
         using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>]
         by (cases "vs!x", auto simp add: bounded_by_def)
 
       with hyp[THEN bspec, OF t] f_c
-      have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some (l1, u1), Some (l2, u2), vs!x]"
+      have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some ivl1, Some ivl2, vs!x]"
         by (auto intro!: bounded_by_Cons)
       from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]]
-      have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \<in> {l .. u}"
+      have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \<in>\<^sub>r ivl"
         by (auto simp add: algebra_simps)
       also have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse (real k) + ?f 0 c =
                (\<Sum> i = 0..<Suc n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) * ?f i c * (xs!x - c)^i) +
@@ -1204,7 +724,7 @@
         unfolding funpow_Suc C_def[symmetric] sum_move0 prod_head_Suc
         by (auto simp add: algebra_simps)
           (simp only: mult.left_commute [of _ "inverse (real k)"] sum_distrib_left [symmetric])
-      finally have "?T \<in> {l .. u}" .
+      finally have "?T \<in>\<^sub>r ivl" .
     }
     thus ?thesis using DERIV by blast
   qed
@@ -1215,11 +735,11 @@
 
 lemma approx_tse:
   assumes "bounded_by xs vs"
-    and bnd_x: "vs ! x = Some (lx, ux)"
-    and bnd_c: "real_of_float c \<in> {lx .. ux}"
+    and bnd_x: "vs ! x = Some ivlx"
+    and bnd_c: "real_of_float c \<in>\<^sub>r ivlx"
     and "x < length vs" and "x < length xs"
-    and ate: "Some (l, u) = approx_tse prec x s c 1 f vs"
-  shows "interpret_floatarith f xs \<in> {l .. u}"
+    and ate: "approx_tse prec x s c 1 f vs = Some ivl"
+  shows "interpret_floatarith f xs \<in>\<^sub>r ivl"
 proof -
   define F where [abs_def]: "F n z =
     interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := z])" for n z
@@ -1231,15 +751,15 @@
 
   from approx_tse_generic[OF \<open>bounded_by xs vs\<close> this bnd_x ate]
   obtain n
-    where DERIV: "\<forall> m z. m < n \<and> real_of_float lx \<le> z \<and> z \<le> real_of_float ux \<longrightarrow> DERIV (F m) z :> F (Suc m) z"
-    and hyp: "\<And> (t::real). t \<in> {lx .. ux} \<Longrightarrow>
+    where DERIV: "\<forall> m z. m < n \<and> real_of_float (lower ivlx) \<le> z \<and> z \<le> real_of_float (upper ivlx) \<longrightarrow> DERIV (F m) z :> F (Suc m) z"
+    and hyp: "\<And> (t::real). t \<in>\<^sub>r ivlx \<Longrightarrow>
            (\<Sum> j = 0..<n. inverse(fact j) * F j c * (xs!x - c)^j) +
              inverse ((fact n)) * F n t * (xs!x - c)^n
-             \<in> {l .. u}" (is "\<And> t. _ \<Longrightarrow> ?taylor t \<in> _")
+             \<in>\<^sub>r ivl" (is "\<And> t. _ \<Longrightarrow> ?taylor t \<in> _")
     unfolding F_def atLeastAtMost_iff[symmetric] prod_fact
-    by blast
+    by (auto simp: set_of_eq Ball_def)
 
-  have bnd_xs: "xs ! x \<in> { lx .. ux }"
+  have bnd_xs: "xs ! x \<in>\<^sub>r ivlx"
     using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x by auto
 
   show ?thesis
@@ -1257,8 +777,9 @@
         by auto
     next
       case False
-      have "lx \<le> real_of_float c" "real_of_float c \<le> ux" "lx \<le> xs!x" "xs!x \<le> ux"
-        using Suc bnd_c \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x by auto
+      have "lower ivlx \<le> real_of_float c" "real_of_float c \<le> upper ivlx" "lower ivlx \<le> xs!x" "xs!x \<le> upper ivlx"
+        using Suc bnd_c \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x
+        by (auto simp: set_of_eq)
       from Taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False]
       obtain t::real where t_bnd: "if xs ! x < c then xs ! x < t \<and> t < c else c < t \<and> t < xs ! x"
         and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) =
@@ -1266,12 +787,12 @@
            F (Suc n') t / (fact (Suc n')) * (xs ! x - c) ^ Suc n'"
         unfolding atLeast0LessThan by blast
 
-      from t_bnd bnd_xs bnd_c have *: "t \<in> {lx .. ux}"
-        by (cases "xs ! x < c") auto
+      from t_bnd bnd_xs bnd_c have *: "t \<in>\<^sub>r ivlx"
+        by (cases "xs ! x < c") (auto simp: set_of_eq)
 
       have "interpret_floatarith f (xs[x := xs ! x]) = ?taylor t"
         unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse)
-      also have "\<dots> \<in> {l .. u}"
+      also have "\<dots> \<in>\<^sub>r ivl"
         using * by (rule hyp)
       finally show ?thesis
         by simp
@@ -1280,119 +801,122 @@
 qed
 
 fun approx_tse_form' where
-"approx_tse_form' prec t f 0 l u cmp =
-  (case approx_tse prec 0 t ((l + u) * Float 1 (- 1)) 1 f [Some (l, u)]
-     of Some (l, u) \<Rightarrow> cmp l u | None \<Rightarrow> False)" |
-"approx_tse_form' prec t f (Suc s) l u cmp =
-  (let m = (l + u) * Float 1 (- 1)
-   in (if approx_tse_form' prec t f s l m cmp then
-      approx_tse_form' prec t f s m u cmp else False))"
+"approx_tse_form' prec t f 0 ivl cmp =
+  (case approx_tse prec 0 t (mid ivl) 1 f [Some ivl]
+     of Some ivl \<Rightarrow> cmp ivl | None \<Rightarrow> False)" |
+"approx_tse_form' prec t f (Suc s) ivl cmp =
+  (let (ivl1, ivl2) = split_float_interval ivl
+   in (if approx_tse_form' prec t f s ivl1 cmp then
+      approx_tse_form' prec t f s ivl2 cmp else False))"
 
 lemma approx_tse_form':
   fixes x :: real
-  assumes "approx_tse_form' prec t f s l u cmp"
-    and "x \<in> {l .. u}"
-  shows "\<exists>l' u' ly uy. x \<in> {l' .. u'} \<and> real_of_float l \<le> l' \<and> u' \<le> real_of_float u \<and> cmp ly uy \<and>
-    approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)"
+  assumes "approx_tse_form' prec t f s ivl cmp"
+    and "x \<in>\<^sub>r ivl"
+  shows "\<exists>ivl' ivly. x \<in>\<^sub>r ivl' \<and> ivl' \<le> ivl \<and> cmp ivly \<and>
+    approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly"
   using assms
-proof (induct s arbitrary: l u)
+proof (induct s arbitrary: ivl)
   case 0
-  then obtain ly uy
-    where *: "approx_tse prec 0 t ((l + u) * Float 1 (- 1)) 1 f [Some (l, u)] = Some (ly, uy)"
-    and **: "cmp ly uy" by (auto elim!: case_optionE)
+  then obtain ivly
+    where *: "approx_tse prec 0 t (mid ivl) 1 f [Some ivl] = Some ivly"
+    and **: "cmp ivly" by (auto elim!: case_optionE)
   with 0 show ?case by auto
 next
   case (Suc s)
-  let ?m = "(l + u) * Float 1 (- 1)"
+  let ?ivl1 = "fst (split_float_interval ivl)"
+  let ?ivl2 = "snd (split_float_interval ivl)"
   from Suc.prems
-  have l: "approx_tse_form' prec t f s l ?m cmp"
-    and u: "approx_tse_form' prec t f s ?m u cmp"
+  have l: "approx_tse_form' prec t f s ?ivl1 cmp"
+    and u: "approx_tse_form' prec t f s ?ivl2 cmp"
     by (auto simp add: Let_def lazy_conj)
 
-  have m_l: "real_of_float l \<le> ?m" and m_u: "?m \<le> real_of_float u"
-    unfolding less_eq_float_def using Suc.prems by auto
-  with \<open>x \<in> { l .. u }\<close> consider "x \<in> { l .. ?m}" | "x \<in> {?m .. u}"
-    by atomize_elim auto
-  thus ?case
+  have subintervals: "?ivl1 \<le> ivl" "?ivl2 \<le> ivl"
+    using mid_le
+    by (auto simp: less_eq_interval_def split_float_interval_bounds)
+
+  from split_float_interval_realD[OF _ \<open>x \<in>\<^sub>r ivl\<close>] consider "x \<in>\<^sub>r ?ivl1" | "x \<in>\<^sub>r ?ivl2"
+    by (auto simp: prod_eq_iff)
+  then show ?case
   proof cases
     case 1
     from Suc.hyps[OF l this]
-    obtain l' u' ly uy where
-      "x \<in> {l' .. u'} \<and> real_of_float l \<le> l' \<and> real_of_float u' \<le> ?m \<and> cmp ly uy \<and>
-        approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)"
+    obtain ivl' ivly where
+      "x \<in>\<^sub>r ivl' \<and> ivl' \<le> fst (split_float_interval ivl) \<and> cmp ivly \<and>
+        approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly"
       by blast
-    with m_u show ?thesis
-      by (auto intro!: exI)
+    then show ?thesis
+      using subintervals
+      by (auto)
   next
     case 2
     from Suc.hyps[OF u this]
-    obtain l' u' ly uy where
-      "x \<in> { l' .. u' } \<and> ?m \<le> real_of_float l' \<and> u' \<le> real_of_float u \<and> cmp ly uy \<and>
-        approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)"
+    obtain ivl' ivly where
+      "x \<in>\<^sub>r ivl' \<and> ivl' \<le> snd (split_float_interval ivl) \<and> cmp ivly \<and>
+        approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly"
       by blast
-    with m_u show ?thesis
-      by (auto intro!: exI)
+    then show ?thesis
+      using subintervals
+      by (auto)
   qed
 qed
 
 lemma approx_tse_form'_less:
   fixes x :: real
-  assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\<lambda> l u. 0 < l)"
-    and x: "x \<in> {l .. u}"
+  assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s ivl (\<lambda> ivl. 0 < lower ivl)"
+    and x: "x \<in>\<^sub>r ivl"
   shows "interpret_floatarith b [x] < interpret_floatarith a [x]"
 proof -
   from approx_tse_form'[OF tse x]
-  obtain l' u' ly uy
-    where x': "x \<in> {l' .. u'}"
-    and "real_of_float l \<le> real_of_float l'"
-    and "real_of_float u' \<le> real_of_float u" and "0 < ly"
-    and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)"
+  obtain ivl' ivly
+    where x': "x \<in>\<^sub>r ivl'"
+    and "ivl' \<le> ivl" and "0 < lower ivly"
+    and tse: "approx_tse prec 0 t (mid ivl') 1 (Add a (Minus b)) [Some ivl'] = Some ivly"
     by blast
 
-  hence "bounded_by [x] [Some (l', u')]"
+  hence "bounded_by [x] [Some ivl']"
     by (auto simp add: bounded_by_def)
-  from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
-  have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
-    by auto
-  from order_less_le_trans[OF _ this, of 0] \<open>0 < ly\<close> show ?thesis
+  from approx_tse[OF this _ _ _ _ tse, of ivl'] x' mid_le
+  have "lower ivly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
+    by (auto simp: set_of_eq)
+  from order_less_le_trans[OF _ this, of 0] \<open>0 < lower ivly\<close> show ?thesis
     by auto
 qed
 
 lemma approx_tse_form'_le:
   fixes x :: real
-  assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\<lambda> l u. 0 \<le> l)"
-    and x: "x \<in> {l .. u}"
+  assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s ivl (\<lambda> ivl. 0 \<le> lower ivl)"
+    and x: "x \<in>\<^sub>r ivl"
   shows "interpret_floatarith b [x] \<le> interpret_floatarith a [x]"
 proof -
   from approx_tse_form'[OF tse x]
-  obtain l' u' ly uy
-    where x': "x \<in> {l' .. u'}"
-    and "l \<le> real_of_float l'"
-    and "real_of_float u' \<le> u" and "0 \<le> ly"
-    and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)"
+  obtain ivl' ivly
+    where x': "x \<in>\<^sub>r ivl'"
+    and "ivl' \<le> ivl" and "0 \<le> lower ivly"
+    and tse: "approx_tse prec 0 t (mid ivl') 1 (Add a (Minus b)) [Some (ivl')] = Some ivly"
     by blast
 
-  hence "bounded_by [x] [Some (l', u')]" by (auto simp add: bounded_by_def)
-  from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
-  have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
-    by auto
-  from order_trans[OF _ this, of 0] \<open>0 \<le> ly\<close> show ?thesis
+  hence "bounded_by [x] [Some ivl']" by (auto simp add: bounded_by_def)
+  from approx_tse[OF this _ _ _ _ tse, of ivl'] x' mid_le
+  have "lower ivly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
+    by (auto simp: set_of_eq)
+  from order_trans[OF _ this, of 0] \<open>0 \<le> lower ivly\<close> show ?thesis
     by auto
 qed
 
 fun approx_tse_concl where
-"approx_tse_concl prec t (Less lf rt) s l u l' u' \<longleftrightarrow>
-    approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)" |
-"approx_tse_concl prec t (LessEqual lf rt) s l u l' u' \<longleftrightarrow>
-    approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)" |
-"approx_tse_concl prec t (AtLeastAtMost x lf rt) s l u l' u' \<longleftrightarrow>
-    (if approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l) then
-      approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l) else False)" |
-"approx_tse_concl prec t (Conj f g) s l u l' u' \<longleftrightarrow>
-    approx_tse_concl prec t f s l u l' u' \<and> approx_tse_concl prec t g s l u l' u'" |
-"approx_tse_concl prec t (Disj f g) s l u l' u' \<longleftrightarrow>
-    approx_tse_concl prec t f s l u l' u' \<or> approx_tse_concl prec t g s l u l' u'" |
-"approx_tse_concl _ _ _ _ _ _ _ _ \<longleftrightarrow> False"
+"approx_tse_concl prec t (Less lf rt) s ivl ivl' \<longleftrightarrow>
+    approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 < lower ivl)" |
+"approx_tse_concl prec t (LessEqual lf rt) s ivl ivl' \<longleftrightarrow>
+    approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)" |
+"approx_tse_concl prec t (AtLeastAtMost x lf rt) s ivl ivl' \<longleftrightarrow>
+    (if approx_tse_form' prec t (Add x (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl) then
+      approx_tse_form' prec t (Add rt (Minus x)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl) else False)" |
+"approx_tse_concl prec t (Conj f g) s ivl ivl' \<longleftrightarrow>
+    approx_tse_concl prec t f s ivl ivl' \<and> approx_tse_concl prec t g s ivl ivl'" |
+"approx_tse_concl prec t (Disj f g) s ivl ivl' \<longleftrightarrow>
+    approx_tse_concl prec t f s ivl ivl' \<or> approx_tse_concl prec t g s ivl ivl'" |
+"approx_tse_concl _ _ _ _ _ _ \<longleftrightarrow> False"
 
 definition
   "approx_tse_form prec t s f =
@@ -1400,7 +924,7 @@
       Bound x a b f \<Rightarrow>
         x = Var 0 \<and>
         (case (approx prec a [None], approx prec b [None]) of
-          (Some (l, u), Some (l', u')) \<Rightarrow> approx_tse_concl prec t f s l u l' u'
+          (Some ivl, Some ivl') \<Rightarrow> approx_tse_concl prec t f s ivl ivl'
         | _ \<Rightarrow> False)
     | _ \<Rightarrow> False)"
 
@@ -1409,9 +933,9 @@
   shows "interpret_form f [x]"
 proof (cases f)
   case f_def: (Bound i a b f')
-  with assms obtain l u l' u'
-    where a: "approx prec a [None] = Some (l, u)"
-    and b: "approx prec b [None] = Some (l', u')"
+  with assms obtain ivl ivl'
+    where a: "approx prec a [None] = Some ivl"
+    and b: "approx prec b [None] = Some ivl'"
     unfolding approx_tse_form_def by (auto elim!: case_optionE)
 
   from f_def assms have "i = Var 0"
@@ -1421,30 +945,31 @@
   {
     let ?f = "\<lambda>z. interpret_floatarith z [x]"
     assume "?f i \<in> { ?f a .. ?f b }"
-    with approx[OF _ a[symmetric], of "[x]"] approx[OF _ b[symmetric], of "[x]"]
-    have bnd: "x \<in> { l .. u'}" unfolding bounded_by_def i by auto
+    with approx[OF _ a, of "[x]"] approx[OF _ b, of "[x]"]
+    have bnd: "x \<in>\<^sub>r sup ivl ivl'" unfolding bounded_by_def i
+      by (auto simp: set_of_eq sup_float_def inf_float_def min_def max_def)
 
     have "interpret_form f' [x]"
       using assms[unfolded f_def]
     proof (induct f')
       case (Less lf rt)
       with a b
-      have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)"
+      have "approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 < lower ivl)"
         unfolding approx_tse_form_def by auto
       from approx_tse_form'_less[OF this bnd]
       show ?case using Less by auto
     next
       case (LessEqual lf rt)
       with f_def a b assms
-      have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
+      have "approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)"
         unfolding approx_tse_form_def by auto
       from approx_tse_form'_le[OF this bnd]
       show ?case using LessEqual by auto
     next
       case (AtLeastAtMost x lf rt)
       with f_def a b assms
-      have "approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l)"
-        and "approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"
+      have "approx_tse_form' prec t (Add rt (Minus x)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)"
+        and "approx_tse_form' prec t (Add x (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)"
         unfolding approx_tse_form_def lazy_conj by (auto split: if_split_asm)
       from approx_tse_form'_le[OF this(1) bnd] approx_tse_form'_le[OF this(2) bnd]
       show ?case using AtLeastAtMost by auto
@@ -1455,14 +980,14 @@
 
 text \<open>\<^term>\<open>approx_form_eval\<close> is only used for the {\tt value}-command.\<close>
 
-fun approx_form_eval :: "nat \<Rightarrow> form \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option list" where
+fun approx_form_eval :: "nat \<Rightarrow> form \<Rightarrow> (float interval) option list \<Rightarrow> (float interval) option list" where
 "approx_form_eval prec (Bound (Var n) a b f) bs =
    (case (approx prec a bs, approx prec b bs)
-   of (Some (l, _), Some (_, u)) \<Rightarrow> approx_form_eval prec f (bs[n := Some (l, u)])
+   of (Some ivl1, Some ivl2) \<Rightarrow> approx_form_eval prec f (bs[n := Some (sup ivl1 ivl2)])
     | _ \<Rightarrow> bs)" |
 "approx_form_eval prec (Assign (Var n) a f) bs =
    (case (approx prec a bs)
-   of (Some (l, u)) \<Rightarrow> approx_form_eval prec f (bs[n := Some (l, u)])
+   of (Some ivl) \<Rightarrow> approx_form_eval prec f (bs[n := Some ivl])
     | _ \<Rightarrow> bs)" |
 "approx_form_eval prec (Less a b) bs = bs @ [approx prec a bs, approx prec b bs]" |
 "approx_form_eval prec (LessEqual a b) bs = bs @ [approx prec a bs, approx prec b bs]" |
@@ -1473,105 +998,71 @@
 
 subsection \<open>Implement proof method \texttt{approximation}\<close>
 
-oracle approximation_oracle = \<open>fn (thy, t) =>
-let
-  fun bad t = error ("Bad term: " ^ Syntax.string_of_term_global thy t);
+ML \<open>
+val _ = \<comment> \<open>Trusting the oracle \<close>@{oracle_name "holds_by_evaluation"}
+signature APPROXIMATION_COMPUTATION = sig
+val approx_bool: Proof.context -> term -> term
+val approx_arith: Proof.context -> term -> term
+val approx_form_eval: Proof.context -> term -> term
+val approx_conv: Proof.context -> conv
+end
+
+structure Approximation_Computation : APPROXIMATION_COMPUTATION = struct
+
+  fun approx_conv ctxt ct =
+    @{computation_check
+      terms: Trueprop
+        "0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc
+          "(+)::nat\<Rightarrow>nat\<Rightarrow>nat" "(-)::nat\<Rightarrow>nat\<Rightarrow>nat" "(*)::nat\<Rightarrow>nat\<Rightarrow>nat"
+        "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int"
+          "(+)::int\<Rightarrow>int\<Rightarrow>int" "(-)::int\<Rightarrow>int\<Rightarrow>int" "(*)::int\<Rightarrow>int\<Rightarrow>int" "uminus::int\<Rightarrow>int"
+        "replicate :: _ \<Rightarrow> (float interval) option \<Rightarrow> _"
+        "interval_of::float\<Rightarrow>float interval"
+        approx'
+        approx_form
+        approx_tse_form
+        approx_form_eval
+      datatypes: bool
+        int
+        nat
+        integer
+        "nat list"
+        float
+        "(float interval) option list"
+        floatarith
+        form
+    } ctxt ct
 
   fun term_of_bool true = \<^term>\<open>True\<close>
     | term_of_bool false = \<^term>\<open>False\<close>;
 
   val mk_int = HOLogic.mk_number \<^typ>\<open>int\<close> o @{code integer_of_int};
-  fun dest_int (\<^term>\<open>int_of_integer\<close> $ j) = @{code int_of_integer} (snd (HOLogic.dest_number j))
-    | dest_int i = @{code int_of_integer} (snd (HOLogic.dest_number i));
 
   fun term_of_float (@{code Float} (k, l)) =
     \<^term>\<open>Float\<close> $ mk_int k $ mk_int l;
 
-  fun term_of_float_float_option NONE = \<^term>\<open>None :: (float \<times> float) option\<close>
-    | term_of_float_float_option (SOME ff) = \<^term>\<open>Some :: float \<times> float \<Rightarrow> _\<close>
-        $ HOLogic.mk_prod (apply2 term_of_float ff);
-
-  val term_of_float_float_option_list =
-    HOLogic.mk_list \<^typ>\<open>(float \<times> float) option\<close> o map term_of_float_float_option;
-
-  fun nat_of_term t = @{code nat_of_integer}
-    (HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t));
-
-  fun float_of_term (\<^term>\<open>Float\<close> $ k $ l) =
-        @{code Float} (dest_int k, dest_int l)
-    | float_of_term t = bad t;
+  fun term_of_float_interval x = @{term "Interval::_\<Rightarrow>float interval"} $
+    HOLogic.mk_prod
+      (apply2 term_of_float (@{code lowerF} x, @{code upperF} x))
 
-  fun floatarith_of_term (\<^term>\<open>Add\<close> $ a $ b) = @{code Add} (floatarith_of_term a, floatarith_of_term b)
-    | floatarith_of_term (\<^term>\<open>Minus\<close> $ a) = @{code Minus} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Mult\<close> $ a $ b) = @{code Mult} (floatarith_of_term a, floatarith_of_term b)
-    | floatarith_of_term (\<^term>\<open>Inverse\<close> $ a) = @{code Inverse} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Cos\<close> $ a) = @{code Cos} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Arctan\<close> $ a) = @{code Arctan} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Abs\<close> $ a) = @{code Abs} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Max\<close> $ a $ b) = @{code Max} (floatarith_of_term a, floatarith_of_term b)
-    | floatarith_of_term (\<^term>\<open>Min\<close> $ a $ b) = @{code Min} (floatarith_of_term a, floatarith_of_term b)
-    | floatarith_of_term \<^term>\<open>Pi\<close> = @{code Pi}
-    | floatarith_of_term (\<^term>\<open>Sqrt\<close> $ a) = @{code Sqrt} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Exp\<close> $ a) = @{code Exp} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Powr\<close> $ a $ b) = @{code Powr} (floatarith_of_term a, floatarith_of_term b)
-    | floatarith_of_term (\<^term>\<open>Ln\<close> $ a) = @{code Ln} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Power\<close> $ a $ n) =
-        @{code Power} (floatarith_of_term a, nat_of_term n)
-    | floatarith_of_term (\<^term>\<open>Floor\<close> $ a) = @{code Floor} (floatarith_of_term a)
-    | floatarith_of_term (\<^term>\<open>Var\<close> $ n) = @{code Var} (nat_of_term n)
-    | floatarith_of_term (\<^term>\<open>Num\<close> $ m) = @{code Num} (float_of_term m)
-    | floatarith_of_term t = bad t;
+  fun term_of_float_interval_option NONE = \<^term>\<open>None :: (float interval) option\<close>
+    | term_of_float_interval_option (SOME ff) = \<^term>\<open>Some :: float interval \<Rightarrow> _\<close>
+        $ (term_of_float_interval ff);
+
+  val term_of_float_interval_option_list =
+    HOLogic.mk_list \<^typ>\<open>(float interval) option\<close> o map term_of_float_interval_option;
 
-  fun form_of_term (\<^term>\<open>Bound\<close> $ a $ b $ c $ p) = @{code Bound}
-        (floatarith_of_term a, floatarith_of_term b, floatarith_of_term c, form_of_term p)
-    | form_of_term (\<^term>\<open>Assign\<close> $ a $ b $ p) = @{code Assign}
-        (floatarith_of_term a, floatarith_of_term b, form_of_term p)
-    | form_of_term (\<^term>\<open>Less\<close> $ a $ b) = @{code Less}
-        (floatarith_of_term a, floatarith_of_term b)
-    | form_of_term (\<^term>\<open>LessEqual\<close> $ a $ b) = @{code LessEqual}
-        (floatarith_of_term a, floatarith_of_term b)
-    | form_of_term (\<^term>\<open>Conj\<close> $ a $ b) = @{code Conj}
-        (form_of_term a, form_of_term b)
-    | form_of_term (\<^term>\<open>Disj\<close> $ a $ b) = @{code Disj}
-        (form_of_term a, form_of_term b)
-    | form_of_term (\<^term>\<open>AtLeastAtMost\<close> $ a $ b $ c) = @{code AtLeastAtMost}
-        (floatarith_of_term a, floatarith_of_term b, floatarith_of_term c)
-    | form_of_term t = bad t;
+  val approx_bool = @{computation bool}
+    (fn _ => fn x => case x of SOME b => term_of_bool b
+      | NONE => error "Computation approx_bool failed.")
+  val approx_arith = @{computation "float interval option"}
+    (fn _ => fn x => case x of SOME ffo => term_of_float_interval_option ffo
+      | NONE => error "Computation approx_arith failed.")
+  val approx_form_eval = @{computation "float interval option list"}
+    (fn _ => fn x => case x of SOME ffos => term_of_float_interval_option_list ffos
+      | NONE => error "Computation approx_form_eval failed.")
 
-  fun float_float_option_of_term \<^term>\<open>None :: (float \<times> float) option\<close> = NONE
-    | float_float_option_of_term (\<^term>\<open>Some :: float \<times> float \<Rightarrow> _\<close> $ ff) =
-        SOME (apply2 float_of_term (HOLogic.dest_prod ff))
-    | float_float_option_of_term (\<^term>\<open>approx'\<close> $ n $ a $ ffs) = @{code approx'}
-        (nat_of_term n) (floatarith_of_term a) (float_float_option_list_of_term ffs)
-    | float_float_option_of_term t = bad t
-  and float_float_option_list_of_term
-        (\<^term>\<open>replicate :: _ \<Rightarrow> (float \<times> float) option \<Rightarrow> _\<close> $ n $ \<^term>\<open>None :: (float \<times> float) option\<close>) =
-          @{code replicate} (nat_of_term n) NONE
-    | float_float_option_list_of_term (\<^term>\<open>approx_form_eval\<close> $ n $ p $ ffs) =
-        @{code approx_form_eval} (nat_of_term n) (form_of_term p) (float_float_option_list_of_term ffs)
-    | float_float_option_list_of_term t = map float_float_option_of_term
-        (HOLogic.dest_list t);
-
-  val nat_list_of_term = map nat_of_term o HOLogic.dest_list ;
-
-  fun bool_of_term (\<^term>\<open>approx_form\<close> $ n $ p $ ffs $ ms) = @{code approx_form}
-        (nat_of_term n) (form_of_term p) (float_float_option_list_of_term ffs) (nat_list_of_term ms)
-    | bool_of_term (\<^term>\<open>approx_tse_form\<close> $ m $ n $ q $ p) =
-        @{code approx_tse_form} (nat_of_term m) (nat_of_term n) (nat_of_term q) (form_of_term p)
-    | bool_of_term t = bad t;
-
-  fun eval t = case fastype_of t
-   of \<^typ>\<open>bool\<close> =>
-        (term_of_bool o bool_of_term) t
-    | \<^typ>\<open>(float \<times> float) option\<close> =>
-        (term_of_float_float_option o float_float_option_of_term) t
-    | \<^typ>\<open>(float \<times> float) option list\<close> =>
-        (term_of_float_float_option_list o float_float_option_list_of_term) t
-    | _ => bad t;
-
-  val normalize = eval o Envir.beta_norm o Envir.eta_long [];
-
-in Thm.global_cterm_of thy (Logic.mk_equals (t, normalize t)) end
+end
 \<close>
 
 lemma intervalE: "a \<le> x \<and> x \<le> b \<Longrightarrow> \<lbrakk> x \<in> { a .. b } \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
--- a/src/HOL/Decision_Procs/Approximation_Bounds.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Decision_Procs/Approximation_Bounds.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -11,13 +11,15 @@
 theory Approximation_Bounds
 imports
   Complex_Main
-  "HOL-Library.Float"
+  "HOL-Library.Interval_Float"
   Dense_Linear_Order
 begin
 
 declare powr_neg_one [simp]
 declare powr_neg_numeral [simp]
 
+context includes interval.lifting begin
+
 section "Horner Scheme"
 
 subsection \<open>Define auxiliary helper \<open>horner\<close> function\<close>
@@ -193,6 +195,44 @@
     l1 \<le> x ^ n \<and> x ^ n \<le> u1"
   using float_power_bnds by auto
 
+lift_definition power_float_interval :: "nat \<Rightarrow> nat \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>p n (l, u). float_power_bnds p n l u"
+  using float_power_bnds
+  by (auto simp: bnds_power dest!: float_power_bnds[OF sym])
+
+lemma lower_power_float_interval:
+  "lower (power_float_interval p n x) = fst (float_power_bnds p n (lower x) (upper x))"
+  by transfer auto
+lemma upper_power_float_interval:
+  "upper (power_float_interval p n x) = snd (float_power_bnds p n (lower x) (upper x))"
+  by transfer auto
+
+lemma power_float_intervalI: "x \<in>\<^sub>r X \<Longrightarrow> x ^ n \<in>\<^sub>r power_float_interval p n X"
+  using float_power_bnds[OF prod.collapse]
+  by (auto simp: set_of_eq lower_power_float_interval upper_power_float_interval)
+
+lemma power_float_interval_mono:
+  "set_of (power_float_interval prec n A)
+    \<subseteq> set_of (power_float_interval prec n B)"
+  if "set_of A \<subseteq> set_of B"
+proof -
+  define la where "la = real_of_float (lower A)"
+  define ua where "ua = real_of_float (upper A)"
+  define lb where "lb = real_of_float (lower B)"
+  define ub where "ub = real_of_float (upper B)"
+  have ineqs: "lb \<le> la" "la \<le> ua" "ua \<le> ub" "lb \<le> ub"
+    using that lower_le_upper[of A] lower_le_upper[of B]
+    by (auto simp: la_def ua_def lb_def ub_def set_of_eq)
+  show ?thesis
+    using ineqs
+    by (simp add: set_of_subset_iff float_power_bnds_def max_def
+        power_down_fl.rep_eq power_up_fl.rep_eq
+        lower_power_float_interval upper_power_float_interval
+        la_def[symmetric] ua_def[symmetric] lb_def[symmetric] ub_def[symmetric])
+      (auto intro!: power_down_mono power_up_mono intro: order_trans[where y=0])
+qed
+
+
 section \<open>Approximation utility functions\<close>
 
 definition bnds_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<times> float" where
@@ -220,6 +260,42 @@
   ultimately show ?thesis by simp
 qed
 
+lift_definition mult_float_interval::"nat \<Rightarrow> float interval \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>prec. \<lambda>(a1, a2). \<lambda>(b1, b2). bnds_mult prec a1 a2 b1 b2"
+  by (auto dest!: bnds_mult[OF sym])
+
+lemma lower_mult_float_interval:
+  "lower (mult_float_interval p x y) = fst (bnds_mult p (lower x) (upper x) (lower y) (upper y))"
+  by transfer auto
+lemma upper_mult_float_interval:
+  "upper (mult_float_interval p x y) = snd (bnds_mult p (lower x) (upper x) (lower y) (upper y))"
+  by transfer auto
+
+lemma mult_float_interval:
+  "set_of (real_interval A) * set_of (real_interval B) \<subseteq>
+    set_of (real_interval (mult_float_interval prec A B))"
+proof -
+  let ?bm = "bnds_mult prec (lower A) (upper A) (lower B) (upper B)"
+  show ?thesis
+    using bnds_mult[of "fst ?bm" "snd ?bm", simplified, OF refl]
+    by (auto simp: set_of_eq set_times_def upper_mult_float_interval lower_mult_float_interval)
+qed
+
+lemma mult_float_intervalI:
+  "x * y \<in>\<^sub>r mult_float_interval prec A B"
+  if "x \<in>\<^sub>i real_interval A" "y \<in>\<^sub>i real_interval B"
+  using mult_float_interval[of A B] that
+  by (auto simp: )
+
+lemma mult_float_interval_mono:
+  "set_of (mult_float_interval prec A B) \<subseteq> set_of (mult_float_interval prec X Y)"
+  if "set_of A \<subseteq> set_of X" "set_of B \<subseteq> set_of Y"
+  using that
+  apply transfer
+  unfolding bnds_mult_def atLeastatMost_subset_iff float_plus_down.rep_eq float_plus_up.rep_eq
+  by (auto simp: float_plus_down.rep_eq float_plus_up.rep_eq mult_float_mono1 mult_float_mono2)
+
+
 definition map_bnds :: "(nat \<Rightarrow> float \<Rightarrow> float) \<Rightarrow> (nat \<Rightarrow> float \<Rightarrow> float) \<Rightarrow>
                            nat \<Rightarrow> (float \<times> float) \<Rightarrow> (float \<times> float)" where
   "map_bnds lb ub prec = (\<lambda>(l,u). (lb prec l, ub prec u))"
@@ -456,6 +532,25 @@
   show "sqrt x \<le> u" unfolding u using bnds_sqrt'[of ux prec] by auto
 qed
 
+lift_definition sqrt_float_interval::"nat \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>prec. \<lambda>(lx, ux). (lb_sqrt prec lx, ub_sqrt prec ux)"
+  using bnds_sqrt'
+  by auto (meson order_trans real_sqrt_le_iff)
+
+lemma lower_float_interval: "lower (sqrt_float_interval prec X) = lb_sqrt prec (lower X)"
+  by transfer auto
+
+lemma upper_float_interval: "upper (sqrt_float_interval prec X) = ub_sqrt prec (upper X)"
+  by transfer auto
+
+lemma sqrt_float_interval:
+  "sqrt ` set_of (real_interval X) \<subseteq> set_of (real_interval (sqrt_float_interval prec X))"
+  using bnds_sqrt
+  by (auto simp: set_of_eq lower_float_interval upper_float_interval)
+
+lemma sqrt_float_intervalI: "sqrt x \<in>\<^sub>r sqrt_float_interval p X" if "x \<in>\<^sub>r X"
+  using sqrt_float_interval[of X p] that
+  by auto
 
 section "Arcus tangens and \<pi>"
 
@@ -711,6 +806,18 @@
   ultimately show ?thesis by auto
 qed
 
+lift_definition pi_float_interval::"nat \<Rightarrow> float interval" is "\<lambda>prec. (lb_pi prec, ub_pi prec)"
+  using pi_boundaries
+  by (auto intro: order_trans)
+
+lemma lower_pi_float_interval: "lower (pi_float_interval prec) = lb_pi prec"
+  by transfer auto
+lemma upper_pi_float_interval: "upper (pi_float_interval prec) = ub_pi prec"
+  by transfer auto
+lemma pi_float_interval: "pi \<in> set_of (real_interval (pi_float_interval prec))"
+  using pi_boundaries
+  by (auto simp: set_of_eq lower_pi_float_interval upper_pi_float_interval)
+
 
 subsection "Compute arcus tangens in the entire domain"
 
@@ -1056,6 +1163,32 @@
   qed
 qed
 
+lemmas [simp del] = lb_arctan.simps ub_arctan.simps
+
+lemma lb_arctan: "arctan (real_of_float x) \<le> y \<Longrightarrow> real_of_float (lb_arctan prec x) \<le> y"
+  and ub_arctan: "y \<le> arctan x \<Longrightarrow> y \<le> ub_arctan prec x"
+  for x::float and y::real
+  using arctan_boundaries[of x prec] by auto
+
+lift_definition arctan_float_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>prec. \<lambda>(lx, ux). (lb_arctan prec lx, ub_arctan prec ux)"
+  by (auto intro!: lb_arctan ub_arctan arctan_monotone')
+
+lemma lower_arctan_float_interval: "lower (arctan_float_interval p x) = lb_arctan p (lower x)"
+  by transfer auto
+lemma upper_arctan_float_interval: "upper (arctan_float_interval p x) = ub_arctan p (upper x)"
+  by transfer auto
+
+lemma arctan_float_interval:
+  "arctan ` set_of (real_interval x) \<subseteq> set_of (real_interval (arctan_float_interval p x))"
+  by (auto simp: set_of_eq lower_arctan_float_interval upper_arctan_float_interval
+      intro!: lb_arctan ub_arctan arctan_monotone')
+
+lemma arctan_float_intervalI:
+  "arctan x \<in>\<^sub>r arctan_float_interval p X" if "x \<in>\<^sub>r X"
+  using arctan_float_interval[of X p] that
+  by auto
+
 
 section "Sinus and Cosinus"
 
@@ -1794,6 +1927,31 @@
   qed
 qed
 
+lemma bnds_cos_lower: "\<And>x. real_of_float xl \<le> x \<Longrightarrow> x \<le> real_of_float xu \<Longrightarrow> cos x \<le> y \<Longrightarrow> real_of_float (fst (bnds_cos prec xl xu)) \<le> y"
+  and bnds_cos_upper: "\<And>x. real_of_float xl \<le> x \<Longrightarrow> x \<le> real_of_float xu \<Longrightarrow> y \<le> cos x \<Longrightarrow> y \<le> real_of_float (snd (bnds_cos prec xl xu))"
+  for xl xu::float and y::real
+  using bnds_cos[of "fst (bnds_cos prec xl xu)" "snd (bnds_cos prec xl xu)" prec]
+  by force+
+
+lift_definition cos_float_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>prec. \<lambda>(lx, ux). bnds_cos prec lx ux"
+  using bnds_cos
+  by auto (metis (full_types) order_refl order_trans)
+
+lemma lower_cos_float_interval: "lower (cos_float_interval p x) = fst (bnds_cos p (lower x) (upper x))"
+  by transfer auto
+lemma upper_cos_float_interval: "upper (cos_float_interval p x) = snd (bnds_cos p (lower x) (upper x))"
+  by transfer auto
+
+lemma cos_float_interval:
+  "cos ` set_of (real_interval x) \<subseteq> set_of (real_interval (cos_float_interval p x))"
+  by (auto simp: set_of_eq bnds_cos_lower bnds_cos_upper lower_cos_float_interval
+      upper_cos_float_interval)
+
+lemma cos_float_intervalI: "cos x \<in>\<^sub>r cos_float_interval p X" if "x \<in>\<^sub>r X"
+  using cos_float_interval[of X p] that
+  by auto
+
 
 section "Exponential function"
 
@@ -2137,6 +2295,31 @@
   qed
 qed
 
+lemmas [simp del] = lb_exp.simps ub_exp.simps
+
+lemma lb_exp: "exp x \<le> y \<Longrightarrow> lb_exp prec x \<le> y"
+  and ub_exp: "y \<le> exp x \<Longrightarrow> y \<le> ub_exp prec x"
+  for x::float and y::real using exp_boundaries[of x prec] by auto
+
+lift_definition exp_float_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>prec. \<lambda>(lx, ux). (lb_exp prec lx, ub_exp prec ux)"
+  by (auto simp: lb_exp ub_exp)
+
+lemma lower_exp_float_interval: "lower (exp_float_interval p x) = lb_exp p (lower x)"
+  by transfer auto
+lemma upper_exp_float_interval: "upper (exp_float_interval p x) = ub_exp p (upper x)"
+  by transfer auto
+
+lemma exp_float_interval:
+  "exp ` set_of (real_interval x) \<subseteq> set_of (real_interval (exp_float_interval p x))"
+  using exp_boundaries
+  by (auto simp: set_of_eq lower_exp_float_interval upper_exp_float_interval lb_exp ub_exp)
+
+lemma exp_float_intervalI:
+  "exp x \<in>\<^sub>r exp_float_interval p X" if "x \<in>\<^sub>r X"
+  using exp_float_interval[of X p] that
+  by auto
+
 
 section "Logarithm"
 
@@ -2211,7 +2394,7 @@
   also have "\<dots> \<le> ?ub"
     unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric]
     unfolding mult.commute[of "real_of_float x"] od
-    using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1",
+    using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2 * od + 1",
       OF \<open>0 \<le> real_of_float x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real_of_float x\<close>
     unfolding real_of_float_power
     by (rule mult_right_mono)
@@ -2703,6 +2886,46 @@
   ultimately show "l \<le> ln x \<and> ln x \<le> u" ..
 qed
 
+lemmas [simp del] = lb_ln.simps ub_ln.simps
+
+lemma lb_lnD:
+  "y \<le> ln x \<and> 0 < real_of_float x" if "lb_ln prec x = Some y"
+  using lb_ln[OF that[symmetric]] by auto
+
+lemma ub_lnD:
+  "ln x \<le> y\<and> 0 < real_of_float x" if "ub_ln prec x = Some y"
+  using ub_ln[OF that[symmetric]] by auto
+
+lift_definition(code_dt) ln_float_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval option"
+  is "\<lambda>prec. \<lambda>(lx, ux).
+    Option.bind (lb_ln prec lx) (\<lambda>l.
+      Option.bind (ub_ln prec ux) (\<lambda>u. Some (l, u)))"
+  by (auto simp: pred_option_def bind_eq_Some_conv ln_le_cancel_iff[symmetric]
+      simp del: ln_le_cancel_iff dest!: lb_lnD ub_lnD)
+
+lemma ln_float_interval_eq_Some_conv:
+  "ln_float_interval p x = Some y \<longleftrightarrow>
+    lb_ln p (lower x) = Some (lower y) \<and> ub_ln p (upper x) = Some (upper y)"
+  by transfer (auto simp: bind_eq_Some_conv)
+
+lemma ln_float_interval: "ln ` set_of (real_interval x) \<subseteq> set_of (real_interval y)"
+  if "ln_float_interval p x = Some y"
+  using that lb_ln[of "lower y" p "lower x"]
+    ub_ln[of "lower y" p "lower x"]
+  apply (auto simp add: set_of_eq ln_float_interval_eq_Some_conv ln_le_cancel_iff)
+   apply (meson less_le_trans ln_less_cancel_iff not_le)
+  by (meson less_le_trans ln_less_cancel_iff not_le ub_lnD)
+
+lemma ln_float_intervalI:
+  "ln x \<in> set_of' (ln_float_interval p X)" if "x \<in>\<^sub>r X"
+  using ln_float_interval[of p X] that
+  by (auto simp: set_of'_def split: option.splits)
+
+lemma ln_float_interval_eqI:
+  "ln x \<in>\<^sub>r IVL" if "ln_float_interval p X = Some IVL" "x \<in>\<^sub>r X"
+  using ln_float_intervalI[of x X p] that
+  by (auto simp: set_of'_def split: option.splits)
+
 
 section \<open>Real power function\<close>
 
@@ -2719,8 +2942,6 @@
        Some (map_bnds lb_exp ub_exp prec 
                (bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2)))"
 
-lemmas [simp del] = lb_exp.simps ub_exp.simps
-
 lemma mono_exp_real: "mono (exp :: real \<Rightarrow> real)"
   by (auto simp: mono_def)
 
@@ -2798,4 +3019,29 @@
   qed
 qed
 
+lift_definition(code_dt) powr_float_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval \<Rightarrow> float interval option"
+  is "\<lambda>prec. \<lambda>(l1, u1). \<lambda>(l2, u2). bnds_powr prec l1 u1 l2 u2"
+  by (auto simp: pred_option_def dest!: bnds_powr[OF sym])
+
+lemma powr_float_interval:
+  "{x powr y | x y. x \<in> set_of (real_interval X) \<and> y \<in> set_of (real_interval Y)}
+    \<subseteq> set_of (real_interval R)"
+  if "powr_float_interval prec X Y = Some R"
+  using that
+  by transfer (auto dest!: bnds_powr[OF sym])
+
+lemma powr_float_intervalI:
+  "x powr y \<in> set_of' (powr_float_interval p X Y)"
+  if "x \<in>\<^sub>r X" "y \<in>\<^sub>r Y"
+  using powr_float_interval[of p X Y] that
+  by (auto simp: set_of'_def split: option.splits)
+
+lemma powr_float_interval_eqI:
+  "x powr y \<in>\<^sub>r IVL"
+  if "powr_float_interval p X Y = Some IVL" "x \<in>\<^sub>r X" "y \<in>\<^sub>r Y"
+  using powr_float_intervalI[of x X y Y p] that
+  by (auto simp: set_of'_def split: option.splits)
+
 end
+
+end
\ No newline at end of file
--- a/src/HOL/Decision_Procs/approximation.ML	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Decision_Procs/approximation.ML	Tue Nov 05 22:56:06 2019 +0100
@@ -10,7 +10,7 @@
   val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic
 end
 
-structure Approximation: APPROXIMATION =
+structure Approximation =
 struct
 
 fun reorder_bounds_tac ctxt prems i =
@@ -42,18 +42,14 @@
     fold prepend_prem order all_tac
   end
 
-fun approximation_conv ctxt ct =
-  approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct);
-
-fun approximate ctxt t =
-  approximation_oracle (Proof_Context.theory_of ctxt, t)
-  |> Thm.prop_of |> Logic.dest_equals |> snd;
-
-(* Should be in HOL.thy ? *)
-fun gen_eval_tac conv ctxt =
-  CONVERSION
-    (Object_Logic.judgment_conv ctxt (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
-  THEN' resolve_tac ctxt [TrueI]
+fun approximate ctxt t = case fastype_of t
+   of \<^typ>\<open>bool\<close> =>
+        Approximation_Computation.approx_bool ctxt t
+    | \<^typ>\<open>(float interval) option\<close> =>
+        Approximation_Computation.approx_arith ctxt t
+    | \<^typ>\<open>(float interval) option list\<close> =>
+        Approximation_Computation.approx_form_eval ctxt t
+    | _ => error ("Bad term: " ^ Syntax.string_of_term ctxt t);
 
 fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let
     fun lookup_splitting (Free (name, _)) =
@@ -125,14 +121,16 @@
   | dest_float t = raise TERM ("dest_float", [t])
 
 
-fun dest_ivl (Const (\<^const_name>\<open>Some\<close>, _) $
-              (Const (\<^const_name>\<open>Pair\<close>, _) $ u $ l)) = SOME (dest_float u, dest_float l)
+fun dest_ivl
+  (Const (\<^const_name>\<open>Some\<close>, _) $
+    (Const (\<^const_name>\<open>Interval\<close>, _) $ ((Const (\<^const_name>\<open>Pair\<close>, _) $ u $ l)))) =
+      SOME (dest_float u, dest_float l)
   | dest_ivl (Const (\<^const_name>\<open>None\<close>, _)) = NONE
   | dest_ivl t = raise TERM ("dest_result", [t])
 
 fun mk_approx' prec t = (\<^const>\<open>approx'\<close>
                        $ HOLogic.mk_number \<^typ>\<open>nat\<close> prec
-                       $ t $ \<^term>\<open>[] :: (float * float) option list\<close>)
+                       $ t $ \<^term>\<open>[] :: (float interval) option list\<close>)
 
 fun mk_approx_form_eval prec t xs = (\<^const>\<open>approx_form_eval\<close>
                        $ HOLogic.mk_number \<^typ>\<open>nat\<close> prec
@@ -241,8 +239,8 @@
          |> HOLogic.dest_Trueprop
          |> dest_interpret_form
          |> (fn (data, xs) =>
-            mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\<open>(float * float) option\<close>
-              (map (fn _ => \<^term>\<open>None :: (float * float) option\<close>) (HOLogic.dest_list xs)))
+            mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\<open>(float interval) option\<close>
+              (map (fn _ => \<^term>\<open>None :: (float interval) option\<close>) (HOLogic.dest_list xs)))
          |> approximate ctxt
          |> HOLogic.dest_list
          |> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term)
@@ -286,10 +284,11 @@
     (opt_modes -- Parse.term
       >> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t)));
 
-fun approximation_tac prec splitting taylor ctxt i =
-  prepare_form_tac ctxt i
-  THEN reify_form_tac ctxt i
-  THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
-  THEN gen_eval_tac (approximation_conv ctxt) ctxt i
+fun approximation_tac prec splitting taylor ctxt =
+  prepare_form_tac ctxt
+  THEN' reify_form_tac ctxt
+  THEN' rewrite_interpret_form_tac ctxt prec splitting taylor
+  THEN' CONVERSION (Approximation_Computation.approx_conv ctxt)
+  THEN' resolve_tac ctxt [TrueI]
 
 end;
--- a/src/HOL/Decision_Procs/approximation_generator.ML	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Decision_Procs/approximation_generator.ML	Tue Nov 05 22:56:06 2019 +0100
@@ -92,6 +92,10 @@
 fun real_of_Float (@{code Float} (m, e)) =
     real_of_man_exp (@{code integer_of_int} m) (@{code integer_of_int} e)
 
+fun term_of_int i = (HOLogic.mk_number @{typ int} (@{code integer_of_int} i))
+
+fun term_of_Float (@{code Float} (m, e)) = @{term Float} $ term_of_int m $ term_of_int e
+
 fun is_True \<^term>\<open>True\<close> = true
   | is_True _ = false
 
@@ -129,8 +133,8 @@
       \<^const>\<open>approx_form\<close> $
         HOLogic.mk_number \<^typ>\<open>nat\<close> prec $
         e $
-        (HOLogic.mk_list \<^typ>\<open>(float * float) option\<close>
-          (map (fn t => mk_Some \<^typ>\<open>float * float\<close> $ HOLogic.mk_prod (t, t)) ts)) $
+        (HOLogic.mk_list \<^typ>\<open>(float interval) option\<close>
+          (map (fn t => mk_Some \<^typ>\<open>float interval\<close> $ (\<^term>\<open>interval_of::float\<Rightarrow>_\<close> $ t)) ts)) $
         \<^term>\<open>[] :: nat list\<close>
   in
     (if
@@ -142,7 +146,7 @@
       | IEEEReal.Unordered => false
     then
       let
-        val ts = map (fn x => snd x ()) rs
+        val ts = map (fn x => term_of_Float (fst x)) rs
         val ts' = map
           (AList.lookup op = (map dest_Free xs ~~ ts)
             #> the_default Term.dummy
--- a/src/HOL/Library/Boolean_Algebra.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Library/Boolean_Algebra.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -211,7 +211,7 @@
 definition xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "\<oplus>" 65)
   where "x \<oplus> y = (x \<^bold>\<sqinter> \<sim> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> y)"
 
-sublocale xor: abel_semigroup xor
+sublocale xor: comm_monoid xor \<zero>
 proof
   fix x y z :: 'a
   let ?t = "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z)"
@@ -222,6 +222,8 @@
       (simp only: conj_disj_distribs conj_ac disj_ac)
   show "x \<oplus> y = y \<oplus> x"
     by (simp only: xor_def conj_commute disj_commute)
+  show "x \<oplus> \<zero> = x"
+    by (simp add: xor_def)
 qed
 
 lemmas xor_assoc = xor.assoc
--- a/src/HOL/Library/Float.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Library/Float.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -1104,6 +1104,201 @@
 
 end
 
+lemma truncate_up_nonneg_mono:
+  assumes "0 \<le> x" "x \<le> y"
+  shows "truncate_up prec x \<le> truncate_up prec y"
+proof -
+  consider "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>" | "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>" "0 < x" | "x \<le> 0"
+    by arith
+  then show ?thesis
+  proof cases
+    case 1
+    then show ?thesis
+      using assms
+      by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono)
+  next
+    case 2
+    from assms \<open>0 < x\<close> have "log 2 x \<le> log 2 y"
+      by auto
+    with \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close>
+    have logless: "log 2 x < log 2 y"
+      by linarith
+    have flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
+      using \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close> \<open>log 2 x \<le> log 2 y\<close> by linarith
+    have "truncate_up prec x =
+      real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor> )\<rceil> * 2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)"
+      using assms by (simp add: truncate_up_def round_up_def)
+    also have "\<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> (2 ^ (Suc prec))"
+    proof (simp only: ceiling_le_iff)
+      have "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le>
+        x * (2 powr real (Suc prec) / (2 powr log 2 x))"
+        using real_of_int_floor_add_one_ge[of "log 2 x"] assms
+        by (auto simp: algebra_simps simp flip: powr_diff intro!: mult_left_mono)
+      then show "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> real_of_int ((2::int) ^ (Suc prec))"
+        using \<open>0 < x\<close> by (simp add: powr_realpow powr_add)
+    qed
+    then have "real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> 2 powr int (Suc prec)"
+      by (auto simp: powr_realpow powr_add)
+        (metis power_Suc of_int_le_numeral_power_cancel_iff)
+    also
+    have "2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)"
+      using logless flogless by (auto intro!: floor_mono)
+    also have "2 powr real_of_int (int (Suc prec)) \<le>
+        2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1))"
+      using assms \<open>0 < x\<close>
+      by (auto simp: algebra_simps)
+    finally have "truncate_up prec x \<le>
+        2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)"
+      by simp
+    also have "\<dots> = 2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor>) - real_of_int (int prec - \<lfloor>log 2 y\<rfloor>))"
+      by (subst powr_add[symmetric]) simp
+    also have "\<dots> = y"
+      using \<open>0 < x\<close> assms
+      by (simp add: powr_add)
+    also have "\<dots> \<le> truncate_up prec y"
+      by (rule truncate_up)
+    finally show ?thesis .
+  next
+    case 3
+    then show ?thesis
+      using assms
+      by (auto intro!: truncate_up_le)
+  qed
+qed
+
+lemma truncate_up_switch_sign_mono:
+  assumes "x \<le> 0" "0 \<le> y"
+  shows "truncate_up prec x \<le> truncate_up prec y"
+proof -
+  note truncate_up_nonpos[OF \<open>x \<le> 0\<close>]
+  also note truncate_up_le[OF \<open>0 \<le> y\<close>]
+  finally show ?thesis .
+qed
+
+lemma truncate_down_switch_sign_mono:
+  assumes "x \<le> 0"
+    and "0 \<le> y"
+    and "x \<le> y"
+  shows "truncate_down prec x \<le> truncate_down prec y"
+proof -
+  note truncate_down_le[OF \<open>x \<le> 0\<close>]
+  also note truncate_down_nonneg[OF \<open>0 \<le> y\<close>]
+  finally show ?thesis .
+qed
+
+lemma truncate_down_nonneg_mono:
+  assumes "0 \<le> x" "x \<le> y"
+  shows "truncate_down prec x \<le> truncate_down prec y"
+proof -
+  consider "x \<le> 0" | "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" |
+    "0 < x" "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
+    by arith
+  then show ?thesis
+  proof cases
+    case 1
+    with assms have "x = 0" "0 \<le> y" by simp_all
+    then show ?thesis
+      by (auto intro!: truncate_down_nonneg)
+  next
+    case 2
+    then show ?thesis
+      using assms
+      by (auto simp: truncate_down_def round_down_def intro!: floor_mono)
+  next
+    case 3
+    from \<open>0 < x\<close> have "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y"
+      using assms by auto
+    with \<open>\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>\<close>
+    have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
+      unfolding atomize_conj abs_of_pos[OF \<open>0 < x\<close>] abs_of_pos[OF \<open>0 < y\<close>]
+      by (metis floor_less_cancel linorder_cases not_le)
+    have "2 powr prec \<le> y * 2 powr real prec / (2 powr log 2 y)"
+      using \<open>0 < y\<close> by simp
+    also have "\<dots> \<le> y * 2 powr real (Suc prec) / (2 powr (real_of_int \<lfloor>log 2 y\<rfloor> + 1))"
+      using \<open>0 \<le> y\<close> \<open>0 \<le> x\<close> assms(2)
+      by (auto intro!: powr_mono divide_left_mono
+          simp: of_nat_diff powr_add powr_diff)
+    also have "\<dots> = y * 2 powr real (Suc prec) / (2 powr real_of_int \<lfloor>log 2 y\<rfloor> * 2)"
+      by (auto simp: powr_add)
+    finally have "(2 ^ prec) \<le> \<lfloor>y * 2 powr real_of_int (int (Suc prec) - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
+      using \<open>0 \<le> y\<close>
+      by (auto simp: powr_diff le_floor_iff powr_realpow powr_add)
+    then have "(2 ^ (prec)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>) \<le> truncate_down prec y"
+      by (auto simp: truncate_down_def round_down_def)
+    moreover have "x \<le> (2 ^ prec) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)"
+    proof -
+      have "x = 2 powr (log 2 \<bar>x\<bar>)" using \<open>0 < x\<close> by simp
+      also have "\<dots> \<le> (2 ^ (Suc prec )) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)"
+        using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"] \<open>0 < x\<close>
+        by (auto simp flip: powr_realpow powr_add simp: algebra_simps powr_mult_base le_powr_iff)
+      also
+      have "2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> + 1)"
+        using logless flogless \<open>x > 0\<close> \<open>y > 0\<close>
+        by (auto intro!: floor_mono)
+      finally show ?thesis
+        by (auto simp flip: powr_realpow simp: powr_diff assms of_nat_diff)
+    qed
+    ultimately show ?thesis
+      by (metis dual_order.trans truncate_down)
+  qed
+qed
+
+lemma truncate_down_eq_truncate_up: "truncate_down p x = - truncate_up p (-x)"
+  and truncate_up_eq_truncate_down: "truncate_up p x = - truncate_down p (-x)"
+  by (auto simp: truncate_up_uminus_eq truncate_down_uminus_eq)
+
+lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y"
+  apply (cases "0 \<le> x")
+  apply (rule truncate_down_nonneg_mono, assumption+)
+  apply (simp add: truncate_down_eq_truncate_up)
+  apply (cases "0 \<le> y")
+  apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono)
+  done
+
+lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y"
+  by (simp add: truncate_up_eq_truncate_down truncate_down_mono)
+
+lemma truncate_up_nonneg: "0 \<le> truncate_up p x" if "0 \<le> x"
+  by (simp add: that truncate_up_le)
+
+lemma truncate_up_pos: "0 < truncate_up p x" if "0 < x"
+  by (meson less_le_trans that truncate_up)
+
+lemma truncate_up_less_zero_iff[simp]: "truncate_up p x < 0 \<longleftrightarrow> x < 0"
+proof -
+  have f1: "\<forall>n r. truncate_up n r + truncate_down n (- 1 * r) = 0"
+    by (simp add: truncate_down_uminus_eq)
+  have f2: "(\<forall>v0 v1. truncate_up v0 v1 + truncate_down v0 (- 1 * v1) = 0) = (\<forall>v0 v1. truncate_up v0 v1 = - 1 * truncate_down v0 (- 1 * v1))"
+    by (auto simp: truncate_up_eq_truncate_down)
+  have f3: "\<forall>x1. ((0::real) < x1) = (\<not> x1 \<le> 0)"
+    by fastforce
+  have "(- 1 * x \<le> 0) = (0 \<le> x)"
+    by force
+  then have "0 \<le> x \<or> \<not> truncate_down p (- 1 * x) \<le> 0"
+    using f3 by (meson truncate_down_pos)
+  then have "(0 \<le> truncate_up p x) \<noteq> (\<not> 0 \<le> x)"
+    using f2 f1 truncate_up_nonneg by force
+  then show ?thesis
+    by linarith
+qed
+
+lemma truncate_up_nonneg_iff[simp]: "truncate_up p x \<ge> 0 \<longleftrightarrow> x \<ge> 0"
+  using truncate_up_less_zero_iff[of p x] truncate_up_nonneg[of x]
+  by linarith
+
+lemma truncate_down_less_zero_iff[simp]: "truncate_down p x < 0 \<longleftrightarrow> x < 0"
+  by (metis le_less_trans not_less_iff_gr_or_eq truncate_down truncate_down_pos truncate_down_zero)
+
+lemma truncate_down_nonneg_iff[simp]: "truncate_down p x \<ge> 0 \<longleftrightarrow> x \<ge> 0"
+  using truncate_down_less_zero_iff[of p x] truncate_down_nonneg[of x p]
+  by linarith
+
+lemma truncate_down_eq_zero_iff[simp]: "truncate_down prec x = 0 \<longleftrightarrow> x = 0"
+  by (metis not_less_iff_gr_or_eq truncate_down_less_zero_iff truncate_down_pos truncate_down_zero)
+
+lemma truncate_up_eq_zero_iff[simp]: "truncate_up prec x = 0 \<longleftrightarrow> x = 0"
+  by (metis not_less_iff_gr_or_eq truncate_up_less_zero_iff truncate_up_pos truncate_up_zero)
+
 
 subsection \<open>Approximation of positive rationals\<close>
 
@@ -1324,109 +1519,6 @@
 end
 
 
-subsection \<open>Approximate Power\<close>
-
-lemma div2_less_self[termination_simp]: "odd n \<Longrightarrow> n div 2 < n" for n :: nat
-  by (simp add: odd_pos)
-
-fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
-where
-  "power_down p x 0 = 1"
-| "power_down p x (Suc n) =
-    (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2)
-     else truncate_down (Suc p) (x * power_down p x n))"
-
-fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
-where
-  "power_up p x 0 = 1"
-| "power_up p x (Suc n) =
-    (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2)
-     else truncate_up p (x * power_up p x n))"
-
-lift_definition power_up_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_up
-  by (induct_tac rule: power_up.induct) simp_all
-
-lift_definition power_down_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_down
-  by (induct_tac rule: power_down.induct) simp_all
-
-lemma power_float_transfer[transfer_rule]:
-  "(rel_fun pcr_float (rel_fun (=) pcr_float)) (^) (^)"
-  unfolding power_def
-  by transfer_prover
-
-lemma compute_power_up_fl[code]:
-    "power_up_fl p x 0 = 1"
-    "power_up_fl p x (Suc n) =
-      (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2)
-       else float_round_up p (x * power_up_fl p x n))"
-  and compute_power_down_fl[code]:
-    "power_down_fl p x 0 = 1"
-    "power_down_fl p x (Suc n) =
-      (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2)
-       else float_round_down (Suc p) (x * power_down_fl p x n))"
-  unfolding atomize_conj by transfer simp
-
-lemma power_down_pos: "0 < x \<Longrightarrow> 0 < power_down p x n"
-  by (induct p x n rule: power_down.induct)
-    (auto simp del: odd_Suc_div_two intro!: truncate_down_pos)
-
-lemma power_down_nonneg: "0 \<le> x \<Longrightarrow> 0 \<le> power_down p x n"
-  by (induct p x n rule: power_down.induct)
-    (auto simp del: odd_Suc_div_two intro!: truncate_down_nonneg mult_nonneg_nonneg)
-
-lemma power_down: "0 \<le> x \<Longrightarrow> power_down p x n \<le> x ^ n"
-proof (induct p x n rule: power_down.induct)
-  case (2 p x n)
-  have ?case if "odd n"
-  proof -
-    from that 2 have "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2"
-      by (auto intro: power_mono power_down_nonneg simp del: odd_Suc_div_two)
-    also have "\<dots> = x ^ (Suc n div 2 * 2)"
-      by (simp flip: power_mult)
-    also have "Suc n div 2 * 2 = Suc n"
-      using \<open>odd n\<close> by presburger
-    finally show ?thesis
-      using that by (auto intro!: truncate_down_le simp del: odd_Suc_div_two)
-  qed
-  then show ?case
-    by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg)
-qed simp
-
-lemma power_up: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up p x n"
-proof (induct p x n rule: power_up.induct)
-  case (2 p x n)
-  have ?case if "odd n"
-  proof -
-    from that even_Suc have "Suc n = Suc n div 2 * 2"
-      by presburger
-    then have "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
-      by (simp flip: power_mult)
-    also from that 2 have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2"
-      by (auto intro: power_mono simp del: odd_Suc_div_two)
-    finally show ?thesis
-      using that by (auto intro!: truncate_up_le simp del: odd_Suc_div_two)
-  qed
-  then show ?case
-    by (auto intro!: truncate_up_le mult_left_mono 2)
-qed simp
-
-lemmas power_up_le = order_trans[OF _ power_up]
-  and power_up_less = less_le_trans[OF _ power_up]
-  and power_down_le = order_trans[OF power_down]
-
-lemma power_down_fl: "0 \<le> x \<Longrightarrow> power_down_fl p x n \<le> x ^ n"
-  by transfer (rule power_down)
-
-lemma power_up_fl: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up_fl p x n"
-  by transfer (rule power_up)
-
-lemma real_power_up_fl: "real_of_float (power_up_fl p x n) = power_up p x n"
-  by transfer simp
-
-lemma real_power_down_fl: "real_of_float (power_down_fl p x n) = power_down p x n"
-  by transfer simp
-
-
 subsection \<open>Approximate Addition\<close>
 
 definition "plus_down prec x y = truncate_down prec (x + y)"
@@ -1794,6 +1886,332 @@
 
 end
 
+lemma plus_down_mono: "plus_down p a b \<le> plus_down p c d" if "a + b \<le> c + d"
+  by (auto simp: plus_down_def intro!: truncate_down_mono that)
+
+lemma plus_up_mono: "plus_up p a b \<le> plus_up p c d" if "a + b \<le> c + d"
+  by (auto simp: plus_up_def intro!: truncate_up_mono that)
+
+subsection \<open>Approximate Multiplication\<close>
+
+lemma mult_mono_nonpos_nonneg: "a * b \<le> c * d"
+  if  "a \<le> c" "a \<le> 0" "0 \<le> d" "d \<le> b" for a b c d::"'a::ordered_ring"
+  by (meson dual_order.trans mult_left_mono_neg mult_right_mono that)
+
+lemma mult_mono_nonneg_nonpos: "b * a \<le> d * c"
+  if  "a \<le> c" "c \<le> 0" "0 \<le> d" "d \<le> b" for a b c d::"'a::ordered_ring"
+  by (meson dual_order.trans mult_right_mono_neg mult_left_mono that)
+
+lemma mult_mono_nonpos_nonpos: "a * b \<le> c * d"
+  if  "a \<ge> c" "a \<le> 0" "b \<ge> d" "d \<le> 0" for a b c d::real
+  by (meson dual_order.trans mult_left_mono_neg mult_right_mono_neg that)
+
+lemma mult_float_mono1:
+  notes mono_rules = plus_down_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono
+  shows "a \<le> b \<Longrightarrow> ab \<le> bb \<Longrightarrow>
+       aa \<le> a \<Longrightarrow>
+       b \<le> ba \<Longrightarrow>
+       ac \<le> ab \<Longrightarrow>
+       bb \<le> bc \<Longrightarrow>
+       plus_down prec (nprt aa * pprt bc)
+        (plus_down prec (nprt ba * nprt bc)
+          (plus_down prec (pprt aa * pprt ac)
+            (pprt ba * nprt ac)))
+       \<le> plus_down prec (nprt a * pprt bb)
+           (plus_down prec (nprt b * nprt bb)
+             (plus_down prec (pprt a * pprt ab)
+               (pprt b * nprt ab)))"
+  apply (rule order_trans)
+   apply (rule mono_rules | assumption)+
+    apply (rule mult_mono_nonpos_nonneg)
+       apply (rule mono_rules | assumption)+
+    apply (rule mult_mono_nonpos_nonpos)
+       apply (rule mono_rules | assumption)+
+    apply (rule mult_mono)
+       apply (rule mono_rules | assumption)+
+   apply (rule mult_mono_nonneg_nonpos)
+      apply (rule mono_rules | assumption)+
+  by (rule order_refl)+
+
+lemma mult_float_mono2:
+  notes mono_rules = plus_up_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono
+  shows "a \<le> b \<Longrightarrow>
+       ab \<le> bb \<Longrightarrow>
+       aa \<le> a \<Longrightarrow>
+       b \<le> ba \<Longrightarrow>
+       ac \<le> ab \<Longrightarrow>
+       bb \<le> bc \<Longrightarrow>
+       plus_up prec (pprt b * pprt bb)
+        (plus_up prec (pprt a * nprt bb)
+          (plus_up prec (nprt b * pprt ab)
+            (nprt a * nprt ab)))
+       \<le> plus_up prec (pprt ba * pprt bc)
+           (plus_up prec (pprt aa * nprt bc)
+             (plus_up prec (nprt ba * pprt ac)
+               (nprt aa * nprt ac)))"
+  apply (rule order_trans)
+   apply (rule mono_rules | assumption)+
+    apply (rule mult_mono)
+       apply (rule mono_rules | assumption)+
+    apply (rule mult_mono_nonneg_nonpos)
+       apply (rule mono_rules | assumption)+
+    apply (rule mult_mono_nonpos_nonneg)
+       apply (rule mono_rules | assumption)+
+   apply (rule mult_mono_nonpos_nonpos)
+      apply (rule mono_rules | assumption)+
+  by (rule order_refl)+
+
+
+subsection \<open>Approximate Power\<close>
+
+lemma div2_less_self[termination_simp]: "odd n \<Longrightarrow> n div 2 < n" for n :: nat
+  by (simp add: odd_pos)
+
+fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
+where
+  "power_down p x 0 = 1"
+| "power_down p x (Suc n) =
+    (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2)
+     else truncate_down (Suc p) (x * power_down p x n))"
+
+fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
+where
+  "power_up p x 0 = 1"
+| "power_up p x (Suc n) =
+    (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2)
+     else truncate_up p (x * power_up p x n))"
+
+lift_definition power_up_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_up
+  by (induct_tac rule: power_up.induct) simp_all
+
+lift_definition power_down_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_down
+  by (induct_tac rule: power_down.induct) simp_all
+
+lemma power_float_transfer[transfer_rule]:
+  "(rel_fun pcr_float (rel_fun (=) pcr_float)) (^) (^)"
+  unfolding power_def
+  by transfer_prover
+
+lemma compute_power_up_fl[code]:
+    "power_up_fl p x 0 = 1"
+    "power_up_fl p x (Suc n) =
+      (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2)
+       else float_round_up p (x * power_up_fl p x n))"
+  and compute_power_down_fl[code]:
+    "power_down_fl p x 0 = 1"
+    "power_down_fl p x (Suc n) =
+      (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2)
+       else float_round_down (Suc p) (x * power_down_fl p x n))"
+  unfolding atomize_conj by transfer simp
+
+lemma power_down_pos: "0 < x \<Longrightarrow> 0 < power_down p x n"
+  by (induct p x n rule: power_down.induct)
+    (auto simp del: odd_Suc_div_two intro!: truncate_down_pos)
+
+lemma power_down_nonneg: "0 \<le> x \<Longrightarrow> 0 \<le> power_down p x n"
+  by (induct p x n rule: power_down.induct)
+    (auto simp del: odd_Suc_div_two intro!: truncate_down_nonneg mult_nonneg_nonneg)
+
+lemma power_down: "0 \<le> x \<Longrightarrow> power_down p x n \<le> x ^ n"
+proof (induct p x n rule: power_down.induct)
+  case (2 p x n)
+  have ?case if "odd n"
+  proof -
+    from that 2 have "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2"
+      by (auto intro: power_mono power_down_nonneg simp del: odd_Suc_div_two)
+    also have "\<dots> = x ^ (Suc n div 2 * 2)"
+      by (simp flip: power_mult)
+    also have "Suc n div 2 * 2 = Suc n"
+      using \<open>odd n\<close> by presburger
+    finally show ?thesis
+      using that by (auto intro!: truncate_down_le simp del: odd_Suc_div_two)
+  qed
+  then show ?case
+    by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg)
+qed simp
+
+lemma power_up: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up p x n"
+proof (induct p x n rule: power_up.induct)
+  case (2 p x n)
+  have ?case if "odd n"
+  proof -
+    from that even_Suc have "Suc n = Suc n div 2 * 2"
+      by presburger
+    then have "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
+      by (simp flip: power_mult)
+    also from that 2 have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2"
+      by (auto intro: power_mono simp del: odd_Suc_div_two)
+    finally show ?thesis
+      using that by (auto intro!: truncate_up_le simp del: odd_Suc_div_two)
+  qed
+  then show ?case
+    by (auto intro!: truncate_up_le mult_left_mono 2)
+qed simp
+
+lemmas power_up_le = order_trans[OF _ power_up]
+  and power_up_less = less_le_trans[OF _ power_up]
+  and power_down_le = order_trans[OF power_down]
+
+lemma power_down_fl: "0 \<le> x \<Longrightarrow> power_down_fl p x n \<le> x ^ n"
+  by transfer (rule power_down)
+
+lemma power_up_fl: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up_fl p x n"
+  by transfer (rule power_up)
+
+lemma real_power_up_fl: "real_of_float (power_up_fl p x n) = power_up p x n"
+  by transfer simp
+
+lemma real_power_down_fl: "real_of_float (power_down_fl p x n) = power_down p x n"
+  by transfer simp
+
+lemmas [simp del] = power_down.simps(2) power_up.simps(2)
+
+lemmas power_down_simp = power_down.simps(2)
+lemmas power_up_simp = power_up.simps(2)
+
+lemma power_down_even_nonneg: "even n \<Longrightarrow> 0 \<le> power_down p x n"
+  by (induct p x n rule: power_down.induct)
+    (auto simp: power_down_simp simp del: odd_Suc_div_two intro!: truncate_down_nonneg )
+
+lemma power_down_eq_zero_iff[simp]: "power_down prec b n = 0 \<longleftrightarrow> b = 0 \<and> n \<noteq> 0"
+proof (induction n arbitrary: b rule: less_induct)
+  case (less x)
+  then show ?case
+    using power_down_simp[of _ _ "x - 1"]
+    by (cases x) (auto simp add: div2_less_self)
+qed
+
+lemma power_down_nonneg_iff[simp]:
+  "power_down prec b n \<ge> 0 \<longleftrightarrow> even n \<or> b \<ge> 0"
+proof (induction n arbitrary: b rule: less_induct)
+  case (less x)
+  show ?case
+    using less(1)[of "x - 1" b] power_down_simp[of _ _ "x - 1"]
+    by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff)
+qed
+
+lemma power_down_neg_iff[simp]:
+  "power_down prec b n < 0 \<longleftrightarrow>
+    b < 0 \<and> odd n"
+  using power_down_nonneg_iff[of prec b n] by (auto simp del: power_down_nonneg_iff)
+
+lemma power_down_nonpos_iff[simp]:
+  notes [simp del] = power_down_neg_iff power_down_eq_zero_iff
+  shows "power_down prec b n \<le> 0 \<longleftrightarrow> b < 0 \<and> odd n \<or> b = 0 \<and> n \<noteq> 0"
+  using power_down_neg_iff[of prec b n] power_down_eq_zero_iff[of prec b n]
+  by auto
+
+lemma power_down_mono:
+  "power_down prec a n \<le> power_down prec b n"
+  if "((0 \<le> a \<and> a \<le> b)\<or>(odd n \<and> a \<le> b) \<or> (even n \<and> a \<le> 0 \<and> b \<le> a))"
+  using that
+proof (induction n arbitrary: a b rule: less_induct)
+  case (less i)
+  show ?case
+  proof (cases i)
+    case j: (Suc j)
+    note IH = less[unfolded j even_Suc not_not]
+    note [simp del] = power_down.simps
+    show ?thesis
+    proof cases
+      assume [simp]: "even j"
+      have "a * power_down prec a j \<le> b * power_down prec b j"
+        by (smt IH(1) IH(2) \<open>even j\<close> lessI mult_mono' mult_mono_nonpos_nonneg power_down_even_nonneg)
+      then have "truncate_down (Suc prec) (a * power_down prec a j) \<le> truncate_down (Suc prec) (b * power_down prec b j)"
+        by (auto intro!: truncate_down_mono simp: abs_le_square_iff[symmetric] abs_real_def)
+      then show ?thesis
+        unfolding j
+        by (simp add: power_down_simp)
+    next
+      assume [simp]: "odd j"
+      have "power_down prec 0 (Suc (j div 2)) \<le> - power_down prec b (Suc (j div 2))"
+        if "b < 0" "even (j div 2)"
+        apply (rule order_trans[where y=0])
+        using IH that by (auto simp: div2_less_self)
+      then have "truncate_down (Suc prec) ((power_down prec a (Suc (j div 2)))\<^sup>2)
+        \<le> truncate_down (Suc prec) ((power_down prec b (Suc (j div 2)))\<^sup>2)"
+        using IH
+        by (auto intro!: truncate_down_mono intro: order_trans[where y=0]
+            simp: abs_le_square_iff[symmetric] abs_real_def
+            div2_less_self)
+      then show ?thesis
+        unfolding j
+        by (simp add: power_down_simp)
+    qed
+  qed simp
+qed
+
+lemma power_up_even_nonneg: "even n \<Longrightarrow> 0 \<le> power_up p x n"
+  by (induct p x n rule: power_up.induct)
+    (auto simp: power_up.simps simp del: odd_Suc_div_two intro!: )
+
+lemma power_up_eq_zero_iff[simp]: "power_up prec b n = 0 \<longleftrightarrow> b = 0 \<and> n \<noteq> 0"
+proof (induction n arbitrary: b rule: less_induct)
+  case (less x)
+  then show ?case
+    using power_up_simp[of _ _ "x - 1"]
+    by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff div2_less_self)
+qed
+
+lemma power_up_nonneg_iff[simp]:
+  "power_up prec b n \<ge> 0 \<longleftrightarrow> even n \<or> b \<ge> 0"
+proof (induction n arbitrary: b rule: less_induct)
+  case (less x)
+  show ?case
+    using less(1)[of "x - 1" b] power_up_simp[of _ _ "x - 1"]
+    by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff)
+qed
+
+lemma power_up_neg_iff[simp]:
+  "power_up prec b n < 0 \<longleftrightarrow> b < 0 \<and> odd n"
+  using power_up_nonneg_iff[of prec b n] by (auto simp del: power_up_nonneg_iff)
+
+lemma power_up_nonpos_iff[simp]:
+  notes [simp del] = power_up_neg_iff power_up_eq_zero_iff
+  shows "power_up prec b n \<le> 0 \<longleftrightarrow> b < 0 \<and> odd n \<or> b = 0 \<and> n \<noteq> 0"
+  using power_up_neg_iff[of prec b n] power_up_eq_zero_iff[of prec b n]
+  by auto
+
+lemma power_up_mono:
+  "power_up prec a n \<le> power_up prec b n"
+  if "((0 \<le> a \<and> a \<le> b)\<or>(odd n \<and> a \<le> b) \<or> (even n \<and> a \<le> 0 \<and> b \<le> a))"
+  using that
+proof (induction n arbitrary: a b rule: less_induct)
+  case (less i)
+  show ?case
+  proof (cases i)
+    case j: (Suc j)
+    note IH = less[unfolded j even_Suc not_not]
+    note [simp del] = power_up.simps
+    show ?thesis
+    proof cases
+      assume [simp]: "even j"
+      have "a * power_up prec a j \<le> b * power_up prec b j"
+        by (smt IH(1) IH(2) \<open>even j\<close> lessI mult_mono' mult_mono_nonpos_nonneg power_up_even_nonneg)
+      then have "truncate_up prec (a * power_up prec a j) \<le> truncate_up prec (b * power_up prec b j)"
+        by (auto intro!: truncate_up_mono simp: abs_le_square_iff[symmetric] abs_real_def)
+      then show ?thesis
+        unfolding j
+        by (simp add: power_up_simp)
+    next
+      assume [simp]: "odd j"
+      have "power_up prec 0 (Suc (j div 2)) \<le> - power_up prec b (Suc (j div 2))"
+        if "b < 0" "even (j div 2)"
+        apply (rule order_trans[where y=0])
+        using IH that by (auto simp: div2_less_self)
+      then have "truncate_up prec ((power_up prec a (Suc (j div 2)))\<^sup>2)
+        \<le> truncate_up prec ((power_up prec b (Suc (j div 2)))\<^sup>2)"
+        using IH
+        by (auto intro!: truncate_up_mono intro: order_trans[where y=0]
+            simp: abs_le_square_iff[symmetric] abs_real_def
+            div2_less_self)
+      then show ?thesis
+        unfolding j
+        by (simp add: power_up_simp)
+    qed
+  qed simp
+qed
+
 
 subsection \<open>Lemmas needed by Approximate\<close>
 
@@ -1948,160 +2366,6 @@
   "0 \<le> real_of_float x \<Longrightarrow> real_of_float y \<le> 0 \<Longrightarrow> real_of_float (float_divr prec x y) \<le> 0"
   by transfer (rule real_divr_nonneg_neg_upper_bound)
 
-lemma truncate_up_nonneg_mono:
-  assumes "0 \<le> x" "x \<le> y"
-  shows "truncate_up prec x \<le> truncate_up prec y"
-proof -
-  consider "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>" | "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>" "0 < x" | "x \<le> 0"
-    by arith
-  then show ?thesis
-  proof cases
-    case 1
-    then show ?thesis
-      using assms
-      by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono)
-  next
-    case 2
-    from assms \<open>0 < x\<close> have "log 2 x \<le> log 2 y"
-      by auto
-    with \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close>
-    have logless: "log 2 x < log 2 y"
-      by linarith
-    have flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
-      using \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close> \<open>log 2 x \<le> log 2 y\<close> by linarith
-    have "truncate_up prec x =
-      real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor> )\<rceil> * 2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)"
-      using assms by (simp add: truncate_up_def round_up_def)
-    also have "\<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> (2 ^ (Suc prec))"
-    proof (simp only: ceiling_le_iff)
-      have "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le>
-        x * (2 powr real (Suc prec) / (2 powr log 2 x))"
-        using real_of_int_floor_add_one_ge[of "log 2 x"] assms
-        by (auto simp: algebra_simps simp flip: powr_diff intro!: mult_left_mono)
-      then show "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> real_of_int ((2::int) ^ (Suc prec))"
-        using \<open>0 < x\<close> by (simp add: powr_realpow powr_add)
-    qed
-    then have "real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> 2 powr int (Suc prec)"
-      by (auto simp: powr_realpow powr_add)
-        (metis power_Suc of_int_le_numeral_power_cancel_iff)
-    also
-    have "2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)"
-      using logless flogless by (auto intro!: floor_mono)
-    also have "2 powr real_of_int (int (Suc prec)) \<le>
-        2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1))"
-      using assms \<open>0 < x\<close>
-      by (auto simp: algebra_simps)
-    finally have "truncate_up prec x \<le>
-        2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)"
-      by simp
-    also have "\<dots> = 2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor>) - real_of_int (int prec - \<lfloor>log 2 y\<rfloor>))"
-      by (subst powr_add[symmetric]) simp
-    also have "\<dots> = y"
-      using \<open>0 < x\<close> assms
-      by (simp add: powr_add)
-    also have "\<dots> \<le> truncate_up prec y"
-      by (rule truncate_up)
-    finally show ?thesis .
-  next
-    case 3
-    then show ?thesis
-      using assms
-      by (auto intro!: truncate_up_le)
-  qed
-qed
-
-lemma truncate_up_switch_sign_mono:
-  assumes "x \<le> 0" "0 \<le> y"
-  shows "truncate_up prec x \<le> truncate_up prec y"
-proof -
-  note truncate_up_nonpos[OF \<open>x \<le> 0\<close>]
-  also note truncate_up_le[OF \<open>0 \<le> y\<close>]
-  finally show ?thesis .
-qed
-
-lemma truncate_down_switch_sign_mono:
-  assumes "x \<le> 0"
-    and "0 \<le> y"
-    and "x \<le> y"
-  shows "truncate_down prec x \<le> truncate_down prec y"
-proof -
-  note truncate_down_le[OF \<open>x \<le> 0\<close>]
-  also note truncate_down_nonneg[OF \<open>0 \<le> y\<close>]
-  finally show ?thesis .
-qed
-
-lemma truncate_down_nonneg_mono:
-  assumes "0 \<le> x" "x \<le> y"
-  shows "truncate_down prec x \<le> truncate_down prec y"
-proof -
-  consider "x \<le> 0" | "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" |
-    "0 < x" "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
-    by arith
-  then show ?thesis
-  proof cases
-    case 1
-    with assms have "x = 0" "0 \<le> y" by simp_all
-    then show ?thesis
-      by (auto intro!: truncate_down_nonneg)
-  next
-    case 2
-    then show ?thesis
-      using assms
-      by (auto simp: truncate_down_def round_down_def intro!: floor_mono)
-  next
-    case 3
-    from \<open>0 < x\<close> have "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y"
-      using assms by auto
-    with \<open>\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>\<close>
-    have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
-      unfolding atomize_conj abs_of_pos[OF \<open>0 < x\<close>] abs_of_pos[OF \<open>0 < y\<close>]
-      by (metis floor_less_cancel linorder_cases not_le)
-    have "2 powr prec \<le> y * 2 powr real prec / (2 powr log 2 y)"
-      using \<open>0 < y\<close> by simp
-    also have "\<dots> \<le> y * 2 powr real (Suc prec) / (2 powr (real_of_int \<lfloor>log 2 y\<rfloor> + 1))"
-      using \<open>0 \<le> y\<close> \<open>0 \<le> x\<close> assms(2)
-      by (auto intro!: powr_mono divide_left_mono
-          simp: of_nat_diff powr_add powr_diff)
-    also have "\<dots> = y * 2 powr real (Suc prec) / (2 powr real_of_int \<lfloor>log 2 y\<rfloor> * 2)"
-      by (auto simp: powr_add)
-    finally have "(2 ^ prec) \<le> \<lfloor>y * 2 powr real_of_int (int (Suc prec) - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
-      using \<open>0 \<le> y\<close>
-      by (auto simp: powr_diff le_floor_iff powr_realpow powr_add)
-    then have "(2 ^ (prec)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>) \<le> truncate_down prec y"
-      by (auto simp: truncate_down_def round_down_def)
-    moreover have "x \<le> (2 ^ prec) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)"
-    proof -
-      have "x = 2 powr (log 2 \<bar>x\<bar>)" using \<open>0 < x\<close> by simp
-      also have "\<dots> \<le> (2 ^ (Suc prec )) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)"
-        using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"] \<open>0 < x\<close>
-        by (auto simp flip: powr_realpow powr_add simp: algebra_simps powr_mult_base le_powr_iff)
-      also
-      have "2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> + 1)"
-        using logless flogless \<open>x > 0\<close> \<open>y > 0\<close>
-        by (auto intro!: floor_mono)
-      finally show ?thesis
-        by (auto simp flip: powr_realpow simp: powr_diff assms of_nat_diff)
-    qed
-    ultimately show ?thesis
-      by (metis dual_order.trans truncate_down)
-  qed
-qed
-
-lemma truncate_down_eq_truncate_up: "truncate_down p x = - truncate_up p (-x)"
-  and truncate_up_eq_truncate_down: "truncate_up p x = - truncate_down p (-x)"
-  by (auto simp: truncate_up_uminus_eq truncate_down_uminus_eq)
-
-lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y"
-  apply (cases "0 \<le> x")
-  apply (rule truncate_down_nonneg_mono, assumption+)
-  apply (simp add: truncate_down_eq_truncate_up)
-  apply (cases "0 \<le> y")
-  apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono)
-  done
-
-lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y"
-  by (simp add: truncate_up_eq_truncate_down truncate_down_mono)
-
 lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0"
   by (auto simp: zero_float_def mult_le_0_iff)
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Interval.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -0,0 +1,854 @@
+(* Title: Interval
+   Author: Christoph Traut, TU Muenchen
+           Fabian Immler, TU Muenchen
+*)
+section \<open>Interval Type\<close>
+theory Interval
+  imports
+    Complex_Main
+    Lattice_Algebras
+    Set_Algebras
+begin
+
+text \<open>A type of non-empty, closed intervals.\<close>
+
+typedef (overloaded) 'a interval =
+  "{(a::'a::preorder, b). a \<le> b}"
+  morphisms bounds_of_interval Interval
+  by auto
+
+setup_lifting type_definition_interval
+
+lift_definition lower::"('a::preorder) interval \<Rightarrow> 'a" is fst .
+
+lift_definition upper::"('a::preorder) interval \<Rightarrow> 'a" is snd .
+
+lemma interval_eq_iff: "a = b \<longleftrightarrow> lower a = lower b \<and> upper a = upper b"
+  by transfer auto
+
+lemma interval_eqI: "lower a = lower b \<Longrightarrow> upper a = upper b \<Longrightarrow> a = b"
+  by (auto simp: interval_eq_iff)
+
+lemma lower_le_upper[simp]: "lower i \<le> upper i"
+  by transfer auto
+
+lift_definition set_of :: "'a::preorder interval \<Rightarrow> 'a set" is "\<lambda>x. {fst x .. snd x}" .
+
+lemma set_of_eq: "set_of x = {lower x .. upper x}"
+  by transfer simp
+
+context notes [[typedef_overloaded]] begin
+
+lift_definition(code_dt) Interval'::"'a::preorder \<Rightarrow> 'a::preorder \<Rightarrow> 'a interval option"
+  is "\<lambda>a b. if a \<le> b then Some (a, b) else None"
+  by auto
+
+lemma Interval'_split:
+  "P (Interval' a b) \<longleftrightarrow>
+    (\<forall>ivl. a \<le> b \<longrightarrow> lower ivl = a \<longrightarrow> upper ivl = b \<longrightarrow> P (Some ivl)) \<and> (\<not>a\<le>b \<longrightarrow> P None)"
+  by transfer auto
+
+lemma Interval'_split_asm:
+  "P (Interval' a b) \<longleftrightarrow>
+    \<not>((\<exists>ivl. a \<le> b \<and> lower ivl = a \<and> upper ivl = b \<and> \<not>P (Some ivl)) \<or> (\<not>a\<le>b \<and> \<not>P None))"
+  unfolding Interval'_split
+  by auto
+
+lemmas Interval'_splits = Interval'_split Interval'_split_asm
+
+lemma Interval'_eq_Some: "Interval' a b = Some i \<Longrightarrow> lower i = a \<and> upper i = b"
+  by (simp split: Interval'_splits)
+
+end
+
+instantiation "interval" :: ("{preorder,equal}") equal
+begin
+
+definition "equal_class.equal a b \<equiv> (lower a = lower b) \<and> (upper a = upper b)"
+
+instance proof qed (simp add: equal_interval_def interval_eq_iff)
+end
+
+instantiation interval :: ("preorder") ord begin
+
+definition less_eq_interval :: "'a interval \<Rightarrow> 'a interval \<Rightarrow> bool"
+  where "less_eq_interval a b \<longleftrightarrow> lower b \<le> lower a \<and> upper a \<le> upper b"
+
+definition less_interval :: "'a interval \<Rightarrow> 'a interval \<Rightarrow> bool"
+  where  "less_interval x y = (x \<le> y \<and> \<not> y \<le> x)"
+
+instance proof qed
+end
+
+instantiation interval :: ("lattice") semilattice_sup
+begin
+
+lift_definition sup_interval :: "'a interval \<Rightarrow> 'a interval \<Rightarrow> 'a interval"
+  is "\<lambda>(a, b) (c, d). (inf a c, sup b d)"
+  by (auto simp: le_infI1 le_supI1)
+
+lemma lower_sup[simp]: "lower (sup A B) = inf (lower A) (lower B)"
+  by transfer auto
+
+lemma upper_sup[simp]: "upper (sup A B) = sup (upper A) (upper B)"
+  by transfer auto
+
+instance proof qed (auto simp: less_eq_interval_def less_interval_def interval_eq_iff)
+end
+
+lemma set_of_interval_union: "set_of A \<union> set_of B \<subseteq> set_of (sup A B)" for A::"'a::lattice interval"
+  by (auto simp: set_of_eq)
+
+lemma interval_union_commute: "sup A B = sup B A" for A::"'a::lattice interval"
+  by (auto simp add: interval_eq_iff inf.commute sup.commute)
+
+lemma interval_union_mono1: "set_of a \<subseteq> set_of (sup a A)" for A :: "'a::lattice interval"
+  using set_of_interval_union by blast
+
+lemma interval_union_mono2: "set_of A \<subseteq> set_of (sup a A)" for A :: "'a::lattice interval"
+  using set_of_interval_union by blast
+
+lift_definition interval_of :: "'a::preorder \<Rightarrow> 'a interval" is "\<lambda>x. (x, x)"
+  by auto
+
+lemma lower_interval_of[simp]: "lower (interval_of a) = a"
+  by transfer auto
+
+lemma upper_interval_of[simp]: "upper (interval_of a) = a"
+  by transfer auto
+
+definition width :: "'a::{preorder,minus} interval \<Rightarrow> 'a"
+  where "width i = upper i - lower i"
+
+
+instantiation "interval" :: ("ordered_ab_semigroup_add") ab_semigroup_add
+begin
+
+lift_definition plus_interval::"'a interval \<Rightarrow> 'a interval \<Rightarrow> 'a interval"
+  is "\<lambda>(a, b). \<lambda>(c, d). (a + c, b + d)"
+  by (auto intro!: add_mono)
+lemma lower_plus[simp]: "lower (plus A B) = plus (lower A) (lower B)"
+  by transfer auto
+lemma upper_plus[simp]: "upper (plus A B) = plus (upper A) (upper B)"
+  by transfer auto
+
+instance proof qed (auto simp: interval_eq_iff less_eq_interval_def ac_simps)
+end
+
+instance "interval" :: ("{ordered_ab_semigroup_add, lattice}") ordered_ab_semigroup_add
+proof qed (auto simp: less_eq_interval_def intro!: add_mono)
+
+instantiation "interval" :: ("{preorder,zero}") zero
+begin
+
+lift_definition zero_interval::"'a interval" is "(0, 0)" by auto
+lemma lower_zero[simp]: "lower 0 = 0"
+  by transfer auto
+lemma upper_zero[simp]: "upper 0 = 0"
+  by transfer auto
+instance proof qed
+end
+
+instance "interval" :: ("{ordered_comm_monoid_add}") comm_monoid_add
+proof qed (auto simp: interval_eq_iff)
+
+instance "interval" :: ("{ordered_comm_monoid_add,lattice}") ordered_comm_monoid_add ..
+
+instantiation "interval" :: ("{ordered_ab_group_add}") uminus
+begin
+
+lift_definition uminus_interval::"'a interval \<Rightarrow> 'a interval" is "\<lambda>(a, b). (-b, -a)" by auto
+lemma lower_uminus[simp]: "lower (- A) = - upper A"
+  by transfer auto
+lemma upper_uminus[simp]: "upper (- A) = - lower A"
+  by transfer auto
+instance ..
+end
+
+instantiation "interval" :: ("{ordered_ab_group_add}") minus
+begin
+
+definition minus_interval::"'a interval \<Rightarrow> 'a interval \<Rightarrow> 'a interval"
+  where "minus_interval a b = a + - b"
+lemma lower_minus[simp]: "lower (minus A B) = minus (lower A) (upper B)"
+  by (auto simp: minus_interval_def)
+lemma upper_minus[simp]: "upper (minus A B) = minus (upper A) (lower B)"
+  by (auto simp: minus_interval_def)
+
+instance ..
+end
+
+instantiation "interval" :: (linordered_semiring) times
+begin
+
+lift_definition times_interval :: "'a interval \<Rightarrow> 'a interval \<Rightarrow> 'a interval"
+  is "\<lambda>(a1, a2). \<lambda>(b1, b2).
+    (let x1 = a1 * b1; x2 = a1 * b2; x3 = a2 * b1; x4 = a2 * b2
+    in (min x1 (min x2 (min x3 x4)), max x1 (max x2 (max x3 x4))))"
+  by (auto simp: Let_def intro!: min.coboundedI1 max.coboundedI1)
+
+lemma lower_times:
+  "lower (times A B) = Min {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}"
+  by transfer (auto simp: Let_def)
+
+lemma upper_times:
+  "upper (times A B) = Max {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}"
+  by transfer (auto simp: Let_def)
+
+instance ..
+end
+
+lemma interval_eq_set_of_iff: "X = Y \<longleftrightarrow> set_of X = set_of Y" for X Y::"'a::order interval"
+  by (auto simp: set_of_eq interval_eq_iff)
+
+
+subsection \<open>Membership\<close>
+
+abbreviation (in preorder) in_interval ("(_/ \<in>\<^sub>i _)" [51, 51] 50)
+  where "in_interval x X \<equiv> x \<in> set_of X"
+
+lemma in_interval_to_interval[intro!]: "a \<in>\<^sub>i interval_of a"
+  by (auto simp: set_of_eq)
+
+lemma plus_in_intervalI:
+  fixes x y :: "'a :: ordered_ab_semigroup_add"
+  shows "x \<in>\<^sub>i X \<Longrightarrow> y \<in>\<^sub>i Y \<Longrightarrow> x + y \<in>\<^sub>i X + Y"
+  by (simp add: add_mono_thms_linordered_semiring(1) set_of_eq)
+
+lemma connected_set_of[intro, simp]:
+  "connected (set_of X)" for X::"'a::linear_continuum_topology interval"
+  by (auto simp: set_of_eq )
+
+lemma ex_sum_in_interval_lemma: "\<exists>xa\<in>{la .. ua}. \<exists>xb\<in>{lb .. ub}. x = xa + xb"
+  if "la \<le> ua" "lb \<le> ub" "la + lb \<le> x" "x \<le> ua + ub"
+    "ua - la \<le> ub - lb"
+  for la b c d::"'a::linordered_ab_group_add"
+proof -
+  define wa where "wa = ua - la"
+  define wb where "wb = ub - lb"
+  define w where "w = wa + wb"
+  define d where "d = x - la - lb"
+  define da where "da = max 0 (min wa (d - wa))"
+  define db where "db = d - da"
+  from that have nonneg: "0 \<le> wa" "0 \<le> wb" "0 \<le> w" "0 \<le> d" "d \<le> w"
+    by (auto simp add: wa_def wb_def w_def d_def add.commute le_diff_eq)
+  have "0 \<le> db"
+    by (auto simp: da_def nonneg db_def intro!: min.coboundedI2)
+  have "x = (la + da) + (lb + db)"
+    by (simp add: da_def db_def d_def)
+  moreover
+  have "x - la - ub \<le> da"
+    using that
+    unfolding da_def
+    by (intro max.coboundedI2) (auto simp: wa_def d_def diff_le_eq diff_add_eq)
+  then have "db \<le> wb"
+    by (auto simp: db_def d_def wb_def algebra_simps)
+  with \<open>0 \<le> db\<close> that nonneg have "lb + db \<in> {lb..ub}"
+    by (auto simp: wb_def algebra_simps)
+  moreover
+  have "da \<le> wa"
+    by (auto simp: da_def nonneg)
+  then have "la + da \<in> {la..ua}"
+    by (auto simp: da_def wa_def algebra_simps)
+  ultimately show ?thesis
+    by force
+qed
+
+
+lemma ex_sum_in_interval: "\<exists>xa\<ge>la. xa \<le> ua \<and> (\<exists>xb\<ge>lb. xb \<le> ub \<and> x = xa + xb)"
+  if a: "la \<le> ua" and b: "lb \<le> ub" and x: "la + lb \<le> x" "x \<le> ua + ub"
+  for la b c d::"'a::linordered_ab_group_add"
+proof -
+  from linear consider "ua - la \<le> ub - lb" | "ub - lb \<le> ua - la"
+    by blast
+  then show ?thesis
+  proof cases
+    case 1
+    from ex_sum_in_interval_lemma[OF that 1]
+    show ?thesis by auto
+  next
+    case 2
+    from x have "lb + la \<le> x" "x \<le> ub + ua" by (simp_all add: ac_simps)
+    from ex_sum_in_interval_lemma[OF b a this 2]
+    show ?thesis by auto
+  qed
+qed
+
+lemma Icc_plus_Icc:
+  "{a .. b} + {c .. d} = {a + c .. b + d}"
+  if "a \<le> b" "c \<le> d"
+  for a b c d::"'a::linordered_ab_group_add"
+  using ex_sum_in_interval[OF that]
+  by (auto intro: add_mono simp: atLeastAtMost_iff Bex_def set_plus_def)
+
+lemma set_of_plus:
+  fixes A :: "'a::linordered_ab_group_add interval"
+  shows "set_of (A + B) = set_of A + set_of B"
+  using Icc_plus_Icc[of "lower A" "upper A" "lower B" "upper B"]
+  by (auto simp: set_of_eq)
+
+lemma plus_in_intervalE:
+  fixes xy :: "'a :: linordered_ab_group_add"
+  assumes "xy \<in>\<^sub>i X + Y"
+  obtains x y where "xy = x + y" "x \<in>\<^sub>i X" "y \<in>\<^sub>i Y"
+  using assms
+  unfolding set_of_plus set_plus_def
+  by auto
+
+lemma set_of_uminus: "set_of (-X) = {- x | x. x \<in> set_of X}"
+  for X :: "'a :: ordered_ab_group_add interval"
+  by (auto simp: set_of_eq simp: le_minus_iff minus_le_iff
+      intro!: exI[where x="-x" for x])
+
+lemma uminus_in_intervalI:
+  fixes x :: "'a :: ordered_ab_group_add"
+  shows "x \<in>\<^sub>i X \<Longrightarrow> -x \<in>\<^sub>i -X"
+  by (auto simp: set_of_uminus)
+
+lemma uminus_in_intervalD:
+  fixes x :: "'a :: ordered_ab_group_add"
+  shows "x \<in>\<^sub>i - X \<Longrightarrow> - x \<in>\<^sub>i X"
+  by (auto simp: set_of_uminus)
+
+lemma minus_in_intervalI:
+  fixes x y :: "'a :: ordered_ab_group_add"
+  shows "x \<in>\<^sub>i X \<Longrightarrow> y \<in>\<^sub>i Y \<Longrightarrow> x - y \<in>\<^sub>i X - Y"
+  by (metis diff_conv_add_uminus minus_interval_def plus_in_intervalI uminus_in_intervalI)
+
+lemma set_of_minus: "set_of (X - Y) = {x - y | x y . x \<in> set_of X \<and> y \<in> set_of Y}"
+  for X Y :: "'a :: linordered_ab_group_add interval"
+  unfolding minus_interval_def set_of_plus set_of_uminus set_plus_def
+  by force
+
+lemma times_in_intervalI:
+  fixes x y::"'a::linordered_ring"
+  assumes "x \<in>\<^sub>i X" "y \<in>\<^sub>i Y"
+  shows "x * y \<in>\<^sub>i X * Y"
+proof -
+  define X1 where "X1 \<equiv> lower X"
+  define X2 where "X2 \<equiv> upper X"
+  define Y1 where "Y1 \<equiv> lower Y"
+  define Y2 where "Y2 \<equiv> upper Y"
+  from assms have assms: "X1 \<le> x" "x \<le> X2" "Y1 \<le> y" "y \<le> Y2"
+    by (auto simp: X1_def X2_def Y1_def Y2_def set_of_eq)
+  have "(X1 * Y1 \<le> x * y \<or> X1 * Y2 \<le> x * y \<or> X2 * Y1 \<le> x * y \<or> X2 * Y2 \<le> x * y) \<and>
+        (X1 * Y1 \<ge> x * y \<or> X1 * Y2 \<ge> x * y \<or> X2 * Y1 \<ge> x * y \<or> X2 * Y2 \<ge> x * y)"
+  proof (cases x "0::'a" rule: linorder_cases)
+    case x0: less
+    show ?thesis
+    proof (cases "y < 0")
+      case y0: True
+      from y0 x0 assms have "x * y \<le> X1 * y" by (intro mult_right_mono_neg, auto)
+      also from x0 y0 assms have "X1 * y \<le> X1 * Y1" by (intro mult_left_mono_neg, auto)
+      finally have 1: "x * y \<le> X1 * Y1".
+      show ?thesis proof(cases "X2 \<le> 0")
+        case True
+        with assms have "X2 * Y2 \<le> X2 * y" by (auto intro: mult_left_mono_neg)
+        also from assms y0 have "... \<le> x * y" by (auto intro: mult_right_mono_neg)
+        finally have "X2 * Y2 \<le> x * y".
+        with 1 show ?thesis by auto
+      next
+        case False
+        with assms have "X2 * Y1 \<le> X2 * y" by (auto intro: mult_left_mono)
+        also from assms y0 have "... \<le> x * y" by (auto intro: mult_right_mono_neg)
+        finally have "X2 * Y1 \<le> x * y".
+        with 1 show ?thesis by auto
+      qed
+    next
+      case False
+      then have y0: "y \<ge> 0" by auto
+      from x0 y0 assms have "X1 * Y2 \<le> x * Y2" by (intro mult_right_mono, auto)
+      also from y0 x0 assms have "... \<le> x * y" by (intro mult_left_mono_neg, auto)
+      finally have 1: "X1 * Y2 \<le> x * y".
+      show ?thesis
+      proof(cases "X2 \<le> 0")
+        case X2: True
+        from assms y0 have "x * y \<le> X2 * y" by (intro mult_right_mono)
+        also from assms X2 have "... \<le> X2 * Y1" by (auto intro: mult_left_mono_neg)
+        finally have "x * y \<le> X2 * Y1".
+        with 1 show ?thesis by auto
+      next
+        case X2: False
+        from assms y0 have "x * y \<le> X2 * y" by (intro mult_right_mono)
+        also from assms X2 have "... \<le> X2 * Y2" by (auto intro: mult_left_mono)
+        finally have "x * y \<le> X2 * Y2".
+        with 1 show ?thesis by auto
+      qed
+    qed
+  next
+    case [simp]: equal
+    with assms show ?thesis by (cases "Y2 \<le> 0", auto intro:mult_sign_intros)
+  next
+    case x0: greater
+    show ?thesis
+    proof (cases "y < 0")
+      case y0: True
+      from x0 y0 assms have "X2 * Y1 \<le> X2 * y" by (intro mult_left_mono, auto)
+      also from y0 x0 assms have "X2 * y \<le> x * y" by (intro mult_right_mono_neg, auto)
+      finally have 1: "X2 * Y1 \<le> x * y".
+      show ?thesis
+      proof(cases "Y2 \<le> 0")
+        case Y2: True
+        from x0 assms have "x * y \<le> x * Y2" by (auto intro: mult_left_mono)
+        also from assms Y2 have "... \<le> X1 * Y2" by (auto intro: mult_right_mono_neg)
+        finally have "x * y \<le> X1 * Y2".
+        with 1 show ?thesis by auto
+      next
+        case Y2: False
+        from x0 assms have "x * y \<le> x * Y2" by (auto intro: mult_left_mono)
+        also from assms Y2 have "... \<le> X2 * Y2" by (auto intro: mult_right_mono)
+        finally have "x * y \<le> X2 * Y2".
+        with 1 show ?thesis by auto
+      qed
+    next
+      case y0: False
+      from x0 y0 assms have "x * y \<le> X2 * y" by (intro mult_right_mono, auto)
+      also from y0 x0 assms have "... \<le> X2 * Y2" by (intro mult_left_mono, auto)
+      finally have 1: "x * y \<le> X2 * Y2".
+      show ?thesis
+      proof(cases "X1 \<le> 0")
+        case True
+        with assms have "X1 * Y2 \<le> X1 * y" by (auto intro: mult_left_mono_neg)
+        also from assms y0 have "... \<le> x * y" by (auto intro: mult_right_mono)
+        finally have "X1 * Y2 \<le> x * y".
+        with 1 show ?thesis by auto
+      next
+        case False
+        with assms have "X1 * Y1 \<le> X1 * y" by (auto intro: mult_left_mono)
+        also from assms y0 have "... \<le> x * y" by (auto intro: mult_right_mono)
+        finally have "X1 * Y1 \<le> x * y".
+        with 1 show ?thesis by auto
+      qed
+    qed
+  qed
+  hence min:"min (X1 * Y1) (min (X1 * Y2) (min (X2 * Y1) (X2 * Y2))) \<le> x * y"
+    and max:"x * y \<le> max (X1 * Y1) (max (X1 * Y2) (max (X2 * Y1) (X2 * Y2)))"
+    by (auto simp:min_le_iff_disj le_max_iff_disj)
+  show ?thesis using min max
+    by (auto simp: Let_def X1_def X2_def Y1_def Y2_def set_of_eq lower_times upper_times)
+qed
+
+lemma times_in_intervalE:
+  fixes xy :: "'a :: {linordered_semiring, real_normed_algebra, linear_continuum_topology}"
+    \<comment> \<open>TODO: linear continuum topology is pretty strong\<close>
+  assumes "xy \<in>\<^sub>i X * Y"
+  obtains x y where "xy = x * y" "x \<in>\<^sub>i X" "y \<in>\<^sub>i Y"
+proof -
+  let ?mult = "\<lambda>(x, y). x * y"
+  let ?XY = "set_of X \<times> set_of Y"
+  have cont: "continuous_on ?XY ?mult"
+    by (auto intro!: tendsto_eq_intros simp: continuous_on_def split_beta')
+  have conn: "connected (?mult ` ?XY)"
+    by (rule connected_continuous_image[OF cont]) auto
+  have "lower (X * Y) \<in> ?mult ` ?XY" "upper (X * Y) \<in> ?mult ` ?XY"
+    by (auto simp: set_of_eq lower_times upper_times min_def max_def split: if_splits)
+  from connectedD_interval[OF conn this, of xy] assms
+  obtain x y where "xy = x * y" "x \<in>\<^sub>i X" "y \<in>\<^sub>i Y" by (auto simp: set_of_eq)
+  then show ?thesis ..
+qed
+
+lemma set_of_times: "set_of (X * Y) = {x * y | x y. x \<in> set_of X \<and> y \<in> set_of Y}"
+  for X Y::"'a :: {linordered_ring, real_normed_algebra, linear_continuum_topology} interval"
+  by (auto intro!: times_in_intervalI elim!: times_in_intervalE)
+
+instance "interval" :: (linordered_idom) cancel_semigroup_add
+proof qed (auto simp: interval_eq_iff)
+
+lemma interval_mul_commute: "A * B = B * A" for A B:: "'a::linordered_idom interval"
+  by (simp add: interval_eq_iff lower_times upper_times ac_simps)
+
+lemma interval_times_zero_right[simp]: "A * 0 = 0" for A :: "'a::linordered_ring interval"
+  by (simp add: interval_eq_iff lower_times upper_times ac_simps)
+
+lemma interval_times_zero_left[simp]:
+  "0 * A = 0" for A :: "'a::linordered_ring interval"
+  by (simp add: interval_eq_iff lower_times upper_times ac_simps)
+
+instantiation "interval" :: ("{preorder,one}") one
+begin
+
+lift_definition one_interval::"'a interval" is "(1, 1)" by auto
+lemma lower_one[simp]: "lower 1 = 1"
+  by transfer auto
+lemma upper_one[simp]: "upper 1 = 1"
+  by transfer auto
+instance proof qed
+end
+
+instance interval :: ("{one, preorder, linordered_semiring}") power
+proof qed
+
+lemma set_of_one[simp]: "set_of (1::'a::{one, order} interval) = {1}"
+  by (auto simp: set_of_eq)
+
+instance "interval" ::
+  ("{linordered_idom,linordered_ring, real_normed_algebra, linear_continuum_topology}") monoid_mult
+  apply standard
+  unfolding interval_eq_set_of_iff set_of_times
+  subgoal
+    by (auto simp: interval_eq_set_of_iff set_of_times; metis mult.assoc)
+  by auto
+
+lemma one_times_ivl_left[simp]: "1 * A = A" for A :: "'a::linordered_idom interval"
+  by (simp add: interval_eq_iff lower_times upper_times ac_simps min_def max_def)
+
+lemma one_times_ivl_right[simp]: "A * 1 = A" for A :: "'a::linordered_idom interval"
+  by (metis interval_mul_commute one_times_ivl_left)
+
+lemma set_of_power_mono: "a^n \<in> set_of (A^n)" if "a \<in> set_of A"
+  for a :: "'a::linordered_idom"
+  using that
+  by (induction n) (auto intro!: times_in_intervalI)
+
+lemma set_of_add_cong:
+  "set_of (A + B) = set_of (A' + B')"
+  if "set_of A = set_of A'" "set_of B = set_of B'"
+  for A :: "'a::linordered_ab_group_add interval"
+  unfolding set_of_plus that ..
+
+lemma set_of_add_inc_left:
+  "set_of (A + B) \<subseteq> set_of (A' + B)"
+  if "set_of A \<subseteq> set_of A'"
+  for A :: "'a::linordered_ab_group_add interval"
+  unfolding set_of_plus using that by (auto simp: set_plus_def)
+
+lemma set_of_add_inc_right:
+  "set_of (A + B) \<subseteq> set_of (A + B')"
+  if "set_of B \<subseteq> set_of B'"
+  for A :: "'a::linordered_ab_group_add interval"
+  using set_of_add_inc_left[OF that]
+  by (simp add: add.commute)
+
+lemma set_of_add_inc:
+  "set_of (A + B) \<subseteq> set_of (A' + B')"
+  if "set_of A \<subseteq> set_of A'" "set_of B \<subseteq> set_of B'"
+  for A :: "'a::linordered_ab_group_add interval"
+  using set_of_add_inc_left[OF that(1)] set_of_add_inc_right[OF that(2)]
+  by auto
+
+lemma set_of_neg_inc:
+  "set_of (-A) \<subseteq> set_of (-A')"
+  if "set_of A \<subseteq> set_of A'"
+  for A :: "'a::ordered_ab_group_add interval"
+  using that
+  unfolding set_of_uminus
+  by auto
+
+lemma set_of_sub_inc_left:
+  "set_of (A - B) \<subseteq> set_of (A' - B)"
+  if "set_of A \<subseteq> set_of A'"
+  for A :: "'a::linordered_ab_group_add interval"
+  using that
+  unfolding set_of_minus
+  by auto
+
+lemma set_of_sub_inc_right:
+  "set_of (A - B) \<subseteq> set_of (A - B')"
+  if "set_of B \<subseteq> set_of B'"
+  for A :: "'a::linordered_ab_group_add interval"
+  using that
+  unfolding set_of_minus
+  by auto
+
+lemma set_of_sub_inc:
+  "set_of (A - B) \<subseteq> set_of (A' - B')"
+  if "set_of A \<subseteq> set_of A'" "set_of B \<subseteq> set_of B'"
+  for A :: "'a::linordered_idom interval"
+  using set_of_sub_inc_left[OF that(1)] set_of_sub_inc_right[OF that(2)]
+  by auto
+
+lemma set_of_mul_inc_right:
+  "set_of (A * B) \<subseteq> set_of (A * B')"
+  if "set_of B \<subseteq> set_of B'"
+  for A :: "'a::linordered_ring interval"
+  using that
+  apply transfer
+  apply (clarsimp simp add: Let_def)
+  apply (intro conjI)
+         apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+        apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+       apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+      apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+     apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+    apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+   apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+  apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans)
+  done
+
+lemma set_of_distrib_left:
+  "set_of (B * (A1 + A2)) \<subseteq> set_of (B * A1 + B * A2)"
+  for A1 :: "'a::linordered_ring interval"
+  apply transfer
+  apply (clarsimp simp: Let_def distrib_left distrib_right)
+  apply (intro conjI)
+         apply (metis add_mono min.cobounded1 min.left_commute)
+        apply (metis add_mono min.cobounded1 min.left_commute)
+       apply (metis add_mono min.cobounded1 min.left_commute)
+      apply (metis add_mono min.assoc min.cobounded2)
+     apply (meson add_mono order.trans max.cobounded1 max.cobounded2)
+    apply (meson add_mono order.trans max.cobounded1 max.cobounded2)
+   apply (meson add_mono order.trans max.cobounded1 max.cobounded2)
+  apply (meson add_mono order.trans max.cobounded1 max.cobounded2)
+  done
+
+lemma set_of_distrib_right:
+  "set_of ((A1 + A2) * B) \<subseteq> set_of (A1 * B + A2 * B)"
+  for A1 A2 B :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval"
+  unfolding set_of_times set_of_plus set_plus_def
+  apply clarsimp
+  subgoal for b a1 a2
+    apply (rule exI[where x="a1 * b"])
+    apply (rule conjI)
+    subgoal by force
+    subgoal
+      apply (rule exI[where x="a2 * b"])
+      apply (rule conjI)
+      subgoal by force
+      subgoal by (simp add: algebra_simps)
+      done
+    done
+  done
+
+lemma set_of_mul_inc_left:
+  "set_of (A * B) \<subseteq> set_of (A' * B)"
+  if "set_of A \<subseteq> set_of A'"
+  for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval"
+  using that
+  unfolding set_of_times
+  by auto
+
+lemma set_of_mul_inc:
+  "set_of (A * B) \<subseteq> set_of (A' * B')"
+  if "set_of A \<subseteq> set_of A'" "set_of B \<subseteq> set_of B'"
+  for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval"
+  using that unfolding set_of_times by auto
+
+lemma set_of_pow_inc:
+  "set_of (A^n) \<subseteq> set_of (A'^n)"
+  if "set_of A \<subseteq> set_of A'"
+  for A :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval"
+  using that
+  by (induction n, simp_all add: set_of_mul_inc)
+
+lemma set_of_distrib_right_left:
+  "set_of ((A1 + A2) * (B1 + B2)) \<subseteq> set_of (A1 * B1 + A1 * B2 + A2 * B1 + A2 * B2)"
+  for A1 :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval"
+proof-
+  have "set_of ((A1 + A2) * (B1 + B2)) \<subseteq> set_of (A1 * (B1 + B2) + A2 * (B1 + B2))"
+    by (rule set_of_distrib_right)
+  also have "... \<subseteq> set_of ((A1 * B1 + A1 * B2) + A2 * (B1 + B2))"
+    by (rule set_of_add_inc_left[OF set_of_distrib_left])
+  also have "... \<subseteq> set_of ((A1 * B1 + A1 * B2) + (A2 * B1 + A2 * B2))"
+    by (rule set_of_add_inc_right[OF set_of_distrib_left])
+  finally show ?thesis
+    by (simp add: add.assoc)
+qed
+
+lemma mult_bounds_enclose_zero1:
+  "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \<le> 0"
+  "0 \<le> max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))"
+  if "la \<le> 0" "0 \<le> ua"
+  for la lb ua ub:: "'a::linordered_idom"
+  subgoal by (metis (no_types, hide_lams) that eq_iff min_le_iff_disj mult_zero_left mult_zero_right
+        zero_le_mult_iff)
+  subgoal by (metis that le_max_iff_disj mult_zero_right order_refl zero_le_mult_iff)
+  done
+
+lemma mult_bounds_enclose_zero2:
+  "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \<le> 0"
+  "0 \<le> max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))"
+  if "lb \<le> 0" "0 \<le> ub"
+  for la lb ua ub:: "'a::linordered_idom"
+  using mult_bounds_enclose_zero1[OF that, of la ua]
+  by (simp_all add: ac_simps)
+
+lemma set_of_mul_contains_zero:
+  "0 \<in> set_of (A * B)"
+  if "0 \<in> set_of A \<or> 0 \<in> set_of B"
+  for A :: "'a::linordered_idom interval"
+  using that
+  by (auto simp: set_of_eq lower_times upper_times algebra_simps mult_le_0_iff
+      mult_bounds_enclose_zero1 mult_bounds_enclose_zero2)
+
+instance "interval" :: (linordered_semiring) mult_zero
+  apply standard
+  subgoal by transfer auto
+  subgoal by transfer auto
+  done
+
+lift_definition min_interval::"'a::linorder interval \<Rightarrow> 'a interval \<Rightarrow> 'a interval" is
+  "\<lambda>(l1, u1). \<lambda>(l2, u2). (min l1 l2, min u1 u2)"
+  by (auto simp: min_def)
+lemma lower_min_interval[simp]: "lower (min_interval x y) = min (lower x) (lower y)"
+  by transfer auto
+lemma upper_min_interval[simp]: "upper (min_interval x y) = min (upper x) (upper y)"
+  by transfer auto
+
+lemma min_intervalI:
+  "a \<in>\<^sub>i A \<Longrightarrow> b \<in>\<^sub>i B \<Longrightarrow> min a b \<in>\<^sub>i min_interval A B"
+  by (auto simp: set_of_eq min_def)
+
+lift_definition max_interval::"'a::linorder interval \<Rightarrow> 'a interval \<Rightarrow> 'a interval" is
+  "\<lambda>(l1, u1). \<lambda>(l2, u2). (max l1 l2, max u1 u2)"
+  by (auto simp: max_def)
+lemma lower_max_interval[simp]: "lower (max_interval x y) = max (lower x) (lower y)"
+  by transfer auto
+lemma upper_max_interval[simp]: "upper (max_interval x y) = max (upper x) (upper y)"
+  by transfer auto
+
+lemma max_intervalI:
+  "a \<in>\<^sub>i A \<Longrightarrow> b \<in>\<^sub>i B \<Longrightarrow> max a b \<in>\<^sub>i max_interval A B"
+  by (auto simp: set_of_eq max_def)
+
+lift_definition abs_interval::"'a::linordered_idom interval \<Rightarrow> 'a interval" is
+  "(\<lambda>(l,u). (if l < 0 \<and> 0 < u then 0 else min \<bar>l\<bar> \<bar>u\<bar>, max \<bar>l\<bar> \<bar>u\<bar>))"
+  by auto
+
+lemma lower_abs_interval[simp]:
+  "lower (abs_interval x) = (if lower x < 0 \<and> 0 < upper x then 0 else min \<bar>lower x\<bar> \<bar>upper x\<bar>)"
+  by transfer auto
+lemma upper_abs_interval[simp]: "upper (abs_interval x) = max \<bar>lower x\<bar> \<bar>upper x\<bar>"
+  by transfer auto
+
+lemma in_abs_intervalI1:
+  "lx < 0 \<Longrightarrow> 0 < ux \<Longrightarrow> 0 \<le> xa \<Longrightarrow> xa \<le> max (- lx) (ux) \<Longrightarrow> xa \<in> abs ` {lx..ux}"
+  for xa::"'a::linordered_idom"
+  by (metis abs_minus_cancel abs_of_nonneg atLeastAtMost_iff image_eqI le_less le_max_iff_disj
+      le_minus_iff neg_le_0_iff_le order_trans)
+
+lemma in_abs_intervalI2:
+  "min (\<bar>lx\<bar>) \<bar>ux\<bar> \<le> xa \<Longrightarrow> xa \<le> max \<bar>lx\<bar> \<bar>ux\<bar> \<Longrightarrow> lx \<le> ux \<Longrightarrow> 0 \<le> lx \<or> ux \<le> 0 \<Longrightarrow>
+    xa \<in> abs ` {lx..ux}"
+  for xa::"'a::linordered_idom"
+  by (force intro: image_eqI[where x="-xa"] image_eqI[where x="xa"])
+
+lemma set_of_abs_interval: "set_of (abs_interval x) = abs ` set_of x"
+  by (auto simp: set_of_eq not_less intro: in_abs_intervalI1 in_abs_intervalI2 cong del: image_cong_simp)
+
+fun split_domain :: "('a::preorder interval \<Rightarrow> 'a interval list) \<Rightarrow> 'a interval list \<Rightarrow> 'a interval list list"
+  where "split_domain split [] = [[]]"
+  | "split_domain split (I#Is) = (
+         let S = split I;
+             D = split_domain split Is
+         in concat (map (\<lambda>d. map (\<lambda>s. s # d) S) D)
+       )"
+
+context notes [[typedef_overloaded]] begin
+lift_definition(code_dt) split_interval::"'a::linorder interval \<Rightarrow> 'a \<Rightarrow> ('a interval \<times> 'a interval)"
+  is "\<lambda>(l, u) x. ((min l x, max l x), (min u x, max u x))"
+  by (auto simp: min_def)
+end
+
+lemma split_domain_nonempty:
+  assumes "\<And>I. split I \<noteq> []"
+  shows "split_domain split I \<noteq> []"
+  using last_in_set assms
+  by (induction I, auto)
+
+lemma lower_split_interval1: "lower (fst (split_interval X m)) = min (lower X) m"
+  and lower_split_interval2: "lower (snd (split_interval X m)) = min (upper X) m"
+  and upper_split_interval1: "upper (fst (split_interval X m)) = max (lower X) m"
+  and upper_split_interval2: "upper (snd (split_interval X m)) = max (upper X) m"
+  subgoal by transfer auto
+  subgoal by transfer (auto simp: min.commute)
+  subgoal by transfer (auto simp: )
+  subgoal by transfer (auto simp: )
+  done
+
+lemma split_intervalD: "split_interval X x = (A, B) \<Longrightarrow> set_of X \<subseteq> set_of A \<union> set_of B"
+  unfolding set_of_eq
+  by transfer (auto simp: min_def max_def split: if_splits)
+
+instantiation interval :: ("{topological_space, preorder}") topological_space
+begin
+
+definition open_interval_def[code del]: "open (X::'a interval set) =
+  (\<forall>x\<in>X.
+      \<exists>A B.
+         open A \<and>
+         open B \<and>
+         lower x \<in> A \<and> upper x \<in> B \<and> Interval ` (A \<times> B) \<subseteq> X)"
+
+instance
+proof
+  show "open (UNIV :: ('a interval) set)"
+    unfolding open_interval_def by auto
+next
+  fix S T :: "('a interval) set"
+  assume "open S" "open T"
+  show "open (S \<inter> T)"
+    unfolding open_interval_def
+  proof (safe)
+    fix x assume "x \<in> S" "x \<in> T"
+    from \<open>x \<in> S\<close> \<open>open S\<close> obtain Sl Su where S:
+      "open Sl" "open Su" "lower x \<in> Sl" "upper x \<in> Su" "Interval ` (Sl \<times> Su) \<subseteq> S"
+      by (auto simp: open_interval_def)
+    from \<open>x \<in> T\<close> \<open>open T\<close> obtain Tl Tu where T:
+      "open Tl" "open Tu" "lower x \<in> Tl" "upper x \<in> Tu" "Interval ` (Tl \<times> Tu) \<subseteq> T"
+      by (auto simp: open_interval_def)
+
+    let ?L = "Sl \<inter> Tl" and ?U = "Su \<inter> Tu" 
+    have "open ?L \<and> open ?U \<and> lower x \<in> ?L \<and> upper x \<in> ?U \<and> Interval ` (?L \<times> ?U) \<subseteq> S \<inter> T"
+      using S T by (auto simp add: open_Int)
+    then show "\<exists>A B. open A \<and> open B \<and> lower x \<in> A \<and> upper x \<in> B \<and> Interval ` (A \<times> B) \<subseteq> S \<inter> T"
+      by fast
+  qed
+qed (unfold open_interval_def, fast)
+
+end
+
+
+subsection \<open>Quickcheck\<close>
+
+lift_definition Ivl::"'a \<Rightarrow> 'a::preorder \<Rightarrow> 'a interval" is "\<lambda>a b. (min a b, b)"
+  by (auto simp: min_def)
+
+instantiation interval :: ("{exhaustive,preorder}") exhaustive
+begin
+
+definition exhaustive_interval::"('a interval \<Rightarrow> (bool \<times> term list) option)
+     \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option"
+  where
+    "exhaustive_interval f d =
+    Quickcheck_Exhaustive.exhaustive (\<lambda>x. Quickcheck_Exhaustive.exhaustive (\<lambda>y. f (Ivl x y)) d) d"
+
+instance ..
+
+end
+
+definition (in term_syntax) [code_unfold]:
+  "valtermify_interval x y = Code_Evaluation.valtermify (Ivl::'a::{preorder,typerep}\<Rightarrow>_) {\<cdot>} x {\<cdot>} y"
+
+instantiation interval :: ("{full_exhaustive,preorder,typerep}") full_exhaustive
+begin
+
+definition full_exhaustive_interval::
+  "('a interval \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option)
+     \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" where
+  "full_exhaustive_interval f d =
+    Quickcheck_Exhaustive.full_exhaustive
+      (\<lambda>x. Quickcheck_Exhaustive.full_exhaustive (\<lambda>y. f (valtermify_interval x y)) d) d"
+
+instance ..
+
+end
+
+instantiation interval :: ("{random,preorder,typerep}") random
+begin
+
+definition random_interval ::
+  "natural
+  \<Rightarrow> natural \<times> natural
+     \<Rightarrow> ('a interval \<times> (unit \<Rightarrow> term)) \<times> natural \<times> natural" where
+  "random_interval i =
+  scomp (Quickcheck_Random.random i)
+    (\<lambda>man. scomp (Quickcheck_Random.random i) (\<lambda>exp. Pair (valtermify_interval man exp)))"
+
+instance ..
+
+end
+
+lifting_update interval.lifting
+lifting_forget interval.lifting
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Interval_Float.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -0,0 +1,357 @@
+section \<open>Approximate Operations on Intervals of Floating Point Numbers\<close>
+theory Interval_Float
+  imports
+    Interval
+    Float
+begin
+
+definition mid :: "float interval \<Rightarrow> float"
+  where "mid i = (lower i + upper i) * Float 1 (-1)"
+
+lemma mid_in_interval: "mid i \<in>\<^sub>i i"
+  using lower_le_upper[of i]
+  by (auto simp: mid_def set_of_eq powr_minus)
+
+lemma mid_le: "lower i \<le> mid i" "mid i \<le> upper i"
+  using mid_in_interval
+  by (auto simp: set_of_eq)
+
+definition centered :: "float interval \<Rightarrow> float interval"
+  where "centered i = i - interval_of (mid i)"
+
+definition "split_float_interval x = split_interval x ((lower x + upper x) * Float 1 (-1))"
+
+lemma split_float_intervalD: "split_float_interval X = (A, B) \<Longrightarrow> set_of X \<subseteq> set_of A \<union> set_of B"
+  by (auto dest!: split_intervalD simp: split_float_interval_def)
+
+lemma split_float_interval_bounds:
+  shows
+    lower_split_float_interval1: "lower (fst (split_float_interval X)) = lower X"
+  and lower_split_float_interval2: "lower (snd (split_float_interval X)) = mid X"
+  and upper_split_float_interval1: "upper (fst (split_float_interval X)) = mid X"
+  and upper_split_float_interval2: "upper (snd (split_float_interval X)) = upper X"
+  using mid_le[of X]
+  by (auto simp: split_float_interval_def mid_def[symmetric] min_def max_def real_of_float_eq
+      lower_split_interval1 lower_split_interval2
+      upper_split_interval1 upper_split_interval2)
+
+lemmas float_round_down_le[intro] = order_trans[OF float_round_down]
+  and float_round_up_ge[intro] = order_trans[OF _ float_round_up]
+
+text \<open>TODO: many of the lemmas should move to theories Float or Approximation
+  (the latter should be based on type @{type interval}.\<close>
+
+subsection "Intervals with Floating Point Bounds"
+
+context includes interval.lifting begin
+
+lift_definition round_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval"
+  is "\<lambda>p. \<lambda>(l, u). (float_round_down p l, float_round_up p u)"
+  by (auto simp: intro!: float_round_down_le float_round_up_le)
+
+lemma lower_round_ivl[simp]: "lower (round_interval p x) = float_round_down p (lower x)"
+  by transfer auto
+lemma upper_round_ivl[simp]: "upper (round_interval p x) = float_round_up p (upper x)"
+  by transfer auto
+
+lemma round_ivl_correct: "set_of A \<subseteq> set_of (round_interval prec A)"
+  by (auto simp: set_of_eq float_round_down_le float_round_up_le)
+
+lift_definition truncate_ivl :: "nat \<Rightarrow> real interval \<Rightarrow> real interval"
+  is "\<lambda>p. \<lambda>(l, u). (truncate_down p l, truncate_up p u)"
+  by (auto intro!: truncate_down_le truncate_up_le)
+
+lemma lower_truncate_ivl[simp]: "lower (truncate_ivl p x) = truncate_down p (lower x)"
+  by transfer auto
+lemma upper_truncate_ivl[simp]: "upper (truncate_ivl p x) = truncate_up p (upper x)"
+  by transfer auto
+
+lemma truncate_ivl_correct: "set_of A \<subseteq> set_of (truncate_ivl prec A)"
+  by (auto simp: set_of_eq intro!: truncate_down_le truncate_up_le)
+
+lift_definition real_interval::"float interval \<Rightarrow> real interval"
+  is "\<lambda>(l, u). (real_of_float l, real_of_float u)"
+  by auto
+
+lemma lower_real_interval[simp]: "lower (real_interval x) = lower x"
+  by transfer auto
+lemma upper_real_interval[simp]: "upper (real_interval x) = upper x"
+  by transfer auto
+
+definition "set_of' x = (case x of None \<Rightarrow> UNIV | Some i \<Rightarrow> set_of (real_interval i))"
+
+lemma real_interval_min_interval[simp]:
+  "real_interval (min_interval a b) = min_interval (real_interval a) (real_interval b)"
+  by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_min)
+
+lemma real_interval_max_interval[simp]:
+  "real_interval (max_interval a b) = max_interval (real_interval a) (real_interval b)"
+  by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max)
+
+lemma in_intervalI:
+  "x \<in>\<^sub>i X" if "lower X \<le> x" "x \<le> upper X"
+  using that by (auto simp: set_of_eq)
+
+abbreviation in_real_interval ("(_/ \<in>\<^sub>r _)" [51, 51] 50) where
+  "x \<in>\<^sub>r X \<equiv> x \<in>\<^sub>i real_interval X"
+
+lemma in_real_intervalI:
+  "x \<in>\<^sub>r X" if "lower X \<le> x" "x \<le> upper X" for x::real and X::"float interval"
+  using that
+  by (intro in_intervalI) auto
+
+subsection \<open>intros for \<open>real_interval\<close>\<close>
+
+lemma in_round_intervalI: "x \<in>\<^sub>r A  \<Longrightarrow> x \<in>\<^sub>r (round_interval prec A)"
+  by (auto simp: set_of_eq float_round_down_le float_round_up_le)
+
+lemma zero_in_float_intervalI: "0 \<in>\<^sub>r 0"
+  by (auto simp: set_of_eq)
+
+lemma plus_in_float_intervalI: "a + b \<in>\<^sub>r A + B" if "a \<in>\<^sub>r A" "b \<in>\<^sub>r B"
+  using that
+  by (auto simp: set_of_eq)
+
+lemma minus_in_float_intervalI: "a - b \<in>\<^sub>r A - B" if "a \<in>\<^sub>r A" "b \<in>\<^sub>r B"
+  using that
+  by (auto simp: set_of_eq)
+
+lemma uminus_in_float_intervalI: "-a \<in>\<^sub>r -A" if "a \<in>\<^sub>r A"
+  using that
+  by (auto simp: set_of_eq)
+
+lemma real_interval_times: "real_interval (A * B) = real_interval A * real_interval B"
+  by (auto simp: interval_eq_iff lower_times upper_times min_def max_def)
+
+lemma times_in_float_intervalI: "a * b \<in>\<^sub>r A * B" if "a \<in>\<^sub>r A" "b \<in>\<^sub>r B"
+  using times_in_intervalI[OF that]
+  by (auto simp: real_interval_times)
+
+lemma real_interval_abs: "real_interval (abs_interval A) = abs_interval (real_interval A)"
+  by (auto simp: interval_eq_iff min_def max_def)
+
+lemma abs_in_float_intervalI: "abs a \<in>\<^sub>r abs_interval A" if "a \<in>\<^sub>r A"
+  by (auto simp: set_of_abs_interval real_interval_abs intro!: imageI that)
+
+lemma interval_of[intro,simp]: "x \<in>\<^sub>r interval_of x"
+  by (auto simp: set_of_eq)
+
+lemma split_float_interval_realD: "split_float_interval X = (A, B) \<Longrightarrow> x \<in>\<^sub>r X \<Longrightarrow> x \<in>\<^sub>r A \<or> x \<in>\<^sub>r B"
+  by (auto simp: set_of_eq prod_eq_iff split_float_interval_bounds)
+
+
+subsection \<open>bounds for lists\<close>
+
+lemma lower_Interval: "lower (Interval x) = fst x"
+  and upper_Interval: "upper (Interval x) = snd x"
+  if "fst x \<le> snd x"
+  using that
+  by (auto simp: lower_def upper_def Interval_inverse split_beta')
+
+definition all_in_i :: "'a::preorder list \<Rightarrow> 'a interval list \<Rightarrow> bool"
+  (infix "(all'_in\<^sub>i)" 50)
+  where "x all_in\<^sub>i I = (length x = length I \<and> (\<forall>i < length I. x ! i \<in>\<^sub>i I ! i))"
+
+definition all_in :: "real list \<Rightarrow> float interval list \<Rightarrow> bool"
+  (infix "(all'_in)" 50)
+  where "x all_in I = (length x = length I \<and> (\<forall>i < length I. x ! i \<in>\<^sub>r I ! i))"
+
+definition all_subset :: "'a::order interval list \<Rightarrow> 'a interval list \<Rightarrow> bool"
+  (infix "(all'_subset)" 50)
+  where "I all_subset J = (length I = length J \<and> (\<forall>i < length I. set_of (I!i) \<subseteq> set_of (J!i)))"
+
+lemmas [simp] = all_in_def all_subset_def
+
+lemma all_subsetD:
+  assumes "I all_subset J"
+  assumes "x all_in I"
+  shows "x all_in J"
+  using assms
+  by (auto simp: set_of_eq; fastforce)
+
+lemma round_interval_mono: "set_of (round_interval prec X) \<subseteq> set_of (round_interval prec Y)"
+  if "set_of X \<subseteq> set_of Y"
+  using that
+  by transfer
+    (auto simp: float_round_down.rep_eq float_round_up.rep_eq truncate_down_mono truncate_up_mono)
+
+lemma Ivl_simps[simp]: "lower (Ivl a b) = min a b" "upper (Ivl a b) = b"
+  subgoal by transfer simp
+  subgoal by transfer simp
+  done
+
+lemma set_of_subset_iff: "set_of X \<subseteq> set_of Y \<longleftrightarrow> lower Y \<le> lower X \<and> upper X \<le> upper Y"
+  for X Y::"'a::linorder interval"
+  by (auto simp: set_of_eq subset_iff)
+
+lemma bounds_of_interval_eq_lower_upper:
+  "bounds_of_interval ivl = (lower ivl, upper ivl)" if "lower ivl \<le> upper ivl"
+  using that
+  by (auto simp: lower.rep_eq upper.rep_eq)
+
+lemma real_interval_Ivl: "real_interval (Ivl a b) = Ivl a b"
+  by transfer (auto simp: min_def)
+
+lemma set_of_mul_contains_real_zero:
+  "0 \<in>\<^sub>r (A * B)" if "0 \<in>\<^sub>r A \<or> 0 \<in>\<^sub>r B"
+  using that set_of_mul_contains_zero[of A B]
+  by (auto simp: set_of_eq)
+
+fun subdivide_interval :: "nat \<Rightarrow> float interval \<Rightarrow> float interval list"
+  where "subdivide_interval 0 I = [I]"
+  | "subdivide_interval (Suc n) I = (
+         let m = mid I
+         in (subdivide_interval n (Ivl (lower I) m)) @ (subdivide_interval n (Ivl m (upper I)))
+       )"
+
+lemma subdivide_interval_length:
+  shows "length (subdivide_interval n I) = 2^n"
+  by(induction n arbitrary: I, simp_all add: Let_def)
+
+lemma lower_le_mid: "lower x \<le> mid x" "real_of_float (lower x) \<le> mid x"
+  and mid_le_upper: "mid x \<le> upper x" "real_of_float (mid x) \<le> upper x"
+  unfolding mid_def
+  subgoal by transfer (auto simp: powr_neg_one)
+  subgoal by transfer (auto simp: powr_neg_one)
+  subgoal by transfer (auto simp: powr_neg_one)
+  subgoal by transfer (auto simp: powr_neg_one)
+  done
+
+lemma subdivide_interval_correct:
+  "list_ex (\<lambda>i. x \<in>\<^sub>r i) (subdivide_interval n I)" if "x \<in>\<^sub>r I" for x::real
+  using that
+proof(induction n arbitrary: x I)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  from \<open>x \<in>\<^sub>r I\<close> consider "x \<in>\<^sub>r Ivl (lower I) (mid I)" | "x \<in>\<^sub>r Ivl (mid I) (upper I)"
+    by (cases "x \<le> real_of_float (mid I)")
+      (auto simp: set_of_eq min_def lower_le_mid mid_le_upper)
+  from this[case_names lower upper] show ?case
+    by cases (use Suc.IH in \<open>auto simp: Let_def\<close>)
+qed
+
+fun interval_list_union :: "'a::lattice interval list \<Rightarrow> 'a interval"
+  where "interval_list_union [] = undefined"
+  | "interval_list_union [I] = I"
+  | "interval_list_union (I#Is) = sup I (interval_list_union Is)"
+
+lemma interval_list_union_correct:
+  assumes "S \<noteq> []"
+  assumes "i < length S"
+  shows "set_of (S!i) \<subseteq> set_of (interval_list_union S)"
+  using assms
+proof(induction S arbitrary: i)
+  case (Cons a S i)
+  thus ?case
+  proof(cases S)
+    fix b S'
+    assume "S = b # S'"
+    hence "S \<noteq> []"
+      by simp
+    show ?thesis
+    proof(cases i)
+      case 0
+      show ?thesis
+        apply(cases S)
+        using interval_union_mono1
+        by (auto simp add: 0)
+    next
+      case (Suc i_prev)
+      hence "i_prev < length S"
+        using Cons(3) by simp
+
+      from Cons(1)[OF \<open>S \<noteq> []\<close> this] Cons(1)
+      have "set_of ((a # S) ! i) \<subseteq> set_of (interval_list_union S)"
+        by (simp add: \<open>i = Suc i_prev\<close>)
+      also have "... \<subseteq> set_of (interval_list_union (a # S))"
+        using \<open>S \<noteq> []\<close>
+        apply(cases S)
+        using interval_union_mono2
+        by auto
+      finally show ?thesis .
+    qed
+  qed simp
+qed simp
+
+lemma split_domain_correct:
+  fixes x :: "real list"
+  assumes "x all_in I"
+  assumes split_correct: "\<And>x a I. x \<in>\<^sub>r I \<Longrightarrow> list_ex (\<lambda>i::float interval. x \<in>\<^sub>r i) (split I)"
+  shows "list_ex (\<lambda>s. x all_in s) (split_domain split I)"
+  using assms(1)
+proof(induction I arbitrary: x)
+  case (Cons I Is x)
+  have "x \<noteq> []"
+    using Cons(2) by auto
+  obtain x' xs where x_decomp: "x = x' # xs"
+    using \<open>x \<noteq> []\<close> list.exhaust by auto
+  hence "x' \<in>\<^sub>r I" "xs all_in Is"
+    using Cons(2)
+    by auto
+  show ?case
+    using Cons(1)[OF \<open>xs all_in Is\<close>]
+      split_correct[OF \<open>x' \<in>\<^sub>r I\<close>]
+    apply (auto simp add: list_ex_iff set_of_eq)
+    by (smt length_Cons less_Suc_eq_0_disj nth_Cons_0 nth_Cons_Suc x_decomp)
+qed simp
+
+
+lift_definition(code_dt) inverse_float_interval::"nat \<Rightarrow> float interval \<Rightarrow> float interval option" is
+  "\<lambda>prec (l, u). if (0 < l \<or> u < 0) then Some (float_divl prec 1 u, float_divr prec 1 l) else None"
+  by (auto intro!: order_trans[OF float_divl] order_trans[OF _ float_divr]
+      simp: divide_simps)
+
+lemma inverse_float_interval_eq_Some_conv:
+  defines "one \<equiv> (1::float)"
+  shows 
+    "inverse_float_interval p X = Some R \<longleftrightarrow>
+    (lower X > 0 \<or> upper X < 0) \<and>
+    lower R = float_divl p one (upper X) \<and>
+    upper R = float_divr p one (lower X)"
+  by clarsimp (transfer fixing: one, force simp: one_def split: if_splits)
+
+lemma inverse_float_interval:
+  "inverse ` set_of (real_interval X) \<subseteq> set_of (real_interval Y)"
+  if "inverse_float_interval p X = Some Y"
+  using that
+  apply (clarsimp simp: set_of_eq inverse_float_interval_eq_Some_conv)
+  by (intro order_trans[OF float_divl] order_trans[OF _ float_divr] conjI)
+    (auto simp: divide_simps)
+
+lemma inverse_float_intervalI:
+  "x \<in>\<^sub>r X \<Longrightarrow> inverse x \<in> set_of' (inverse_float_interval p X)"
+  using inverse_float_interval[of p X]
+  by (auto simp: set_of'_def split: option.splits)
+
+lemma inverse_float_interval_eqI: "inverse_float_interval p X = Some IVL \<Longrightarrow> x \<in>\<^sub>r X \<Longrightarrow> inverse x \<in>\<^sub>r IVL"
+  using inverse_float_intervalI[of x X p]
+  by (auto simp: set_of'_def)
+
+lemma real_interval_abs_interval[simp]:
+  "real_interval (abs_interval x) = abs_interval (real_interval x)"
+  by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max real_of_float_min)
+
+lift_definition floor_float_interval::"float interval \<Rightarrow> float interval" is
+  "\<lambda>(l, u). (floor_fl l, floor_fl u)"
+  by (auto intro!: floor_mono simp: floor_fl.rep_eq)
+
+lemma lower_floor_float_interval[simp]: "lower (floor_float_interval x) = floor_fl (lower x)"
+  by transfer auto
+lemma upper_floor_float_interval[simp]: "upper (floor_float_interval x) = floor_fl (upper x)"
+  by transfer auto
+
+lemma floor_float_intervalI: "\<lfloor>x\<rfloor> \<in>\<^sub>r floor_float_interval X" if "x \<in>\<^sub>r X"
+  using that by (auto simp: set_of_eq floor_fl_def floor_mono)
+
+end
+
+
+subsection \<open>constants for code generation\<close>
+
+definition lowerF::"float interval \<Rightarrow> float" where "lowerF = lower"
+definition upperF::"float interval \<Rightarrow> float" where "upperF = upper"
+
+
+end
\ No newline at end of file
--- a/src/HOL/Library/Library.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Library/Library.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -38,6 +38,8 @@
   Groups_Big_Fun
   Indicator_Function
   Infinite_Set
+  Interval
+  Interval_Float
   IArray
   Landau_Symbols
   Lattice_Algebras
--- a/src/HOL/ROOT	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/ROOT	Tue Nov 05 22:56:06 2019 +0100
@@ -650,7 +650,7 @@
     Triangular_Numbers
     Unification
     While_Combinator_Example
-    Word_Type
+    Word
     veriT_Preprocessing
   theories [skip_proofs = false]
     SAT_Examples
--- a/src/HOL/Real.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/Real.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -1313,6 +1313,96 @@
   by simp
 
 
+subsection \<open>Archimedean properties and useful consequences\<close>
+
+text\<open>Bernoulli's inequality\<close>
+proposition Bernoulli_inequality:
+  fixes x :: real
+  assumes "-1 \<le> x"
+    shows "1 + n * x \<le> (1 + x) ^ n"
+proof (induct n)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  have "1 + Suc n * x \<le> 1 + (Suc n)*x + n * x^2"
+    by (simp add: algebra_simps)
+  also have "... = (1 + x) * (1 + n*x)"
+    by (auto simp: power2_eq_square algebra_simps)
+  also have "... \<le> (1 + x) ^ Suc n"
+    using Suc.hyps assms mult_left_mono by fastforce
+  finally show ?case .
+qed
+
+corollary Bernoulli_inequality_even:
+  fixes x :: real
+  assumes "even n"
+    shows "1 + n * x \<le> (1 + x) ^ n"
+proof (cases "-1 \<le> x \<or> n=0")
+  case True
+  then show ?thesis
+    by (auto simp: Bernoulli_inequality)
+next
+  case False
+  then have "real n \<ge> 1"
+    by simp
+  with False have "n * x \<le> -1"
+    by (metis linear minus_zero mult.commute mult.left_neutral mult_left_mono_neg neg_le_iff_le order_trans zero_le_one)
+  then have "1 + n * x \<le> 0"
+    by auto
+  also have "... \<le> (1 + x) ^ n"
+    using assms
+    using zero_le_even_power by blast
+  finally show ?thesis .
+qed
+
+corollary real_arch_pow:
+  fixes x :: real
+  assumes x: "1 < x"
+  shows "\<exists>n. y < x^n"
+proof -
+  from x have x0: "x - 1 > 0"
+    by arith
+  from reals_Archimedean3[OF x0, rule_format, of y]
+  obtain n :: nat where n: "y < real n * (x - 1)" by metis
+  from x0 have x00: "x- 1 \<ge> -1" by arith
+  from Bernoulli_inequality[OF x00, of n] n
+  have "y < x^n" by auto
+  then show ?thesis by metis
+qed
+
+corollary real_arch_pow_inv:
+  fixes x y :: real
+  assumes y: "y > 0"
+    and x1: "x < 1"
+  shows "\<exists>n. x^n < y"
+proof (cases "x > 0")
+  case True
+  with x1 have ix: "1 < 1/x" by (simp add: field_simps)
+  from real_arch_pow[OF ix, of "1/y"]
+  obtain n where n: "1/y < (1/x)^n" by blast
+  then show ?thesis using y \<open>x > 0\<close>
+    by (auto simp add: field_simps)
+next
+  case False
+  with y x1 show ?thesis
+    by (metis less_le_trans not_less power_one_right)
+qed
+
+lemma forall_pos_mono:
+  "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
+    (\<And>n::nat. n \<noteq> 0 \<Longrightarrow> P (inverse (real n))) \<Longrightarrow> (\<And>e. 0 < e \<Longrightarrow> P e)"
+  by (metis real_arch_inverse)
+
+lemma forall_pos_mono_1:
+  "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow>
+    (\<And>n. P (inverse (real (Suc n)))) \<Longrightarrow> 0 < e \<Longrightarrow> P e"
+  apply (rule forall_pos_mono)
+  apply auto
+  apply (metis Suc_pred of_nat_Suc)
+  done
+
+
 subsection \<open>Floor and Ceiling Functions from the Reals to the Integers\<close>
 
 (* FIXME: theorems for negative numerals. Many duplicates, e.g. from Archimedean_Field.thy. *)
--- a/src/HOL/ex/Bit_Lists.thy	Tue Nov 05 22:55:50 2019 +0100
+++ b/src/HOL/ex/Bit_Lists.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -5,31 +5,10 @@
 
 theory Bit_Lists
   imports
-    Main
-    "HOL-Library.Boolean_Algebra"
-begin
-
-context ab_group_add
+    Bit_Operations
 begin
 
-lemma minus_diff_commute:
-  "- b - a = - a - b"
-  by (simp only: diff_conv_add_uminus add.commute)
-
-end
-
-
-subsection \<open>Bit representations\<close>
-
-subsubsection \<open>Mere syntactic bit representation\<close>
-
-class bit_representation =
-  fixes bits_of :: "'a \<Rightarrow> bool list"
-    and of_bits :: "bool list \<Rightarrow> 'a"
-  assumes of_bits_of [simp]: "of_bits (bits_of a) = a"
-
-
-subsubsection \<open>Algebraic bit representation\<close>
+subsection \<open>Fragments of algebraic bit representations\<close>
 
 context comm_semiring_1
 begin
@@ -57,7 +36,7 @@
     simp_all add: algebra_simps)
 
 lemma unsigned_of_bits_take [simp]:
-  "unsigned_of_bits (take n bs) = take_bit n (unsigned_of_bits bs)"
+  "unsigned_of_bits (take n bs) = Parity.take_bit n (unsigned_of_bits bs)"
 proof (induction bs arbitrary: n)
   case Nil
   then show ?case
@@ -69,7 +48,7 @@
 qed
 
 lemma unsigned_of_bits_drop [simp]:
-  "unsigned_of_bits (drop n bs) = drop_bit n (unsigned_of_bits bs)"
+  "unsigned_of_bits (drop n bs) = Parity.drop_bit n (unsigned_of_bits bs)"
 proof (induction bs arbitrary: n)
   case Nil
   then show ?case
@@ -80,10 +59,44 @@
     by (cases n) simp_all
 qed
 
+primrec n_bits_of :: "nat \<Rightarrow> 'a \<Rightarrow> bool list"
+  where
+    "n_bits_of 0 a = []"
+  | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)"
+
+lemma n_bits_of_eq_iff:
+  "n_bits_of n a = n_bits_of n b \<longleftrightarrow> Parity.take_bit n a = Parity.take_bit n b"
+  apply (induction n arbitrary: a b)
+   apply (auto elim!: evenE oddE)
+   apply (metis dvd_triv_right even_plus_one_iff)
+  apply (metis dvd_triv_right even_plus_one_iff)
+  done
+
+lemma take_n_bits_of [simp]:
+  "take m (n_bits_of n a) = n_bits_of (min m n) a"
+proof -
+  define q and v and w where "q = min m n" and "v = m - q" and "w = n - q"
+  then have "v = 0 \<or> w = 0"
+    by auto
+  then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a"
+    by (induction q arbitrary: a) auto
+  with q_def v_def w_def show ?thesis
+    by simp
+qed
+
+lemma unsigned_of_bits_n_bits_of [simp]:
+  "unsigned_of_bits (n_bits_of n a) = Parity.take_bit n a"
+  by (induction n arbitrary: a) (simp_all add: ac_simps)
+
 end
 
 
-subsubsection \<open>Instances\<close>
+subsection \<open>Syntactic bit representation\<close>
+
+class bit_representation =
+  fixes bits_of :: "'a \<Rightarrow> bool list"
+    and of_bits :: "bool list \<Rightarrow> 'a"
+  assumes of_bits_of [simp]: "of_bits (bits_of a) = a"
 
 text \<open>Unclear whether a \<^typ>\<open>bool\<close> instantiation is needed or not\<close>
 
@@ -216,7 +229,7 @@
   by (auto simp add: of_bits_int_def)
 
 lemma of_bits_drop [simp]:
-  "of_bits (drop n bs) = drop_bit n (of_bits bs :: int)"
+  "of_bits (drop n bs) = Parity.drop_bit n (of_bits bs :: int)"
     if "n < length bs"
 using that proof (induction bs arbitrary: n)
   case Nil
@@ -240,140 +253,47 @@
 
 end
 
-
-subsection \<open>Syntactic bit operations\<close>
+lemma unsigned_of_bits_eq_of_bits:
+  "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)"
+  by (simp add: of_bits_int_def)
 
-class bit_operations = bit_representation +
-  fixes not :: "'a \<Rightarrow> 'a"  ("NOT")
-    and "and" :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "AND" 64)
-    and or :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "OR"  59)
-    and xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "XOR" 59)
-    and shift_left :: "'a \<Rightarrow> nat \<Rightarrow> 'a"  (infixl "<<" 55)
-    and shift_right :: "'a \<Rightarrow> nat \<Rightarrow> 'a"  (infixl ">>" 55)
-  assumes not_eq: "not = of_bits \<circ> map Not \<circ> bits_of"
-    and and_eq: "length bs = length cs \<Longrightarrow>
-      of_bits bs AND of_bits cs = of_bits (map2 (\<and>) bs cs)"
-    and semilattice_and: "semilattice (AND)"
-    and or_eq: "length bs = length cs \<Longrightarrow>
-      of_bits bs OR of_bits cs = of_bits (map2 (\<or>) bs cs)"
-    and semilattice_or: "semilattice (OR)"
-    and xor_eq: "length bs = length cs \<Longrightarrow>
-      of_bits bs XOR of_bits cs = of_bits (map2 (\<noteq>) bs cs)"
-    and abel_semigroup_xor: "abel_semigroup (XOR)"
-    and shift_right_eq: "a << n = of_bits (replicate n False @ bits_of a)"
-    and shift_left_eq: "n < length (bits_of a) \<Longrightarrow> a >> n = of_bits (drop n (bits_of a))"
+instantiation word :: (len) bit_representation
 begin
 
-text \<open>
-  We want the bitwise operations to bind slightly weaker
-  than \<open>+\<close> and \<open>-\<close>, but \<open>~~\<close> to
-  bind slightly stronger than \<open>*\<close>.
-\<close>
+lift_definition bits_of_word :: "'a word \<Rightarrow> bool list"
+  is "n_bits_of LENGTH('a)"
+  by (simp add: n_bits_of_eq_iff)
 
-sublocale "and": semilattice "(AND)"
-  by (fact semilattice_and)
+lift_definition of_bits_word :: "bool list \<Rightarrow> 'a word"
+  is unsigned_of_bits .
 
-sublocale or: semilattice "(OR)"
-  by (fact semilattice_or)
-
-sublocale xor: abel_semigroup "(XOR)"
-  by (fact abel_semigroup_xor)
+instance proof
+  fix a :: "'a word"
+  show "of_bits (bits_of a) = a"
+    by transfer simp
+qed
 
 end
 
 
-subsubsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
-
-locale zip_nat = single: abel_semigroup f
-    for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70) +
-  assumes end_of_bits: "\<not> False \<^bold>* False"
-begin
-
-lemma False_P_imp:
-  "False \<^bold>* True \<and> P" if "False \<^bold>* P"
-  using that end_of_bits by (cases P) simp_all
-
-function F :: "nat \<Rightarrow> nat \<Rightarrow> nat"  (infixl "\<^bold>\<times>" 70)
-  where "m \<^bold>\<times> n = (if m = 0 \<and> n = 0 then 0
-    else of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2)"
-  by auto
-
-termination
-  by (relation "measure (case_prod (+))") auto
-
-lemma zero_left_eq:
-  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
-  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
-
-lemma zero_right_eq:
-  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
-  by (induction m rule: nat_bit_induct) (simp_all add: end_of_bits)
-
-lemma simps [simp]:
-  "0 \<^bold>\<times> 0 = 0"
-  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
-  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
-  "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
-  by (simp_all only: zero_left_eq zero_right_eq) simp
-
-lemma rec:
-  "m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
-  by (cases "m = 0 \<and> n = 0") (auto simp add: end_of_bits)
-
-declare F.simps [simp del]
+subsection \<open>Bit representations with bit operations\<close>
 
-sublocale abel_semigroup F
-proof
-  show "m \<^bold>\<times> n \<^bold>\<times> q = m \<^bold>\<times> (n \<^bold>\<times> q)" for m n q :: nat
-  proof (induction m arbitrary: n q rule: nat_bit_induct)
-    case zero
-    show ?case
-      by simp
-  next
-    case (even m)
-    with rec [of "2 * m"] rec [of _ q] show ?case
-      by (cases "even n") (auto dest: False_P_imp)
-  next
-    case (odd m)
-    with rec [of "Suc (2 * m)"] rec [of _ q] show ?case
-      by (cases "even n"; cases "even q")
-        (auto dest: False_P_imp simp add: ac_simps)
-  qed
-  show "m \<^bold>\<times> n = n \<^bold>\<times> m" for m n :: nat
-  proof (induction m arbitrary: n rule: nat_bit_induct)
-    case zero
-    show ?case
-      by (simp add: ac_simps)
-  next
-    case (even m)
-    with rec [of "2 * m" n] rec [of n "2 * m"] show ?case
-      by (simp add: ac_simps)
-  next
-    case (odd m)
-    with rec [of "Suc (2 * m)" n] rec [of n "Suc (2 * m)"] show ?case
-      by (simp add: ac_simps)
-  qed
-qed
+class semiring_bit_representation = semiring_bit_operations + bit_representation +
+  assumes and_eq: "length bs = length cs \<Longrightarrow>
+      of_bits bs AND of_bits cs = of_bits (map2 (\<and>) bs cs)"
+    and or_eq: "length bs = length cs \<Longrightarrow>
+      of_bits bs OR of_bits cs = of_bits (map2 (\<or>) bs cs)"
+    and xor_eq: "length bs = length cs \<Longrightarrow>
+      of_bits bs XOR of_bits cs = of_bits (map2 (\<noteq>) bs cs)"
+    and shift_bit_eq: "shift_bit n a = of_bits (replicate n False @ bits_of a)"
+    and drop_bit_eq: "n < length (bits_of a) \<Longrightarrow> drop_bit n a = of_bits (drop n (bits_of a))"
 
-lemma self [simp]:
-  "n \<^bold>\<times> n = of_bool (True \<^bold>* True) * n"
-  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
+class ring_bit_representation = ring_bit_operations + semiring_bit_representation +
+  assumes not_eq: "not = of_bits \<circ> map Not \<circ> bits_of"
 
-lemma even_iff [simp]:
-  "even (m \<^bold>\<times> n) \<longleftrightarrow> \<not> (odd m \<^bold>* odd n)"
-proof (induction m arbitrary: n rule: nat_bit_induct)
-  case zero
-  show ?case
-    by (cases "even n") (simp_all add: end_of_bits)
-next
-  case (even m)
-  then show ?case
-    by (simp add: rec [of "2 * m"])
-next
-  case (odd m)
-  then show ?case
-    by (simp add: rec [of "Suc (2 * m)"])
-qed
+
+context zip_nat
+begin
 
 lemma of_bits:
   "of_bits bs \<^bold>\<times> of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: nat)"
@@ -391,366 +311,15 @@
 
 end
 
-instantiation nat :: bit_operations
-begin
-
-definition not_nat :: "nat \<Rightarrow> nat"
-  where "not_nat = of_bits \<circ> map Not \<circ> bits_of"
-
-global_interpretation and_nat: zip_nat "(\<and>)"
-  defines and_nat = and_nat.F
-  by standard auto
-
-global_interpretation and_nat: semilattice "(AND) :: nat \<Rightarrow> nat \<Rightarrow> nat"
-proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard)
-  show "n AND n = n" for n :: nat
-    by (simp add: and_nat.self)
-qed
-
-declare and_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
-  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
-
-lemma zero_nat_and_eq [simp]:
-  "0 AND n = 0" for n :: nat
-  by simp
-
-lemma and_zero_nat_eq [simp]:
-  "n AND 0 = 0" for n :: nat
-  by simp
-
-global_interpretation or_nat: zip_nat "(\<or>)"
-  defines or_nat = or_nat.F
-  by standard auto
-
-global_interpretation or_nat: semilattice "(OR) :: nat \<Rightarrow> nat \<Rightarrow> nat"
-proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard)
-  show "n OR n = n" for n :: nat
-    by (simp add: or_nat.self)
-qed
-
-declare or_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
-  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
-
-lemma zero_nat_or_eq [simp]:
-  "0 OR n = n" for n :: nat
-  by simp
-
-lemma or_zero_nat_eq [simp]:
-  "n OR 0 = n" for n :: nat
-  by simp
-
-global_interpretation xor_nat: zip_nat "(\<noteq>)"
-  defines xor_nat = xor_nat.F
-  by standard auto
-
-declare xor_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
-  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
-
-lemma zero_nat_xor_eq [simp]:
-  "0 XOR n = n" for n :: nat
-  by simp
-
-lemma xor_zero_nat_eq [simp]:
-  "n XOR 0 = n" for n :: nat
-  by simp
-
-definition shift_left_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  where [simp]: "m << n = push_bit n m" for m :: nat
-
-definition shift_right_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  where [simp]: "m >> n = drop_bit n m" for m :: nat
+instance nat :: semiring_bit_representation
+  apply standard
+      apply (simp_all only: and_nat.of_bits or_nat.of_bits xor_nat.of_bits)
+   apply (simp_all add: drop_bit_eq_div Parity.drop_bit_eq_div shift_bit_eq_mult push_bit_eq_mult)
+  done
 
-instance proof
-  show "semilattice ((AND) :: nat \<Rightarrow> _)"
-    by (fact and_nat.semilattice_axioms)
-  show "semilattice ((OR):: nat \<Rightarrow> _)"
-    by (fact or_nat.semilattice_axioms)
-  show "abel_semigroup ((XOR):: nat \<Rightarrow> _)"
-    by (fact xor_nat.abel_semigroup_axioms)
-  show "(not :: nat \<Rightarrow> _) = of_bits \<circ> map Not \<circ> bits_of"
-    by (fact not_nat_def)
-  show "of_bits bs AND of_bits cs = (of_bits (map2 (\<and>) bs cs) :: nat)"
-    if "length bs = length cs" for bs cs
-    using that by (fact and_nat.of_bits)
-  show "of_bits bs OR of_bits cs = (of_bits (map2 (\<or>) bs cs) :: nat)"
-    if "length bs = length cs" for bs cs
-    using that by (fact or_nat.of_bits)
-  show "of_bits bs XOR of_bits cs = (of_bits (map2 (\<noteq>) bs cs) :: nat)"
-    if "length bs = length cs" for bs cs
-    using that by (fact xor_nat.of_bits)
-  show "m << n = of_bits (replicate n False @ bits_of m)"
-    for m n :: nat
-    by simp
-  show "m >> n = of_bits (drop n (bits_of m))"
-    for m n :: nat
-    by simp
-qed
-
-end
-
-global_interpretation or_nat: semilattice_neutr "(OR)" "0 :: nat"
-  by standard simp
-
-global_interpretation xor_nat: comm_monoid "(XOR)" "0 :: nat"
-  by standard simp
-
-lemma not_nat_simps [simp]:
-  "NOT 0 = (0::nat)"
-  "n > 0 \<Longrightarrow> NOT n = of_bool (even n) + 2 * NOT (n div 2)" for n :: nat
-  by (simp_all add: not_eq)
-
-lemma not_1_nat [simp]:
-  "NOT 1 = (0::nat)"
-  by simp
-
-lemma not_Suc_0 [simp]:
-  "NOT (Suc 0) = (0::nat)"
-  by simp
-
-lemma Suc_0_and_eq [simp]:
-  "Suc 0 AND n = n mod 2"
-  by (cases n) auto
-
-lemma and_Suc_0_eq [simp]:
-  "n AND Suc 0 = n mod 2"
-  using Suc_0_and_eq [of n] by (simp add: ac_simps)
-
-lemma Suc_0_or_eq [simp]:
-  "Suc 0 OR n = n + of_bool (even n)"
-  by (cases n) (simp_all add: ac_simps)
-
-lemma or_Suc_0_eq [simp]:
-  "n OR Suc 0 = n + of_bool (even n)"
-  using Suc_0_or_eq [of n] by (simp add: ac_simps)
-
-lemma Suc_0_xor_eq [simp]:
-  "Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)"
-  by (cases n) (simp_all add: ac_simps)
-
-lemma xor_Suc_0_eq [simp]:
-  "n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)"
-  using Suc_0_xor_eq [of n] by (simp add: ac_simps)
-
-
-subsubsection \<open>Instance \<^typ>\<open>int\<close>\<close>
-
-abbreviation (input) complement :: "int \<Rightarrow> int"
-  where "complement k \<equiv> - k - 1"
-
-lemma complement_half:
-  "complement (k * 2) div 2 = complement k"
-  by simp
-
-lemma complement_div_2:
-  "complement (k div 2) = complement k div 2"
-  by linarith
-
-locale zip_int = single: abel_semigroup f
-  for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70)
+context zip_int
 begin
  
-lemma False_False_imp_True_True:
-  "True \<^bold>* True" if "False \<^bold>* False"
-proof (rule ccontr)
-  assume "\<not> True \<^bold>* True"
-  with that show False
-    using single.assoc [of False True True]
-    by (cases "False \<^bold>* True") simp_all
-qed
-
-function F :: "int \<Rightarrow> int \<Rightarrow> int"  (infixl "\<^bold>\<times>" 70)
-  where "k \<^bold>\<times> l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
-    then - of_bool (odd k \<^bold>* odd l)
-    else of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2)"
-  by auto
-
-termination
-  by (relation "measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))") auto
-
-lemma zero_left_eq:
-  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> l
-     | (True, False) \<Rightarrow> complement l
-     | (True, True) \<Rightarrow> - 1)"
-  by (induction l rule: int_bit_induct)
-   (simp_all split: bool.split) 
-
-lemma minus_left_eq:
-  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> l
-     | (True, False) \<Rightarrow> complement l
-     | (True, True) \<Rightarrow> - 1)"
-  by (induction l rule: int_bit_induct)
-   (simp_all split: bool.split) 
-
-lemma zero_right_eq:
-  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> k
-     | (True, False) \<Rightarrow> complement k
-     | (True, True) \<Rightarrow> - 1)"
-  by (induction k rule: int_bit_induct)
-    (simp_all add: ac_simps split: bool.split)
-
-lemma minus_right_eq:
-  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> k
-     | (True, False) \<Rightarrow> complement k
-     | (True, True) \<Rightarrow> - 1)"
-  by (induction k rule: int_bit_induct)
-    (simp_all add: ac_simps split: bool.split)
-
-lemma simps [simp]:
-  "0 \<^bold>\<times> 0 = - of_bool (False \<^bold>* False)"
-  "- 1 \<^bold>\<times> 0 = - of_bool (True \<^bold>* False)"
-  "0 \<^bold>\<times> - 1 = - of_bool (False \<^bold>* True)"
-  "- 1 \<^bold>\<times> - 1 = - of_bool (True \<^bold>* True)"
-  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> l
-     | (True, False) \<Rightarrow> complement l
-     | (True, True) \<Rightarrow> - 1)"
-  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> l
-     | (True, False) \<Rightarrow> complement l
-     | (True, True) \<Rightarrow> - 1)"
-  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> k
-     | (True, False) \<Rightarrow> complement k
-     | (True, True) \<Rightarrow> - 1)"
-  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> k
-     | (True, False) \<Rightarrow> complement k
-     | (True, True) \<Rightarrow> - 1)"
-  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> l \<noteq> - 1
-    \<Longrightarrow> k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
-  by simp_all[4] (simp_all only: zero_left_eq minus_left_eq zero_right_eq minus_right_eq, simp)
-
-declare F.simps [simp del]
-
-lemma rec:
-  "k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
-  by (cases "k \<in> {0, - 1} \<and> l \<in> {0, - 1}") (auto simp add: ac_simps F.simps [of k l] split: bool.split)
-
-sublocale abel_semigroup F
-proof
-  show "k \<^bold>\<times> l \<^bold>\<times> r = k \<^bold>\<times> (l \<^bold>\<times> r)" for k l r :: int
-  proof (induction k arbitrary: l r rule: int_bit_induct)
-    case zero
-    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "False \<^bold>* False" "\<not> False \<^bold>* True"
-    proof (induction l arbitrary: r rule: int_bit_induct)
-      case zero
-      from that show ?case
-        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-    next
-      case minus
-      from that show ?case
-        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-    next
-      case (even l)
-      with that rec [of _ r] show ?case
-        by (cases "even r")
-          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
-    next
-      case (odd l)
-      moreover have "- l - 1 = - 1 - l"
-        by simp
-      ultimately show ?case
-        using that rec [of _ r]
-        by (cases "even r")
-          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-    qed
-    then show ?case
-      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-  next
-    case minus
-    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "\<not> True \<^bold>* True" "False \<^bold>* True"
-    proof (induction l arbitrary: r rule: int_bit_induct)
-      case zero
-      from that show ?case
-        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-    next
-      case minus
-      from that show ?case
-        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-    next
-      case (even l)
-      with that rec [of _ r] show ?case
-        by (cases "even r")
-          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
-    next
-      case (odd l)
-      moreover have "- l - 1 = - 1 - l"
-        by simp
-      ultimately show ?case
-        using that rec [of _ r]
-        by (cases "even r")
-          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-    qed
-    then show ?case
-      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
-  next
-    case (even k)
-    with rec [of "k * 2"] rec [of _ r] show ?case
-      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
-  next
-    case (odd k)
-    with rec [of "1 + k * 2"] rec [of _ r] show ?case
-      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
-  qed
-  show "k \<^bold>\<times> l = l \<^bold>\<times> k" for k l :: int
-  proof (induction k arbitrary: l rule: int_bit_induct)
-    case zero
-    show ?case
-      by simp
-  next
-    case minus
-    show ?case
-      by simp
-  next
-    case (even k)
-    with rec [of "k * 2" l] rec [of l "k * 2"] show ?case
-      by (simp add: ac_simps)
-  next
-    case (odd k)
-    with rec [of "k * 2 + 1" l] rec [of l "k * 2 + 1"] show ?case
-      by (simp add: ac_simps)
-  qed
-qed
-
-lemma self [simp]:
-  "k \<^bold>\<times> k = (case (False \<^bold>* False, True \<^bold>* True)
-    of (False, False) \<Rightarrow> 0
-     | (False, True) \<Rightarrow> k
-     | (True, True) \<Rightarrow> - 1)"
-  by (induction k rule: int_bit_induct) (auto simp add: False_False_imp_True_True split: bool.split)
-
-lemma even_iff [simp]:
-  "even (k \<^bold>\<times> l) \<longleftrightarrow> \<not> (odd k \<^bold>* odd l)"
-proof (induction k arbitrary: l rule: int_bit_induct)
-  case zero
-  show ?case
-    by (cases "even l") (simp_all split: bool.splits)
-next
-  case minus
-  show ?case
-    by (cases "even l") (simp_all split: bool.splits)
-next
-  case (even k)
-  then show ?case
-    by (simp add: rec [of "k * 2"])
-next
-  case (odd k)
-  then show ?case
-    by (simp add: rec [of "1 + k * 2"])
-qed
-
 lemma of_bits:
   "of_bits bs \<^bold>\<times> of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: int)"
     if "length bs = length cs" and "\<not> False \<^bold>* False" for bs cs
@@ -782,265 +351,18 @@
 
 end
 
-instantiation int :: bit_operations
-begin
-
-definition not_int :: "int \<Rightarrow> int"
-  where "not_int = complement"
-
-global_interpretation and_int: zip_int "(\<and>)"
-  defines and_int = and_int.F
-  by standard
-
-declare and_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
-  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
-
-global_interpretation and_int: semilattice "(AND) :: int \<Rightarrow> int \<Rightarrow> int"
-proof (rule semilattice.intro, fact and_int.abel_semigroup_axioms, standard)
-  show "k AND k = k" for k :: int
-    by (simp add: and_int.self)
-qed
-
-lemma zero_int_and_eq [simp]:
-  "0 AND k = 0" for k :: int
-  by simp
-
-lemma and_zero_int_eq [simp]:
-  "k AND 0 = 0" for k :: int
-  by simp
-
-lemma minus_int_and_eq [simp]:
-  "- 1 AND k = k" for k :: int
-  by simp
-
-lemma and_minus_int_eq [simp]:
-  "k AND - 1 = k" for k :: int
-  by simp
-
-global_interpretation or_int: zip_int "(\<or>)"
-  defines or_int = or_int.F
-  by standard
-
-declare or_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
-  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
-
-global_interpretation or_int: semilattice "(OR) :: int \<Rightarrow> int \<Rightarrow> int"
-proof (rule semilattice.intro, fact or_int.abel_semigroup_axioms, standard)
-  show "k OR k = k" for k :: int
-    by (simp add: or_int.self)
-qed
-
-lemma zero_int_or_eq [simp]:
-  "0 OR k = k" for k :: int
-  by simp
-
-lemma and_zero_or_eq [simp]:
-  "k OR 0 = k" for k :: int
-  by simp
-
-lemma minus_int_or_eq [simp]:
-  "- 1 OR k = - 1" for k :: int
-  by simp
-
-lemma or_minus_int_eq [simp]:
-  "k OR - 1 = - 1" for k :: int
-  by simp
-
-global_interpretation xor_int: zip_int "(\<noteq>)"
-  defines xor_int = xor_int.F
-  by standard
-
-declare xor_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
-  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
-
-lemma zero_int_xor_eq [simp]:
-  "0 XOR k = k" for k :: int
-  by simp
-
-lemma and_zero_xor_eq [simp]:
-  "k XOR 0 = k" for k :: int
-  by simp
-
-lemma minus_int_xor_eq [simp]:
-  "- 1 XOR k = complement k" for k :: int
-  by simp
-
-lemma xor_minus_int_eq [simp]:
-  "k XOR - 1 = complement k" for k :: int
-  by simp
-
-definition shift_left_int :: "int \<Rightarrow> nat \<Rightarrow> int"
-  where [simp]: "k << n = push_bit n k" for k :: int
-
-definition shift_right_int :: "int \<Rightarrow> nat \<Rightarrow> int"
-  where [simp]: "k >> n = drop_bit n k" for k :: int
-
-instance proof
-  show "semilattice ((AND) :: int \<Rightarrow> _)"
-    by (fact and_int.semilattice_axioms)
-  show "semilattice ((OR) :: int \<Rightarrow> _)"
-    by (fact or_int.semilattice_axioms)
-  show "abel_semigroup ((XOR) :: int \<Rightarrow> _)"
-    by (fact xor_int.abel_semigroup_axioms)
+instance int :: ring_bit_representation
+proof
   show "(not :: int \<Rightarrow> _) = of_bits \<circ> map Not \<circ> bits_of"
   proof (rule sym, rule ext)
     fix k :: int
     show "(of_bits \<circ> map Not \<circ> bits_of) k = NOT k"
       by (induction k rule: int_bit_induct) (simp_all add: not_int_def)
   qed
-  show "of_bits bs AND of_bits cs = (of_bits (map2 (\<and>) bs cs) :: int)"
-    if "length bs = length cs" for bs cs
-    using that by (rule and_int.of_bits) simp
-  show "of_bits bs OR of_bits cs = (of_bits (map2 (\<or>) bs cs) :: int)"
-    if "length bs = length cs" for bs cs
-    using that by (rule or_int.of_bits) simp
-  show "of_bits bs XOR of_bits cs = (of_bits (map2 (\<noteq>) bs cs) :: int)"
-    if "length bs = length cs" for bs cs
-    using that by (rule xor_int.of_bits) simp
-  show "k << n = of_bits (replicate n False @ bits_of k)"
+  show "shift_bit n k = of_bits (replicate n False @ bits_of k)"
     for k :: int and n :: nat
-    by (cases "n = 0") simp_all
-  show "k >> n = of_bits (drop n (bits_of k))"
-    if "n < length (bits_of k)"
-    for k :: int and n :: nat
-    using that by simp
-qed
+    by (cases "n = 0") (simp_all add: shift_bit_eq_push_bit)
+qed (simp_all add: and_int.of_bits or_int.of_bits xor_int.of_bits
+  drop_bit_eq_drop_bit)
 
 end
-
-global_interpretation and_int: semilattice_neutr "(AND)" "- 1 :: int"
-  by standard simp
-
-global_interpretation or_int: semilattice_neutr "(OR)" "0 :: int"
-  by standard simp
-
-global_interpretation xor_int: comm_monoid "(XOR)" "0 :: int"
-  by standard simp
-
-lemma not_int_simps [simp]:
-  "NOT 0 = (- 1 :: int)"
-  "NOT (- 1) = (0 :: int)"
-  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
-  by (auto simp add: not_int_def elim: oddE)
-
-lemma not_one_int [simp]:
-  "NOT 1 = (- 2 :: int)"
-  by simp
-
-lemma even_not_iff [simp]:
-  "even (NOT k) \<longleftrightarrow> odd k"
-  for k :: int
-  by (simp add: not_int_def)
-
-lemma not_div_2:
-  "NOT k div 2 = NOT (k div 2)"
-  for k :: int
-  by (simp add: complement_div_2 not_int_def)
-
-lemma one_and_int_eq [simp]:
-  "1 AND k = k mod 2" for k :: int
-  using and_int.rec [of 1] by (simp add: mod2_eq_if)
-
-lemma and_one_int_eq [simp]:
-  "k AND 1 = k mod 2" for k :: int
-  using one_and_int_eq [of 1] by (simp add: ac_simps)
-
-lemma one_or_int_eq [simp]:
-  "1 OR k = k + of_bool (even k)" for k :: int
-  using or_int.rec [of 1] by (auto elim: oddE)
-
-lemma or_one_int_eq [simp]:
-  "k OR 1 = k + of_bool (even k)" for k :: int
-  using one_or_int_eq [of k] by (simp add: ac_simps)
-
-lemma one_xor_int_eq [simp]:
-  "1 XOR k = k + of_bool (even k) - of_bool (odd k)" for k :: int
-  using xor_int.rec [of 1] by (auto elim: oddE)
-
-lemma xor_one_int_eq [simp]:
-  "k XOR 1 = k + of_bool (even k) - of_bool (odd k)" for k :: int
-  using one_xor_int_eq [of k] by (simp add: ac_simps)
-
-global_interpretation bit_int: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: int"
-  rewrites "bit_int.xor = ((XOR) :: int \<Rightarrow> _)"
-proof -
-  interpret bit_int: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: int"
-  proof
-    show "k AND (l OR r) = k AND l OR k AND r"
-      for k l r :: int
-    proof (induction k arbitrary: l r rule: int_bit_induct)
-      case zero
-      show ?case
-        by simp
-    next
-      case minus
-      show ?case
-        by simp
-    next
-      case (even k)
-      then show ?case by (simp add: and_int.rec [of "k * 2"]
-        or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of l])
-    next
-      case (odd k)
-      then show ?case by (simp add: and_int.rec [of "1 + k * 2"]
-        or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of "1 + (k AND l div 2) * 2"] or_int.rec [of l])
-    qed
-    show "k OR l AND r = (k OR l) AND (k OR r)"
-      for k l r :: int
-    proof (induction k arbitrary: l r rule: int_bit_induct)
-      case zero
-      then show ?case
-        by simp
-    next
-      case minus
-      then show ?case
-        by simp
-    next
-      case (even k)
-      then show ?case by (simp add: or_int.rec [of "k * 2"]
-        and_int.rec [of "(k OR l div 2) * 2"] and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l])
-    next
-      case (odd k)
-      then show ?case by (simp add: or_int.rec [of "1 + k * 2"]
-        and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l])
-    qed
-    show "k AND NOT k = 0" for k :: int
-      by (induction k rule: int_bit_induct)
-        (simp_all add: not_int_def complement_half minus_diff_commute [of 1])
-    show "k OR NOT k = - 1" for k :: int
-      by (induction k rule: int_bit_induct)
-        (simp_all add: not_int_def complement_half minus_diff_commute [of 1])
-  qed simp_all
-  show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: int)"
-    by (fact bit_int.boolean_algebra_axioms)
-  show "bit_int.xor = ((XOR) :: int \<Rightarrow> _)"
-  proof (rule ext)+
-    fix k l :: int
-    have "k XOR l = k AND NOT l OR NOT k AND l"
-    proof (induction k arbitrary: l rule: int_bit_induct)
-      case zero
-      show ?case
-        by simp
-    next
-      case minus
-      show ?case
-        by (simp add: not_int_def)
-    next
-      case (even k)
-      then show ?case
-        by (simp add: xor_int.rec [of "k * 2"] and_int.rec [of "k * 2"]
-          or_int.rec [of _ "1 + 2 * NOT k AND l"] not_div_2)
-          (simp add: and_int.rec [of _ l])
-    next
-      case (odd k)
-      then show ?case
-        by (simp add: xor_int.rec [of "1 + k * 2"] and_int.rec [of "1 + k * 2"]
-          or_int.rec [of _ "2 * NOT k AND l"] not_div_2)
-          (simp add: and_int.rec [of _ l])
-    qed
-    then show "bit_int.xor k l = k XOR l"
-      by (simp add: bit_int.xor_def)
-  qed
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Bit_Operations.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -0,0 +1,1033 @@
+(*  Author:  Florian Haftmann, TUM
+*)
+
+section \<open>Proof of concept for purely algebraically founded lists of bits\<close>
+
+theory Bit_Operations
+  imports
+    "HOL-Library.Boolean_Algebra"
+    Word
+begin
+
+hide_const (open) drop_bit take_bit
+
+subsection \<open>Algebraic structures with bits\<close>
+
+class semiring_bits = semiring_parity +
+  assumes bit_split_eq: \<open>\<And>a. of_bool (odd a) + 2 * (a div 2) = a\<close>
+    and bit_eq_rec: \<open>\<And>a b. a = b \<longleftrightarrow> (even a = even b) \<and> a div 2 = b div 2\<close>
+    and bit_induct [case_names stable rec]:
+    \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a)
+     \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a))
+        \<Longrightarrow> P a\<close>
+
+
+subsubsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
+
+instance nat :: semiring_bits
+proof
+  show \<open>of_bool (odd n) + 2 * (n div 2) = n\<close>
+    for n :: nat
+    by simp
+  show \<open>m = n \<longleftrightarrow> (even m \<longleftrightarrow> even n) \<and> m div 2 = n div 2\<close>
+    for m n :: nat
+    by (auto dest: odd_two_times_div_two_succ)
+  show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close>
+    and rec: \<open>\<And>n b. P n \<Longrightarrow> (of_bool b + 2 * n) div 2 = n \<Longrightarrow> P (of_bool b + 2 * n)\<close>
+    for P and n :: nat
+  proof (induction n rule: nat_bit_induct)
+    case zero
+    from stable [of 0] show ?case
+      by simp
+  next
+    case (even n)
+    with rec [of n False] show ?case
+      by simp
+  next
+    case (odd n)
+    with rec [of n True] show ?case
+      by simp
+  qed
+qed
+
+
+subsubsection \<open>Instance \<^typ>\<open>int\<close>\<close>
+
+instance int :: semiring_bits
+proof
+  show \<open>of_bool (odd k) + 2 * (k div 2) = k\<close>
+    for k :: int
+    by (auto elim: oddE)
+  show \<open>k = l \<longleftrightarrow> (even k \<longleftrightarrow> even l) \<and> k div 2 = l div 2\<close>
+    for k l :: int
+    by (auto dest: odd_two_times_div_two_succ)
+  show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close>
+    and rec: \<open>\<And>k b. P k \<Longrightarrow> (of_bool b + 2 * k) div 2 = k \<Longrightarrow> P (of_bool b + 2 * k)\<close>
+    for P and k :: int
+  proof (induction k rule: int_bit_induct)
+    case zero
+    from stable [of 0] show ?case
+      by simp
+  next
+    case minus
+    from stable [of \<open>- 1\<close>] show ?case
+      by simp
+  next
+    case (even k)
+    with rec [of k False] show ?case
+      by (simp add: ac_simps)
+  next
+    case (odd k)
+    with rec [of k True] show ?case
+      by (simp add: ac_simps)
+  qed
+qed
+
+
+subsubsection \<open>Instance \<^typ>\<open>'a word\<close>\<close>
+
+instance word :: (len) semiring_bits
+proof
+  show \<open>of_bool (odd a) + 2 * (a div 2) = a\<close>
+    for a :: \<open>'a word\<close>
+    apply (cases \<open>even a\<close>; simp, transfer; cases rule: length_cases [where ?'a = 'a])
+      apply auto
+    apply (auto simp add: take_bit_eq_mod)
+    apply (metis add.commute even_take_bit_eq len_not_eq_0 mod_mod_trivial odd_two_times_div_two_succ take_bit_eq_mod)
+    done
+  show \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close>
+    for a b :: \<open>'a word\<close>
+    apply transfer
+    apply (cases rule: length_cases [where ?'a = 'a])
+     apply auto
+       apply (metis even_take_bit_eq len_not_eq_0)
+      apply (metis even_take_bit_eq len_not_eq_0)
+    apply (metis (no_types, hide_lams) diff_add_cancel dvd_div_mult_self even_take_bit_eq mult_2_right take_bit_add take_bit_minus)
+    apply (metis bit_ident drop_bit_eq_div drop_bit_half even_take_bit_eq even_two_times_div_two mod_div_trivial odd_two_times_div_two_succ take_bit_eq_mod)
+    done
+  show \<open>P a\<close> if stable: \<open>\<And>a. a div 2 = a \<Longrightarrow> P a\<close>
+    and rec: \<open>\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)\<close>
+  for P and a :: \<open>'a word\<close>
+  proof (induction a rule: word_bit_induct)
+    case zero
+    from stable [of 0] show ?case
+      by simp
+  next
+    case (even a)
+    with rec [of a False] show ?case
+      using bit_word_half_eq [of a False] by (simp add: ac_simps)
+  next
+    case (odd a)
+    with rec [of a True] show ?case
+      using bit_word_half_eq [of a True] by (simp add: ac_simps)
+  qed
+qed
+
+
+subsection \<open>Bit shifts in suitable algebraic structures\<close>
+
+class semiring_bit_shifts = semiring_bits +
+  fixes shift_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  assumes shift_bit_eq_mult: \<open>shift_bit n a = a * 2 ^ n\<close>
+  fixes drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  assumes drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close>
+begin
+
+definition take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  where take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close>
+
+text \<open>
+  Logically, \<^const>\<open>shift_bit\<close>,
+  \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them
+  as separate operations makes proofs easier, otherwise proof automation
+  would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> in a way obfuscating the basic
+  algebraic relationships between those operations; having
+  \<^const>\<open>push_bit\<close> and \<^const>\<open>drop_bit\<close> as definitional class operations
+  takes into account that specific instances of these can be implemented
+  differently wrt. code generation.
+\<close>
+
+end
+
+subsubsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
+
+instantiation nat :: semiring_bit_shifts
+begin
+
+definition shift_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+  where \<open>shift_bit_nat n m = m * 2 ^ n\<close>
+
+definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+  where \<open>drop_bit_nat n m = m div 2 ^ n\<close>
+
+instance proof
+  show \<open>shift_bit n m = m * 2 ^ n\<close> for n m :: nat
+    by (simp add: shift_bit_nat_def)
+  show \<open>drop_bit n m = m div 2 ^ n\<close> for n m :: nat
+    by (simp add: drop_bit_nat_def)
+qed
+
+end
+
+
+subsubsection \<open>Instance \<^typ>\<open>int\<close>\<close>
+
+instantiation int :: semiring_bit_shifts
+begin
+
+definition shift_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+  where \<open>shift_bit_int n k = k * 2 ^ n\<close>
+
+definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+  where \<open>drop_bit_int n k = k div 2 ^ n\<close>
+
+instance proof
+  show \<open>shift_bit n k = k * 2 ^ n\<close> for n :: nat and k :: int
+    by (simp add: shift_bit_int_def)
+  show \<open>drop_bit n k = k div 2 ^ n\<close> for n :: nat and k :: int
+    by (simp add: drop_bit_int_def)
+qed
+
+end
+
+lemma shift_bit_eq_push_bit:
+  \<open>shift_bit = (push_bit :: nat \<Rightarrow> int \<Rightarrow> int)\<close>
+  by (simp add: fun_eq_iff push_bit_eq_mult shift_bit_eq_mult)
+
+lemma drop_bit_eq_drop_bit:
+  \<open>drop_bit = (Parity.drop_bit :: nat \<Rightarrow> int \<Rightarrow> int)\<close>
+  by (simp add: fun_eq_iff drop_bit_eq_div Parity.drop_bit_eq_div)
+
+lemma take_bit_eq_take_bit:
+  \<open>take_bit = (Parity.take_bit :: nat \<Rightarrow> int \<Rightarrow> int)\<close>
+  by (simp add: fun_eq_iff take_bit_eq_mod Parity.take_bit_eq_mod)
+
+
+subsubsection \<open>Instance \<^typ>\<open>'a word\<close>\<close>
+
+instantiation word :: (len) semiring_bit_shifts
+begin
+
+lift_definition shift_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
+  is shift_bit
+proof -
+  show \<open>Parity.take_bit LENGTH('a) (shift_bit n k) = Parity.take_bit LENGTH('a) (shift_bit n l)\<close>
+    if \<open>Parity.take_bit LENGTH('a) k = Parity.take_bit LENGTH('a) l\<close> for k l :: int and n :: nat
+  proof -
+    from that
+    have \<open>Parity.take_bit (LENGTH('a) - n) (Parity.take_bit LENGTH('a) k)
+      = Parity.take_bit (LENGTH('a) - n) (Parity.take_bit LENGTH('a) l)\<close>
+      by simp
+    moreover have \<open>min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\<close>
+      by simp
+    ultimately show ?thesis
+      by (simp add: shift_bit_eq_push_bit take_bit_push_bit)
+  qed
+qed
+
+lift_definition drop_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>
+  is \<open>\<lambda>n. drop_bit n \<circ> take_bit LENGTH('a)\<close>
+  by (simp add: take_bit_eq_mod Parity.take_bit_eq_mod)
+
+instance proof
+  show \<open>shift_bit n a = a * 2 ^ n\<close> for n :: nat and a :: "'a word"
+    by transfer (simp add: shift_bit_eq_mult)
+  show \<open>drop_bit n a = a div 2 ^ n\<close> for n :: nat and a :: "'a word"
+  proof (cases \<open>n < LENGTH('a)\<close>)
+    case True
+    then show ?thesis
+      by transfer
+        (simp add: Parity.take_bit_eq_mod take_bit_eq_mod drop_bit_eq_div)
+  next
+    case False
+    then obtain m where n: \<open>n = LENGTH('a) + m\<close>
+      by (auto simp add: not_less dest: le_Suc_ex)
+    then show ?thesis
+      by transfer
+        (simp add: Parity.take_bit_eq_mod take_bit_eq_mod drop_bit_eq_div power_add zdiv_zmult2_eq)
+  qed
+qed
+
+end
+
+
+subsection \<open>Bit operations in suitable algebraic structures\<close>
+
+class semiring_bit_operations = semiring_bit_shifts +
+  fixes "and" :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "AND" 64)
+    and or :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "OR"  59)
+    and xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "XOR" 59)
+begin
+
+text \<open>
+  We want the bitwise operations to bind slightly weaker
+  than \<open>+\<close> and \<open>-\<close>, but \<open>~~\<close> to
+  bind slightly stronger than \<open>*\<close>.
+  For the sake of code generation
+  the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close>
+  are specified as definitional class operations.
+
+\<close>
+
+definition bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close>
+  where \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close>
+
+definition map_bit :: \<open>nat \<Rightarrow> (bool \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  where \<open>map_bit n f a = take_bit n a + shift_bit n (of_bool (f (bit a n)) + drop_bit (Suc n) a)\<close>
+
+definition set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  where \<open>set_bit n = map_bit n top\<close>
+
+definition unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  where \<open>unset_bit n = map_bit n bot\<close>
+
+definition flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+  where \<open>flip_bit n = map_bit n Not\<close>
+
+text \<open>
+  The logical core are \<^const>\<open>bit\<close> and \<^const>\<open>map_bit\<close>; having 
+  <^const>\<open>set_bit\<close>, \<^const>\<open>unset_bit\<close> and \<^const>\<open>flip_bit\<close> as separate
+  operations allows to implement them using bit masks later.
+\<close>
+
+end
+
+class ring_bit_operations = semiring_bit_operations + ring_parity +
+  fixes not :: \<open>'a \<Rightarrow> 'a\<close>  (\<open>NOT\<close>)
+  assumes boolean_algebra: \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
+    and boolean_algebra_xor_eq: \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
+begin
+
+sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
+  rewrites \<open>bit.xor = (XOR)\<close>
+proof -
+  interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
+    by (fact boolean_algebra)
+  show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
+    by standard
+  show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>  
+    by (fact boolean_algebra_xor_eq)
+qed
+
+text \<open>
+  For the sake of code generation \<^const>\<open>not\<close> is specified as
+  definitional class operation.  Note that \<^const>\<open>not\<close> has no
+  sensible definition for unlimited but only positive bit strings
+  (type \<^typ>\<open>nat\<close>).
+\<close>
+
+end
+
+
+subsubsection \<open>Instance \<^typ>\<open>nat\<close>\<close>
+
+locale zip_nat = single: abel_semigroup f
+    for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70) +
+  assumes end_of_bits: "\<not> False \<^bold>* False"
+begin
+
+lemma False_P_imp:
+  "False \<^bold>* True \<and> P" if "False \<^bold>* P"
+  using that end_of_bits by (cases P) simp_all
+
+function F :: "nat \<Rightarrow> nat \<Rightarrow> nat"  (infixl "\<^bold>\<times>" 70)
+  where "m \<^bold>\<times> n = (if m = 0 \<and> n = 0 then 0
+    else of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2)"
+  by auto
+
+termination
+  by (relation "measure (case_prod (+))") auto
+
+lemma zero_left_eq:
+  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
+  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
+
+lemma zero_right_eq:
+  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
+  by (induction m rule: nat_bit_induct) (simp_all add: end_of_bits)
+
+lemma simps [simp]:
+  "0 \<^bold>\<times> 0 = 0"
+  "0 \<^bold>\<times> n = of_bool (False \<^bold>* True) * n"
+  "m \<^bold>\<times> 0 = of_bool (True \<^bold>* False) * m"
+  "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
+  by (simp_all only: zero_left_eq zero_right_eq) simp
+
+lemma rec:
+  "m \<^bold>\<times> n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\<times> (n div 2) * 2"
+  by (cases "m = 0 \<and> n = 0") (auto simp add: end_of_bits)
+
+declare F.simps [simp del]
+
+sublocale abel_semigroup F
+proof
+  show "m \<^bold>\<times> n \<^bold>\<times> q = m \<^bold>\<times> (n \<^bold>\<times> q)" for m n q :: nat
+  proof (induction m arbitrary: n q rule: nat_bit_induct)
+    case zero
+    show ?case
+      by simp
+  next
+    case (even m)
+    with rec [of "2 * m"] rec [of _ q] show ?case
+      by (cases "even n") (auto dest: False_P_imp)
+  next
+    case (odd m)
+    with rec [of "Suc (2 * m)"] rec [of _ q] show ?case
+      by (cases "even n"; cases "even q")
+        (auto dest: False_P_imp simp add: ac_simps)
+  qed
+  show "m \<^bold>\<times> n = n \<^bold>\<times> m" for m n :: nat
+  proof (induction m arbitrary: n rule: nat_bit_induct)
+    case zero
+    show ?case
+      by (simp add: ac_simps)
+  next
+    case (even m)
+    with rec [of "2 * m" n] rec [of n "2 * m"] show ?case
+      by (simp add: ac_simps)
+  next
+    case (odd m)
+    with rec [of "Suc (2 * m)" n] rec [of n "Suc (2 * m)"] show ?case
+      by (simp add: ac_simps)
+  qed
+qed
+
+lemma self [simp]:
+  "n \<^bold>\<times> n = of_bool (True \<^bold>* True) * n"
+  by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits)
+
+lemma even_iff [simp]:
+  "even (m \<^bold>\<times> n) \<longleftrightarrow> \<not> (odd m \<^bold>* odd n)"
+proof (induction m arbitrary: n rule: nat_bit_induct)
+  case zero
+  show ?case
+    by (cases "even n") (simp_all add: end_of_bits)
+next
+  case (even m)
+  then show ?case
+    by (simp add: rec [of "2 * m"])
+next
+  case (odd m)
+  then show ?case
+    by (simp add: rec [of "Suc (2 * m)"])
+qed
+
+end
+
+instantiation nat :: semiring_bit_operations
+begin
+
+global_interpretation and_nat: zip_nat "(\<and>)"
+  defines and_nat = and_nat.F
+  by standard auto
+
+global_interpretation and_nat: semilattice "(AND) :: nat \<Rightarrow> nat \<Rightarrow> nat"
+proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard)
+  show "n AND n = n" for n :: nat
+    by (simp add: and_nat.self)
+qed
+
+declare and_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
+  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
+
+lemma zero_nat_and_eq [simp]:
+  "0 AND n = 0" for n :: nat
+  by simp
+
+lemma and_zero_nat_eq [simp]:
+  "n AND 0 = 0" for n :: nat
+  by simp
+
+global_interpretation or_nat: zip_nat "(\<or>)"
+  defines or_nat = or_nat.F
+  by standard auto
+
+global_interpretation or_nat: semilattice "(OR) :: nat \<Rightarrow> nat \<Rightarrow> nat"
+proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard)
+  show "n OR n = n" for n :: nat
+    by (simp add: or_nat.self)
+qed
+
+declare or_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
+  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
+
+lemma zero_nat_or_eq [simp]:
+  "0 OR n = n" for n :: nat
+  by simp
+
+lemma or_zero_nat_eq [simp]:
+  "n OR 0 = n" for n :: nat
+  by simp
+
+global_interpretation xor_nat: zip_nat "(\<noteq>)"
+  defines xor_nat = xor_nat.F
+  by standard auto
+
+declare xor_nat.simps [simp] \<comment> \<open>inconsistent declaration handling by
+  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
+
+lemma zero_nat_xor_eq [simp]:
+  "0 XOR n = n" for n :: nat
+  by simp
+
+lemma xor_zero_nat_eq [simp]:
+  "n XOR 0 = n" for n :: nat
+  by simp
+
+instance ..
+
+end
+
+global_interpretation or_nat: semilattice_neutr "(OR)" "0 :: nat"
+  by standard simp
+
+global_interpretation xor_nat: comm_monoid "(XOR)" "0 :: nat"
+  by standard simp
+
+lemma Suc_0_and_eq [simp]:
+  "Suc 0 AND n = n mod 2"
+  by (cases n) auto
+
+lemma and_Suc_0_eq [simp]:
+  "n AND Suc 0 = n mod 2"
+  using Suc_0_and_eq [of n] by (simp add: ac_simps)
+
+lemma Suc_0_or_eq [simp]:
+  "Suc 0 OR n = n + of_bool (even n)"
+  by (cases n) (simp_all add: ac_simps)
+
+lemma or_Suc_0_eq [simp]:
+  "n OR Suc 0 = n + of_bool (even n)"
+  using Suc_0_or_eq [of n] by (simp add: ac_simps)
+
+lemma Suc_0_xor_eq [simp]:
+  "Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)"
+  by (cases n) (simp_all add: ac_simps)
+
+lemma xor_Suc_0_eq [simp]:
+  "n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)"
+  using Suc_0_xor_eq [of n] by (simp add: ac_simps)
+
+
+subsubsection \<open>Instance \<^typ>\<open>int\<close>\<close>
+
+abbreviation (input) complement :: "int \<Rightarrow> int"
+  where "complement k \<equiv> - k - 1"
+
+lemma complement_half:
+  "complement (k * 2) div 2 = complement k"
+  by simp
+
+lemma complement_div_2:
+  "complement (k div 2) = complement k div 2"
+  by linarith
+
+locale zip_int = single: abel_semigroup f
+  for f :: "bool \<Rightarrow> bool \<Rightarrow> bool"  (infixl "\<^bold>*" 70)
+begin
+ 
+lemma False_False_imp_True_True:
+  "True \<^bold>* True" if "False \<^bold>* False"
+proof (rule ccontr)
+  assume "\<not> True \<^bold>* True"
+  with that show False
+    using single.assoc [of False True True]
+    by (cases "False \<^bold>* True") simp_all
+qed
+
+function F :: "int \<Rightarrow> int \<Rightarrow> int"  (infixl "\<^bold>\<times>" 70)
+  where "k \<^bold>\<times> l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
+    then - of_bool (odd k \<^bold>* odd l)
+    else of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2)"
+  by auto
+
+termination
+  by (relation "measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))") auto
+
+lemma zero_left_eq:
+  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> l
+     | (True, False) \<Rightarrow> complement l
+     | (True, True) \<Rightarrow> - 1)"
+  by (induction l rule: int_bit_induct)
+   (simp_all split: bool.split) 
+
+lemma minus_left_eq:
+  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> l
+     | (True, False) \<Rightarrow> complement l
+     | (True, True) \<Rightarrow> - 1)"
+  by (induction l rule: int_bit_induct)
+   (simp_all split: bool.split) 
+
+lemma zero_right_eq:
+  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> k
+     | (True, False) \<Rightarrow> complement k
+     | (True, True) \<Rightarrow> - 1)"
+  by (induction k rule: int_bit_induct)
+    (simp_all add: ac_simps split: bool.split)
+
+lemma minus_right_eq:
+  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> k
+     | (True, False) \<Rightarrow> complement k
+     | (True, True) \<Rightarrow> - 1)"
+  by (induction k rule: int_bit_induct)
+    (simp_all add: ac_simps split: bool.split)
+
+lemma simps [simp]:
+  "0 \<^bold>\<times> 0 = - of_bool (False \<^bold>* False)"
+  "- 1 \<^bold>\<times> 0 = - of_bool (True \<^bold>* False)"
+  "0 \<^bold>\<times> - 1 = - of_bool (False \<^bold>* True)"
+  "- 1 \<^bold>\<times> - 1 = - of_bool (True \<^bold>* True)"
+  "0 \<^bold>\<times> l = (case (False \<^bold>* False, False \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> l
+     | (True, False) \<Rightarrow> complement l
+     | (True, True) \<Rightarrow> - 1)"
+  "- 1 \<^bold>\<times> l = (case (True \<^bold>* False, True \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> l
+     | (True, False) \<Rightarrow> complement l
+     | (True, True) \<Rightarrow> - 1)"
+  "k \<^bold>\<times> 0 = (case (False \<^bold>* False, False \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> k
+     | (True, False) \<Rightarrow> complement k
+     | (True, True) \<Rightarrow> - 1)"
+  "k \<^bold>\<times> - 1 = (case (True \<^bold>* False, True \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> k
+     | (True, False) \<Rightarrow> complement k
+     | (True, True) \<Rightarrow> - 1)"
+  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> l \<noteq> - 1
+    \<Longrightarrow> k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
+  by simp_all[4] (simp_all only: zero_left_eq minus_left_eq zero_right_eq minus_right_eq, simp)
+
+declare F.simps [simp del]
+
+lemma rec:
+  "k \<^bold>\<times> l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\<times> (l div 2) * 2"
+  by (cases "k \<in> {0, - 1} \<and> l \<in> {0, - 1}") (auto simp add: ac_simps F.simps [of k l] split: bool.split)
+
+sublocale abel_semigroup F
+proof
+  show "k \<^bold>\<times> l \<^bold>\<times> r = k \<^bold>\<times> (l \<^bold>\<times> r)" for k l r :: int
+  proof (induction k arbitrary: l r rule: int_bit_induct)
+    case zero
+    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "False \<^bold>* False" "\<not> False \<^bold>* True"
+    proof (induction l arbitrary: r rule: int_bit_induct)
+      case zero
+      from that show ?case
+        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+    next
+      case minus
+      from that show ?case
+        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+    next
+      case (even l)
+      with that rec [of _ r] show ?case
+        by (cases "even r")
+          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
+    next
+      case (odd l)
+      moreover have "- l - 1 = - 1 - l"
+        by simp
+      ultimately show ?case
+        using that rec [of _ r]
+        by (cases "even r")
+          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+    qed
+    then show ?case
+      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+  next
+    case minus
+    have "complement l \<^bold>\<times> r = complement (l \<^bold>\<times> r)" if "\<not> True \<^bold>* True" "False \<^bold>* True"
+    proof (induction l arbitrary: r rule: int_bit_induct)
+      case zero
+      from that show ?case
+        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+    next
+      case minus
+      from that show ?case
+        by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+    next
+      case (even l)
+      with that rec [of _ r] show ?case
+        by (cases "even r")
+          (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits)
+    next
+      case (odd l)
+      moreover have "- l - 1 = - 1 - l"
+        by simp
+      ultimately show ?case
+        using that rec [of _ r]
+        by (cases "even r")
+          (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+    qed
+    then show ?case
+      by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits)
+  next
+    case (even k)
+    with rec [of "k * 2"] rec [of _ r] show ?case
+      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
+  next
+    case (odd k)
+    with rec [of "1 + k * 2"] rec [of _ r] show ?case
+      by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True)
+  qed
+  show "k \<^bold>\<times> l = l \<^bold>\<times> k" for k l :: int
+  proof (induction k arbitrary: l rule: int_bit_induct)
+    case zero
+    show ?case
+      by simp
+  next
+    case minus
+    show ?case
+      by simp
+  next
+    case (even k)
+    with rec [of "k * 2" l] rec [of l "k * 2"] show ?case
+      by (simp add: ac_simps)
+  next
+    case (odd k)
+    with rec [of "k * 2 + 1" l] rec [of l "k * 2 + 1"] show ?case
+      by (simp add: ac_simps)
+  qed
+qed
+
+lemma self [simp]:
+  "k \<^bold>\<times> k = (case (False \<^bold>* False, True \<^bold>* True)
+    of (False, False) \<Rightarrow> 0
+     | (False, True) \<Rightarrow> k
+     | (True, True) \<Rightarrow> - 1)"
+  by (induction k rule: int_bit_induct) (auto simp add: False_False_imp_True_True split: bool.split)
+
+lemma even_iff [simp]:
+  "even (k \<^bold>\<times> l) \<longleftrightarrow> \<not> (odd k \<^bold>* odd l)"
+proof (induction k arbitrary: l rule: int_bit_induct)
+  case zero
+  show ?case
+    by (cases "even l") (simp_all split: bool.splits)
+next
+  case minus
+  show ?case
+    by (cases "even l") (simp_all split: bool.splits)
+next
+  case (even k)
+  then show ?case
+    by (simp add: rec [of "k * 2"])
+next
+  case (odd k)
+  then show ?case
+    by (simp add: rec [of "1 + k * 2"])
+qed
+
+end
+
+instantiation int :: ring_bit_operations
+begin
+
+definition not_int :: "int \<Rightarrow> int"
+  where "not_int = complement"
+
+global_interpretation and_int: zip_int "(\<and>)"
+  defines and_int = and_int.F
+  by standard
+
+declare and_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
+  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
+
+global_interpretation and_int: semilattice "(AND) :: int \<Rightarrow> int \<Rightarrow> int"
+proof (rule semilattice.intro, fact and_int.abel_semigroup_axioms, standard)
+  show "k AND k = k" for k :: int
+    by (simp add: and_int.self)
+qed
+
+lemma zero_int_and_eq [simp]:
+  "0 AND k = 0" for k :: int
+  by simp
+
+lemma and_zero_int_eq [simp]:
+  "k AND 0 = 0" for k :: int
+  by simp
+
+lemma minus_int_and_eq [simp]:
+  "- 1 AND k = k" for k :: int
+  by simp
+
+lemma and_minus_int_eq [simp]:
+  "k AND - 1 = k" for k :: int
+  by simp
+
+global_interpretation or_int: zip_int "(\<or>)"
+  defines or_int = or_int.F
+  by standard
+
+declare or_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
+  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
+
+global_interpretation or_int: semilattice "(OR) :: int \<Rightarrow> int \<Rightarrow> int"
+proof (rule semilattice.intro, fact or_int.abel_semigroup_axioms, standard)
+  show "k OR k = k" for k :: int
+    by (simp add: or_int.self)
+qed
+
+lemma zero_int_or_eq [simp]:
+  "0 OR k = k" for k :: int
+  by simp
+
+lemma and_zero_or_eq [simp]:
+  "k OR 0 = k" for k :: int
+  by simp
+
+lemma minus_int_or_eq [simp]:
+  "- 1 OR k = - 1" for k :: int
+  by simp
+
+lemma or_minus_int_eq [simp]:
+  "k OR - 1 = - 1" for k :: int
+  by simp
+
+global_interpretation xor_int: zip_int "(\<noteq>)"
+  defines xor_int = xor_int.F
+  by standard
+
+declare xor_int.simps [simp] \<comment> \<open>inconsistent declaration handling by
+  \<open>global_interpretation\<close> in \<open>instantiation\<close>\<close>
+
+lemma zero_int_xor_eq [simp]:
+  "0 XOR k = k" for k :: int
+  by simp
+
+lemma and_zero_xor_eq [simp]:
+  "k XOR 0 = k" for k :: int
+  by simp
+
+lemma minus_int_xor_eq [simp]:
+  "- 1 XOR k = complement k" for k :: int
+  by simp
+
+lemma xor_minus_int_eq [simp]:
+  "k XOR - 1 = complement k" for k :: int
+  by simp
+
+lemma not_div_2:
+  "NOT k div 2 = NOT (k div 2)"
+  for k :: int
+  by (simp add: complement_div_2 not_int_def)
+
+lemma not_int_simps [simp]:
+  "NOT 0 = (- 1 :: int)"
+  "NOT (- 1) = (0 :: int)"
+  "k \<noteq> 0 \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int
+  by (auto simp add: not_int_def elim: oddE)
+
+lemma not_one_int [simp]:
+  "NOT 1 = (- 2 :: int)"
+  by simp
+
+lemma even_not_iff [simp]:
+  "even (NOT k) \<longleftrightarrow> odd k"
+  for k :: int
+  by (simp add: not_int_def)
+
+instance proof
+  interpret bit_int: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: int"
+  proof
+    show "k AND (l OR r) = k AND l OR k AND r"
+      for k l r :: int
+    proof (induction k arbitrary: l r rule: int_bit_induct)
+      case zero
+      show ?case
+        by simp
+    next
+      case minus
+      show ?case
+        by simp
+    next
+      case (even k)
+      then show ?case by (simp add: and_int.rec [of "k * 2"]
+        or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of l])
+    next
+      case (odd k)
+      then show ?case by (simp add: and_int.rec [of "1 + k * 2"]
+        or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of "1 + (k AND l div 2) * 2"] or_int.rec [of l])
+    qed
+    show "k OR l AND r = (k OR l) AND (k OR r)"
+      for k l r :: int
+    proof (induction k arbitrary: l r rule: int_bit_induct)
+      case zero
+      then show ?case
+        by simp
+    next
+      case minus
+      then show ?case
+        by simp
+    next
+      case (even k)
+      then show ?case by (simp add: or_int.rec [of "k * 2"]
+        and_int.rec [of "(k OR l div 2) * 2"] and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l])
+    next
+      case (odd k)
+      then show ?case by (simp add: or_int.rec [of "1 + k * 2"]
+        and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l])
+    qed
+    show "k AND NOT k = 0" for k :: int
+      by (induction k rule: int_bit_induct)
+        (simp_all add: not_int_def complement_half minus_diff_commute [of 1])
+    show "k OR NOT k = - 1" for k :: int
+      by (induction k rule: int_bit_induct)
+        (simp_all add: not_int_def complement_half minus_diff_commute [of 1])
+  qed (simp_all add: and_int.assoc or_int.assoc,
+    simp_all add: and_int.commute or_int.commute)
+  show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: int)"
+    by (fact bit_int.boolean_algebra_axioms)
+  show "bit_int.xor = ((XOR) :: int \<Rightarrow> _)"
+  proof (rule ext)+
+    fix k l :: int
+    have "k XOR l = k AND NOT l OR NOT k AND l"
+    proof (induction k arbitrary: l rule: int_bit_induct)
+      case zero
+      show ?case
+        by simp
+    next
+      case minus
+      show ?case
+        by (simp add: not_int_def)
+    next
+      case (even k)
+      then show ?case
+        by (simp add: xor_int.rec [of "k * 2"] and_int.rec [of "k * 2"]
+          or_int.rec [of _ "1 + 2 * NOT k AND l"] not_div_2)
+          (simp add: and_int.rec [of _ l])
+    next
+      case (odd k)
+      then show ?case
+        by (simp add: xor_int.rec [of "1 + k * 2"] and_int.rec [of "1 + k * 2"]
+          or_int.rec [of _ "2 * NOT k AND l"] not_div_2)
+          (simp add: and_int.rec [of _ l])
+    qed
+    then show "bit_int.xor k l = k XOR l"
+      by (simp add: bit_int.xor_def)
+  qed
+qed
+
+end
+
+lemma one_and_int_eq [simp]:
+  "1 AND k = k mod 2" for k :: int
+  using and_int.rec [of 1] by (simp add: mod2_eq_if)
+
+lemma and_one_int_eq [simp]:
+  "k AND 1 = k mod 2" for k :: int
+  using one_and_int_eq [of 1] by (simp add: ac_simps)
+
+lemma one_or_int_eq [simp]:
+  "1 OR k = k + of_bool (even k)" for k :: int
+  using or_int.rec [of 1] by (auto elim: oddE)
+
+lemma or_one_int_eq [simp]:
+  "k OR 1 = k + of_bool (even k)" for k :: int
+  using one_or_int_eq [of k] by (simp add: ac_simps)
+
+lemma one_xor_int_eq [simp]:
+  "1 XOR k = k + of_bool (even k) - of_bool (odd k)" for k :: int
+  using xor_int.rec [of 1] by (auto elim: oddE)
+
+lemma xor_one_int_eq [simp]:
+  "k XOR 1 = k + of_bool (even k) - of_bool (odd k)" for k :: int
+  using one_xor_int_eq [of k] by (simp add: ac_simps)
+
+lemma take_bit_complement_iff:
+  "Parity.take_bit n (complement k) = Parity.take_bit n (complement l) \<longleftrightarrow> Parity.take_bit n k = Parity.take_bit n l"
+  for k l :: int
+  by (simp add: Parity.take_bit_eq_mod mod_eq_dvd_iff dvd_diff_commute)
+
+lemma take_bit_not_iff:
+  "Parity.take_bit n (NOT k) = Parity.take_bit n (NOT l) \<longleftrightarrow> Parity.take_bit n k = Parity.take_bit n l"
+  for k l :: int
+  by (simp add: not_int_def take_bit_complement_iff)
+
+lemma take_bit_and [simp]:
+  "Parity.take_bit n (k AND l) = Parity.take_bit n k AND Parity.take_bit n l"
+  for k l :: int
+  apply (induction n arbitrary: k l)
+   apply simp
+  apply (subst and_int.rec)
+  apply (subst (2) and_int.rec)
+  apply simp
+  done
+
+lemma take_bit_or [simp]:
+  "Parity.take_bit n (k OR l) = Parity.take_bit n k OR Parity.take_bit n l"
+  for k l :: int
+  apply (induction n arbitrary: k l)
+   apply simp
+  apply (subst or_int.rec)
+  apply (subst (2) or_int.rec)
+  apply simp
+  done
+
+lemma take_bit_xor [simp]:
+  "Parity.take_bit n (k XOR l) = Parity.take_bit n k XOR Parity.take_bit n l"
+  for k l :: int
+  apply (induction n arbitrary: k l)
+   apply simp
+  apply (subst xor_int.rec)
+  apply (subst (2) xor_int.rec)
+  apply simp
+  done
+
+
+subsubsection \<open>Instance \<^typ>\<open>'a word\<close>\<close>
+
+instantiation word :: (len) ring_bit_operations
+begin
+
+lift_definition not_word :: "'a word \<Rightarrow> 'a word"
+  is not
+  by (simp add: take_bit_not_iff)
+
+lift_definition and_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is "and"
+  by simp
+
+lift_definition or_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is or
+  by simp
+
+lift_definition xor_word ::  "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is xor
+  by simp
+
+instance proof
+  interpret bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a word"
+  proof
+    show "a AND (b OR c) = a AND b OR a AND c"
+      for a b c :: "'a word"
+      by transfer (simp add: bit.conj_disj_distrib)
+    show "a OR b AND c = (a OR b) AND (a OR c)"
+      for a b c :: "'a word"
+      by transfer (simp add: bit.disj_conj_distrib)
+  qed (transfer; simp add: ac_simps)+
+  show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: 'a word)"
+    by (fact bit_word.boolean_algebra_axioms)
+  show "bit_word.xor = ((XOR) :: 'a word \<Rightarrow> _)"
+  proof (rule ext)+
+    fix a b :: "'a word"
+    have "a XOR b = a AND NOT b OR NOT a AND b"
+      by transfer (simp add: bit.xor_def)
+    then show "bit_word.xor a b = a XOR b"
+      by (simp add: bit_word.xor_def)
+  qed
+qed
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Word.thy	Tue Nov 05 22:56:06 2019 +0100
@@ -0,0 +1,574 @@
+(*  Author:  Florian Haftmann, TUM
+*)
+
+section \<open>Proof of concept for algebraically founded bit word types\<close>
+
+theory Word
+  imports
+    Main
+    "HOL-Library.Type_Length"
+begin
+
+subsection \<open>Preliminaries\<close>
+
+context ab_group_add
+begin
+
+lemma minus_diff_commute:
+  "- b - a = - a - b"
+  by (simp only: diff_conv_add_uminus add.commute)
+
+end
+
+lemma take_bit_uminus:
+  "take_bit n (- (take_bit n k)) = take_bit n (- k)" for k :: int
+  by (simp add: take_bit_eq_mod mod_minus_eq)
+
+lemma take_bit_minus:
+  "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" for k l :: int
+  by (simp add: take_bit_eq_mod mod_diff_eq)
+
+lemma take_bit_nonnegative [simp]:
+  "take_bit n k \<ge> 0" for k :: int
+  by (simp add: take_bit_eq_mod)
+
+definition signed_take_bit :: "nat \<Rightarrow> int \<Rightarrow> int"
+  where signed_take_bit_eq_take_bit:
+    "signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n"
+
+lemma signed_take_bit_eq_take_bit':
+  "signed_take_bit (n - Suc 0) k = take_bit n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)" if "n > 0"
+  using that by (simp add: signed_take_bit_eq_take_bit)
+
+lemma signed_take_bit_0 [simp]:
+  "signed_take_bit 0 k = - (k mod 2)"
+proof (cases "even k")
+  case True
+  then have "odd (k + 1)"
+    by simp
+  then have "(k + 1) mod 2 = 1"
+    by (simp add: even_iff_mod_2_eq_zero)
+  with True show ?thesis
+    by (simp add: signed_take_bit_eq_take_bit)
+next
+  case False
+  then show ?thesis
+    by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one)
+qed
+
+lemma signed_take_bit_Suc [simp]:
+  "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2"
+  by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps)
+
+lemma signed_take_bit_of_0 [simp]:
+  "signed_take_bit n 0 = 0"
+  by (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod)
+
+lemma signed_take_bit_of_minus_1 [simp]:
+  "signed_take_bit n (- 1) = - 1"
+  by (induct n) simp_all
+
+lemma signed_take_bit_eq_iff_take_bit_eq:
+  "signed_take_bit (n - Suc 0) k = signed_take_bit (n - Suc 0) l \<longleftrightarrow> take_bit n k = take_bit n l" (is "?P \<longleftrightarrow> ?Q")
+  if "n > 0"
+proof -
+  from that obtain m where m: "n = Suc m"
+    by (cases n) auto
+  show ?thesis
+  proof
+    assume ?Q
+    have "take_bit (Suc m) (k + 2 ^ m) =
+      take_bit (Suc m) (take_bit (Suc m) k + take_bit (Suc m) (2 ^ m))"
+      by (simp only: take_bit_add)
+    also have "\<dots> =
+      take_bit (Suc m) (take_bit (Suc m) l + take_bit (Suc m) (2 ^ m))"
+      by (simp only: \<open>?Q\<close> m [symmetric])
+    also have "\<dots> = take_bit (Suc m) (l + 2 ^ m)"
+      by (simp only: take_bit_add)
+    finally show ?P
+      by (simp only: signed_take_bit_eq_take_bit m) simp
+  next
+    assume ?P
+    with that have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n"
+      by (simp add: signed_take_bit_eq_take_bit' take_bit_eq_mod)
+    then have "(i + (k + 2 ^ (n - Suc 0))) mod 2 ^ n = (i + (l + 2 ^ (n - Suc 0))) mod 2 ^ n" for i
+      by (metis mod_add_eq)
+    then have "k mod 2 ^ n = l mod 2 ^ n"
+      by (metis add_diff_cancel_right' uminus_add_conv_diff)
+    then show ?Q
+      by (simp add: take_bit_eq_mod)
+  qed
+qed
+
+
+subsection \<open>Bit strings as quotient type\<close>
+
+subsubsection \<open>Basic properties\<close>
+
+quotient_type (overloaded) 'a word = int / "\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len0) l"
+  by (auto intro!: equivpI reflpI sympI transpI)
+
+instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}"
+begin
+
+lift_definition zero_word :: "'a word"
+  is 0
+  .
+
+lift_definition one_word :: "'a word"
+  is 1
+  .
+
+lift_definition plus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is plus
+  by (subst take_bit_add [symmetric]) (simp add: take_bit_add)
+
+lift_definition uminus_word :: "'a word \<Rightarrow> 'a word"
+  is uminus
+  by (subst take_bit_uminus [symmetric]) (simp add: take_bit_uminus)
+
+lift_definition minus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is minus
+  by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus)
+
+lift_definition times_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is times
+  by (auto simp add: take_bit_eq_mod intro: mod_mult_cong)
+
+instance
+  by standard (transfer; simp add: algebra_simps)+
+
+end
+
+instance word :: (len) comm_ring_1
+  by standard (transfer; simp)+
+
+quickcheck_generator word
+  constructors:
+    "zero_class.zero :: ('a::len0) word",
+    "numeral :: num \<Rightarrow> ('a::len0) word",
+    "uminus :: ('a::len0) word \<Rightarrow> ('a::len0) word"
+
+context
+  includes lifting_syntax
+  notes power_transfer [transfer_rule]
+begin
+
+lemma power_transfer_word [transfer_rule]:
+  \<open>(pcr_word ===> (=) ===> pcr_word) (^) (^)\<close>
+  by transfer_prover
+
+end
+
+
+subsubsection \<open>Conversions\<close>
+
+context
+  includes lifting_syntax
+  notes transfer_rule_numeral [transfer_rule]
+    transfer_rule_of_nat [transfer_rule]
+    transfer_rule_of_int [transfer_rule]
+begin
+
+lemma [transfer_rule]:
+  "((=) ===> (pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool)) numeral numeral"
+  by transfer_prover
+
+lemma [transfer_rule]:
+  "((=) ===> pcr_word) int of_nat"
+  by transfer_prover
+
+lemma [transfer_rule]:
+  "((=) ===> pcr_word) (\<lambda>k. k) of_int"
+proof -
+  have "((=) ===> pcr_word) of_int of_int"
+    by transfer_prover
+  then show ?thesis by (simp add: id_def)
+qed
+
+end
+
+lemma abs_word_eq:
+  "abs_word = of_int"
+  by (rule ext) (transfer, rule)
+
+context semiring_1
+begin
+
+lift_definition unsigned :: "'b::len0 word \<Rightarrow> 'a"
+  is "of_nat \<circ> nat \<circ> take_bit LENGTH('b)"
+  by simp
+
+lemma unsigned_0 [simp]:
+  "unsigned 0 = 0"
+  by transfer simp
+
+end
+
+context semiring_char_0
+begin
+
+lemma word_eq_iff_unsigned:
+  "a = b \<longleftrightarrow> unsigned a = unsigned b"
+  by safe (transfer; simp add: eq_nat_nat_iff)
+
+end
+
+instantiation word :: (len0) equal
+begin
+
+definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
+  where "equal_word a b \<longleftrightarrow> (unsigned a :: int) = unsigned b"
+
+instance proof
+  fix a b :: "'a word"
+  show "HOL.equal a b \<longleftrightarrow> a = b"
+    using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def)
+qed
+
+end
+
+context ring_1
+begin
+
+lift_definition signed :: "'b::len word \<Rightarrow> 'a"
+  is "of_int \<circ> signed_take_bit (LENGTH('b) - 1)"
+  by (simp add: signed_take_bit_eq_iff_take_bit_eq [symmetric])
+
+lemma signed_0 [simp]:
+  "signed 0 = 0"
+  by transfer simp
+
+end
+
+lemma unsigned_of_nat [simp]:
+  "unsigned (of_nat n :: 'a word) = take_bit LENGTH('a::len) n"
+  by transfer (simp add: nat_eq_iff take_bit_eq_mod zmod_int)
+
+lemma of_nat_unsigned [simp]:
+  "of_nat (unsigned a) = a"
+  by transfer simp
+
+lemma of_int_unsigned [simp]:
+  "of_int (unsigned a) = a"
+  by transfer simp
+
+lemma unsigned_nat_less:
+  \<open>unsigned a < (2 ^ LENGTH('a) :: nat)\<close> for a :: \<open>'a::len0 word\<close>
+  by transfer (simp add: take_bit_eq_mod)
+
+lemma unsigned_int_less:
+  \<open>unsigned a < (2 ^ LENGTH('a) :: int)\<close> for a :: \<open>'a::len0 word\<close>
+  by transfer (simp add: take_bit_eq_mod)
+
+context ring_char_0
+begin
+
+lemma word_eq_iff_signed:
+  "a = b \<longleftrightarrow> signed a = signed b"
+  by safe (transfer; auto simp add: signed_take_bit_eq_iff_take_bit_eq)
+
+end
+
+lemma signed_of_int [simp]:
+  "signed (of_int k :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) k"
+  by transfer simp
+
+lemma of_int_signed [simp]:
+  "of_int (signed a) = a"
+  by transfer (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod mod_simps)
+
+
+subsubsection \<open>Properties\<close>
+
+lemma length_cases:
+  obtains (triv) "LENGTH('a::len) = 1" "take_bit LENGTH('a) 2 = (0 :: int)"
+    | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)"
+proof (cases "LENGTH('a) \<ge> 2")
+  case False
+  then have "LENGTH('a) = 1"
+    by (auto simp add: not_le dest: less_2_cases)
+  then have "take_bit LENGTH('a) 2 = (0 :: int)"
+    by simp
+  with \<open>LENGTH('a) = 1\<close> triv show ?thesis
+    by simp
+next
+  case True
+  then obtain n where "LENGTH('a) = Suc (Suc n)"
+    by (auto dest: le_Suc_ex)
+  then have "take_bit LENGTH('a) 2 = (2 :: int)"
+    by simp
+  with take_bit_2 show ?thesis
+    by simp
+qed
+
+
+subsubsection \<open>Division\<close>
+
+instantiation word :: (len0) modulo
+begin
+
+lift_definition divide_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is "\<lambda>a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b"
+  by simp
+
+lift_definition modulo_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
+  is "\<lambda>a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b"
+  by simp
+
+instance ..
+
+end
+
+lemma zero_word_div_eq [simp]:
+  \<open>0 div a = 0\<close> for a :: \<open>'a::len0 word\<close>
+  by transfer simp
+
+lemma div_zero_word_eq [simp]:
+  \<open>a div 0 = 0\<close> for a :: \<open>'a::len0 word\<close>
+  by transfer simp
+
+context
+  includes lifting_syntax
+begin
+
+lemma [transfer_rule]:
+  "(pcr_word ===> (\<longleftrightarrow>)) even ((dvd) 2 :: 'a::len word \<Rightarrow> bool)"
+proof -
+  have even_word_unfold: "even k \<longleftrightarrow> (\<exists>l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \<longleftrightarrow> ?Q")
+    for k :: int
+  proof
+    assume ?P
+    then show ?Q
+      by auto
+  next
+    assume ?Q
+    then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" ..
+    then have "even (take_bit LENGTH('a) k)"
+      by simp
+    then show ?P
+      by simp
+  qed
+  show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def])
+    transfer_prover
+qed
+
+end
+
+instance word :: (len) semiring_modulo
+proof
+  show "a div b * b + a mod b = a" for a b :: "'a word"
+  proof transfer
+    fix k l :: int
+    define r :: int where "r = 2 ^ LENGTH('a)"
+    then have r: "take_bit LENGTH('a) k = k mod r" for k
+      by (simp add: take_bit_eq_mod)
+    have "k mod r = ((k mod r) div (l mod r) * (l mod r)
+      + (k mod r) mod (l mod r)) mod r"
+      by (simp add: div_mult_mod_eq)
+    also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r
+      + (k mod r) mod (l mod r)) mod r"
+      by (simp add: mod_add_left_eq)
+    also have "... = (((k mod r) div (l mod r) * l) mod r
+      + (k mod r) mod (l mod r)) mod r"
+      by (simp add: mod_mult_right_eq)
+    finally have "k mod r = ((k mod r) div (l mod r) * l
+      + (k mod r) mod (l mod r)) mod r"
+      by (simp add: mod_simps)
+    with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l
+      + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k"
+      by simp
+  qed
+qed
+
+instance word :: (len) semiring_parity
+proof
+  show "\<not> 2 dvd (1::'a word)"
+    by transfer simp
+  show even_iff_mod_2_eq_0: "2 dvd a \<longleftrightarrow> a mod 2 = 0"
+    for a :: "'a word"
+    by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd)
+  show "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1"
+    for a :: "'a word"
+    by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd)
+qed
+
+
+subsubsection \<open>Orderings\<close>
+
+instantiation word :: (len0) linorder
+begin
+
+lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
+  is "\<lambda>a b. take_bit LENGTH('a) a \<le> take_bit LENGTH('a) b"
+  by simp
+
+lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
+  is "\<lambda>a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b"
+  by simp
+
+instance
+  by standard (transfer; auto)+
+
+end
+
+context linordered_semidom
+begin
+
+lemma word_less_eq_iff_unsigned:
+  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
+  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
+
+lemma word_less_iff_unsigned:
+  "a < b \<longleftrightarrow> unsigned a < unsigned b"
+  by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative])
+
+end
+
+lemma word_greater_zero_iff:
+  \<open>a > 0 \<longleftrightarrow> a \<noteq> 0\<close> for a :: \<open>'a::len0 word\<close>
+  by transfer (simp add: less_le)
+
+lemma of_nat_word_eq_iff:
+  \<open>of_nat m = (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m = take_bit LENGTH('a) n\<close>
+  by transfer (simp add: take_bit_of_nat)
+
+lemma of_nat_word_less_eq_iff:
+  \<open>of_nat m \<le> (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
+  by transfer (simp add: take_bit_of_nat)
+
+lemma of_nat_word_less_iff:
+  \<open>of_nat m < (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
+  by transfer (simp add: take_bit_of_nat)
+
+lemma of_nat_word_eq_0_iff:
+  \<open>of_nat n = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd n\<close>
+  using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff)
+
+lemma of_int_word_eq_iff:
+  \<open>of_int k = (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
+  by transfer rule
+
+lemma of_int_word_less_eq_iff:
+  \<open>of_int k \<le> (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
+  by transfer rule
+
+lemma of_int_word_less_iff:
+  \<open>of_int k < (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
+  by transfer rule
+
+lemma of_int_word_eq_0_iff:
+  \<open>of_int k = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd k\<close>
+  using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff)
+
+
+subsection \<open>Bit structure on \<^typ>\<open>'a word\<close>\<close>
+
+lemma word_bit_induct [case_names zero even odd]:
+  \<open>P a\<close> if word_zero: \<open>P 0\<close>
+    and word_even: \<open>\<And>a. P a \<Longrightarrow> 0 < a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (2 * a)\<close>
+    and word_odd: \<open>\<And>a. P a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (1 + 2 * a)\<close>
+  for P and a :: \<open>'a::len word\<close>
+proof -
+  define m :: nat where \<open>m = LENGTH('a) - 1\<close>
+  then have l: \<open>LENGTH('a) = Suc m\<close>
+    by simp
+  define n :: nat where \<open>n = unsigned a\<close>
+  then have \<open>n < 2 ^ LENGTH('a)\<close>
+    by (simp add: unsigned_nat_less)
+  then have \<open>n < 2 * 2 ^ m\<close>
+    by (simp add: l)
+  then have \<open>P (of_nat n)\<close>
+  proof (induction n rule: nat_bit_induct)
+    case zero
+    show ?case
+      by simp (rule word_zero)
+  next
+    case (even n)
+    then have \<open>n < 2 ^ m\<close>
+      by simp
+    with even.IH have \<open>P (of_nat n)\<close>
+      by simp
+    moreover from \<open>n < 2 ^ m\<close> even.hyps have \<open>0 < (of_nat n :: 'a word)\<close>
+      by (auto simp add: word_greater_zero_iff of_nat_word_eq_0_iff l)
+    moreover from \<open>n < 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close>
+      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
+      by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l)
+    ultimately have \<open>P (2 * of_nat n)\<close>
+      by (rule word_even)
+    then show ?case
+      by simp
+  next
+    case (odd n)
+    then have \<open>Suc n \<le> 2 ^ m\<close>
+      by simp
+    with odd.IH have \<open>P (of_nat n)\<close>
+      by simp
+    moreover from \<open>Suc n \<le> 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close>
+      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
+      by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l)
+    ultimately have \<open>P (1 + 2 * of_nat n)\<close>
+      by (rule word_odd)
+    then show ?case
+      by simp
+  qed
+  then show ?thesis
+    by (simp add: n_def)
+qed
+
+lemma bit_word_half_eq:
+  \<open>(of_bool b + a * 2) div 2 = a\<close>
+    if \<open>a < 2 ^ (LENGTH('a) - Suc 0)\<close>
+    for a :: \<open>'a::len word\<close>
+proof (cases rule: length_cases [where ?'a = 'a])
+  case triv
+  have \<open>of_bool (odd k) < (1 :: int) \<longleftrightarrow> even k\<close> for k :: int
+    by auto
+  with triv that show ?thesis
+    by (auto; transfer) simp_all
+next
+  case take_bit_2
+  obtain n where length: \<open>LENGTH('a) = Suc n\<close>
+    by (cases \<open>LENGTH('a)\<close>) simp_all
+  show ?thesis proof (cases b)
+    case False
+    moreover have \<open>a * 2 div 2 = a\<close>
+    using that proof transfer
+      fix k :: int
+      from length have \<open>k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\<close>
+        by simp
+      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
+      with \<open>LENGTH('a) = Suc n\<close>
+      have \<open>k mod 2 ^ LENGTH('a) = k mod 2 ^ n\<close>
+        by (simp add: take_bit_eq_mod divmod_digit_0)
+      ultimately have \<open>take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\<close>
+        by (simp add: take_bit_eq_mod)
+      with take_bit_2 show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2)
+        = take_bit LENGTH('a) k\<close>
+        by simp
+    qed
+    ultimately show ?thesis
+      by simp
+  next
+    case True
+    moreover have \<open>(1 + a * 2) div 2 = a\<close>
+    using that proof transfer
+      fix k :: int
+      from length have \<open>(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\<close>
+        using pos_zmod_mult_2 [of \<open>2 ^ n\<close> k] by (simp add: ac_simps)
+      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
+      with \<open>LENGTH('a) = Suc n\<close>
+      have \<open>k mod 2 ^ LENGTH('a) = k mod 2 ^ n\<close>
+        by (simp add: take_bit_eq_mod divmod_digit_0)
+      ultimately have \<open>take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\<close>
+        by (simp add: take_bit_eq_mod)
+      with take_bit_2 show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2)
+        = take_bit LENGTH('a) k\<close>
+        by simp
+    qed
+    ultimately show ?thesis
+      by simp
+  qed
+qed
+
+end
--- a/src/HOL/ex/Word_Type.thy	Tue Nov 05 22:55:50 2019 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,730 +0,0 @@
-(*  Author:  Florian Haftmann, TUM
-*)
-
-section \<open>Proof of concept for algebraically founded bit word types\<close>
-
-theory Word_Type
-  imports
-    Main
-    "HOL-ex.Bit_Lists"
-    "HOL-Library.Type_Length"
-begin
-
-subsection \<open>Preliminaries\<close>
-
-lemma take_bit_uminus:
-  "take_bit n (- (take_bit n k)) = take_bit n (- k)" for k :: int
-  by (simp add: take_bit_eq_mod mod_minus_eq)
-
-lemma take_bit_minus:
-  "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" for k l :: int
-  by (simp add: take_bit_eq_mod mod_diff_eq)
-
-lemma take_bit_nonnegative [simp]:
-  "take_bit n k \<ge> 0" for k :: int
-  by (simp add: take_bit_eq_mod)
-
-definition signed_take_bit :: "nat \<Rightarrow> int \<Rightarrow> int"
-  where signed_take_bit_eq_take_bit:
-    "signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n"
-
-lemma signed_take_bit_eq_take_bit':
-  "signed_take_bit (n - Suc 0) k = take_bit n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)" if "n > 0"
-  using that by (simp add: signed_take_bit_eq_take_bit)
-  
-lemma signed_take_bit_0 [simp]:
-  "signed_take_bit 0 k = - (k mod 2)"
-proof (cases "even k")
-  case True
-  then have "odd (k + 1)"
-    by simp
-  then have "(k + 1) mod 2 = 1"
-    by (simp add: even_iff_mod_2_eq_zero)
-  with True show ?thesis
-    by (simp add: signed_take_bit_eq_take_bit)
-next
-  case False
-  then show ?thesis
-    by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one)
-qed
-
-lemma signed_take_bit_Suc [simp]:
-  "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2"
-  by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps)
-
-lemma signed_take_bit_of_0 [simp]:
-  "signed_take_bit n 0 = 0"
-  by (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod)
-
-lemma signed_take_bit_of_minus_1 [simp]:
-  "signed_take_bit n (- 1) = - 1"
-  by (induct n) simp_all
-
-lemma signed_take_bit_eq_iff_take_bit_eq:
-  "signed_take_bit (n - Suc 0) k = signed_take_bit (n - Suc 0) l \<longleftrightarrow> take_bit n k = take_bit n l" (is "?P \<longleftrightarrow> ?Q")
-  if "n > 0"
-proof -
-  from that obtain m where m: "n = Suc m"
-    by (cases n) auto
-  show ?thesis
-  proof 
-    assume ?Q
-    have "take_bit (Suc m) (k + 2 ^ m) =
-      take_bit (Suc m) (take_bit (Suc m) k + take_bit (Suc m) (2 ^ m))"
-      by (simp only: take_bit_add)
-    also have "\<dots> =
-      take_bit (Suc m) (take_bit (Suc m) l + take_bit (Suc m) (2 ^ m))"
-      by (simp only: \<open>?Q\<close> m [symmetric])
-    also have "\<dots> = take_bit (Suc m) (l + 2 ^ m)"
-      by (simp only: take_bit_add)
-    finally show ?P
-      by (simp only: signed_take_bit_eq_take_bit m) simp
-  next
-    assume ?P
-    with that have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n"
-      by (simp add: signed_take_bit_eq_take_bit' take_bit_eq_mod)
-    then have "(i + (k + 2 ^ (n - Suc 0))) mod 2 ^ n = (i + (l + 2 ^ (n - Suc 0))) mod 2 ^ n" for i
-      by (metis mod_add_eq)
-    then have "k mod 2 ^ n = l mod 2 ^ n"
-      by (metis add_diff_cancel_right' uminus_add_conv_diff)
-    then show ?Q
-      by (simp add: take_bit_eq_mod)
-  qed
-qed 
-
-
-subsection \<open>Bit strings as quotient type\<close>
-
-subsubsection \<open>Basic properties\<close>
-
-quotient_type (overloaded) 'a word = int / "\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len0) l"
-  by (auto intro!: equivpI reflpI sympI transpI)
-
-instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}"
-begin
-
-lift_definition zero_word :: "'a word"
-  is 0
-  .
-
-lift_definition one_word :: "'a word"
-  is 1
-  .
-
-lift_definition plus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is plus
-  by (subst take_bit_add [symmetric]) (simp add: take_bit_add)
-
-lift_definition uminus_word :: "'a word \<Rightarrow> 'a word"
-  is uminus
-  by (subst take_bit_uminus [symmetric]) (simp add: take_bit_uminus)
-
-lift_definition minus_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is minus
-  by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus)
-
-lift_definition times_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is times
-  by (auto simp add: take_bit_eq_mod intro: mod_mult_cong)
-
-instance
-  by standard (transfer; simp add: algebra_simps)+
-
-end
-
-instance word :: (len) comm_ring_1
-  by standard (transfer; simp)+
-
-quickcheck_generator word
-  constructors:
-    "zero_class.zero :: ('a::len0) word",
-    "numeral :: num \<Rightarrow> ('a::len0) word",
-    "uminus :: ('a::len0) word \<Rightarrow> ('a::len0) word"
-
-context
-  includes lifting_syntax
-  notes power_transfer [transfer_rule]
-begin
-
-lemma power_transfer_word [transfer_rule]:
-  \<open>(pcr_word ===> (=) ===> pcr_word) (^) (^)\<close>
-  by transfer_prover
-
-end
-
-
-subsubsection \<open>Conversions\<close>
-
-context
-  includes lifting_syntax
-  notes transfer_rule_numeral [transfer_rule]
-    transfer_rule_of_nat [transfer_rule]
-    transfer_rule_of_int [transfer_rule]
-begin
-
-lemma [transfer_rule]:
-  "((=) ===> (pcr_word :: int \<Rightarrow> 'a::len word \<Rightarrow> bool)) numeral numeral"
-  by transfer_prover
-
-lemma [transfer_rule]:
-  "((=) ===> pcr_word) int of_nat"
-  by transfer_prover
-  
-lemma [transfer_rule]:
-  "((=) ===> pcr_word) (\<lambda>k. k) of_int"
-proof -
-  have "((=) ===> pcr_word) of_int of_int"
-    by transfer_prover
-  then show ?thesis by (simp add: id_def)
-qed
-
-end
-
-lemma abs_word_eq:
-  "abs_word = of_int"
-  by (rule ext) (transfer, rule)
-
-context semiring_1
-begin
-
-lift_definition unsigned :: "'b::len0 word \<Rightarrow> 'a"
-  is "of_nat \<circ> nat \<circ> take_bit LENGTH('b)"
-  by simp
-
-lemma unsigned_0 [simp]:
-  "unsigned 0 = 0"
-  by transfer simp
-
-end
-
-context semiring_char_0
-begin
-
-lemma word_eq_iff_unsigned:
-  "a = b \<longleftrightarrow> unsigned a = unsigned b"
-  by safe (transfer; simp add: eq_nat_nat_iff)
-
-end
-
-instantiation word :: (len0) equal
-begin
-
-definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
-  where "equal_word a b \<longleftrightarrow> (unsigned a :: int) = unsigned b"
-
-instance proof
-  fix a b :: "'a word"
-  show "HOL.equal a b \<longleftrightarrow> a = b"
-    using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def)
-qed
-
-end
-
-context ring_1
-begin
-
-lift_definition signed :: "'b::len word \<Rightarrow> 'a"
-  is "of_int \<circ> signed_take_bit (LENGTH('b) - 1)"
-  by (simp add: signed_take_bit_eq_iff_take_bit_eq [symmetric])
-
-lemma signed_0 [simp]:
-  "signed 0 = 0"
-  by transfer simp
-
-end
-
-lemma unsigned_of_nat [simp]:
-  "unsigned (of_nat n :: 'a word) = take_bit LENGTH('a::len) n"
-  by transfer (simp add: nat_eq_iff take_bit_eq_mod zmod_int)
-
-lemma of_nat_unsigned [simp]:
-  "of_nat (unsigned a) = a"
-  by transfer simp
-
-lemma of_int_unsigned [simp]:
-  "of_int (unsigned a) = a"
-  by transfer simp
-
-lemma unsigned_nat_less:
-  \<open>unsigned a < (2 ^ LENGTH('a) :: nat)\<close> for a :: \<open>'a::len0 word\<close>
-  by transfer (simp add: take_bit_eq_mod)
-
-lemma unsigned_int_less:
-  \<open>unsigned a < (2 ^ LENGTH('a) :: int)\<close> for a :: \<open>'a::len0 word\<close>
-  by transfer (simp add: take_bit_eq_mod)
-
-context ring_char_0
-begin
-
-lemma word_eq_iff_signed:
-  "a = b \<longleftrightarrow> signed a = signed b"
-  by safe (transfer; auto simp add: signed_take_bit_eq_iff_take_bit_eq)
-
-end
-
-lemma signed_of_int [simp]:
-  "signed (of_int k :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) k"
-  by transfer simp
-
-lemma of_int_signed [simp]:
-  "of_int (signed a) = a"
-  by transfer (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod mod_simps)
-
-
-subsubsection \<open>Properties\<close>
-
-lemma length_cases:
-  obtains (triv) "LENGTH('a::len) = 1" "take_bit LENGTH('a) 2 = (0 :: int)"
-    | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)"
-proof (cases "LENGTH('a) \<ge> 2")
-  case False
-  then have "LENGTH('a) = 1"
-    by (auto simp add: not_le dest: less_2_cases)
-  then have "take_bit LENGTH('a) 2 = (0 :: int)"
-    by simp
-  with \<open>LENGTH('a) = 1\<close> triv show ?thesis
-    by simp
-next
-  case True
-  then obtain n where "LENGTH('a) = Suc (Suc n)"
-    by (auto dest: le_Suc_ex)
-  then have "take_bit LENGTH('a) 2 = (2 :: int)"
-    by simp
-  with take_bit_2 show ?thesis
-    by simp
-qed
-
-
-subsubsection \<open>Division\<close>
-
-instantiation word :: (len0) modulo
-begin
-
-lift_definition divide_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is "\<lambda>a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b"
-  by simp
-
-lift_definition modulo_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is "\<lambda>a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b"
-  by simp
-
-instance ..
-
-end
-
-lemma zero_word_div_eq [simp]:
-  \<open>0 div a = 0\<close> for a :: \<open>'a::len0 word\<close>
-  by transfer simp
-
-lemma div_zero_word_eq [simp]:
-  \<open>a div 0 = 0\<close> for a :: \<open>'a::len0 word\<close>
-  by transfer simp
-
-context
-  includes lifting_syntax
-begin
-
-lemma [transfer_rule]:
-  "(pcr_word ===> (\<longleftrightarrow>)) even ((dvd) 2 :: 'a::len word \<Rightarrow> bool)"
-proof -
-  have even_word_unfold: "even k \<longleftrightarrow> (\<exists>l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \<longleftrightarrow> ?Q")
-    for k :: int
-  proof
-    assume ?P
-    then show ?Q
-      by auto
-  next
-    assume ?Q
-    then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" ..
-    then have "even (take_bit LENGTH('a) k)"
-      by simp
-    then show ?P
-      by simp
-  qed
-  show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def])
-    transfer_prover
-qed
-
-end
-
-instance word :: (len) semiring_modulo
-proof
-  show "a div b * b + a mod b = a" for a b :: "'a word"
-  proof transfer
-    fix k l :: int
-    define r :: int where "r = 2 ^ LENGTH('a)"
-    then have r: "take_bit LENGTH('a) k = k mod r" for k
-      by (simp add: take_bit_eq_mod)
-    have "k mod r = ((k mod r) div (l mod r) * (l mod r)
-      + (k mod r) mod (l mod r)) mod r"
-      by (simp add: div_mult_mod_eq)
-    also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r
-      + (k mod r) mod (l mod r)) mod r"
-      by (simp add: mod_add_left_eq)
-    also have "... = (((k mod r) div (l mod r) * l) mod r
-      + (k mod r) mod (l mod r)) mod r"
-      by (simp add: mod_mult_right_eq)
-    finally have "k mod r = ((k mod r) div (l mod r) * l
-      + (k mod r) mod (l mod r)) mod r"
-      by (simp add: mod_simps)
-    with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l
-      + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k"
-      by simp
-  qed
-qed
-
-instance word :: (len) semiring_parity
-proof
-  show "\<not> 2 dvd (1::'a word)"
-    by transfer simp
-  show even_iff_mod_2_eq_0: "2 dvd a \<longleftrightarrow> a mod 2 = 0"
-    for a :: "'a word"
-    by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd)
-  show "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1"
-    for a :: "'a word"
-    by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd)
-qed
-
-
-subsubsection \<open>Orderings\<close>
-
-instantiation word :: (len0) linorder
-begin
-
-lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
-  is "\<lambda>a b. take_bit LENGTH('a) a \<le> take_bit LENGTH('a) b"
-  by simp
-
-lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"
-  is "\<lambda>a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b"
-  by simp
-
-instance
-  by standard (transfer; auto)+
-
-end
-
-context linordered_semidom
-begin
-
-lemma word_less_eq_iff_unsigned:
-  "a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"
-  by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)
-
-lemma word_less_iff_unsigned:
-  "a < b \<longleftrightarrow> unsigned a < unsigned b"
-  by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative])
-
-end
-
-lemma word_greater_zero_iff:
-  \<open>a > 0 \<longleftrightarrow> a \<noteq> 0\<close> for a :: \<open>'a::len0 word\<close>
-  by transfer (simp add: less_le)
-
-lemma of_nat_word_eq_iff:
-  \<open>of_nat m = (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m = take_bit LENGTH('a) n\<close>
-  by transfer (simp add: take_bit_of_nat)
-
-lemma of_nat_word_less_eq_iff:
-  \<open>of_nat m \<le> (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
-  by transfer (simp add: take_bit_of_nat)
-
-lemma of_nat_word_less_iff:
-  \<open>of_nat m < (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
-  by transfer (simp add: take_bit_of_nat)
-
-lemma of_nat_word_eq_0_iff:
-  \<open>of_nat n = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd n\<close>
-  using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff)
-
-lemma of_int_word_eq_iff:
-  \<open>of_int k = (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
-  by transfer rule
-
-lemma of_int_word_less_eq_iff:
-  \<open>of_int k \<le> (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
-  by transfer rule
-
-lemma of_int_word_less_iff:
-  \<open>of_int k < (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
-  by transfer rule
-
-lemma of_int_word_eq_0_iff:
-  \<open>of_int k = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd k\<close>
-  using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff)
-
-
-subsection \<open>Bit operation on \<^typ>\<open>'a word\<close>\<close>
-
-context unique_euclidean_semiring_with_nat
-begin
-
-primrec n_bits_of :: "nat \<Rightarrow> 'a \<Rightarrow> bool list"
-  where
-    "n_bits_of 0 a = []"
-  | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)"
-
-lemma n_bits_of_eq_iff:
-  "n_bits_of n a = n_bits_of n b \<longleftrightarrow> take_bit n a = take_bit n b"
-  apply (induction n arbitrary: a b)
-   apply (auto elim!: evenE oddE)
-   apply (metis dvd_triv_right even_plus_one_iff)
-  apply (metis dvd_triv_right even_plus_one_iff)
-  done
-
-lemma take_n_bits_of [simp]:
-  "take m (n_bits_of n a) = n_bits_of (min m n) a"
-proof -
-  define q and v and w where "q = min m n" and "v = m - q" and "w = n - q"
-  then have "v = 0 \<or> w = 0"
-    by auto
-  then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a"
-    by (induction q arbitrary: a) auto
-  with q_def v_def w_def show ?thesis
-    by simp
-qed
-
-lemma unsigned_of_bits_n_bits_of [simp]:
-  "unsigned_of_bits (n_bits_of n a) = take_bit n a"
-  by (induction n arbitrary: a) (simp_all add: ac_simps)
-
-end
-
-lemma unsigned_of_bits_eq_of_bits:
-  "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)"
-  by (simp add: of_bits_int_def)
-
-
-instantiation word :: (len) bit_representation
-begin
-
-lift_definition bits_of_word :: "'a word \<Rightarrow> bool list"
-  is "n_bits_of LENGTH('a)"
-  by (simp add: n_bits_of_eq_iff)
-
-lift_definition of_bits_word :: "bool list \<Rightarrow> 'a word"
-  is unsigned_of_bits .
-
-instance proof
-  fix a :: "'a word"
-  show "of_bits (bits_of a) = a"
-    by transfer simp
-qed
-
-end
-
-lemma take_bit_complement_iff:
-  "take_bit n (complement k) = take_bit n (complement l) \<longleftrightarrow> take_bit n k = take_bit n l"
-  for k l :: int
-  by (simp add: take_bit_eq_mod mod_eq_dvd_iff dvd_diff_commute)
-
-lemma take_bit_not_iff:
-  "take_bit n (NOT k) = take_bit n (NOT l) \<longleftrightarrow> take_bit n k = take_bit n l"
-  for k l :: int
-  by (simp add: not_int_def take_bit_complement_iff)
-
-lemma n_bits_of_not:
-  "n_bits_of n (NOT k) = map Not (n_bits_of n k)"
-  for k :: int
-  by (induction n arbitrary: k) (simp_all add: not_div_2)
-
-lemma take_bit_and [simp]:
-  "take_bit n (k AND l) = take_bit n k AND take_bit n l"
-  for k l :: int
-  apply (induction n arbitrary: k l)
-   apply simp
-  apply (subst and_int.rec)
-  apply (subst (2) and_int.rec)
-  apply simp
-  done
-
-lemma take_bit_or [simp]:
-  "take_bit n (k OR l) = take_bit n k OR take_bit n l"
-  for k l :: int
-  apply (induction n arbitrary: k l)
-   apply simp
-  apply (subst or_int.rec)
-  apply (subst (2) or_int.rec)
-  apply simp
-  done
-
-lemma take_bit_xor [simp]:
-  "take_bit n (k XOR l) = take_bit n k XOR take_bit n l"
-  for k l :: int
-  apply (induction n arbitrary: k l)
-   apply simp
-  apply (subst xor_int.rec)
-  apply (subst (2) xor_int.rec)
-  apply simp
-  done
-
-instantiation word :: (len) bit_operations
-begin
-
-lift_definition not_word :: "'a word \<Rightarrow> 'a word"
-  is not
-  by (simp add: take_bit_not_iff)
-
-lift_definition and_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is "and"
-  by simp
-
-lift_definition or_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is or
-  by simp
-
-lift_definition xor_word ::  "'a word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-  is xor
-  by simp
-
-lift_definition shift_left_word :: "'a word \<Rightarrow> nat \<Rightarrow> 'a word"
-  is shift_left
-proof -
-  show "take_bit LENGTH('a) (k << n) = take_bit LENGTH('a) (l << n)"
-    if "take_bit LENGTH('a) k = take_bit LENGTH('a) l" for k l :: int and n :: nat
-  proof -
-    from that
-    have "take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k)
-      = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)"
-      by simp
-    moreover have "min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n"
-      by simp
-    ultimately show ?thesis by (simp add: take_bit_push_bit)
-  qed
-qed
-
-lift_definition shift_right_word :: "'a word \<Rightarrow> nat \<Rightarrow> 'a word"
-  is "\<lambda>k n. drop_bit n (take_bit LENGTH('a) k)"
-  by simp
-
-instance proof
-  show "semilattice ((AND) :: 'a word \<Rightarrow> _)"
-    by standard (transfer; simp add: ac_simps)+
-  show "semilattice ((OR) :: 'a word \<Rightarrow> _)"
-    by standard (transfer; simp add: ac_simps)+
-  show "abel_semigroup ((XOR) :: 'a word \<Rightarrow> _)"
-    by standard (transfer; simp add: ac_simps)+
-  show "not = (of_bits \<circ> map Not \<circ> bits_of :: 'a word \<Rightarrow> 'a word)"
-  proof
-    fix a :: "'a word"
-    have "NOT a = of_bits (map Not (bits_of a))"
-      by transfer (simp flip: unsigned_of_bits_take n_bits_of_not add: take_map)
-    then show "NOT a = (of_bits \<circ> map Not \<circ> bits_of) a"
-      by simp
-  qed
-  show "of_bits bs AND of_bits cs = (of_bits (map2 (\<and>) bs cs) :: 'a word)"
-    if "length bs = length cs" for bs cs
-    using that apply transfer
-    apply (simp only: unsigned_of_bits_eq_of_bits)
-    apply (subst and_eq)
-    apply simp_all
-    done
-  show "of_bits bs OR of_bits cs = (of_bits (map2 (\<or>) bs cs) :: 'a word)"
-    if "length bs = length cs" for bs cs
-    using that apply transfer
-    apply (simp only: unsigned_of_bits_eq_of_bits)
-    apply (subst or_eq)
-    apply simp_all
-    done
-  show "of_bits bs XOR of_bits cs = (of_bits (map2 (\<noteq>) bs cs) :: 'a word)"
-    if "length bs = length cs" for bs cs
-    using that apply transfer
-    apply (simp only: unsigned_of_bits_eq_of_bits)
-    apply (subst xor_eq)
-    apply simp_all
-    done
-  show "a << n = of_bits (replicate n False @ bits_of a)"
-    for a :: "'a word" and n :: nat
-    by transfer (simp add: push_bit_take_bit)
-  show "a >> n = of_bits (drop n (bits_of a))"
-    if "n < length (bits_of a)"
-    for a :: "'a word" and n :: nat
-    using that by transfer simp
-qed
-
-
-subsection \<open>Bit structure on \<^typ>\<open>'a word\<close>\<close>
-
-lemma word_bit_induct [case_names zero even odd]:
-  \<open>P a\<close> if word_zero: \<open>P 0\<close>
-    and word_even: \<open>\<And>a. P a \<Longrightarrow> 0 < a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (2 * a)\<close>
-    and word_odd: \<open>\<And>a. P a \<Longrightarrow> a < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> P (1 + 2 * a)\<close>
-  for P and a :: \<open>'a::len word\<close>
-proof -
-  define m :: nat where \<open>m = LENGTH('a) - 1\<close>
-  then have l: \<open>LENGTH('a) = Suc m\<close>
-    by simp
-  define n :: nat where \<open>n = unsigned a\<close>
-  then have \<open>n < 2 ^ LENGTH('a)\<close>
-    by (simp add: unsigned_nat_less)
-  then have \<open>n < 2 * 2 ^ m\<close>
-    by (simp add: l)
-  then have \<open>P (of_nat n)\<close>
-  proof (induction n rule: nat_bit_induct)
-    case zero
-    show ?case
-      by simp (rule word_zero)
-  next
-    case (even n)
-    then have \<open>n < 2 ^ m\<close>
-      by simp
-    with even.IH have \<open>P (of_nat n)\<close>
-      by simp
-    moreover from \<open>n < 2 ^ m\<close> even.hyps have \<open>0 < (of_nat n :: 'a word)\<close>
-      by (auto simp add: word_greater_zero_iff of_nat_word_eq_0_iff l)
-    moreover from \<open>n < 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close>
-      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
-      by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l)
-    ultimately have \<open>P (2 * of_nat n)\<close>
-      by (rule word_even)
-    then show ?case
-      by simp
-  next
-    case (odd n)
-    then have \<open>Suc n \<le> 2 ^ m\<close>
-      by simp
-    with odd.IH have \<open>P (of_nat n)\<close>
-      by simp
-    moreover from \<open>Suc n \<le> 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\<close>
-      using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]
-      by (cases \<open>m = 0\<close>) (simp_all add: not_less take_bit_eq_self ac_simps l)
-    ultimately have \<open>P (1 + 2 * of_nat n)\<close>
-      by (rule word_odd)
-    then show ?case
-      by simp
-  qed
-  then show ?thesis
-    by (simp add: n_def)
-qed
-
-end
-
-global_interpretation bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a::len word"
-  rewrites "bit_word.xor = ((XOR) :: 'a word \<Rightarrow> _)"
-proof -
-  interpret bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a word"
-  proof
-    show "a AND (b OR c) = a AND b OR a AND c"
-      for a b c :: "'a word"
-      by transfer (simp add: bit_int.conj_disj_distrib)
-    show "a OR b AND c = (a OR b) AND (a OR c)"
-      for a b c :: "'a word"
-      by transfer (simp add: bit_int.disj_conj_distrib)
-    show "a AND NOT a = 0" for a :: "'a word"
-      by transfer simp
-    show "a OR NOT a = - 1" for a :: "'a word"
-      by transfer simp
-  qed (transfer; simp)+
-  show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: 'a word)"
-    by (fact bit_word.boolean_algebra_axioms)
-  show "bit_word.xor = ((XOR) :: 'a word \<Rightarrow> _)"
-  proof (rule ext)+
-    fix a b :: "'a word"
-    have "a XOR b = a AND NOT b OR NOT a AND b"
-      by transfer (simp add: bit_int.xor_def)
-    then show "bit_word.xor a b = a XOR b"
-      by (simp add: bit_word.xor_def)
-  qed
-qed
-
-end