misc tuning and modernization;
authorwenzelm
Sun, 03 Dec 2017 22:28:19 +0100
changeset 67123 3fe40ff1b921
parent 67122 85b40f300fab
child 67124 335ed2834ebc
child 67125 361b3ef643a7
misc tuning and modernization;
src/HOL/Decision_Procs/Algebra_Aux.thy
src/HOL/Decision_Procs/Commutative_Ring.thy
src/HOL/Decision_Procs/Commutative_Ring_Complete.thy
src/HOL/Decision_Procs/Cooper.thy
src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
src/HOL/Decision_Procs/Rat_Pair.thy
src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy
src/HOL/Decision_Procs/Reflective_Field.thy
src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy
--- a/src/HOL/Decision_Procs/Algebra_Aux.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Algebra_Aux.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -5,16 +5,17 @@
 section \<open>Things that can be added to the Algebra library\<close>
 
 theory Algebra_Aux
-imports "HOL-Algebra.Ring"
+  imports "HOL-Algebra.Ring"
 begin
 
-definition of_natural :: "('a, 'm) ring_scheme \<Rightarrow> nat \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<^sub>\<nat>\<index>") where
-  "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> = (op \<oplus>\<^bsub>R\<^esub> \<one>\<^bsub>R\<^esub> ^^ n) \<zero>\<^bsub>R\<^esub>"
+definition of_natural :: "('a, 'm) ring_scheme \<Rightarrow> nat \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<^sub>\<nat>\<index>")
+  where "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> = (op \<oplus>\<^bsub>R\<^esub> \<one>\<^bsub>R\<^esub> ^^ n) \<zero>\<^bsub>R\<^esub>"
 
-definition of_integer :: "('a, 'm) ring_scheme \<Rightarrow> int \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<index>") where
-  "\<guillemotleft>i\<guillemotright>\<^bsub>R\<^esub> = (if 0 \<le> i then \<guillemotleft>nat i\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> else \<ominus>\<^bsub>R\<^esub> \<guillemotleft>nat (- i)\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub>)"
+definition of_integer :: "('a, 'm) ring_scheme \<Rightarrow> int \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<index>")
+  where "\<guillemotleft>i\<guillemotright>\<^bsub>R\<^esub> = (if 0 \<le> i then \<guillemotleft>nat i\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> else \<ominus>\<^bsub>R\<^esub> \<guillemotleft>nat (- i)\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub>)"
 
-context ring begin
+context ring
+begin
 
 lemma of_nat_0 [simp]: "\<guillemotleft>0\<guillemotright>\<^sub>\<nat> = \<zero>"
   by (simp add: of_natural_def)
@@ -39,10 +40,16 @@
 
 lemma of_nat_diff [simp]: "n \<le> m \<Longrightarrow> \<guillemotleft>m - n\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>"
 proof (induct m arbitrary: n)
-  case (Suc m)
-  note Suc' = this
+  case 0
+  then show ?case by (simp add: minus_eq)
+next
+  case Suc': (Suc m)
   show ?case
   proof (cases n)
+    case 0
+    then show ?thesis
+      by (simp add: minus_eq)
+  next
     case (Suc k)
     with Suc' have "\<guillemotleft>Suc m - Suc k\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>k\<guillemotright>\<^sub>\<nat>" by simp
     also have "\<dots> = \<one> \<oplus> \<ominus> \<one> \<oplus> (\<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>k\<guillemotright>\<^sub>\<nat>)"
@@ -50,8 +57,8 @@
     also have "\<dots> = \<guillemotleft>Suc m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>Suc k\<guillemotright>\<^sub>\<nat>"
       by (simp add: minus_eq minus_add a_ac)
     finally show ?thesis using Suc by simp
-  qed (simp add: minus_eq)
-qed (simp add: minus_eq)
+  qed
+qed
 
 lemma of_int_add [simp]: "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<oplus> \<guillemotleft>j\<guillemotright>"
 proof (cases "0 \<le> i")
@@ -59,7 +66,8 @@
   show ?thesis
   proof (cases "0 \<le> j")
     case True
-    with \<open>0 \<le> i\<close> show ?thesis by (simp add: of_integer_def nat_add_distrib)
+    with \<open>0 \<le> i\<close> show ?thesis
+      by (simp add: of_integer_def nat_add_distrib)
   next
     case False
     show ?thesis
@@ -112,7 +120,7 @@
     case False
     with \<open>\<not> 0 \<le> i\<close> show ?thesis
       by (simp add: of_integer_def nat_add_distrib minus_add diff_conv_add_uminus
-        del: add_uminus_conv_diff uminus_add_conv_diff)
+          del: add_uminus_conv_diff uminus_add_conv_diff)
   qed
 qed
 
@@ -171,7 +179,7 @@
 
 lemma eq_diff0:
   assumes "x \<in> carrier R" "y \<in> carrier R"
-  shows "(x \<ominus> y = \<zero>) = (x = y)"
+  shows "x \<ominus> y = \<zero> \<longleftrightarrow> x = y"
 proof
   assume "x \<ominus> y = \<zero>"
   with assms have "x \<oplus> (\<ominus> y \<oplus> y) = y"
@@ -187,7 +195,7 @@
 
 lemma eq_neg_iff_add_eq_0:
   assumes "x \<in> carrier R" "y \<in> carrier R"
-  shows "(x = \<ominus> y) = (x \<oplus> y = \<zero>)"
+  shows "x = \<ominus> y \<longleftrightarrow> x \<oplus> y = \<zero>"
 proof
   assume "x = \<ominus> y"
   with assms show "x \<oplus> y = \<zero>" by (simp add: l_neg)
@@ -201,7 +209,7 @@
 
 lemma neg_equal_iff_equal:
   assumes x: "x \<in> carrier R" and y: "y \<in> carrier R"
-  shows "(\<ominus> x = \<ominus> y) = (x = y)"
+  shows "\<ominus> x = \<ominus> y \<longleftrightarrow> x = y"
 proof
   assume "\<ominus> x = \<ominus> y"
   then have "\<ominus> (\<ominus> x) = \<ominus> (\<ominus> y)" by simp
@@ -225,8 +233,8 @@
 lemma (in cring) of_int_power [simp]: "\<guillemotleft>i ^ n\<guillemotright> = \<guillemotleft>i\<guillemotright> (^) n"
   by (induct n) (simp_all add: m_ac)
 
-definition cring_class_ops :: "'a::comm_ring_1 ring" where
-  "cring_class_ops \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>"
+definition cring_class_ops :: "'a::comm_ring_1 ring"
+  where "cring_class_ops \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>"
 
 lemma cring_class: "cring cring_class_ops"
   apply unfold_locales
@@ -274,17 +282,16 @@
   times_class power_class of_nat_class of_int_class carrier_class
 
 interpretation cring_class: cring "cring_class_ops::'a::comm_ring_1 ring"
-  rewrites
-    "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0" and
-    "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1" and
-    "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y" and
-    "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y" and
-    "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x" and
-    "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y" and
-    "(x::'a) (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n" and
-    "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n" and
-    "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i" and
-    "(Int.of_int (numeral m)::'a) = numeral m"
+  rewrites "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0"
+    and "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1"
+    and "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y"
+    and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
+    and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x"
+    and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
+    and "(x::'a) (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n"
+    and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
+    and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i"
+    and "(Int.of_int (numeral m)::'a) = numeral m"
   by (simp_all add: cring_class class_simps)
 
 lemma (in domain) nat_pow_eq_0_iff [simp]:
@@ -302,12 +309,12 @@
     by (simp add: integral_iff eq_neg_iff_add_eq_0 eq_diff0 r_neg)
 next
   assume "x = y \<or> x = \<ominus> y"
-  with assms show "x \<otimes> x = y \<otimes> y" by (auto simp add: l_minus r_minus)
+  with assms show "x \<otimes> x = y \<otimes> y"
+    by (auto simp add: l_minus r_minus)
 qed
 
-definition
-  m_div :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oslash>\<index>" 70) where
-  "x \<oslash>\<^bsub>G\<^esub> y = (if y = \<zero>\<^bsub>G\<^esub> then \<zero>\<^bsub>G\<^esub> else x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)"
+definition m_div :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oslash>\<index>" 70)
+  where "x \<oslash>\<^bsub>G\<^esub> y = (if y = \<zero>\<^bsub>G\<^esub> then \<zero>\<^bsub>G\<^esub> else x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)"
 
 context field
 begin
@@ -323,58 +330,62 @@
 
 lemma inverse_unique:
   assumes a: "a \<in> carrier R"
-  and b: "b \<in> carrier R"
-  and ab: "a \<otimes> b = \<one>"
+    and b: "b \<in> carrier R"
+    and ab: "a \<otimes> b = \<one>"
   shows "inv a = b"
 proof -
-  have "a \<noteq> \<zero>" using ab b by (cases "a = \<zero>") simp_all
-  moreover with a have "inv a \<otimes> (a \<otimes> b) = inv a" by (simp add: ab)
-  ultimately show ?thesis using a b by (simp add: m_assoc [symmetric])
+  from ab b have *: "a \<noteq> \<zero>"
+    by (cases "a = \<zero>") simp_all
+  with a have "inv a \<otimes> (a \<otimes> b) = inv a"
+    by (simp add: ab)
+  with a b * show ?thesis
+    by (simp add: m_assoc [symmetric])
 qed
 
-lemma nonzero_inverse_inverse_eq:
-  "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> inv (inv a) = a"
+lemma nonzero_inverse_inverse_eq: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> inv (inv a) = a"
   by (rule inverse_unique) simp_all
 
 lemma inv_1 [simp]: "inv \<one> = \<one>"
   by (rule inverse_unique) simp_all
 
 lemma nonzero_inverse_mult_distrib:
-  assumes "a \<in> carrier R" and "b \<in> carrier R" and "a \<noteq> \<zero>" and "b \<noteq> \<zero>"
+  assumes "a \<in> carrier R" "b \<in> carrier R"
+    and "a \<noteq> \<zero>" "b \<noteq> \<zero>"
   shows "inv (a \<otimes> b) = inv b \<otimes> inv a"
 proof -
-  have "a \<otimes> (b \<otimes> inv b) \<otimes> inv a = \<one>" using assms by simp
-  hence eq: "a \<otimes> b \<otimes> (inv b \<otimes> inv a) = \<one>" using assms
+  from assms have "a \<otimes> (b \<otimes> inv b) \<otimes> inv a = \<one>"
+    by simp
+  with assms have eq: "a \<otimes> b \<otimes> (inv b \<otimes> inv a) = \<one>"
     by (simp only: m_assoc m_closed inv_closed assms)
-  from inverse_unique [OF _ _ eq] assms
-  show ?thesis by simp
+  from assms show ?thesis
+    using inverse_unique [OF _ _ eq] by simp
 qed
 
 lemma nonzero_imp_inverse_nonzero:
   assumes "a \<in> carrier R" and "a \<noteq> \<zero>"
   shows "inv a \<noteq> \<zero>"
 proof
-  assume ianz: "inv a = \<zero>"
-  from assms
-  have "\<one> = a \<otimes> inv a" by simp
-  also with assms have "... = \<zero>" by (simp add: ianz)
+  assume *: "inv a = \<zero>"
+  from assms have **: "\<one> = a \<otimes> inv a"
+    by simp
+  also from assms have "\<dots> = \<zero>" by (simp add: *)
   finally have "\<one> = \<zero>" .
-  thus False by (simp add: eq_commute)
+  then show False by (simp add: eq_commute)
 qed
 
 lemma nonzero_divide_divide_eq_left:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow>
-   a \<oslash> b \<oslash> c = a \<oslash> (b \<otimes> c)"
+    a \<oslash> b \<oslash> c = a \<oslash> (b \<otimes> c)"
   by (simp add: m_div_def nonzero_inverse_mult_distrib m_ac integral_iff)
 
 lemma nonzero_times_divide_eq:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> d \<in> carrier R \<Longrightarrow>
-   b \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<otimes> (c \<oslash> d) = (a \<otimes> c) \<oslash> (b \<otimes> d)"
+    b \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<otimes> (c \<oslash> d) = (a \<otimes> c) \<oslash> (b \<otimes> d)"
   by (simp add: m_div_def nonzero_inverse_mult_distrib m_ac integral_iff)
 
 lemma nonzero_divide_divide_eq:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> d \<in> carrier R \<Longrightarrow>
-   b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<oslash> (c \<oslash> d) = (a \<otimes> d) \<oslash> (b \<otimes> c)"
+    b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<oslash> (c \<oslash> d) = (a \<otimes> d) \<oslash> (b \<otimes> c)"
   by (simp add: m_div_def nonzero_inverse_mult_distrib
     nonzero_imp_inverse_nonzero nonzero_inverse_inverse_eq m_ac integral_iff)
 
@@ -382,8 +393,8 @@
   by (simp add: m_div_def)
 
 lemma add_frac_eq:
-  assumes "x \<in> carrier R" and "y \<in> carrier R" and "z \<in> carrier R" and "w \<in> carrier R"
-  and "y \<noteq> \<zero>" and "z \<noteq> \<zero>"
+  assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" "w \<in> carrier R"
+    and "y \<noteq> \<zero>" "z \<noteq> \<zero>"
   shows "x \<oslash> y \<oplus> w \<oslash> z = (x \<otimes> z \<oplus> w \<otimes> y) \<oslash> (y \<otimes> z)"
 proof -
   from assms
@@ -403,57 +414,56 @@
   by (simp add: m_div_def l_minus)
 
 lemma diff_frac_eq:
-  assumes "x \<in> carrier R" and "y \<in> carrier R" and "z \<in> carrier R" and "w \<in> carrier R"
-  and "y \<noteq> \<zero>" and "z \<noteq> \<zero>"
+  assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" "w \<in> carrier R"
+    and "y \<noteq> \<zero>" "z \<noteq> \<zero>"
   shows "x \<oslash> y \<ominus> w \<oslash> z = (x \<otimes> z \<ominus> w \<otimes> y) \<oslash> (y \<otimes> z)"
-  using assms
-  by (simp add: minus_eq l_minus add_frac_eq)
+  using assms by (simp add: minus_eq l_minus add_frac_eq)
 
 lemma nonzero_mult_divide_mult_cancel_left [simp]:
-  assumes "a \<in> carrier R" and "b \<in> carrier R" and "c \<in> carrier R"
-  and "b \<noteq> \<zero>" and "c \<noteq> \<zero>"
+  assumes "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
+    and "b \<noteq> \<zero>" "c \<noteq> \<zero>"
   shows "(c \<otimes> a) \<oslash> (c \<otimes> b) = a \<oslash> b"
 proof -
   from assms have "(c \<otimes> a) \<oslash> (c \<otimes> b) = c \<otimes> a \<otimes> (inv b \<otimes> inv c)"
     by (simp add: m_div_def nonzero_inverse_mult_distrib integral_iff)
   also from assms have "\<dots> =  a \<otimes> inv b \<otimes> (inv c \<otimes> c)"
     by (simp add: m_ac)
-  also from assms have "\<dots> =  a \<otimes> inv b" by simp
-  finally show ?thesis using assms by (simp add: m_div_def)
+  also from assms have "\<dots> =  a \<otimes> inv b"
+    by simp
+  finally show ?thesis
+    using assms by (simp add: m_div_def)
 qed
 
 lemma times_divide_eq_left [simp]:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow>
-   (b \<oslash> c) \<otimes> a = b \<otimes> a \<oslash> c"
+    (b \<oslash> c) \<otimes> a = b \<otimes> a \<oslash> c"
   by (simp add: m_div_def m_ac)
 
 lemma times_divide_eq_right [simp]:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow>
-   a \<otimes> (b \<oslash> c) = a \<otimes> b \<oslash> c"
+    a \<otimes> (b \<oslash> c) = a \<otimes> b \<oslash> c"
   by (simp add: m_div_def m_ac)
 
 lemma nonzero_power_divide:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow>
-   (a \<oslash> b) (^) (n::nat) = a (^) n \<oslash> b (^) n"
+    (a \<oslash> b) (^) (n::nat) = a (^) n \<oslash> b (^) n"
   by (induct n) (simp_all add: nonzero_divide_divide_eq_left)
 
 lemma r_diff_distr:
   "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> z \<in> carrier R \<Longrightarrow>
-   z \<otimes> (x \<ominus> y) = z \<otimes> x \<ominus> z \<otimes> y"
+    z \<otimes> (x \<ominus> y) = z \<otimes> x \<ominus> z \<otimes> y"
   by (simp add: minus_eq r_distr r_minus)
 
-lemma divide_zero_left [simp]:
-  "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> \<zero> \<oslash> a = \<zero>"
+lemma divide_zero_left [simp]: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> \<zero> \<oslash> a = \<zero>"
   by (simp add: m_div_def)
 
 lemma divide_self: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> a \<oslash> a = \<one>"
   by (simp add: m_div_def)
 
 lemma divide_eq_0_iff:
-  assumes "a \<in> carrier R"
-  and "b \<in> carrier R"
-  and "b \<noteq> \<zero>"
-  shows "(a \<oslash> b = \<zero>) = (a = \<zero>)"
+  assumes "a \<in> carrier R" "b \<in> carrier R"
+    and "b \<noteq> \<zero>"
+  shows "a \<oslash> b = \<zero> \<longleftrightarrow> a = \<zero>"
 proof
   assume "a = \<zero>"
   with assms show "a \<oslash> b = \<zero>" by simp
@@ -471,7 +481,7 @@
 
 lemma field_class: "field (cring_class_ops::'a::field ring)"
   apply unfold_locales
-  apply (simp_all add: cring_class_ops_def)
+    apply (simp_all add: cring_class_ops_def)
   apply (auto simp add: Units_def)
   apply (rule_tac x="1 / x" in exI)
   apply simp
@@ -481,25 +491,24 @@
   apply (simp add: m_div_def m_inv_def class_simps)
   apply (rule impI)
   apply (rule ssubst [OF the_equality, of _ "1 / y"])
-  apply simp_all
+    apply simp_all
   apply (drule conjunct2)
   apply (drule_tac f="\<lambda>x. x / y" in arg_cong)
   apply simp
   done
 
 interpretation field_class: field "cring_class_ops::'a::field ring"
-  rewrites
-    "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0" and
-    "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1" and
-    "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y" and
-    "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y" and
-    "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x" and
-    "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y" and
-    "(x::'a) (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n" and
-    "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n" and
-    "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i" and
-    "(x::'a) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y" and
-    "(Int.of_int (numeral m)::'a) = numeral m"
+  rewrites "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0"
+    and "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1"
+    and "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y"
+    and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
+    and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x"
+    and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
+    and "(x::'a) (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n"
+    and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
+    and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i"
+    and "(x::'a) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y"
+    and "(Int.of_int (numeral m)::'a) = numeral m"
   by (simp_all add: field_class class_simps div_class)
 
 end
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -7,10 +7,10 @@
 section \<open>Proving equalities in commutative rings\<close>
 
 theory Commutative_Ring
-imports
-  Conversions
-  Algebra_Aux
-  "HOL-Library.Code_Target_Numeral"
+  imports
+    Conversions
+    Algebra_Aux
+    "HOL-Library.Code_Target_Numeral"
 begin
 
 text \<open>Syntax of multivariate polynomials (pol) and polynomial expressions.\<close>
@@ -31,10 +31,11 @@
 
 text \<open>Interpretation functions for the shadow syntax.\<close>
 
-context cring begin
+context cring
+begin
 
-definition in_carrier :: "'a list \<Rightarrow> bool" where
-  "in_carrier xs = (\<forall>x\<in>set xs. x \<in> carrier R)"
+definition in_carrier :: "'a list \<Rightarrow> bool"
+  where "in_carrier xs \<longleftrightarrow> (\<forall>x\<in>set xs. x \<in> carrier R)"
 
 lemma in_carrier_Nil: "in_carrier []"
   by (simp add: in_carrier_def)
@@ -43,11 +44,10 @@
   by (simp add: in_carrier_def)
 
 lemma drop_in_carrier [simp]: "in_carrier xs \<Longrightarrow> in_carrier (drop n xs)"
-  using set_drop_subset [of n xs]
-  by (auto simp add: in_carrier_def)
+  using set_drop_subset [of n xs] by (auto simp add: in_carrier_def)
 
 primrec head :: "'a list \<Rightarrow> 'a"
-where
+  where
     "head [] = \<zero>"
   | "head (x # xs) = x"
 
@@ -55,7 +55,7 @@
   by (cases xs) (simp_all add: in_carrier_def)
 
 primrec Ipol :: "'a list \<Rightarrow> pol \<Rightarrow> 'a"
-where
+  where
     "Ipol l (Pc c) = \<guillemotleft>c\<guillemotright>"
   | "Ipol l (Pinj i P) = Ipol (drop i l) P"
   | "Ipol l (PX P x Q) = Ipol l P \<otimes> head l (^) x \<oplus> Ipol (drop 1 l) Q"
@@ -67,12 +67,11 @@
   "Ipol l (Pc (- numeral n)) = \<ominus> \<guillemotleft>numeral n\<guillemotright>"
   by simp_all
 
-lemma Ipol_closed [simp]:
-  "in_carrier l \<Longrightarrow> Ipol l p \<in> carrier R"
+lemma Ipol_closed [simp]: "in_carrier l \<Longrightarrow> Ipol l p \<in> carrier R"
   by (induct p arbitrary: l) simp_all
 
 primrec Ipolex :: "'a list \<Rightarrow> polex \<Rightarrow> 'a"
-where
+  where
     "Ipolex l (Var n) = head (drop n l)"
   | "Ipolex l (Const i) = \<guillemotleft>i\<guillemotright>"
   | "Ipolex l (Add P Q) = Ipolex l P \<oplus> Ipolex l Q"
@@ -92,16 +91,14 @@
 text \<open>Create polynomial normalized polynomials given normalized inputs.\<close>
 
 definition mkPinj :: "nat \<Rightarrow> pol \<Rightarrow> pol"
-where
-  "mkPinj x P =
+  where "mkPinj x P =
     (case P of
       Pc c \<Rightarrow> Pc c
     | Pinj y P \<Rightarrow> Pinj (x + y) P
     | PX p1 y p2 \<Rightarrow> Pinj x P)"
 
 definition mkPX :: "pol \<Rightarrow> nat \<Rightarrow> pol \<Rightarrow> pol"
-where
-  "mkPX P i Q =
+  where "mkPX P i Q =
     (case P of
       Pc c \<Rightarrow> if c = 0 then mkPinj 1 Q else PX P i Q
     | Pinj j R \<Rightarrow> PX P i Q
@@ -110,7 +107,7 @@
 text \<open>Defining the basic ring operations on normalized polynomials\<close>
 
 function add :: "pol \<Rightarrow> pol \<Rightarrow> pol"  (infixl "\<langle>+\<rangle>" 65)
-where
+  where
     "Pc a \<langle>+\<rangle> Pc b = Pc (a + b)"
   | "Pc c \<langle>+\<rangle> Pinj i P = Pinj i (P \<langle>+\<rangle> Pc c)"
   | "Pinj i P \<langle>+\<rangle> Pc c = Pinj i (P \<langle>+\<rangle> Pc c)"
@@ -132,11 +129,11 @@
       (if x = y then mkPX (P1 \<langle>+\<rangle> Q1) x (P2 \<langle>+\<rangle> Q2)
        else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<langle>+\<rangle> Q1) y (P2 \<langle>+\<rangle> Q2)
          else mkPX (PX Q1 (y - x) (Pc 0) \<langle>+\<rangle> P1) x (P2 \<langle>+\<rangle> Q2)))"
-by pat_completeness auto
+  by pat_completeness auto
 termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto
 
 function mul :: "pol \<Rightarrow> pol \<Rightarrow> pol"  (infixl "\<langle>*\<rangle>" 70)
-where
+  where
     "Pc a \<langle>*\<rangle> Pc b = Pc (a * b)"
   | "Pc c \<langle>*\<rangle> Pinj i P =
       (if c = 0 then Pc 0 else mkPinj i (P \<langle>*\<rangle> Pc c))"
@@ -165,35 +162,32 @@
       mkPX (P1 \<langle>*\<rangle> Q1) (x + y) (P2 \<langle>*\<rangle> Q2) \<langle>+\<rangle>
         (mkPX (P1 \<langle>*\<rangle> mkPinj 1 Q2) x (Pc 0) \<langle>+\<rangle>
           (mkPX (Q1 \<langle>*\<rangle> mkPinj 1 P2) y (Pc 0)))"
-by pat_completeness auto
+  by pat_completeness auto
 termination by (relation "measure (\<lambda>(x, y). size x + size y)")
   (auto simp add: mkPinj_def split: pol.split)
 
 text \<open>Negation\<close>
 primrec neg :: "pol \<Rightarrow> pol"
-where
+  where
     "neg (Pc c) = Pc (- c)"
   | "neg (Pinj i P) = Pinj i (neg P)"
   | "neg (PX P x Q) = PX (neg P) x (neg Q)"
 
 text \<open>Subtraction\<close>
 definition sub :: "pol \<Rightarrow> pol \<Rightarrow> pol"  (infixl "\<langle>-\<rangle>" 65)
-where
-  "sub P Q = P \<langle>+\<rangle> neg Q"
+  where "sub P Q = P \<langle>+\<rangle> neg Q"
 
 text \<open>Square for Fast Exponentiation\<close>
 primrec sqr :: "pol \<Rightarrow> pol"
-where
+  where
     "sqr (Pc c) = Pc (c * c)"
   | "sqr (Pinj i P) = mkPinj i (sqr P)"
-  | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \<langle>+\<rangle>
-      mkPX (Pc 2 \<langle>*\<rangle> A \<langle>*\<rangle> mkPinj 1 B) x (Pc 0)"
+  | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \<langle>+\<rangle> mkPX (Pc 2 \<langle>*\<rangle> A \<langle>*\<rangle> mkPinj 1 B) x (Pc 0)"
 
 text \<open>Fast Exponentiation\<close>
 
 fun pow :: "nat \<Rightarrow> pol \<Rightarrow> pol"
-where
-  pow_if [simp del]: "pow n P =
+  where pow_if [simp del]: "pow n P =
    (if n = 0 then Pc 1
     else if even n then pow (n div 2) (sqr P)
     else P \<langle>*\<rangle> pow (n div 2) (sqr P))"
@@ -214,7 +208,7 @@
 text \<open>Normalization of polynomial expressions\<close>
 
 primrec norm :: "polex \<Rightarrow> pol"
-where
+  where
     "norm (Var n) =
        (if n = 0 then PX (Pc 1) 1 (Pc 0)
         else Pinj n (PX (Pc 1) 1 (Pc 0)))"
@@ -360,57 +354,52 @@
   | Minj nat mon
   | MX nat mon
 
-primrec (in cring)
-  Imon :: "'a list \<Rightarrow> mon \<Rightarrow> 'a"
-where
+primrec (in cring) Imon :: "'a list \<Rightarrow> mon \<Rightarrow> 'a"
+  where
     "Imon l (Mc c) = \<guillemotleft>c\<guillemotright>"
   | "Imon l (Minj i M) = Imon (drop i l) M"
   | "Imon l (MX x M) = Imon (drop 1 l) M \<otimes> head l (^) x"
 
-lemma (in cring) Imon_closed [simp]:
-  "in_carrier l \<Longrightarrow> Imon l m \<in> carrier R"
+lemma (in cring) Imon_closed [simp]: "in_carrier l \<Longrightarrow> Imon l m \<in> carrier R"
   by (induct m arbitrary: l) simp_all
 
-definition
-  mkMinj :: "nat \<Rightarrow> mon \<Rightarrow> mon" where
-  "mkMinj i M = (case M of
-       Mc c \<Rightarrow> Mc c
-     | Minj j M \<Rightarrow> Minj (i + j) M
-     | _ \<Rightarrow> Minj i M)"
+definition mkMinj :: "nat \<Rightarrow> mon \<Rightarrow> mon"
+  where "mkMinj i M =
+    (case M of
+      Mc c \<Rightarrow> Mc c
+    | Minj j M \<Rightarrow> Minj (i + j) M
+    | _ \<Rightarrow> Minj i M)"
 
-definition
-  Minj_pred :: "nat \<Rightarrow> mon \<Rightarrow> mon" where
-  "Minj_pred i M = (if i = 1 then M else mkMinj (i - 1) M)"
+definition Minj_pred :: "nat \<Rightarrow> mon \<Rightarrow> mon"
+  where "Minj_pred i M = (if i = 1 then M else mkMinj (i - 1) M)"
 
 primrec mkMX :: "nat \<Rightarrow> mon \<Rightarrow> mon"
-where
-  "mkMX i (Mc c) = MX i (Mc c)"
-| "mkMX i (Minj j M) = (if j = 0 then mkMX i M else MX i (Minj_pred j M))"
-| "mkMX i (MX j M) = MX (i + j) M"
+  where
+    "mkMX i (Mc c) = MX i (Mc c)"
+  | "mkMX i (Minj j M) = (if j = 0 then mkMX i M else MX i (Minj_pred j M))"
+  | "mkMX i (MX j M) = MX (i + j) M"
 
-lemma (in cring) mkMinj_correct:
-  "Imon l (mkMinj i M) = Imon l (Minj i M)"
+lemma (in cring) mkMinj_correct: "Imon l (mkMinj i M) = Imon l (Minj i M)"
   by (simp add: mkMinj_def add.commute split: mon.split)
 
-lemma (in cring) Minj_pred_correct:
-  "0 < i \<Longrightarrow> Imon (drop 1 l) (Minj_pred i M) = Imon l (Minj i M)"
+lemma (in cring) Minj_pred_correct: "0 < i \<Longrightarrow> Imon (drop 1 l) (Minj_pred i M) = Imon l (Minj i M)"
   by (simp add: Minj_pred_def mkMinj_correct)
 
-lemma (in cring) mkMX_correct:
-  "in_carrier l \<Longrightarrow> Imon l (mkMX i M) = Imon l M \<otimes> head l (^) i"
-  by (induct M) (simp_all add: Minj_pred_correct [simplified] nat_pow_mult [symmetric] m_ac split: mon.split)
+lemma (in cring) mkMX_correct: "in_carrier l \<Longrightarrow> Imon l (mkMX i M) = Imon l M \<otimes> head l (^) i"
+  by (induct M)
+    (simp_all add: Minj_pred_correct [simplified] nat_pow_mult [symmetric] m_ac split: mon.split)
 
 fun cfactor :: "pol \<Rightarrow> int \<Rightarrow> pol \<times> pol"
-where
-  "cfactor (Pc c') c = (Pc (c' mod c), Pc (c' div c))"
-| "cfactor (Pinj i P) c =
-     (let (R, S) = cfactor P c
-      in (mkPinj i R, mkPinj i S))"
-| "cfactor (PX P i Q) c =
-     (let
-        (R1, S1) = cfactor P c;
-        (R2, S2) = cfactor Q c
-      in (mkPX R1 i R2, mkPX S1 i S2))"
+  where
+    "cfactor (Pc c') c = (Pc (c' mod c), Pc (c' div c))"
+  | "cfactor (Pinj i P) c =
+       (let (R, S) = cfactor P c
+        in (mkPinj i R, mkPinj i S))"
+  | "cfactor (PX P i Q) c =
+       (let
+          (R1, S1) = cfactor P c;
+          (R2, S2) = cfactor Q c
+        in (mkPX R1 i R2, mkPX S1 i S2))"
 
 lemma (in cring) cfactor_correct:
   "in_carrier l \<Longrightarrow> Ipol l P = Ipol l (fst (cfactor P c)) \<oplus> \<guillemotleft>c\<guillemotright> \<otimes> Ipol l (snd (cfactor P c))"
@@ -430,35 +419,35 @@
 qed
 
 fun mfactor :: "pol \<Rightarrow> mon \<Rightarrow> pol \<times> pol"
-where
-  "mfactor P (Mc c) = (if c = 1 then (Pc 0, P) else cfactor P c)"
-| "mfactor (Pc d) M = (Pc d, Pc 0)"
-| "mfactor (Pinj i P) (Minj j M) =
-     (if i = j then
-        let (R, S) = mfactor P M
-        in (mkPinj i R, mkPinj i S)
-      else if i < j then
-        let (R, S) = mfactor P (Minj (j - i) M)
-        in (mkPinj i R, mkPinj i S)
-      else (Pinj i P, Pc 0))"
-| "mfactor (Pinj i P) (MX j M) = (Pinj i P, Pc 0)"
-| "mfactor (PX P i Q) (Minj j M) =
-     (if j = 0 then mfactor (PX P i Q) M
-      else
-        let
-          (R1, S1) = mfactor P (Minj j M);
-          (R2, S2) = mfactor Q (Minj_pred j M)
-        in (mkPX R1 i R2, mkPX S1 i S2))"
-| "mfactor (PX P i Q) (MX j M) =
-     (if i = j then
-        let (R, S) = mfactor P (mkMinj 1 M)
-        in (mkPX R i Q, S)
-      else if i < j then
-        let (R, S) = mfactor P (MX (j - i) M)
-        in (mkPX R i Q, S)
-      else
-        let (R, S) = mfactor P (mkMinj 1 M)
-        in (mkPX R i Q, mkPX S (i - j) (Pc 0)))"
+  where
+    "mfactor P (Mc c) = (if c = 1 then (Pc 0, P) else cfactor P c)"
+  | "mfactor (Pc d) M = (Pc d, Pc 0)"
+  | "mfactor (Pinj i P) (Minj j M) =
+       (if i = j then
+          let (R, S) = mfactor P M
+          in (mkPinj i R, mkPinj i S)
+        else if i < j then
+          let (R, S) = mfactor P (Minj (j - i) M)
+          in (mkPinj i R, mkPinj i S)
+        else (Pinj i P, Pc 0))"
+  | "mfactor (Pinj i P) (MX j M) = (Pinj i P, Pc 0)"
+  | "mfactor (PX P i Q) (Minj j M) =
+       (if j = 0 then mfactor (PX P i Q) M
+        else
+          let
+            (R1, S1) = mfactor P (Minj j M);
+            (R2, S2) = mfactor Q (Minj_pred j M)
+          in (mkPX R1 i R2, mkPX S1 i S2))"
+  | "mfactor (PX P i Q) (MX j M) =
+       (if i = j then
+          let (R, S) = mfactor P (mkMinj 1 M)
+          in (mkPX R i Q, S)
+        else if i < j then
+          let (R, S) = mfactor P (MX (j - i) M)
+          in (mkPX R i Q, S)
+        else
+          let (R, S) = mfactor P (mkMinj 1 M)
+          in (mkPX R i Q, mkPX S (i - j) (Pc 0)))"
 
 lemmas mfactor_induct = mfactor.induct
   [case_names Mc Pc_Minj Pc_MX Pinj_Minj Pinj_MX PX_Minj PX_MX]
@@ -515,20 +504,20 @@
 qed
 
 primrec mon_of_pol :: "pol \<Rightarrow> mon option"
-where
-  "mon_of_pol (Pc c) = Some (Mc c)"
-| "mon_of_pol (Pinj i P) = (case mon_of_pol P of
-       None \<Rightarrow> None
-     | Some M \<Rightarrow> Some (mkMinj i M))"
-| "mon_of_pol (PX P i Q) =
-     (if Q = Pc 0 then (case mon_of_pol P of
-          None \<Rightarrow> None
-        | Some M \<Rightarrow> Some (mkMX i M))
-      else None)"
+  where
+    "mon_of_pol (Pc c) = Some (Mc c)"
+  | "mon_of_pol (Pinj i P) = (case mon_of_pol P of
+         None \<Rightarrow> None
+       | Some M \<Rightarrow> Some (mkMinj i M))"
+  | "mon_of_pol (PX P i Q) =
+       (if Q = Pc 0 then (case mon_of_pol P of
+            None \<Rightarrow> None
+          | Some M \<Rightarrow> Some (mkMX i M))
+        else None)"
 
 lemma (in cring) mon_of_pol_correct:
   assumes "in_carrier l"
-  and "mon_of_pol P = Some M"
+    and "mon_of_pol P = Some M"
   shows "Ipol l P = Imon l M"
   using assms
 proof (induct P arbitrary: M l)
@@ -539,81 +528,85 @@
 qed (auto simp add: mkMinj_correct split: option.split_asm)
 
 fun (in cring) Ipolex_polex_list :: "'a list \<Rightarrow> (polex \<times> polex) list \<Rightarrow> bool"
-where
-  "Ipolex_polex_list l [] = True"
-| "Ipolex_polex_list l ((P, Q) # pps) = ((Ipolex l P = Ipolex l Q) \<and> Ipolex_polex_list l pps)"
+  where
+    "Ipolex_polex_list l [] = True"
+  | "Ipolex_polex_list l ((P, Q) # pps) = ((Ipolex l P = Ipolex l Q) \<and> Ipolex_polex_list l pps)"
 
 fun (in cring) Imon_pol_list :: "'a list \<Rightarrow> (mon \<times> pol) list \<Rightarrow> bool"
-where
-  "Imon_pol_list l [] = True"
-| "Imon_pol_list l ((M, P) # mps) = ((Imon l M = Ipol l P) \<and> Imon_pol_list l mps)"
+  where
+    "Imon_pol_list l [] = True"
+  | "Imon_pol_list l ((M, P) # mps) = ((Imon l M = Ipol l P) \<and> Imon_pol_list l mps)"
 
 fun mk_monpol_list :: "(polex \<times> polex) list \<Rightarrow> (mon \<times> pol) list"
-where
-  "mk_monpol_list [] = []"
-| "mk_monpol_list ((P, Q) # pps) =
-     (case mon_of_pol (norm P) of
-        None \<Rightarrow> mk_monpol_list pps
-      | Some M \<Rightarrow> (M, norm Q) # mk_monpol_list pps)"
+  where
+    "mk_monpol_list [] = []"
+  | "mk_monpol_list ((P, Q) # pps) =
+       (case mon_of_pol (norm P) of
+          None \<Rightarrow> mk_monpol_list pps
+        | Some M \<Rightarrow> (M, norm Q) # mk_monpol_list pps)"
 
 lemma (in cring) mk_monpol_list_correct:
   "in_carrier l \<Longrightarrow> Ipolex_polex_list l pps \<Longrightarrow> Imon_pol_list l (mk_monpol_list pps)"
   by (induct pps rule: mk_monpol_list.induct)
-    (auto split: option.split
-       simp add: norm_ci [symmetric] mon_of_pol_correct [symmetric])
+    (auto split: option.split simp add: norm_ci [symmetric] mon_of_pol_correct [symmetric])
 
-definition ponesubst :: "pol \<Rightarrow> mon \<Rightarrow> pol \<Rightarrow> pol option" where
-  "ponesubst P1 M P2 =
-     (let (Q, R) = mfactor P1 M
-      in case R of
-          Pc c \<Rightarrow> if c = 0 then None else Some (add Q (mul P2 R))
-        | _ \<Rightarrow> Some (add Q (mul P2 R)))"
+definition ponesubst :: "pol \<Rightarrow> mon \<Rightarrow> pol \<Rightarrow> pol option"
+  where "ponesubst P1 M P2 =
+   (let (Q, R) = mfactor P1 M in
+    (case R of
+      Pc c \<Rightarrow> if c = 0 then None else Some (add Q (mul P2 R))
+    | _ \<Rightarrow> Some (add Q (mul P2 R))))"
 
 fun pnsubst1 :: "pol \<Rightarrow> mon \<Rightarrow> pol \<Rightarrow> nat \<Rightarrow> pol"
-where
-  "pnsubst1 P1 M P2 n = (case ponesubst P1 M P2 of
-       None \<Rightarrow> P1
-     | Some P3 \<Rightarrow> if n = 0 then P3 else pnsubst1 P3 M P2 (n - 1))"
+  where "pnsubst1 P1 M P2 n =
+    (case ponesubst P1 M P2 of
+      None \<Rightarrow> P1
+    | Some P3 \<Rightarrow> if n = 0 then P3 else pnsubst1 P3 M P2 (n - 1))"
 
 lemma pnsubst1_0 [simp]: "pnsubst1 P1 M P2 0 = (case ponesubst P1 M P2 of
   None \<Rightarrow> P1 | Some P3 \<Rightarrow> P3)"
   by (simp split: option.split)
 
-lemma pnsubst1_Suc [simp]: "pnsubst1 P1 M P2 (Suc n) = (case ponesubst P1 M P2 of
-  None \<Rightarrow> P1 | Some P3 \<Rightarrow> pnsubst1 P3 M P2 n)"
+lemma pnsubst1_Suc [simp]:
+  "pnsubst1 P1 M P2 (Suc n) =
+    (case ponesubst P1 M P2 of
+      None \<Rightarrow> P1
+    | Some P3 \<Rightarrow> pnsubst1 P3 M P2 n)"
   by (simp split: option.split)
 
 declare pnsubst1.simps [simp del]
 
-definition pnsubst :: "pol \<Rightarrow> mon \<Rightarrow> pol \<Rightarrow> nat \<Rightarrow> pol option" where
-  "pnsubst P1 M P2 n = (case ponesubst P1 M P2 of
-       None \<Rightarrow> None
-     | Some P3 \<Rightarrow> Some (pnsubst1 P3 M P2 n))"
+definition pnsubst :: "pol \<Rightarrow> mon \<Rightarrow> pol \<Rightarrow> nat \<Rightarrow> pol option"
+  where "pnsubst P1 M P2 n =
+    (case ponesubst P1 M P2 of
+      None \<Rightarrow> None
+    | Some P3 \<Rightarrow> Some (pnsubst1 P3 M P2 n))"
 
 fun psubstl1 :: "pol \<Rightarrow> (mon \<times> pol) list \<Rightarrow> nat \<Rightarrow> pol"
-where
-  "psubstl1 P1 [] n = P1"
-| "psubstl1 P1 ((M, P2) # mps) n = psubstl1 (pnsubst1 P1 M P2 n) mps n"
+  where
+    "psubstl1 P1 [] n = P1"
+  | "psubstl1 P1 ((M, P2) # mps) n = psubstl1 (pnsubst1 P1 M P2 n) mps n"
 
 fun psubstl :: "pol \<Rightarrow> (mon \<times> pol) list \<Rightarrow> nat \<Rightarrow> pol option"
-where
-  "psubstl P1 [] n = None"
-| "psubstl P1 ((M, P2) # mps) n = (case pnsubst P1 M P2 n of
-       None \<Rightarrow> psubstl P1 mps n
-     | Some P3 \<Rightarrow> Some (psubstl1 P3 mps n))"
+  where
+    "psubstl P1 [] n = None"
+  | "psubstl P1 ((M, P2) # mps) n =
+      (case pnsubst P1 M P2 n of
+        None \<Rightarrow> psubstl P1 mps n
+      | Some P3 \<Rightarrow> Some (psubstl1 P3 mps n))"
 
 fun pnsubstl :: "pol \<Rightarrow> (mon \<times> pol) list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> pol"
-where
-  "pnsubstl P1 mps m n = (case psubstl P1 mps n of
-       None \<Rightarrow> P1
-     | Some P3 \<Rightarrow> if m = 0 then P3 else pnsubstl P3 mps (m - 1) n)"
+  where "pnsubstl P1 mps m n =
+    (case psubstl P1 mps n of
+      None \<Rightarrow> P1
+    | Some P3 \<Rightarrow> if m = 0 then P3 else pnsubstl P3 mps (m - 1) n)"
 
-lemma pnsubstl_0 [simp]: "pnsubstl P1 mps 0 n = (case psubstl P1 mps n of
-  None \<Rightarrow> P1 | Some P3 \<Rightarrow> P3)"
+lemma pnsubstl_0 [simp]:
+  "pnsubstl P1 mps 0 n = (case psubstl P1 mps n of None \<Rightarrow> P1 | Some P3 \<Rightarrow> P3)"
   by (simp split: option.split)
 
-lemma pnsubstl_Suc [simp]: "pnsubstl P1 mps (Suc m) n = (case psubstl P1 mps n of
-  None \<Rightarrow> P1 | Some P3 \<Rightarrow> pnsubstl P3 mps m n)"
+lemma pnsubstl_Suc [simp]:
+  "pnsubstl P1 mps (Suc m) n = (case psubstl P1 mps n of None \<Rightarrow> P1 | Some P3 \<Rightarrow> pnsubstl P3 mps m n)"
   by (simp split: option.split)
 
 declare pnsubstl.simps [simp del]
@@ -961,7 +954,8 @@
 end
 \<close>
 
-context cring begin
+context cring
+begin
 
 local_setup \<open>
 Local_Theory.declaration {syntax = false, pervasive = false}
--- a/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -8,23 +8,25 @@
 section \<open>Proof of the relative completeness of method comm-ring\<close>
 
 theory Commutative_Ring_Complete
-imports Commutative_Ring
+  imports Commutative_Ring
 begin
 
 text \<open>Formalization of normal form\<close>
 fun isnorm :: "pol \<Rightarrow> bool"
-where
-  "isnorm (Pc c) \<longleftrightarrow> True"
-| "isnorm (Pinj i (Pc c)) \<longleftrightarrow> False"
-| "isnorm (Pinj i (Pinj j Q)) \<longleftrightarrow> False"
-| "isnorm (Pinj 0 P) \<longleftrightarrow> False"
-| "isnorm (Pinj i (PX Q1 j Q2)) \<longleftrightarrow> isnorm (PX Q1 j Q2)"
-| "isnorm (PX P 0 Q) \<longleftrightarrow> False"
-| "isnorm (PX (Pc c) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm Q"
-| "isnorm (PX (PX P1 j (Pc c)) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm (PX P1 j (Pc c)) \<and> isnorm Q"
-| "isnorm (PX P i Q) \<longleftrightarrow> isnorm P \<and> isnorm Q"
+  where
+    "isnorm (Pc c) \<longleftrightarrow> True"
+  | "isnorm (Pinj i (Pc c)) \<longleftrightarrow> False"
+  | "isnorm (Pinj i (Pinj j Q)) \<longleftrightarrow> False"
+  | "isnorm (Pinj 0 P) \<longleftrightarrow> False"
+  | "isnorm (Pinj i (PX Q1 j Q2)) \<longleftrightarrow> isnorm (PX Q1 j Q2)"
+  | "isnorm (PX P 0 Q) \<longleftrightarrow> False"
+  | "isnorm (PX (Pc c) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm Q"
+  | "isnorm (PX (PX P1 j (Pc c)) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm (PX P1 j (Pc c)) \<and> isnorm Q"
+  | "isnorm (PX P i Q) \<longleftrightarrow> isnorm P \<and> isnorm Q"
 
-(* Some helpful lemmas *)
+
+subsection \<open>Some helpful lemmas\<close>
+
 lemma norm_Pinj_0_False: "isnorm (Pinj 0 P) = False"
   by (cases P) auto
 
@@ -36,24 +38,24 @@
 
 lemma norm_PX2: "isnorm (PX P i Q) \<Longrightarrow> isnorm Q"
   apply (cases i)
-  apply auto
+   apply auto
   apply (cases P)
-  apply auto
+    apply auto
   subgoal for \<dots> pol2 by (cases pol2) auto
   done
 
 lemma norm_PX1: "isnorm (PX P i Q) \<Longrightarrow> isnorm P"
   apply (cases i)
-  apply auto
+   apply auto
   apply (cases P)
-  apply auto
+    apply auto
   subgoal for \<dots> pol2 by (cases pol2) auto
   done
 
 lemma mkPinj_cn: "y \<noteq> 0 \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (mkPinj y Q)"
   apply (auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split)
-  apply (rename_tac nat a, case_tac nat, auto simp add: norm_Pinj_0_False)
-  apply (rename_tac pol a, case_tac pol, auto)
+   apply (rename_tac nat a, case_tac nat, auto simp add: norm_Pinj_0_False)
+   apply (rename_tac pol a, case_tac pol, auto)
   apply (case_tac y, auto)
   done
 
@@ -65,9 +67,9 @@
   case (PX p1 y p2)
   with assms show ?thesis
     apply (cases x)
-    apply auto
+     apply auto
     apply (cases p2)
-    apply auto
+      apply auto
     done
 next
   case Pc
@@ -87,9 +89,9 @@
   case (PX p1 y p2)
   with assms show ?thesis
     apply (cases x)
-    apply auto
+     apply auto
     apply (cases p2)
-    apply auto
+      apply auto
     done
 next
   case Pc
@@ -101,7 +103,7 @@
     by (cases x) auto
 qed
 
-text \<open>mkPX conserves normalizedness (\<open>_cn\<close>)\<close>
+text \<open>\<open>mkPX\<close> conserves normalizedness (\<open>_cn\<close>)\<close>
 lemma mkPX_cn:
   assumes "x \<noteq> 0"
     and "isnorm P"
@@ -123,13 +125,13 @@
     by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2])
   from assms PX Y0 show ?thesis
     apply (cases x)
-    apply (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _])
+     apply (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _])
     apply (cases P2)
-    apply auto
+      apply auto
     done
 qed
 
-text \<open>add conserves normalizedness\<close>
+text \<open>\<open>add\<close> conserves normalizedness\<close>
 lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>+\<rangle> Q)"
 proof (induct P Q rule: add.induct)
   case 1
@@ -138,17 +140,17 @@
   case (2 c i P2)
   then show ?case
     apply (cases P2)
-    apply simp_all
+      apply simp_all
     apply (cases i)
-    apply simp_all
+     apply simp_all
     done
 next
   case (3 i P2 c)
   then show ?case
     apply (cases P2)
-    apply simp_all
+      apply simp_all
     apply (cases i)
-    apply simp_all
+     apply simp_all
     done
 next
   case (4 c P2 i Q2)
@@ -156,9 +158,9 @@
     by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
   with 4 show ?case
     apply (cases i)
-    apply simp
+     apply simp
     apply (cases P2)
-    apply auto
+      apply auto
     subgoal for \<dots> pol2 by (cases pol2) auto
     done
 next
@@ -167,9 +169,9 @@
     by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
   with 5 show ?case
     apply (cases i)
-    apply simp
+     apply simp
     apply (cases P2)
-    apply auto
+      apply auto
     subgoal for \<dots> pol2 by (cases pol2) auto
     done
 next
@@ -326,7 +328,7 @@
   qed
 qed
 
-text \<open>mul concerves normalizedness\<close>
+text \<open>\<open>mul\<close> concerves normalizedness\<close>
 lemma mul_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>*\<rangle> Q)"
 proof (induct P Q rule: mul.induct)
   case 1
@@ -335,17 +337,17 @@
   case (2 c i P2)
   then show ?case
     apply (cases P2)
-    apply simp_all
+      apply simp_all
     apply (cases i)
-    apply (simp_all add: mkPinj_cn)
+     apply (simp_all add: mkPinj_cn)
     done
 next
   case (3 i P2 c)
   then show ?case
     apply (cases P2)
-    apply simp_all
+      apply simp_all
     apply (cases i)
-    apply (simp_all add: mkPinj_cn)
+     apply (simp_all add: mkPinj_cn)
     done
 next
   case (4 c P2 i Q2)
@@ -353,9 +355,9 @@
     by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
   with 4 show ?case
     apply (cases "c = 0")
-    apply simp_all
+     apply simp_all
     apply (cases "i = 0")
-    apply (simp_all add: mkPX_cn)
+     apply (simp_all add: mkPX_cn)
     done
 next
   case (5 P2 i Q2 c)
@@ -363,9 +365,9 @@
     by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
   with 5 show ?case
     apply (cases "c = 0")
-    apply simp_all
+     apply simp_all
     apply (cases "i = 0")
-    apply (simp_all add: mkPX_cn)
+     apply (simp_all add: mkPX_cn)
     done
 next
   case (6 x P2 y Q2)
@@ -381,9 +383,9 @@
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
     from 6 xy y have "isnorm (Pinj d Q2)"
       apply (cases d)
-      apply simp
+       apply simp
       apply (cases Q2)
-      apply auto
+        apply auto
       done
     with 6 * ** y show ?thesis
       by (simp add: mkPinj_cn)
@@ -531,17 +533,14 @@
   case (Pinj i Q)
   then show ?case
     apply (cases Q)
-    apply (auto simp add: mkPX_cn mkPinj_cn)
+      apply (auto simp add: mkPX_cn mkPinj_cn)
     apply (cases i)
-    apply (auto simp add: mkPX_cn mkPinj_cn)
+     apply (auto simp add: mkPX_cn mkPinj_cn)
     done
 next
   case (PX P1 x P2)
   then have "x + x \<noteq> 0" "isnorm P2" "isnorm P1"
-    apply (cases x)
-    using PX
-    apply (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2])
-    done
+    by (cases x) (use PX in \<open>auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]\<close>)
   with PX have "isnorm (mkPX (Pc (1 + 1) \<langle>*\<rangle> P1 \<langle>*\<rangle> mkPinj (Suc 0) P2) x (Pc 0))"
     and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))"
     by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
@@ -549,7 +548,7 @@
     by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
 qed
 
-text \<open>pow conserves normalizedness\<close>
+text \<open>\<open>pow\<close> conserves normalizedness\<close>
 lemma pow_cn: "isnorm P \<Longrightarrow> isnorm (pow n P)"
 proof (induct n arbitrary: P rule: less_induct)
   case (less k)
--- a/src/HOL/Decision_Procs/Cooper.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -8,11 +8,9 @@
   "HOL-Library.Code_Target_Numeral"
 begin
 
-(* Periodicity of dvd *)
+section \<open>Periodicity of \<open>dvd\<close>\<close>
 
-(*********************************************************************************)
-(****                            SHADOW SYNTAX AND SEMANTICS                  ****)
-(*********************************************************************************)
+subsection \<open>Shadow syntax and semantics\<close>
 
 datatype (plugins del: size) num = C int | Bound nat | CN nat int num
   | Neg num | Add num num | Sub num num
@@ -22,28 +20,28 @@
 begin
 
 primrec size_num :: "num \<Rightarrow> nat"
-where
-  "size_num (C c) = 1"
-| "size_num (Bound n) = 1"
-| "size_num (Neg a) = 1 + size_num a"
-| "size_num (Add a b) = 1 + size_num a + size_num b"
-| "size_num (Sub a b) = 3 + size_num a + size_num b"
-| "size_num (CN n c a) = 4 + size_num a"
-| "size_num (Mul c a) = 1 + size_num a"
+  where
+    "size_num (C c) = 1"
+  | "size_num (Bound n) = 1"
+  | "size_num (Neg a) = 1 + size_num a"
+  | "size_num (Add a b) = 1 + size_num a + size_num b"
+  | "size_num (Sub a b) = 3 + size_num a + size_num b"
+  | "size_num (CN n c a) = 4 + size_num a"
+  | "size_num (Mul c a) = 1 + size_num a"
 
 instance ..
 
 end
 
 primrec Inum :: "int list \<Rightarrow> num \<Rightarrow> int"
-where
-  "Inum bs (C c) = c"
-| "Inum bs (Bound n) = bs ! n"
-| "Inum bs (CN n c a) = c * (bs ! n) + Inum bs a"
-| "Inum bs (Neg a) = - Inum bs a"
-| "Inum bs (Add a b) = Inum bs a + Inum bs b"
-| "Inum bs (Sub a b) = Inum bs a - Inum bs b"
-| "Inum bs (Mul c a) = c * Inum bs a"
+  where
+    "Inum bs (C c) = c"
+  | "Inum bs (Bound n) = bs ! n"
+  | "Inum bs (CN n c a) = c * (bs ! n) + Inum bs a"
+  | "Inum bs (Neg a) = - Inum bs a"
+  | "Inum bs (Add a b) = Inum bs a + Inum bs b"
+  | "Inum bs (Sub a b) = Inum bs a - Inum bs b"
+  | "Inum bs (Mul c a) = c * Inum bs a"
 
 datatype (plugins del: size) fm = T | F
   | Lt num | Le num | Gt num | Ge num | Eq num | NEq num
@@ -55,180 +53,180 @@
 begin
 
 primrec size_fm :: "fm \<Rightarrow> nat"
-where
-  "size_fm (NOT p) = 1 + size_fm p"
-| "size_fm (And p q) = 1 + size_fm p + size_fm q"
-| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
-| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
-| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
-| "size_fm (E p) = 1 + size_fm p"
-| "size_fm (A p) = 4 + size_fm p"
-| "size_fm (Dvd i t) = 2"
-| "size_fm (NDvd i t) = 2"
-| "size_fm T = 1" 
-| "size_fm F = 1"
-| "size_fm (Lt _) = 1" 
-| "size_fm (Le _) = 1" 
-| "size_fm (Gt _) = 1" 
-| "size_fm (Ge _) = 1" 
-| "size_fm (Eq _) = 1" 
-| "size_fm (NEq _) = 1" 
-| "size_fm (Closed _) = 1" 
-| "size_fm (NClosed _) = 1"
+  where
+    "size_fm (NOT p) = 1 + size_fm p"
+  | "size_fm (And p q) = 1 + size_fm p + size_fm q"
+  | "size_fm (Or p q) = 1 + size_fm p + size_fm q"
+  | "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
+  | "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
+  | "size_fm (E p) = 1 + size_fm p"
+  | "size_fm (A p) = 4 + size_fm p"
+  | "size_fm (Dvd i t) = 2"
+  | "size_fm (NDvd i t) = 2"
+  | "size_fm T = 1"
+  | "size_fm F = 1"
+  | "size_fm (Lt _) = 1"
+  | "size_fm (Le _) = 1"
+  | "size_fm (Gt _) = 1"
+  | "size_fm (Ge _) = 1"
+  | "size_fm (Eq _) = 1"
+  | "size_fm (NEq _) = 1"
+  | "size_fm (Closed _) = 1"
+  | "size_fm (NClosed _) = 1"
 
 instance ..
 
 end
 
-lemma fmsize_pos [simp]: "size p > 0" for p :: fm
+lemma fmsize_pos [simp]: "size p > 0"
+  for p :: fm
   by (induct p) simp_all
 
-primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool"  \<comment> \<open>Semantics of formulae (fm)\<close>
-where
-  "Ifm bbs bs T \<longleftrightarrow> True"
-| "Ifm bbs bs F \<longleftrightarrow> False"
-| "Ifm bbs bs (Lt a) \<longleftrightarrow> Inum bs a < 0"
-| "Ifm bbs bs (Gt a) \<longleftrightarrow> Inum bs a > 0"
-| "Ifm bbs bs (Le a) \<longleftrightarrow> Inum bs a \<le> 0"
-| "Ifm bbs bs (Ge a) \<longleftrightarrow> Inum bs a \<ge> 0"
-| "Ifm bbs bs (Eq a) \<longleftrightarrow> Inum bs a = 0"
-| "Ifm bbs bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"
-| "Ifm bbs bs (Dvd i b) \<longleftrightarrow> i dvd Inum bs b"
-| "Ifm bbs bs (NDvd i b) \<longleftrightarrow> \<not> i dvd Inum bs b"
-| "Ifm bbs bs (NOT p) \<longleftrightarrow> \<not> Ifm bbs bs p"
-| "Ifm bbs bs (And p q) \<longleftrightarrow> Ifm bbs bs p \<and> Ifm bbs bs q"
-| "Ifm bbs bs (Or p q) \<longleftrightarrow> Ifm bbs bs p \<or> Ifm bbs bs q"
-| "Ifm bbs bs (Imp p q) \<longleftrightarrow> (Ifm bbs bs p \<longrightarrow> Ifm bbs bs q)"
-| "Ifm bbs bs (Iff p q) \<longleftrightarrow> Ifm bbs bs p = Ifm bbs bs q"
-| "Ifm bbs bs (E p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) p)"
-| "Ifm bbs bs (A p) \<longleftrightarrow> (\<forall>x. Ifm bbs (x # bs) p)"
-| "Ifm bbs bs (Closed n) \<longleftrightarrow> bbs ! n"
-| "Ifm bbs bs (NClosed n) \<longleftrightarrow> \<not> bbs ! n"
+primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool"  \<comment> \<open>Semantics of formulae (\<open>fm\<close>)\<close>
+  where
+    "Ifm bbs bs T \<longleftrightarrow> True"
+  | "Ifm bbs bs F \<longleftrightarrow> False"
+  | "Ifm bbs bs (Lt a) \<longleftrightarrow> Inum bs a < 0"
+  | "Ifm bbs bs (Gt a) \<longleftrightarrow> Inum bs a > 0"
+  | "Ifm bbs bs (Le a) \<longleftrightarrow> Inum bs a \<le> 0"
+  | "Ifm bbs bs (Ge a) \<longleftrightarrow> Inum bs a \<ge> 0"
+  | "Ifm bbs bs (Eq a) \<longleftrightarrow> Inum bs a = 0"
+  | "Ifm bbs bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"
+  | "Ifm bbs bs (Dvd i b) \<longleftrightarrow> i dvd Inum bs b"
+  | "Ifm bbs bs (NDvd i b) \<longleftrightarrow> \<not> i dvd Inum bs b"
+  | "Ifm bbs bs (NOT p) \<longleftrightarrow> \<not> Ifm bbs bs p"
+  | "Ifm bbs bs (And p q) \<longleftrightarrow> Ifm bbs bs p \<and> Ifm bbs bs q"
+  | "Ifm bbs bs (Or p q) \<longleftrightarrow> Ifm bbs bs p \<or> Ifm bbs bs q"
+  | "Ifm bbs bs (Imp p q) \<longleftrightarrow> (Ifm bbs bs p \<longrightarrow> Ifm bbs bs q)"
+  | "Ifm bbs bs (Iff p q) \<longleftrightarrow> Ifm bbs bs p = Ifm bbs bs q"
+  | "Ifm bbs bs (E p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) p)"
+  | "Ifm bbs bs (A p) \<longleftrightarrow> (\<forall>x. Ifm bbs (x # bs) p)"
+  | "Ifm bbs bs (Closed n) \<longleftrightarrow> bbs ! n"
+  | "Ifm bbs bs (NClosed n) \<longleftrightarrow> \<not> bbs ! n"
 
 fun prep :: "fm \<Rightarrow> fm"
-where
-  "prep (E T) = T"
-| "prep (E F) = F"
-| "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
-| "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
-| "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
-| "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
-| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
-| "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
-| "prep (E p) = E (prep p)"
-| "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
-| "prep (A p) = prep (NOT (E (NOT p)))"
-| "prep (NOT (NOT p)) = prep p"
-| "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (A p)) = prep (E (NOT p))"
-| "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
-| "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
-| "prep (NOT p) = NOT (prep p)"
-| "prep (Or p q) = Or (prep p) (prep q)"
-| "prep (And p q) = And (prep p) (prep q)"
-| "prep (Imp p q) = prep (Or (NOT p) q)"
-| "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
-| "prep p = p"
+  where
+    "prep (E T) = T"
+  | "prep (E F) = F"
+  | "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
+  | "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
+  | "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
+  | "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
+  | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
+  | "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+  | "prep (E p) = E (prep p)"
+  | "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
+  | "prep (A p) = prep (NOT (E (NOT p)))"
+  | "prep (NOT (NOT p)) = prep p"
+  | "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
+  | "prep (NOT (A p)) = prep (E (NOT p))"
+  | "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
+  | "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
+  | "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
+  | "prep (NOT p) = NOT (prep p)"
+  | "prep (Or p q) = Or (prep p) (prep q)"
+  | "prep (And p q) = And (prep p) (prep q)"
+  | "prep (Imp p q) = prep (Or (NOT p) q)"
+  | "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+  | "prep p = p"
 
 lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p"
   by (induct p arbitrary: bs rule: prep.induct) auto
 
 
 fun qfree :: "fm \<Rightarrow> bool"  \<comment> \<open>Quantifier freeness\<close>
-where
-  "qfree (E p) \<longleftrightarrow> False"
-| "qfree (A p) \<longleftrightarrow> False"
-| "qfree (NOT p) \<longleftrightarrow> qfree p"
-| "qfree (And p q) \<longleftrightarrow> qfree p \<and> qfree q"
-| "qfree (Or  p q) \<longleftrightarrow> qfree p \<and> qfree q"
-| "qfree (Imp p q) \<longleftrightarrow> qfree p \<and> qfree q"
-| "qfree (Iff p q) \<longleftrightarrow> qfree p \<and> qfree q"
-| "qfree p \<longleftrightarrow> True"
+  where
+    "qfree (E p) \<longleftrightarrow> False"
+  | "qfree (A p) \<longleftrightarrow> False"
+  | "qfree (NOT p) \<longleftrightarrow> qfree p"
+  | "qfree (And p q) \<longleftrightarrow> qfree p \<and> qfree q"
+  | "qfree (Or  p q) \<longleftrightarrow> qfree p \<and> qfree q"
+  | "qfree (Imp p q) \<longleftrightarrow> qfree p \<and> qfree q"
+  | "qfree (Iff p q) \<longleftrightarrow> qfree p \<and> qfree q"
+  | "qfree p \<longleftrightarrow> True"
 
 
 text \<open>Boundedness and substitution\<close>
 
-primrec numbound0 :: "num \<Rightarrow> bool"  \<comment> \<open>a num is INDEPENDENT of Bound 0\<close>
-where
-  "numbound0 (C c) \<longleftrightarrow> True"
-| "numbound0 (Bound n) \<longleftrightarrow> n > 0"
-| "numbound0 (CN n i a) \<longleftrightarrow> n > 0 \<and> numbound0 a"
-| "numbound0 (Neg a) \<longleftrightarrow> numbound0 a"
-| "numbound0 (Add a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b"
-| "numbound0 (Sub a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b"
-| "numbound0 (Mul i a) \<longleftrightarrow> numbound0 a"
+primrec numbound0 :: "num \<Rightarrow> bool"  \<comment> \<open>a \<open>num\<close> is \<^emph>\<open>independent\<close> of Bound 0\<close>
+  where
+    "numbound0 (C c) \<longleftrightarrow> True"
+  | "numbound0 (Bound n) \<longleftrightarrow> n > 0"
+  | "numbound0 (CN n i a) \<longleftrightarrow> n > 0 \<and> numbound0 a"
+  | "numbound0 (Neg a) \<longleftrightarrow> numbound0 a"
+  | "numbound0 (Add a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b"
+  | "numbound0 (Sub a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b"
+  | "numbound0 (Mul i a) \<longleftrightarrow> numbound0 a"
 
 lemma numbound0_I:
-  assumes nb: "numbound0 a"
+  assumes "numbound0 a"
   shows "Inum (b # bs) a = Inum (b' # bs) a"
-  using nb by (induct a rule: num.induct) (auto simp add: gr0_conv_Suc)
+  using assms by (induct a rule: num.induct) (auto simp add: gr0_conv_Suc)
 
-primrec bound0 :: "fm \<Rightarrow> bool" \<comment> \<open>A Formula is independent of Bound 0\<close>
-where
-  "bound0 T \<longleftrightarrow> True"
-| "bound0 F \<longleftrightarrow> True"
-| "bound0 (Lt a) \<longleftrightarrow> numbound0 a"
-| "bound0 (Le a) \<longleftrightarrow> numbound0 a"
-| "bound0 (Gt a) \<longleftrightarrow> numbound0 a"
-| "bound0 (Ge a) \<longleftrightarrow> numbound0 a"
-| "bound0 (Eq a) \<longleftrightarrow> numbound0 a"
-| "bound0 (NEq a) \<longleftrightarrow> numbound0 a"
-| "bound0 (Dvd i a) \<longleftrightarrow> numbound0 a"
-| "bound0 (NDvd i a) \<longleftrightarrow> numbound0 a"
-| "bound0 (NOT p) \<longleftrightarrow> bound0 p"
-| "bound0 (And p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
-| "bound0 (Or p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
-| "bound0 (Imp p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
-| "bound0 (Iff p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
-| "bound0 (E p) \<longleftrightarrow> False"
-| "bound0 (A p) \<longleftrightarrow> False"
-| "bound0 (Closed P) \<longleftrightarrow> True"
-| "bound0 (NClosed P) \<longleftrightarrow> True"
+primrec bound0 :: "fm \<Rightarrow> bool" \<comment> \<open>a formula is independent of Bound 0\<close>
+  where
+    "bound0 T \<longleftrightarrow> True"
+  | "bound0 F \<longleftrightarrow> True"
+  | "bound0 (Lt a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (Le a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (Gt a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (Ge a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (Eq a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (NEq a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (Dvd i a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (NDvd i a) \<longleftrightarrow> numbound0 a"
+  | "bound0 (NOT p) \<longleftrightarrow> bound0 p"
+  | "bound0 (And p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
+  | "bound0 (Or p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
+  | "bound0 (Imp p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
+  | "bound0 (Iff p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
+  | "bound0 (E p) \<longleftrightarrow> False"
+  | "bound0 (A p) \<longleftrightarrow> False"
+  | "bound0 (Closed P) \<longleftrightarrow> True"
+  | "bound0 (NClosed P) \<longleftrightarrow> True"
 
 lemma bound0_I:
-  assumes bp: "bound0 p"
+  assumes "bound0 p"
   shows "Ifm bbs (b # bs) p = Ifm bbs (b' # bs) p"
-  using bp numbound0_I[where b="b" and bs="bs" and b'="b'"]
+  using assms numbound0_I[where b="b" and bs="bs" and b'="b'"]
   by (induct p rule: fm.induct) (auto simp add: gr0_conv_Suc)
 
 fun numsubst0 :: "num \<Rightarrow> num \<Rightarrow> num"
-where
-  "numsubst0 t (C c) = (C c)"
-| "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
-| "numsubst0 t (CN 0 i a) = Add (Mul i t) (numsubst0 t a)"
-| "numsubst0 t (CN n i a) = CN n i (numsubst0 t a)"
-| "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
-| "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
-| "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"
-| "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
+  where
+    "numsubst0 t (C c) = (C c)"
+  | "numsubst0 t (Bound n) = (if n = 0 then t else Bound n)"
+  | "numsubst0 t (CN 0 i a) = Add (Mul i t) (numsubst0 t a)"
+  | "numsubst0 t (CN n i a) = CN n i (numsubst0 t a)"
+  | "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
+  | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
+  | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"
+  | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
 
-lemma numsubst0_I: "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
+lemma numsubst0_I: "Inum (b # bs) (numsubst0 a t) = Inum ((Inum (b # bs) a) # bs) t"
   by (induct t rule: numsubst0.induct) (auto simp: nth_Cons')
 
-lemma numsubst0_I':
-  "numbound0 a \<Longrightarrow> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t"
+lemma numsubst0_I': "numbound0 a \<Longrightarrow> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t"
   by (induct t rule: numsubst0.induct) (auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"])
 
-primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm"  \<comment> \<open>substitue a num into a formula for Bound 0\<close>
-where
-  "subst0 t T = T"
-| "subst0 t F = F"
-| "subst0 t (Lt a) = Lt (numsubst0 t a)"
-| "subst0 t (Le a) = Le (numsubst0 t a)"
-| "subst0 t (Gt a) = Gt (numsubst0 t a)"
-| "subst0 t (Ge a) = Ge (numsubst0 t a)"
-| "subst0 t (Eq a) = Eq (numsubst0 t a)"
-| "subst0 t (NEq a) = NEq (numsubst0 t a)"
-| "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
-| "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
-| "subst0 t (NOT p) = NOT (subst0 t p)"
-| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
-| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
-| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
-| "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
-| "subst0 t (Closed P) = (Closed P)"
-| "subst0 t (NClosed P) = (NClosed P)"
+primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm"  \<comment> \<open>substitute a \<open>num\<close> into a formula for Bound 0\<close>
+  where
+    "subst0 t T = T"
+  | "subst0 t F = F"
+  | "subst0 t (Lt a) = Lt (numsubst0 t a)"
+  | "subst0 t (Le a) = Le (numsubst0 t a)"
+  | "subst0 t (Gt a) = Gt (numsubst0 t a)"
+  | "subst0 t (Ge a) = Ge (numsubst0 t a)"
+  | "subst0 t (Eq a) = Eq (numsubst0 t a)"
+  | "subst0 t (NEq a) = NEq (numsubst0 t a)"
+  | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
+  | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
+  | "subst0 t (NOT p) = NOT (subst0 t p)"
+  | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
+  | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
+  | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
+  | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
+  | "subst0 t (Closed P) = (Closed P)"
+  | "subst0 t (NClosed P) = (NClosed P)"
 
 lemma subst0_I:
   assumes "qfree p"
@@ -237,67 +235,67 @@
   by (induct p) (simp_all add: gr0_conv_Suc)
 
 fun decrnum:: "num \<Rightarrow> num"
-where
-  "decrnum (Bound n) = Bound (n - 1)"
-| "decrnum (Neg a) = Neg (decrnum a)"
-| "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
-| "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
-| "decrnum (Mul c a) = Mul c (decrnum a)"
-| "decrnum (CN n i a) = (CN (n - 1) i (decrnum a))"
-| "decrnum a = a"
+  where
+    "decrnum (Bound n) = Bound (n - 1)"
+  | "decrnum (Neg a) = Neg (decrnum a)"
+  | "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
+  | "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
+  | "decrnum (Mul c a) = Mul c (decrnum a)"
+  | "decrnum (CN n i a) = (CN (n - 1) i (decrnum a))"
+  | "decrnum a = a"
 
 fun decr :: "fm \<Rightarrow> fm"
-where
-  "decr (Lt a) = Lt (decrnum a)"
-| "decr (Le a) = Le (decrnum a)"
-| "decr (Gt a) = Gt (decrnum a)"
-| "decr (Ge a) = Ge (decrnum a)"
-| "decr (Eq a) = Eq (decrnum a)"
-| "decr (NEq a) = NEq (decrnum a)"
-| "decr (Dvd i a) = Dvd i (decrnum a)"
-| "decr (NDvd i a) = NDvd i (decrnum a)"
-| "decr (NOT p) = NOT (decr p)"
-| "decr (And p q) = And (decr p) (decr q)"
-| "decr (Or p q) = Or (decr p) (decr q)"
-| "decr (Imp p q) = Imp (decr p) (decr q)"
-| "decr (Iff p q) = Iff (decr p) (decr q)"
-| "decr p = p"
+  where
+    "decr (Lt a) = Lt (decrnum a)"
+  | "decr (Le a) = Le (decrnum a)"
+  | "decr (Gt a) = Gt (decrnum a)"
+  | "decr (Ge a) = Ge (decrnum a)"
+  | "decr (Eq a) = Eq (decrnum a)"
+  | "decr (NEq a) = NEq (decrnum a)"
+  | "decr (Dvd i a) = Dvd i (decrnum a)"
+  | "decr (NDvd i a) = NDvd i (decrnum a)"
+  | "decr (NOT p) = NOT (decr p)"
+  | "decr (And p q) = And (decr p) (decr q)"
+  | "decr (Or p q) = Or (decr p) (decr q)"
+  | "decr (Imp p q) = Imp (decr p) (decr q)"
+  | "decr (Iff p q) = Iff (decr p) (decr q)"
+  | "decr p = p"
 
 lemma decrnum:
-  assumes nb: "numbound0 t"
+  assumes "numbound0 t"
   shows "Inum (x # bs) t = Inum bs (decrnum t)"
-  using nb by (induct t rule: decrnum.induct) (auto simp add: gr0_conv_Suc)
+  using assms by (induct t rule: decrnum.induct) (auto simp add: gr0_conv_Suc)
 
 lemma decr:
-  assumes nb: "bound0 p"
+  assumes assms: "bound0 p"
   shows "Ifm bbs (x # bs) p = Ifm bbs bs (decr p)"
-  using nb by (induct p rule: decr.induct) (simp_all add: gr0_conv_Suc decrnum)
+  using assms by (induct p rule: decr.induct) (simp_all add: gr0_conv_Suc decrnum)
 
 lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
   by (induct p) simp_all
 
 fun isatom :: "fm \<Rightarrow> bool"  \<comment> \<open>test for atomicity\<close>
-where
-  "isatom T \<longleftrightarrow> True"
-| "isatom F \<longleftrightarrow> True"
-| "isatom (Lt a) \<longleftrightarrow> True"
-| "isatom (Le a) \<longleftrightarrow> True"
-| "isatom (Gt a) \<longleftrightarrow> True"
-| "isatom (Ge a) \<longleftrightarrow> True"
-| "isatom (Eq a) \<longleftrightarrow> True"
-| "isatom (NEq a) \<longleftrightarrow> True"
-| "isatom (Dvd i b) \<longleftrightarrow> True"
-| "isatom (NDvd i b) \<longleftrightarrow> True"
-| "isatom (Closed P) \<longleftrightarrow> True"
-| "isatom (NClosed P) \<longleftrightarrow> True"
-| "isatom p \<longleftrightarrow> False"
+  where
+    "isatom T \<longleftrightarrow> True"
+  | "isatom F \<longleftrightarrow> True"
+  | "isatom (Lt a) \<longleftrightarrow> True"
+  | "isatom (Le a) \<longleftrightarrow> True"
+  | "isatom (Gt a) \<longleftrightarrow> True"
+  | "isatom (Ge a) \<longleftrightarrow> True"
+  | "isatom (Eq a) \<longleftrightarrow> True"
+  | "isatom (NEq a) \<longleftrightarrow> True"
+  | "isatom (Dvd i b) \<longleftrightarrow> True"
+  | "isatom (NDvd i b) \<longleftrightarrow> True"
+  | "isatom (Closed P) \<longleftrightarrow> True"
+  | "isatom (NClosed P) \<longleftrightarrow> True"
+  | "isatom p \<longleftrightarrow> False"
 
 lemma numsubst0_numbound0:
   assumes "numbound0 t"
   shows "numbound0 (numsubst0 t a)"
   using assms
 proof (induct a)
-  case (CN n _ _)
+  case (CN n)
   then show ?case by (cases n) simp_all
 qed simp_all
 
@@ -341,10 +339,10 @@
   using nb by (induct xs) (auto simp add: evaldjf_def djf_def Let_def, case_tac "f a", auto)
 
 fun disjuncts :: "fm \<Rightarrow> fm list"
-where
-  "disjuncts (Or p q) = disjuncts p @ disjuncts q"
-| "disjuncts F = []"
-| "disjuncts p = [p]"
+  where
+    "disjuncts (Or p q) = disjuncts p @ disjuncts q"
+  | "disjuncts F = []"
+  | "disjuncts p = [p]"
 
 lemma disjuncts: "(\<exists>q \<in> set (disjuncts p). Ifm bbs bs q) \<longleftrightarrow> Ifm bbs bs p"
   by (induct p rule: disjuncts.induct) auto
@@ -425,36 +423,36 @@
 text \<open>Algebraic simplifications for nums\<close>
 
 fun bnds :: "num \<Rightarrow> nat list"
-where
-  "bnds (Bound n) = [n]"
-| "bnds (CN n c a) = n # bnds a"
-| "bnds (Neg a) = bnds a"
-| "bnds (Add a b) = bnds a @ bnds b"
-| "bnds (Sub a b) = bnds a @ bnds b"
-| "bnds (Mul i a) = bnds a"
-| "bnds a = []"
+  where
+    "bnds (Bound n) = [n]"
+  | "bnds (CN n c a) = n # bnds a"
+  | "bnds (Neg a) = bnds a"
+  | "bnds (Add a b) = bnds a @ bnds b"
+  | "bnds (Sub a b) = bnds a @ bnds b"
+  | "bnds (Mul i a) = bnds a"
+  | "bnds a = []"
 
 fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
-where
-  "lex_ns [] ms \<longleftrightarrow> True"
-| "lex_ns ns [] \<longleftrightarrow> False"
-| "lex_ns (n # ns) (m # ms) \<longleftrightarrow> n < m \<or> (n = m \<and> lex_ns ns ms)"
+  where
+    "lex_ns [] ms \<longleftrightarrow> True"
+  | "lex_ns ns [] \<longleftrightarrow> False"
+  | "lex_ns (n # ns) (m # ms) \<longleftrightarrow> n < m \<or> (n = m \<and> lex_ns ns ms)"
 
 definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
   where "lex_bnd t s = lex_ns (bnds t) (bnds s)"
 
 fun numadd:: "num \<Rightarrow> num \<Rightarrow> num"
-where
-  "numadd (CN n1 c1 r1) (CN n2 c2 r2) =
-    (if n1 = n2 then
-       let c = c1 + c2
-       in if c = 0 then numadd r1 r2 else CN n1 c (numadd r1 r2)
-     else if n1 \<le> n2 then CN n1 c1 (numadd r1 (Add (Mul c2 (Bound n2)) r2))
-     else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1) r2))"
-| "numadd (CN n1 c1 r1) t = CN n1 c1 (numadd r1 t)"
-| "numadd t (CN n2 c2 r2) = CN n2 c2 (numadd t r2)"
-| "numadd (C b1) (C b2) = C (b1 + b2)"
-| "numadd a b = Add a b"
+  where
+    "numadd (CN n1 c1 r1) (CN n2 c2 r2) =
+      (if n1 = n2 then
+         let c = c1 + c2
+         in if c = 0 then numadd r1 r2 else CN n1 c (numadd r1 r2)
+       else if n1 \<le> n2 then CN n1 c1 (numadd r1 (Add (Mul c2 (Bound n2)) r2))
+       else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1) r2))"
+  | "numadd (CN n1 c1 r1) t = CN n1 c1 (numadd r1 t)"
+  | "numadd t (CN n2 c2 r2) = CN n2 c2 (numadd t r2)"
+  | "numadd (C b1) (C b2) = C (b1 + b2)"
+  | "numadd a b = Add a b"
 
 lemma numadd: "Inum bs (numadd t s) = Inum bs (Add t s)"
   by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)
@@ -463,10 +461,10 @@
   by (induct t s rule: numadd.induct) (simp_all add: Let_def)
 
 fun nummul :: "int \<Rightarrow> num \<Rightarrow> num"
-where
-  "nummul i (C j) = C (i * j)"
-| "nummul i (CN n c t) = CN n (c * i) (nummul i t)"
-| "nummul i t = Mul i t"
+  where
+    "nummul i (C j) = C (i * j)"
+  | "nummul i (CN n c t) = CN n (c * i) (nummul i t)"
+  | "nummul i t = Mul i t"
 
 lemma nummul: "Inum bs (nummul i t) = Inum bs (Mul i t)"
   by (induct t arbitrary: i rule: nummul.induct) (simp_all add: algebra_simps)
@@ -493,14 +491,14 @@
   using numsub_def numadd_nb numneg_nb by simp
 
 fun simpnum :: "num \<Rightarrow> num"
-where
-  "simpnum (C j) = C j"
-| "simpnum (Bound n) = CN n 1 (C 0)"
-| "simpnum (Neg t) = numneg (simpnum t)"
-| "simpnum (Add t s) = numadd (simpnum t) (simpnum s)"
-| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
-| "simpnum (Mul i t) = (if i = 0 then C 0 else nummul i (simpnum t))"
-| "simpnum t = t"
+  where
+    "simpnum (C j) = C j"
+  | "simpnum (Bound n) = CN n 1 (C 0)"
+  | "simpnum (Neg t) = numneg (simpnum t)"
+  | "simpnum (Add t s) = numadd (simpnum t) (simpnum s)"
+  | "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
+  | "simpnum (Mul i t) = (if i = 0 then C 0 else nummul i (simpnum t))"
+  | "simpnum t = t"
 
 lemma simpnum_ci: "Inum bs (simpnum t) = Inum bs t"
   by (induct t rule: simpnum.induct) (auto simp add: numneg numadd numsub nummul)
@@ -509,11 +507,11 @@
   by (induct t rule: simpnum.induct) (auto simp add: numadd_nb numsub_nb nummul_nb numneg_nb)
 
 fun not :: "fm \<Rightarrow> fm"
-where
-  "not (NOT p) = p"
-| "not T = F"
-| "not F = T"
-| "not p = NOT p"
+  where
+    "not (NOT p) = p"
+  | "not T = F"
+  | "not F = T"
+  | "not p = NOT p"
 
 lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (NOT p)"
   by (cases p) auto
@@ -525,8 +523,7 @@
   by (cases p) auto
 
 definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "conj p q =
+  where "conj p q =
     (if p = F \<or> q = F then F
      else if p = T then q
      else if q = T then p
@@ -542,8 +539,7 @@
   using conj_def by auto
 
 definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "disj p q =
+  where "disj p q =
     (if p = T \<or> q = T then T
      else if p = F then q
      else if q = F then p
@@ -559,8 +555,7 @@
   using disj_def by auto
 
 definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "imp p q =
+  where "imp p q =
     (if p = F \<or> q = T then T
      else if p = T then q
      else if q = F then not p
@@ -576,8 +571,7 @@
   using imp_def by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) simp_all
 
 definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "iff p q =
+  where "iff p q =
     (if p = q then T
      else if p = not q \<or> not p = q then F
      else if p = F then not q
@@ -597,27 +591,27 @@
   using iff_def by (unfold iff_def, cases "p = q", auto simp add: not_bn)
 
 fun simpfm :: "fm \<Rightarrow> fm"
-where
-  "simpfm (And p q) = conj (simpfm p) (simpfm q)"
-| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
-| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
-| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
-| "simpfm (NOT p) = not (simpfm p)"
-| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v < 0 then T else F | _ \<Rightarrow> Lt a')"
-| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<le> 0 then T else F | _ \<Rightarrow> Le a')"
-| "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v > 0 then T else F | _ \<Rightarrow> Gt a')"
-| "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<ge> 0 then T else F | _ \<Rightarrow> Ge a')"
-| "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v = 0 then T else F | _ \<Rightarrow> Eq a')"
-| "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<noteq> 0 then T else F | _ \<Rightarrow> NEq a')"
-| "simpfm (Dvd i a) =
-    (if i = 0 then simpfm (Eq a)
-     else if \<bar>i\<bar> = 1 then T
-     else let a' = simpnum a in case a' of C v \<Rightarrow> if i dvd v then T else F | _ \<Rightarrow> Dvd i a')"
-| "simpfm (NDvd i a) =
-    (if i = 0 then simpfm (NEq a)
-     else if \<bar>i\<bar> = 1 then F
-     else let a' = simpnum a in case a' of C v \<Rightarrow> if \<not>( i dvd v) then T else F | _ \<Rightarrow> NDvd i a')"
-| "simpfm p = p"
+  where
+    "simpfm (And p q) = conj (simpfm p) (simpfm q)"
+  | "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
+  | "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
+  | "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
+  | "simpfm (NOT p) = not (simpfm p)"
+  | "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v < 0 then T else F | _ \<Rightarrow> Lt a')"
+  | "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<le> 0 then T else F | _ \<Rightarrow> Le a')"
+  | "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v > 0 then T else F | _ \<Rightarrow> Gt a')"
+  | "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<ge> 0 then T else F | _ \<Rightarrow> Ge a')"
+  | "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v = 0 then T else F | _ \<Rightarrow> Eq a')"
+  | "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<noteq> 0 then T else F | _ \<Rightarrow> NEq a')"
+  | "simpfm (Dvd i a) =
+      (if i = 0 then simpfm (Eq a)
+       else if \<bar>i\<bar> = 1 then T
+       else let a' = simpnum a in case a' of C v \<Rightarrow> if i dvd v then T else F | _ \<Rightarrow> Dvd i a')"
+  | "simpfm (NDvd i a) =
+      (if i = 0 then simpfm (NEq a)
+       else if \<bar>i\<bar> = 1 then F
+       else let a' = simpnum a in case a' of C v \<Rightarrow> if \<not>( i dvd v) then T else F | _ \<Rightarrow> NDvd i a')"
+  | "simpfm p = p"
 
 lemma simpfm: "Ifm bbs bs (simpfm p) = Ifm bbs bs p"
 proof (induct p rule: simpfm.induct)
@@ -827,15 +821,15 @@
 
 text \<open>Generic quantifier elimination\<close>
 fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
-where
-  "qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
-| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
-| "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
-| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
-| "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
-| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
-| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
-| "qelim p = (\<lambda>y. simpfm p)"
+  where
+    "qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
+  | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
+  | "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
+  | "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
+  | "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
+  | "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
+  | "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
+  | "qelim p = (\<lambda>y. simpfm p)"
 
 lemma qelim_ci:
   assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> Ifm bbs bs (qe p) = Ifm bbs bs (E p)"
@@ -848,24 +842,24 @@
 text \<open>Linearity for fm where Bound 0 ranges over \<open>\<int>\<close>\<close>
 
 fun zsplit0 :: "num \<Rightarrow> int \<times> num"  \<comment> \<open>splits the bounded from the unbounded part\<close>
-where
-  "zsplit0 (C c) = (0, C c)"
-| "zsplit0 (Bound n) = (if n = 0 then (1, C 0) else (0, Bound n))"
-| "zsplit0 (CN n i a) =
-    (let (i', a') =  zsplit0 a
-     in if n = 0 then (i + i', a') else (i', CN n i a'))"
-| "zsplit0 (Neg a) = (let (i', a') = zsplit0 a in (-i', Neg a'))"
-| "zsplit0 (Add a b) =
-    (let
-      (ia, a') = zsplit0 a;
-      (ib, b') = zsplit0 b
-     in (ia + ib, Add a' b'))"
-| "zsplit0 (Sub a b) =
-    (let
-      (ia, a') = zsplit0 a;
-      (ib, b') = zsplit0 b
-     in (ia - ib, Sub a' b'))"
-| "zsplit0 (Mul i a) = (let (i', a') = zsplit0 a in (i*i', Mul i a'))"
+  where
+    "zsplit0 (C c) = (0, C c)"
+  | "zsplit0 (Bound n) = (if n = 0 then (1, C 0) else (0, Bound n))"
+  | "zsplit0 (CN n i a) =
+      (let (i', a') =  zsplit0 a
+       in if n = 0 then (i + i', a') else (i', CN n i a'))"
+  | "zsplit0 (Neg a) = (let (i', a') = zsplit0 a in (-i', Neg a'))"
+  | "zsplit0 (Add a b) =
+      (let
+        (ia, a') = zsplit0 a;
+        (ib, b') = zsplit0 b
+       in (ia + ib, Add a' b'))"
+  | "zsplit0 (Sub a b) =
+      (let
+        (ia, a') = zsplit0 a;
+        (ib, b') = zsplit0 b
+       in (ia - ib, Sub a' b'))"
+  | "zsplit0 (Mul i a) = (let (i', a') = zsplit0 a in (i*i', Mul i a'))"
 
 lemma zsplit0_I:
   "\<And>n a. zsplit0 t = (n, a) \<Longrightarrow>
@@ -978,91 +972,91 @@
     by simp
 qed
 
-fun iszlfm :: "fm \<Rightarrow> bool"  \<comment> \<open>Linearity test for fm\<close>
-where
-  "iszlfm (And p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"
-| "iszlfm (Or p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"
-| "iszlfm (Eq  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
-| "iszlfm (NEq (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
-| "iszlfm (Lt  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
-| "iszlfm (Le  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
-| "iszlfm (Gt  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
-| "iszlfm (Ge  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
-| "iszlfm (Dvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e"
-| "iszlfm (NDvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e"
-| "iszlfm p \<longleftrightarrow> isatom p \<and> bound0 p"
+fun iszlfm :: "fm \<Rightarrow> bool"  \<comment> \<open>linearity test for fm\<close>
+  where
+    "iszlfm (And p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"
+  | "iszlfm (Or p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"
+  | "iszlfm (Eq  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
+  | "iszlfm (NEq (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
+  | "iszlfm (Lt  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
+  | "iszlfm (Le  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
+  | "iszlfm (Gt  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
+  | "iszlfm (Ge  (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"
+  | "iszlfm (Dvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e"
+  | "iszlfm (NDvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e"
+  | "iszlfm p \<longleftrightarrow> isatom p \<and> bound0 p"
 
 lemma zlin_qfree: "iszlfm p \<Longrightarrow> qfree p"
   by (induct p rule: iszlfm.induct) auto
 
-fun zlfm :: "fm \<Rightarrow> fm"  \<comment> \<open>Linearity transformation for fm\<close>
-where
-  "zlfm (And p q) = And (zlfm p) (zlfm q)"
-| "zlfm (Or p q) = Or (zlfm p) (zlfm q)"
-| "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)"
-| "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (NOT p)) (zlfm (NOT q)))"
-| "zlfm (Lt a) =
-    (let (c, r) = zsplit0 a in
-      if c = 0 then Lt r else
-      if c > 0 then (Lt (CN 0 c r))
-      else Gt (CN 0 (- c) (Neg r)))"
-| "zlfm (Le a) =
-    (let (c, r) = zsplit0 a in
-      if c = 0 then Le r
-      else if c > 0 then Le (CN 0 c r)
-      else Ge (CN 0 (- c) (Neg r)))"
-| "zlfm (Gt a) =
-    (let (c, r) = zsplit0 a in
-      if c = 0 then Gt r else
-      if c > 0 then Gt (CN 0 c r)
-      else Lt (CN 0 (- c) (Neg r)))"
-| "zlfm (Ge a) =
-    (let (c, r) = zsplit0 a in
-      if c = 0 then Ge r
-      else if c > 0 then Ge (CN 0 c r)
-      else Le (CN 0 (- c) (Neg r)))"
-| "zlfm (Eq a) =
-    (let (c, r) = zsplit0 a in
-      if c = 0 then Eq r
-      else if c > 0 then Eq (CN 0 c r)
-      else Eq (CN 0 (- c) (Neg r)))"
-| "zlfm (NEq a) =
-    (let (c, r) = zsplit0 a in
-      if c = 0 then NEq r
-      else if c > 0 then NEq (CN 0 c r)
-      else NEq (CN 0 (- c) (Neg r)))"
-| "zlfm (Dvd i a) =
-    (if i = 0 then zlfm (Eq a)
-     else
-      let (c, r) = zsplit0 a in
-        if c = 0 then Dvd \<bar>i\<bar> r
-        else if c > 0 then Dvd \<bar>i\<bar> (CN 0 c r)
-        else Dvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"
-| "zlfm (NDvd i a) =
-    (if i = 0 then zlfm (NEq a)
-     else
-      let (c, r) = zsplit0 a in
-        if c = 0 then NDvd \<bar>i\<bar> r
-        else if c > 0 then NDvd \<bar>i\<bar> (CN 0 c r)
-        else NDvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"
-| "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))"
-| "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))"
-| "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))"
-| "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))"
-| "zlfm (NOT (NOT p)) = zlfm p"
-| "zlfm (NOT T) = F"
-| "zlfm (NOT F) = T"
-| "zlfm (NOT (Lt a)) = zlfm (Ge a)"
-| "zlfm (NOT (Le a)) = zlfm (Gt a)"
-| "zlfm (NOT (Gt a)) = zlfm (Le a)"
-| "zlfm (NOT (Ge a)) = zlfm (Lt a)"
-| "zlfm (NOT (Eq a)) = zlfm (NEq a)"
-| "zlfm (NOT (NEq a)) = zlfm (Eq a)"
-| "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
-| "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
-| "zlfm (NOT (Closed P)) = NClosed P"
-| "zlfm (NOT (NClosed P)) = Closed P"
-| "zlfm p = p"
+fun zlfm :: "fm \<Rightarrow> fm"  \<comment> \<open>linearity transformation for fm\<close>
+  where
+    "zlfm (And p q) = And (zlfm p) (zlfm q)"
+  | "zlfm (Or p q) = Or (zlfm p) (zlfm q)"
+  | "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)"
+  | "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (NOT p)) (zlfm (NOT q)))"
+  | "zlfm (Lt a) =
+      (let (c, r) = zsplit0 a in
+        if c = 0 then Lt r else
+        if c > 0 then (Lt (CN 0 c r))
+        else Gt (CN 0 (- c) (Neg r)))"
+  | "zlfm (Le a) =
+      (let (c, r) = zsplit0 a in
+        if c = 0 then Le r
+        else if c > 0 then Le (CN 0 c r)
+        else Ge (CN 0 (- c) (Neg r)))"
+  | "zlfm (Gt a) =
+      (let (c, r) = zsplit0 a in
+        if c = 0 then Gt r else
+        if c > 0 then Gt (CN 0 c r)
+        else Lt (CN 0 (- c) (Neg r)))"
+  | "zlfm (Ge a) =
+      (let (c, r) = zsplit0 a in
+        if c = 0 then Ge r
+        else if c > 0 then Ge (CN 0 c r)
+        else Le (CN 0 (- c) (Neg r)))"
+  | "zlfm (Eq a) =
+      (let (c, r) = zsplit0 a in
+        if c = 0 then Eq r
+        else if c > 0 then Eq (CN 0 c r)
+        else Eq (CN 0 (- c) (Neg r)))"
+  | "zlfm (NEq a) =
+      (let (c, r) = zsplit0 a in
+        if c = 0 then NEq r
+        else if c > 0 then NEq (CN 0 c r)
+        else NEq (CN 0 (- c) (Neg r)))"
+  | "zlfm (Dvd i a) =
+      (if i = 0 then zlfm (Eq a)
+       else
+        let (c, r) = zsplit0 a in
+          if c = 0 then Dvd \<bar>i\<bar> r
+          else if c > 0 then Dvd \<bar>i\<bar> (CN 0 c r)
+          else Dvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"
+  | "zlfm (NDvd i a) =
+      (if i = 0 then zlfm (NEq a)
+       else
+        let (c, r) = zsplit0 a in
+          if c = 0 then NDvd \<bar>i\<bar> r
+          else if c > 0 then NDvd \<bar>i\<bar> (CN 0 c r)
+          else NDvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"
+  | "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))"
+  | "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))"
+  | "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))"
+  | "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))"
+  | "zlfm (NOT (NOT p)) = zlfm p"
+  | "zlfm (NOT T) = F"
+  | "zlfm (NOT F) = T"
+  | "zlfm (NOT (Lt a)) = zlfm (Ge a)"
+  | "zlfm (NOT (Le a)) = zlfm (Gt a)"
+  | "zlfm (NOT (Gt a)) = zlfm (Le a)"
+  | "zlfm (NOT (Ge a)) = zlfm (Lt a)"
+  | "zlfm (NOT (Eq a)) = zlfm (NEq a)"
+  | "zlfm (NOT (NEq a)) = zlfm (Eq a)"
+  | "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
+  | "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
+  | "zlfm (NOT (Closed P)) = NClosed P"
+  | "zlfm (NOT (NClosed P)) = Closed P"
+  | "zlfm p = p"
 
 lemma zlfm_I:
   assumes qfp: "qfree p"
@@ -1212,7 +1206,7 @@
   have spl: "zsplit0 a = (?c, ?r)"
     by simp
   from zsplit0_I[OF spl, where x="i" and bs="bs"]
-  have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"
+  have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"
     by auto
   let ?N = "\<lambda>t. Inum (i # bs) t"
   consider "j = 0" | "j \<noteq> 0" "?c = 0" | "j \<noteq> 0" "?c > 0" | "j \<noteq> 0" "?c < 0"
@@ -1247,48 +1241,48 @@
   qed
 qed auto
 
-fun minusinf :: "fm \<Rightarrow> fm" \<comment> \<open>Virtual substitution of \<open>-\<infinity>\<close>\<close>
-where
-  "minusinf (And p q) = And (minusinf p) (minusinf q)"
-| "minusinf (Or p q) = Or (minusinf p) (minusinf q)"
-| "minusinf (Eq  (CN 0 c e)) = F"
-| "minusinf (NEq (CN 0 c e)) = T"
-| "minusinf (Lt  (CN 0 c e)) = T"
-| "minusinf (Le  (CN 0 c e)) = T"
-| "minusinf (Gt  (CN 0 c e)) = F"
-| "minusinf (Ge  (CN 0 c e)) = F"
-| "minusinf p = p"
+fun minusinf :: "fm \<Rightarrow> fm" \<comment> \<open>virtual substitution of \<open>-\<infinity>\<close>\<close>
+  where
+    "minusinf (And p q) = And (minusinf p) (minusinf q)"
+  | "minusinf (Or p q) = Or (minusinf p) (minusinf q)"
+  | "minusinf (Eq  (CN 0 c e)) = F"
+  | "minusinf (NEq (CN 0 c e)) = T"
+  | "minusinf (Lt  (CN 0 c e)) = T"
+  | "minusinf (Le  (CN 0 c e)) = T"
+  | "minusinf (Gt  (CN 0 c e)) = F"
+  | "minusinf (Ge  (CN 0 c e)) = F"
+  | "minusinf p = p"
 
 lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
   by (induct p rule: minusinf.induct) auto
 
-fun plusinf :: "fm \<Rightarrow> fm"  \<comment> \<open>Virtual substitution of \<open>+\<infinity>\<close>\<close>
-where
-  "plusinf (And p q) = And (plusinf p) (plusinf q)"
-| "plusinf (Or p q) = Or (plusinf p) (plusinf q)"
-| "plusinf (Eq  (CN 0 c e)) = F"
-| "plusinf (NEq (CN 0 c e)) = T"
-| "plusinf (Lt  (CN 0 c e)) = F"
-| "plusinf (Le  (CN 0 c e)) = F"
-| "plusinf (Gt  (CN 0 c e)) = T"
-| "plusinf (Ge  (CN 0 c e)) = T"
-| "plusinf p = p"
+fun plusinf :: "fm \<Rightarrow> fm"  \<comment> \<open>virtual substitution of \<open>+\<infinity>\<close>\<close>
+  where
+    "plusinf (And p q) = And (plusinf p) (plusinf q)"
+  | "plusinf (Or p q) = Or (plusinf p) (plusinf q)"
+  | "plusinf (Eq  (CN 0 c e)) = F"
+  | "plusinf (NEq (CN 0 c e)) = T"
+  | "plusinf (Lt  (CN 0 c e)) = F"
+  | "plusinf (Le  (CN 0 c e)) = F"
+  | "plusinf (Gt  (CN 0 c e)) = T"
+  | "plusinf (Ge  (CN 0 c e)) = T"
+  | "plusinf p = p"
 
-fun \<delta> :: "fm \<Rightarrow> int"  \<comment> \<open>Compute \<open>lcm {d| N\<^sup>? Dvd c*x+t \<in> p}\<close>\<close>
-where
-  "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
-| "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
-| "\<delta> (Dvd i (CN 0 c e)) = i"
-| "\<delta> (NDvd i (CN 0 c e)) = i"
-| "\<delta> p = 1"
+fun \<delta> :: "fm \<Rightarrow> int"  \<comment> \<open>compute \<open>lcm {d| N\<^sup>? Dvd c*x+t \<in> p}\<close>\<close>
+  where
+    "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
+  | "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
+  | "\<delta> (Dvd i (CN 0 c e)) = i"
+  | "\<delta> (NDvd i (CN 0 c e)) = i"
+  | "\<delta> p = 1"
 
-fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  \<comment> \<open>check if a given l divides all the ds above\<close>
-where
-  "d_\<delta> (And p q) d \<longleftrightarrow> d_\<delta> p d \<and> d_\<delta> q d"
-| "d_\<delta> (Or p q) d \<longleftrightarrow> d_\<delta> p d \<and> d_\<delta> q d"
-| "d_\<delta> (Dvd i (CN 0 c e)) d \<longleftrightarrow> i dvd d"
-| "d_\<delta> (NDvd i (CN 0 c e)) d \<longleftrightarrow> i dvd d"
-| "d_\<delta> p d \<longleftrightarrow> True"
+fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  \<comment> \<open>check if a given \<open>l\<close> divides all the \<open>ds\<close> above\<close>
+  where
+    "d_\<delta> (And p q) d \<longleftrightarrow> d_\<delta> p d \<and> d_\<delta> q d"
+  | "d_\<delta> (Or p q) d \<longleftrightarrow> d_\<delta> p d \<and> d_\<delta> q d"
+  | "d_\<delta> (Dvd i (CN 0 c e)) d \<longleftrightarrow> i dvd d"
+  | "d_\<delta> (NDvd i (CN 0 c e)) d \<longleftrightarrow> i dvd d"
+  | "d_\<delta> p d \<longleftrightarrow> True"
 
 lemma delta_mono:
   assumes lin: "iszlfm p"
@@ -1310,93 +1304,92 @@
   assumes lin: "iszlfm p"
   shows "d_\<delta> p (\<delta> p) \<and> \<delta> p >0"
   using lin
-  by (induct p rule: iszlfm.induct)  (auto intro: delta_mono simp add: lcm_pos_int)
+  by (induct p rule: iszlfm.induct) (auto intro: delta_mono simp add: lcm_pos_int)
 
 fun a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm"  \<comment> \<open>adjust the coefficients of a formula\<close>
-where
-  "a_\<beta> (And p q) k = And (a_\<beta> p k) (a_\<beta> q k)"
-| "a_\<beta> (Or p q) k = Or (a_\<beta> p k) (a_\<beta> q k)"
-| "a_\<beta> (Eq  (CN 0 c e)) k = Eq (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (NEq (CN 0 c e)) k = NEq (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (Lt  (CN 0 c e)) k = Lt (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (Le  (CN 0 c e)) k = Le (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (Gt  (CN 0 c e)) k = Gt (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (Ge  (CN 0 c e)) k = Ge (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (Dvd i (CN 0 c e)) k = Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> (NDvd i (CN 0 c e)) k = NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e))"
-| "a_\<beta> p k = p"
+  where
+    "a_\<beta> (And p q) k = And (a_\<beta> p k) (a_\<beta> q k)"
+  | "a_\<beta> (Or p q) k = Or (a_\<beta> p k) (a_\<beta> q k)"
+  | "a_\<beta> (Eq  (CN 0 c e)) k = Eq (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (NEq (CN 0 c e)) k = NEq (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (Lt  (CN 0 c e)) k = Lt (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (Le  (CN 0 c e)) k = Le (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (Gt  (CN 0 c e)) k = Gt (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (Ge  (CN 0 c e)) k = Ge (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (Dvd i (CN 0 c e)) k = Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> (NDvd i (CN 0 c e)) k = NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e))"
+  | "a_\<beta> p k = p"
 
-fun d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  \<comment> \<open>test if all coeffs c of c divide a given l\<close>
-where
-  "d_\<beta> (And p q) k \<longleftrightarrow> d_\<beta> p k \<and> d_\<beta> q k"
-| "d_\<beta> (Or p q) k \<longleftrightarrow> d_\<beta> p k \<and> d_\<beta> q k"
-| "d_\<beta> (Eq  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (NEq (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (Lt  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (Le  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (Gt  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (Ge  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (Dvd i (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> (NDvd i (CN 0 c e)) k \<longleftrightarrow> c dvd k"
-| "d_\<beta> p k \<longleftrightarrow> True"
+fun d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool"  \<comment> \<open>test if all coeffs of \<open>c\<close> divide a given \<open>l\<close>\<close>
+  where
+    "d_\<beta> (And p q) k \<longleftrightarrow> d_\<beta> p k \<and> d_\<beta> q k"
+  | "d_\<beta> (Or p q) k \<longleftrightarrow> d_\<beta> p k \<and> d_\<beta> q k"
+  | "d_\<beta> (Eq  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (NEq (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (Lt  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (Le  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (Gt  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (Ge  (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (Dvd i (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> (NDvd i (CN 0 c e)) k \<longleftrightarrow> c dvd k"
+  | "d_\<beta> p k \<longleftrightarrow> True"
 
-fun \<zeta> :: "fm \<Rightarrow> int"  \<comment> \<open>computes the lcm of all coefficients of x\<close>
-where
-  "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
-| "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
-| "\<zeta> (Eq  (CN 0 c e)) = c"
-| "\<zeta> (NEq (CN 0 c e)) = c"
-| "\<zeta> (Lt  (CN 0 c e)) = c"
-| "\<zeta> (Le  (CN 0 c e)) = c"
-| "\<zeta> (Gt  (CN 0 c e)) = c"
-| "\<zeta> (Ge  (CN 0 c e)) = c"
-| "\<zeta> (Dvd i (CN 0 c e)) = c"
-| "\<zeta> (NDvd i (CN 0 c e))= c"
-| "\<zeta> p = 1"
+fun \<zeta> :: "fm \<Rightarrow> int"  \<comment> \<open>computes the lcm of all coefficients of \<open>x\<close>\<close>
+  where
+    "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
+  | "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
+  | "\<zeta> (Eq  (CN 0 c e)) = c"
+  | "\<zeta> (NEq (CN 0 c e)) = c"
+  | "\<zeta> (Lt  (CN 0 c e)) = c"
+  | "\<zeta> (Le  (CN 0 c e)) = c"
+  | "\<zeta> (Gt  (CN 0 c e)) = c"
+  | "\<zeta> (Ge  (CN 0 c e)) = c"
+  | "\<zeta> (Dvd i (CN 0 c e)) = c"
+  | "\<zeta> (NDvd i (CN 0 c e))= c"
+  | "\<zeta> p = 1"
 
 fun \<beta> :: "fm \<Rightarrow> num list"
-where
-  "\<beta> (And p q) = (\<beta> p @ \<beta> q)"
-| "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"
-| "\<beta> (Eq  (CN 0 c e)) = [Sub (C (- 1)) e]"
-| "\<beta> (NEq (CN 0 c e)) = [Neg e]"
-| "\<beta> (Lt  (CN 0 c e)) = []"
-| "\<beta> (Le  (CN 0 c e)) = []"
-| "\<beta> (Gt  (CN 0 c e)) = [Neg e]"
-| "\<beta> (Ge  (CN 0 c e)) = [Sub (C (- 1)) e]"
-| "\<beta> p = []"
+  where
+    "\<beta> (And p q) = (\<beta> p @ \<beta> q)"
+  | "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"
+  | "\<beta> (Eq  (CN 0 c e)) = [Sub (C (- 1)) e]"
+  | "\<beta> (NEq (CN 0 c e)) = [Neg e]"
+  | "\<beta> (Lt  (CN 0 c e)) = []"
+  | "\<beta> (Le  (CN 0 c e)) = []"
+  | "\<beta> (Gt  (CN 0 c e)) = [Neg e]"
+  | "\<beta> (Ge  (CN 0 c e)) = [Sub (C (- 1)) e]"
+  | "\<beta> p = []"
 
 fun \<alpha> :: "fm \<Rightarrow> num list"
-where
-  "\<alpha> (And p q) = \<alpha> p @ \<alpha> q"
-| "\<alpha> (Or p q) = \<alpha> p @ \<alpha> q"
-| "\<alpha> (Eq  (CN 0 c e)) = [Add (C (- 1)) e]"
-| "\<alpha> (NEq (CN 0 c e)) = [e]"
-| "\<alpha> (Lt  (CN 0 c e)) = [e]"
-| "\<alpha> (Le  (CN 0 c e)) = [Add (C (- 1)) e]"
-| "\<alpha> (Gt  (CN 0 c e)) = []"
-| "\<alpha> (Ge  (CN 0 c e)) = []"
-| "\<alpha> p = []"
+  where
+    "\<alpha> (And p q) = \<alpha> p @ \<alpha> q"
+  | "\<alpha> (Or p q) = \<alpha> p @ \<alpha> q"
+  | "\<alpha> (Eq  (CN 0 c e)) = [Add (C (- 1)) e]"
+  | "\<alpha> (NEq (CN 0 c e)) = [e]"
+  | "\<alpha> (Lt  (CN 0 c e)) = [e]"
+  | "\<alpha> (Le  (CN 0 c e)) = [Add (C (- 1)) e]"
+  | "\<alpha> (Gt  (CN 0 c e)) = []"
+  | "\<alpha> (Ge  (CN 0 c e)) = []"
+  | "\<alpha> p = []"
 
 fun mirror :: "fm \<Rightarrow> fm"
-where
-  "mirror (And p q) = And (mirror p) (mirror q)"
-| "mirror (Or p q) = Or (mirror p) (mirror q)"
-| "mirror (Eq  (CN 0 c e)) = Eq (CN 0 c (Neg e))"
-| "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
-| "mirror (Lt  (CN 0 c e)) = Gt (CN 0 c (Neg e))"
-| "mirror (Le  (CN 0 c e)) = Ge (CN 0 c (Neg e))"
-| "mirror (Gt  (CN 0 c e)) = Lt (CN 0 c (Neg e))"
-| "mirror (Ge  (CN 0 c e)) = Le (CN 0 c (Neg e))"
-| "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
-| "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
-| "mirror p = p"
+  where
+    "mirror (And p q) = And (mirror p) (mirror q)"
+  | "mirror (Or p q) = Or (mirror p) (mirror q)"
+  | "mirror (Eq  (CN 0 c e)) = Eq (CN 0 c (Neg e))"
+  | "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
+  | "mirror (Lt  (CN 0 c e)) = Gt (CN 0 c (Neg e))"
+  | "mirror (Le  (CN 0 c e)) = Ge (CN 0 c (Neg e))"
+  | "mirror (Gt  (CN 0 c e)) = Lt (CN 0 c (Neg e))"
+  | "mirror (Ge  (CN 0 c e)) = Le (CN 0 c (Neg e))"
+  | "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
+  | "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
+  | "mirror p = p"
 
 text \<open>Lemmas for the correctness of \<open>\<sigma>_\<rho>\<close>\<close>
 
-lemma dvd1_eq1:
-  fixes x :: int
-  shows "x > 0 \<Longrightarrow> x dvd 1 \<longleftrightarrow> x = 1"
+lemma dvd1_eq1: "x > 0 \<Longrightarrow> x dvd 1 \<longleftrightarrow> x = 1"
+  for x :: int
   by simp
 
 lemma minusinf_inf:
@@ -2141,7 +2134,8 @@
     by blast
 qed
 
-(* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *)
+text \<open>Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff.\<close>
+
 lemma mirror_ex:
   assumes "iszlfm p"
   shows "(\<exists>x. Ifm bbs (x#bs) (mirror p)) \<longleftrightarrow> (\<exists>x. Ifm bbs (x#bs) p)"
@@ -2378,10 +2372,9 @@
   using qelim_ci cooper prep by (auto simp add: pa_def)
 
 definition cooper_test :: "unit \<Rightarrow> fm"
-  where
-    "cooper_test u =
-      pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1)))
-        (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) (Bound 2))))))))"
+  where "cooper_test u =
+    pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1)))
+      (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) (Bound 2))))))))"
 
 ML_val \<open>@{code cooper_test} ()\<close>
 
@@ -2505,10 +2498,10 @@
   let
     val is_op =
       member (op =) [@{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies},
-      @{term "op = :: bool => _"},
-      @{term "op = :: int => _"}, @{term "op < :: int => _"},
-      @{term "op <= :: int => _"}, @{term "Not"}, @{term "All :: (int => _) => _"},
-      @{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}]
+      @{term "op = :: bool \<Rightarrow> _"},
+      @{term "op = :: int \<Rightarrow> _"}, @{term "op < :: int \<Rightarrow> _"},
+      @{term "op \<le> :: int \<Rightarrow> _"}, @{term "Not"}, @{term "All :: (int \<Rightarrow> _) \<Rightarrow> _"},
+      @{term "Ex :: (int \<Rightarrow> _) \<Rightarrow> _"}, @{term "True"}, @{term "False"}]
     fun is_ty t = not (fastype_of t = HOLogic.boolT)
   in
     (case t of
@@ -2629,7 +2622,7 @@
 theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y"
   by cooper
 
-theorem "\<exists>(x::int) y. 0 < x  & 0 \<le> y  & 3 * x - 5 * y = 1"
+theorem "\<exists>(x::int) y. 0 < x  \<and> 0 \<le> y \<and> 3 * x - 5 * y = 1"
   by cooper
 
 theorem "\<not> (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)"
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -21,14 +21,14 @@
 begin
 
 primrec size_tm :: "tm \<Rightarrow> nat"
-where
-  "size_tm (CP c) = polysize c"
-| "size_tm (Bound n) = 1"
-| "size_tm (Neg a) = 1 + size_tm a"
-| "size_tm (Add a b) = 1 + size_tm a + size_tm b"
-| "size_tm (Sub a b) = 3 + size_tm a + size_tm b"
-| "size_tm (Mul c a) = 1 + polysize c + size_tm a"
-| "size_tm (CNP n c a) = 3 + polysize c + size_tm a "
+  where
+    "size_tm (CP c) = polysize c"
+  | "size_tm (Bound n) = 1"
+  | "size_tm (Neg a) = 1 + size_tm a"
+  | "size_tm (Add a b) = 1 + size_tm a + size_tm b"
+  | "size_tm (Sub a b) = 3 + size_tm a + size_tm b"
+  | "size_tm (Mul c a) = 1 + polysize c + size_tm a"
+  | "size_tm (CNP n c a) = 3 + polysize c + size_tm a "
 
 instance ..
 
@@ -36,60 +36,59 @@
 
 text \<open>Semantics of terms tm.\<close>
 primrec Itm :: "'a::{field_char_0,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
-where
-  "Itm vs bs (CP c) = (Ipoly vs c)"
-| "Itm vs bs (Bound n) = bs!n"
-| "Itm vs bs (Neg a) = -(Itm vs bs a)"
-| "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
-| "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
-| "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
-| "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"
+  where
+    "Itm vs bs (CP c) = (Ipoly vs c)"
+  | "Itm vs bs (Bound n) = bs!n"
+  | "Itm vs bs (Neg a) = -(Itm vs bs a)"
+  | "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
+  | "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
+  | "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
+  | "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"
 
 fun allpolys :: "(poly \<Rightarrow> bool) \<Rightarrow> tm \<Rightarrow> bool"
-where
-  "allpolys P (CP c) = P c"
-| "allpolys P (CNP n c p) = (P c \<and> allpolys P p)"
-| "allpolys P (Mul c p) = (P c \<and> allpolys P p)"
-| "allpolys P (Neg p) = allpolys P p"
-| "allpolys P (Add p q) = (allpolys P p \<and> allpolys P q)"
-| "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
-| "allpolys P p = True"
+  where
+    "allpolys P (CP c) = P c"
+  | "allpolys P (CNP n c p) = (P c \<and> allpolys P p)"
+  | "allpolys P (Mul c p) = (P c \<and> allpolys P p)"
+  | "allpolys P (Neg p) = allpolys P p"
+  | "allpolys P (Add p q) = (allpolys P p \<and> allpolys P q)"
+  | "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
+  | "allpolys P p = True"
 
 primrec tmboundslt :: "nat \<Rightarrow> tm \<Rightarrow> bool"
-where
-  "tmboundslt n (CP c) = True"
-| "tmboundslt n (Bound m) = (m < n)"
-| "tmboundslt n (CNP m c a) = (m < n \<and> tmboundslt n a)"
-| "tmboundslt n (Neg a) = tmboundslt n a"
-| "tmboundslt n (Add a b) = (tmboundslt n a \<and> tmboundslt n b)"
-| "tmboundslt n (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)"
-| "tmboundslt n (Mul i a) = tmboundslt n a"
-
-primrec tmbound0 :: "tm \<Rightarrow> bool"  \<comment> \<open>a tm is INDEPENDENT of Bound 0\<close>
-where
-  "tmbound0 (CP c) = True"
-| "tmbound0 (Bound n) = (n>0)"
-| "tmbound0 (CNP n c a) = (n\<noteq>0 \<and> tmbound0 a)"
-| "tmbound0 (Neg a) = tmbound0 a"
-| "tmbound0 (Add a b) = (tmbound0 a \<and> tmbound0 b)"
-| "tmbound0 (Sub a b) = (tmbound0 a \<and> tmbound0 b)"
-| "tmbound0 (Mul i a) = tmbound0 a"
+  where
+    "tmboundslt n (CP c) = True"
+  | "tmboundslt n (Bound m) = (m < n)"
+  | "tmboundslt n (CNP m c a) = (m < n \<and> tmboundslt n a)"
+  | "tmboundslt n (Neg a) = tmboundslt n a"
+  | "tmboundslt n (Add a b) = (tmboundslt n a \<and> tmboundslt n b)"
+  | "tmboundslt n (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)"
+  | "tmboundslt n (Mul i a) = tmboundslt n a"
+
+primrec tmbound0 :: "tm \<Rightarrow> bool"  \<comment> \<open>a \<open>tm\<close> is \<^emph>\<open>independent\<close> of Bound 0\<close>
+  where
+    "tmbound0 (CP c) = True"
+  | "tmbound0 (Bound n) = (n>0)"
+  | "tmbound0 (CNP n c a) = (n\<noteq>0 \<and> tmbound0 a)"
+  | "tmbound0 (Neg a) = tmbound0 a"
+  | "tmbound0 (Add a b) = (tmbound0 a \<and> tmbound0 b)"
+  | "tmbound0 (Sub a b) = (tmbound0 a \<and> tmbound0 b)"
+  | "tmbound0 (Mul i a) = tmbound0 a"
 
 lemma tmbound0_I:
-  assumes nb: "tmbound0 a"
+  assumes "tmbound0 a"
   shows "Itm vs (b#bs) a = Itm vs (b'#bs) a"
-  using nb
-  by (induct a rule: tm.induct) auto
-
-primrec tmbound :: "nat \<Rightarrow> tm \<Rightarrow> bool"  \<comment> \<open>a tm is INDEPENDENT of Bound n\<close>
-where
-  "tmbound n (CP c) = True"
-| "tmbound n (Bound m) = (n \<noteq> m)"
-| "tmbound n (CNP m c a) = (n\<noteq>m \<and> tmbound n a)"
-| "tmbound n (Neg a) = tmbound n a"
-| "tmbound n (Add a b) = (tmbound n a \<and> tmbound n b)"
-| "tmbound n (Sub a b) = (tmbound n a \<and> tmbound n b)"
-| "tmbound n (Mul i a) = tmbound n a"
+  using assms by (induct a rule: tm.induct) auto
+
+primrec tmbound :: "nat \<Rightarrow> tm \<Rightarrow> bool"  \<comment> \<open>a \<open>tm\<close> is \<^emph>\<open>independent\<close> of Bound n\<close>
+  where
+    "tmbound n (CP c) = True"
+  | "tmbound n (Bound m) = (n \<noteq> m)"
+  | "tmbound n (CNP m c a) = (n\<noteq>m \<and> tmbound n a)"
+  | "tmbound n (Neg a) = tmbound n a"
+  | "tmbound n (Add a b) = (tmbound n a \<and> tmbound n b)"
+  | "tmbound n (Sub a b) = (tmbound n a \<and> tmbound n b)"
+  | "tmbound n (Mul i a) = tmbound n a"
 
 lemma tmbound0_tmbound_iff: "tmbound 0 t = tmbound0 t"
   by (induct t) auto
@@ -103,24 +102,24 @@
   by (induct t rule: tm.induct) auto
 
 fun decrtm0 :: "tm \<Rightarrow> tm"
-where
-  "decrtm0 (Bound n) = Bound (n - 1)"
-| "decrtm0 (Neg a) = Neg (decrtm0 a)"
-| "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)"
-| "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)"
-| "decrtm0 (Mul c a) = Mul c (decrtm0 a)"
-| "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
-| "decrtm0 a = a"
+  where
+    "decrtm0 (Bound n) = Bound (n - 1)"
+  | "decrtm0 (Neg a) = Neg (decrtm0 a)"
+  | "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)"
+  | "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)"
+  | "decrtm0 (Mul c a) = Mul c (decrtm0 a)"
+  | "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
+  | "decrtm0 a = a"
 
 fun incrtm0 :: "tm \<Rightarrow> tm"
-where
-  "incrtm0 (Bound n) = Bound (n + 1)"
-| "incrtm0 (Neg a) = Neg (incrtm0 a)"
-| "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)"
-| "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)"
-| "incrtm0 (Mul c a) = Mul c (incrtm0 a)"
-| "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
-| "incrtm0 a = a"
+  where
+    "incrtm0 (Bound n) = Bound (n + 1)"
+  | "incrtm0 (Neg a) = Neg (incrtm0 a)"
+  | "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)"
+  | "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)"
+  | "incrtm0 (Mul c a) = Mul c (incrtm0 a)"
+  | "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
+  | "incrtm0 a = a"
 
 lemma decrtm0:
   assumes nb: "tmbound0 t"
@@ -131,19 +130,19 @@
   by (induct t rule: decrtm0.induct) simp_all
 
 primrec decrtm :: "nat \<Rightarrow> tm \<Rightarrow> tm"
-where
-  "decrtm m (CP c) = (CP c)"
-| "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))"
-| "decrtm m (Neg a) = Neg (decrtm m a)"
-| "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)"
-| "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)"
-| "decrtm m (Mul c a) = Mul c (decrtm m a)"
-| "decrtm m (CNP n c a) = (if n < m then CNP n c (decrtm m a) else CNP (n - 1) c (decrtm m a))"
+  where
+    "decrtm m (CP c) = (CP c)"
+  | "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))"
+  | "decrtm m (Neg a) = Neg (decrtm m a)"
+  | "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)"
+  | "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)"
+  | "decrtm m (Mul c a) = Mul c (decrtm m a)"
+  | "decrtm m (CNP n c a) = (if n < m then CNP n c (decrtm m a) else CNP (n - 1) c (decrtm m a))"
 
 primrec removen :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "removen n [] = []"
-| "removen n (x#xs) = (if n=0 then xs else (x#(removen (n - 1) xs)))"
+  where
+    "removen n [] = []"
+  | "removen n (x#xs) = (if n=0 then xs else (x#(removen (n - 1) xs)))"
 
 lemma removen_same: "n \<ge> length xs \<Longrightarrow> removen n xs = xs"
   by (induct xs arbitrary: n) auto
@@ -152,7 +151,7 @@
   by (induct xs arbitrary: n) auto
 
 lemma removen_length: "length (removen n xs) = (if n \<ge> length xs then length xs else length xs - 1)"
-  by (induct xs arbitrary: n, auto)
+  by (induct xs arbitrary: n) auto
 
 lemma removen_nth:
   "(removen n xs)!m =
@@ -212,14 +211,14 @@
   using bnd nb nle by (induct t rule: tm.induct) (auto simp add: removen_nth)
 
 primrec tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm"
-where
-  "tmsubst0 t (CP c) = CP c"
-| "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)"
-| "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))"
-| "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)"
-| "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)"
-| "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)"
-| "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
+  where
+    "tmsubst0 t (CP c) = CP c"
+  | "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)"
+  | "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))"
+  | "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)"
+  | "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)"
+  | "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)"
+  | "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
 
 lemma tmsubst0: "Itm vs (x # bs) (tmsubst0 t a) = Itm vs (Itm vs (x # bs) t # bs) a"
   by (induct a rule: tm.induct) auto
@@ -228,15 +227,15 @@
   by (induct a rule: tm.induct) auto
 
 primrec tmsubst:: "nat \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm"
-where
-  "tmsubst n t (CP c) = CP c"
-| "tmsubst n t (Bound m) = (if n=m then t else Bound m)"
-| "tmsubst n t (CNP m c a) =
-    (if n = m then Add (Mul c t) (tmsubst n t a) else CNP m c (tmsubst n t a))"
-| "tmsubst n t (Neg a) = Neg (tmsubst n t a)"
-| "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)"
-| "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)"
-| "tmsubst n t (Mul i a) = Mul i (tmsubst n t a)"
+  where
+    "tmsubst n t (CP c) = CP c"
+  | "tmsubst n t (Bound m) = (if n=m then t else Bound m)"
+  | "tmsubst n t (CNP m c a) =
+      (if n = m then Add (Mul c t) (tmsubst n t a) else CNP m c (tmsubst n t a))"
+  | "tmsubst n t (Neg a) = Neg (tmsubst n t a)"
+  | "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)"
+  | "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)"
+  | "tmsubst n t (Mul i a) = Mul i (tmsubst n t a)"
 
 lemma tmsubst:
   assumes nb: "tmboundslt (length bs) a"
@@ -264,26 +263,26 @@
 text \<open>Simplification.\<close>
 
 fun tmadd:: "tm \<Rightarrow> tm \<Rightarrow> tm"
-where
-  "tmadd (CNP n1 c1 r1) (CNP n2 c2 r2) =
-    (if n1 = n2 then
-      let c = c1 +\<^sub>p c2
-      in if c = 0\<^sub>p then tmadd r1 r2 else CNP n1 c (tmadd r1 r2)
-    else if n1 \<le> n2 then (CNP n1 c1 (tmadd r1 (CNP n2 c2 r2)))
-    else (CNP n2 c2 (tmadd (CNP n1 c1 r1) r2)))"
-| "tmadd (CNP n1 c1 r1) t = CNP n1 c1 (tmadd r1 t)"
-| "tmadd t (CNP n2 c2 r2) = CNP n2 c2 (tmadd t r2)"
-| "tmadd (CP b1) (CP b2) = CP (b1 +\<^sub>p b2)"
-| "tmadd a b = Add a b"
+  where
+    "tmadd (CNP n1 c1 r1) (CNP n2 c2 r2) =
+      (if n1 = n2 then
+        let c = c1 +\<^sub>p c2
+        in if c = 0\<^sub>p then tmadd r1 r2 else CNP n1 c (tmadd r1 r2)
+      else if n1 \<le> n2 then (CNP n1 c1 (tmadd r1 (CNP n2 c2 r2)))
+      else (CNP n2 c2 (tmadd (CNP n1 c1 r1) r2)))"
+  | "tmadd (CNP n1 c1 r1) t = CNP n1 c1 (tmadd r1 t)"
+  | "tmadd t (CNP n2 c2 r2) = CNP n2 c2 (tmadd t r2)"
+  | "tmadd (CP b1) (CP b2) = CP (b1 +\<^sub>p b2)"
+  | "tmadd a b = Add a b"
 
 lemma tmadd [simp]: "Itm vs bs (tmadd t s) = Itm vs bs (Add t s)"
   apply (induct t s rule: tmadd.induct)
-  apply (simp_all add: Let_def)
+                      apply (simp_all add: Let_def)
   apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p")
-  apply (case_tac "n1 \<le> n2")
-  apply simp_all
-  apply (case_tac "n1 = n2")
-  apply (simp_all add: algebra_simps)
+   apply (case_tac "n1 \<le> n2")
+    apply simp_all
+   apply (case_tac "n1 = n2")
+    apply (simp_all add: algebra_simps)
   apply (simp only: distrib_left [symmetric] polyadd [symmetric])
   apply simp
   done
@@ -302,10 +301,10 @@
   by (induct t s rule: tmadd.induct) (simp_all add: Let_def polyadd_norm)
 
 fun tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm"
-where
-  "tmmul (CP j) = (\<lambda>i. CP (i *\<^sub>p j))"
-| "tmmul (CNP n c a) = (\<lambda>i. CNP n (i *\<^sub>p c) (tmmul a i))"
-| "tmmul t = (\<lambda>i. Mul i t)"
+  where
+    "tmmul (CP j) = (\<lambda>i. CP (i *\<^sub>p j))"
+  | "tmmul (CNP n c a) = (\<lambda>i. CNP n (i *\<^sub>p c) (tmmul a i))"
+  | "tmmul t = (\<lambda>i. Mul i t)"
 
 lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)"
   by (induct t arbitrary: i rule: tmmul.induct) (simp_all add: field_simps)
@@ -343,12 +342,12 @@
   using tmneg_def by simp
 
 lemma [simp]: "isnpoly (C (-1, 1))"
-  unfolding isnpoly_def by simp
+  by (simp add: isnpoly_def)
 
 lemma tmneg_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)"
-  unfolding tmneg_def by auto
+  by (auto simp: tmneg_def)
 
 lemma tmsub[simp]: "Itm vs bs (tmsub a b) = Itm vs bs (Sub a b)"
   using tmsub_def by simp
@@ -365,19 +364,19 @@
 lemma tmsub_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)"
-  unfolding tmsub_def by (simp add: isnpoly_def)
+  by (simp add: tmsub_def isnpoly_def)
 
 fun simptm :: "tm \<Rightarrow> tm"
-where
-  "simptm (CP j) = CP (polynate j)"
-| "simptm (Bound n) = CNP n (1)\<^sub>p (CP 0\<^sub>p)"
-| "simptm (Neg t) = tmneg (simptm t)"
-| "simptm (Add t s) = tmadd (simptm t) (simptm s)"
-| "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
-| "simptm (Mul i t) =
-    (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
-| "simptm (CNP n c t) =
-    (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p)) (simptm t))"
+  where
+    "simptm (CP j) = CP (polynate j)"
+  | "simptm (Bound n) = CNP n (1)\<^sub>p (CP 0\<^sub>p)"
+  | "simptm (Neg t) = tmneg (simptm t)"
+  | "simptm (Add t s) = tmadd (simptm t) (simptm s)"
+  | "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
+  | "simptm (Mul i t) =
+      (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
+  | "simptm (CNP n c t) =
+      (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p)) (simptm t))"
 
 lemma polynate_stupid:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
@@ -410,15 +409,15 @@
 declare let_cong[fundef_cong del]
 
 fun split0 :: "tm \<Rightarrow> poly \<times> tm"
-where
-  "split0 (Bound 0) = ((1)\<^sub>p, CP 0\<^sub>p)"
-| "split0 (CNP 0 c t) = (let (c', t') = split0 t in (c +\<^sub>p c', t'))"
-| "split0 (Neg t) = (let (c, t') = split0 t in (~\<^sub>p c, Neg t'))"
-| "split0 (CNP n c t) = (let (c', t') = split0 t in (c', CNP n c t'))"
-| "split0 (Add s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))"
-| "split0 (Sub s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))"
-| "split0 (Mul c t) = (let (c', t') = split0 t in (c *\<^sub>p c', Mul c t'))"
-| "split0 t = (0\<^sub>p, t)"
+  where
+    "split0 (Bound 0) = ((1)\<^sub>p, CP 0\<^sub>p)"
+  | "split0 (CNP 0 c t) = (let (c', t') = split0 t in (c +\<^sub>p c', t'))"
+  | "split0 (Neg t) = (let (c, t') = split0 t in (~\<^sub>p c, Neg t'))"
+  | "split0 (CNP n c t) = (let (c', t') = split0 t in (c', CNP n c t'))"
+  | "split0 (Add s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))"
+  | "split0 (Sub s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))"
+  | "split0 (Mul c t) = (let (c', t') = split0 t in (c *\<^sub>p c', Mul c t'))"
+  | "split0 t = (0\<^sub>p, t)"
 
 declare let_cong[fundef_cong]
 
@@ -431,14 +430,14 @@
 lemma split0:
   "tmbound 0 (snd (split0 t)) \<and> Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t"
   apply (induct t rule: split0.induct)
-  apply simp
-  apply (simp add: Let_def split_def field_simps)
-  apply (simp add: Let_def split_def field_simps)
-  apply (simp add: Let_def split_def field_simps)
-  apply (simp add: Let_def split_def field_simps)
-  apply (simp add: Let_def split_def field_simps)
-  apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric])
-  apply (simp add: Let_def split_def field_simps)
+          apply simp
+         apply (simp add: Let_def split_def field_simps)
+        apply (simp add: Let_def split_def field_simps)
+       apply (simp add: Let_def split_def field_simps)
+      apply (simp add: Let_def split_def field_simps)
+     apply (simp add: Let_def split_def field_simps)
+    apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric])
+   apply (simp add: Let_def split_def field_simps)
   apply (simp add: Let_def split_def field_simps)
   done
 
@@ -446,10 +445,9 @@
 proof -
   fix c' t'
   assume "split0 t = (c', t')"
-  then have "c' = fst (split0 t)" and "t' = snd (split0 t)"
+  then have "c' = fst (split0 t)" "t' = snd (split0 t)"
     by auto
-  with split0[where t="t" and bs="bs"]
-  show "Itm vs bs t = Itm vs bs (CNP 0 c' t')"
+  with split0[where t="t" and bs="bs"] show "Itm vs bs t = Itm vs bs (CNP 0 c' t')"
     by simp
 qed
 
@@ -459,7 +457,7 @@
 proof -
   fix c' t'
   assume "split0 t = (c', t')"
-  then have "c' = fst (split0 t)" and "t' = snd (split0 t)"
+  then have "c' = fst (split0 t)" "t' = snd (split0 t)"
     by auto
   with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'"
     by simp
@@ -509,20 +507,20 @@
 begin
 
 primrec size_fm :: "fm \<Rightarrow> nat"
-where
-  "size_fm (NOT p) = 1 + size_fm p"
-| "size_fm (And p q) = 1 + size_fm p + size_fm q"
-| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
-| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
-| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
-| "size_fm (E p) = 1 + size_fm p"
-| "size_fm (A p) = 4 + size_fm p"
-| "size_fm T = 1"
-| "size_fm F = 1"
-| "size_fm (Le _) = 1"
-| "size_fm (Lt _) = 1"
-| "size_fm (Eq _) = 1"
-| "size_fm (NEq _) = 1"
+  where
+    "size_fm (NOT p) = 1 + size_fm p"
+  | "size_fm (And p q) = 1 + size_fm p + size_fm q"
+  | "size_fm (Or p q) = 1 + size_fm p + size_fm q"
+  | "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
+  | "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
+  | "size_fm (E p) = 1 + size_fm p"
+  | "size_fm (A p) = 4 + size_fm p"
+  | "size_fm T = 1"
+  | "size_fm F = 1"
+  | "size_fm (Le _) = 1"
+  | "size_fm (Lt _) = 1"
+  | "size_fm (Eq _) = 1"
+  | "size_fm (NEq _) = 1"
 
 instance ..
 
@@ -533,39 +531,38 @@
 
 text \<open>Semantics of formulae (fm).\<close>
 primrec Ifm ::"'a::linordered_field list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
-where
-  "Ifm vs bs T = True"
-| "Ifm vs bs F = False"
-| "Ifm vs bs (Lt a) = (Itm vs bs a < 0)"
-| "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
-| "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
-| "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
-| "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
-| "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
-| "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
-| "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
-| "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)"
-| "Ifm vs bs (E p) = (\<exists>x. Ifm vs (x#bs) p)"
-| "Ifm vs bs (A p) = (\<forall>x. Ifm vs (x#bs) p)"
+  where
+    "Ifm vs bs T = True"
+  | "Ifm vs bs F = False"
+  | "Ifm vs bs (Lt a) = (Itm vs bs a < 0)"
+  | "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
+  | "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
+  | "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
+  | "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
+  | "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
+  | "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
+  | "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
+  | "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)"
+  | "Ifm vs bs (E p) = (\<exists>x. Ifm vs (x#bs) p)"
+  | "Ifm vs bs (A p) = (\<forall>x. Ifm vs (x#bs) p)"
 
 fun not:: "fm \<Rightarrow> fm"
-where
-  "not (NOT (NOT p)) = not p"
-| "not (NOT p) = p"
-| "not T = F"
-| "not F = T"
-| "not (Lt t) = Le (tmneg t)"
-| "not (Le t) = Lt (tmneg t)"
-| "not (Eq t) = NEq t"
-| "not (NEq t) = Eq t"
-| "not p = NOT p"
+  where
+    "not (NOT (NOT p)) = not p"
+  | "not (NOT p) = p"
+  | "not T = F"
+  | "not F = T"
+  | "not (Lt t) = Le (tmneg t)"
+  | "not (Le t) = Lt (tmneg t)"
+  | "not (Eq t) = NEq t"
+  | "not (NEq t) = Eq t"
+  | "not p = NOT p"
 
 lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
   by (induct p rule: not.induct) auto
 
 definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "conj p q \<equiv>
+  where "conj p q \<equiv>
     (if p = F \<or> q = F then F
      else if p = T then q
      else if q = T then p
@@ -576,8 +573,7 @@
   by (cases "p=F \<or> q=F", simp_all add: conj_def) (cases p, simp_all)
 
 definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "disj p q \<equiv>
+  where "disj p q \<equiv>
     (if (p = T \<or> q = T) then T
      else if p = F then q
      else if q = F then p
@@ -588,8 +584,7 @@
   by (cases "p = T \<or> q = T", simp_all add: disj_def) (cases p, simp_all)
 
 definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "imp p q \<equiv>
+  where "imp p q \<equiv>
     (if p = F \<or> q = T \<or> p = q then T
      else if p = T then q
      else if q = F then not p
@@ -599,8 +594,7 @@
   by (cases "p = F \<or> q = T") (simp_all add: imp_def)
 
 definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "iff p q \<equiv>
+  where "iff p q \<equiv>
    (if p = q then T
     else if p = NOT q \<or> NOT p = q then F
     else if p = F then not q
@@ -614,47 +608,47 @@
 
 text \<open>Quantifier freeness.\<close>
 fun qfree:: "fm \<Rightarrow> bool"
-where
-  "qfree (E p) = False"
-| "qfree (A p) = False"
-| "qfree (NOT p) = qfree p"
-| "qfree (And p q) = (qfree p \<and> qfree q)"
-| "qfree (Or  p q) = (qfree p \<and> qfree q)"
-| "qfree (Imp p q) = (qfree p \<and> qfree q)"
-| "qfree (Iff p q) = (qfree p \<and> qfree q)"
-| "qfree p = True"
+  where
+    "qfree (E p) = False"
+  | "qfree (A p) = False"
+  | "qfree (NOT p) = qfree p"
+  | "qfree (And p q) = (qfree p \<and> qfree q)"
+  | "qfree (Or  p q) = (qfree p \<and> qfree q)"
+  | "qfree (Imp p q) = (qfree p \<and> qfree q)"
+  | "qfree (Iff p q) = (qfree p \<and> qfree q)"
+  | "qfree p = True"
 
 text \<open>Boundedness and substitution.\<close>
 primrec boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool"
-where
-  "boundslt n T = True"
-| "boundslt n F = True"
-| "boundslt n (Lt t) = tmboundslt n t"
-| "boundslt n (Le t) = tmboundslt n t"
-| "boundslt n (Eq t) = tmboundslt n t"
-| "boundslt n (NEq t) = tmboundslt n t"
-| "boundslt n (NOT p) = boundslt n p"
-| "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
-| "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
-| "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
-| "boundslt n (Iff p q) = (boundslt n p \<and> boundslt n q)"
-| "boundslt n (E p) = boundslt (Suc n) p"
-| "boundslt n (A p) = boundslt (Suc n) p"
-
-fun bound0:: "fm \<Rightarrow> bool"  \<comment> \<open>a Formula is independent of Bound 0\<close>
-where
-  "bound0 T = True"
-| "bound0 F = True"
-| "bound0 (Lt a) = tmbound0 a"
-| "bound0 (Le a) = tmbound0 a"
-| "bound0 (Eq a) = tmbound0 a"
-| "bound0 (NEq a) = tmbound0 a"
-| "bound0 (NOT p) = bound0 p"
-| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
-| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
-| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
-| "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
-| "bound0 p = False"
+  where
+    "boundslt n T = True"
+  | "boundslt n F = True"
+  | "boundslt n (Lt t) = tmboundslt n t"
+  | "boundslt n (Le t) = tmboundslt n t"
+  | "boundslt n (Eq t) = tmboundslt n t"
+  | "boundslt n (NEq t) = tmboundslt n t"
+  | "boundslt n (NOT p) = boundslt n p"
+  | "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
+  | "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
+  | "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
+  | "boundslt n (Iff p q) = (boundslt n p \<and> boundslt n q)"
+  | "boundslt n (E p) = boundslt (Suc n) p"
+  | "boundslt n (A p) = boundslt (Suc n) p"
+
+fun bound0:: "fm \<Rightarrow> bool"  \<comment> \<open>a formula is independent of Bound 0\<close>
+  where
+    "bound0 T = True"
+  | "bound0 F = True"
+  | "bound0 (Lt a) = tmbound0 a"
+  | "bound0 (Le a) = tmbound0 a"
+  | "bound0 (Eq a) = tmbound0 a"
+  | "bound0 (NEq a) = tmbound0 a"
+  | "bound0 (NOT p) = bound0 p"
+  | "bound0 (And p q) = (bound0 p \<and> bound0 q)"
+  | "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
+  | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
+  | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
+  | "bound0 p = False"
 
 lemma bound0_I:
   assumes bp: "bound0 p"
@@ -662,21 +656,21 @@
   using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"]
   by (induct p rule: bound0.induct) auto
 
-primrec bound:: "nat \<Rightarrow> fm \<Rightarrow> bool"  \<comment> \<open>a Formula is independent of Bound n\<close>
-where
-  "bound m T = True"
-| "bound m F = True"
-| "bound m (Lt t) = tmbound m t"
-| "bound m (Le t) = tmbound m t"
-| "bound m (Eq t) = tmbound m t"
-| "bound m (NEq t) = tmbound m t"
-| "bound m (NOT p) = bound m p"
-| "bound m (And p q) = (bound m p \<and> bound m q)"
-| "bound m (Or p q) = (bound m p \<and> bound m q)"
-| "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
-| "bound m (Iff p q) = (bound m p \<and> bound m q)"
-| "bound m (E p) = bound (Suc m) p"
-| "bound m (A p) = bound (Suc m) p"
+primrec bound:: "nat \<Rightarrow> fm \<Rightarrow> bool"  \<comment> \<open>a formula is independent of Bound n\<close>
+  where
+    "bound m T = True"
+  | "bound m F = True"
+  | "bound m (Lt t) = tmbound m t"
+  | "bound m (Le t) = tmbound m t"
+  | "bound m (Eq t) = tmbound m t"
+  | "bound m (NEq t) = tmbound m t"
+  | "bound m (NOT p) = bound m p"
+  | "bound m (And p q) = (bound m p \<and> bound m q)"
+  | "bound m (Or p q) = (bound m p \<and> bound m q)"
+  | "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
+  | "bound m (Iff p q) = (bound m p \<and> bound m q)"
+  | "bound m (E p) = bound (Suc m) p"
+  | "bound m (A p) = bound (Suc m) p"
 
 lemma bound_I:
   assumes bnd: "boundslt (length bs) p"
@@ -707,39 +701,38 @@
 qed auto
 
 fun decr0 :: "fm \<Rightarrow> fm"
-where
-  "decr0 (Lt a) = Lt (decrtm0 a)"
-| "decr0 (Le a) = Le (decrtm0 a)"
-| "decr0 (Eq a) = Eq (decrtm0 a)"
-| "decr0 (NEq a) = NEq (decrtm0 a)"
-| "decr0 (NOT p) = NOT (decr0 p)"
-| "decr0 (And p q) = conj (decr0 p) (decr0 q)"
-| "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
-| "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
-| "decr0 (Iff p q) = iff (decr0 p) (decr0 q)"
-| "decr0 p = p"
+  where
+    "decr0 (Lt a) = Lt (decrtm0 a)"
+  | "decr0 (Le a) = Le (decrtm0 a)"
+  | "decr0 (Eq a) = Eq (decrtm0 a)"
+  | "decr0 (NEq a) = NEq (decrtm0 a)"
+  | "decr0 (NOT p) = NOT (decr0 p)"
+  | "decr0 (And p q) = conj (decr0 p) (decr0 q)"
+  | "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
+  | "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
+  | "decr0 (Iff p q) = iff (decr0 p) (decr0 q)"
+  | "decr0 p = p"
 
 lemma decr0:
-  assumes nb: "bound0 p"
+  assumes "bound0 p"
   shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)"
-  using nb
-  by (induct p rule: decr0.induct) (simp_all add: decrtm0)
+  using assms by (induct p rule: decr0.induct) (simp_all add: decrtm0)
 
 primrec decr :: "nat \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "decr m T = T"
-| "decr m F = F"
-| "decr m (Lt t) = (Lt (decrtm m t))"
-| "decr m (Le t) = (Le (decrtm m t))"
-| "decr m (Eq t) = (Eq (decrtm m t))"
-| "decr m (NEq t) = (NEq (decrtm m t))"
-| "decr m (NOT p) = NOT (decr m p)"
-| "decr m (And p q) = conj (decr m p) (decr m q)"
-| "decr m (Or p q) = disj (decr m p) (decr m q)"
-| "decr m (Imp p q) = imp (decr m p) (decr m q)"
-| "decr m (Iff p q) = iff (decr m p) (decr m q)"
-| "decr m (E p) = E (decr (Suc m) p)"
-| "decr m (A p) = A (decr (Suc m) p)"
+  where
+    "decr m T = T"
+  | "decr m F = F"
+  | "decr m (Lt t) = (Lt (decrtm m t))"
+  | "decr m (Le t) = (Le (decrtm m t))"
+  | "decr m (Eq t) = (Eq (decrtm m t))"
+  | "decr m (NEq t) = (NEq (decrtm m t))"
+  | "decr m (NOT p) = NOT (decr m p)"
+  | "decr m (And p q) = conj (decr m p) (decr m q)"
+  | "decr m (Or p q) = disj (decr m p) (decr m q)"
+  | "decr m (Imp p q) = imp (decr m p) (decr m q)"
+  | "decr m (Iff p q) = iff (decr m p) (decr m q)"
+  | "decr m (E p) = E (decr (Suc m) p)"
+  | "decr m (A p) = A (decr (Suc m) p)"
 
 lemma decr:
   assumes bnd: "boundslt (length bs) p"
@@ -774,20 +767,20 @@
 qed (auto simp add: decrtm removen_nth)
 
 primrec subst0 :: "tm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "subst0 t T = T"
-| "subst0 t F = F"
-| "subst0 t (Lt a) = Lt (tmsubst0 t a)"
-| "subst0 t (Le a) = Le (tmsubst0 t a)"
-| "subst0 t (Eq a) = Eq (tmsubst0 t a)"
-| "subst0 t (NEq a) = NEq (tmsubst0 t a)"
-| "subst0 t (NOT p) = NOT (subst0 t p)"
-| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
-| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
-| "subst0 t (Imp p q) = Imp (subst0 t p)  (subst0 t q)"
-| "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
-| "subst0 t (E p) = E p"
-| "subst0 t (A p) = A p"
+  where
+    "subst0 t T = T"
+  | "subst0 t F = F"
+  | "subst0 t (Lt a) = Lt (tmsubst0 t a)"
+  | "subst0 t (Le a) = Le (tmsubst0 t a)"
+  | "subst0 t (Eq a) = Eq (tmsubst0 t a)"
+  | "subst0 t (NEq a) = NEq (tmsubst0 t a)"
+  | "subst0 t (NOT p) = NOT (subst0 t p)"
+  | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
+  | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
+  | "subst0 t (Imp p q) = Imp (subst0 t p)  (subst0 t q)"
+  | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
+  | "subst0 t (E p) = E p"
+  | "subst0 t (A p) = A p"
 
 lemma subst0:
   assumes qf: "qfree p"
@@ -799,24 +792,23 @@
   assumes bp: "tmbound0 t"
     and qf: "qfree p"
   shows "bound0 (subst0 t p)"
-  using qf tmsubst0_nb[OF bp] bp
-  by (induct p rule: fm.induct) auto
+  using qf tmsubst0_nb[OF bp] bp by (induct p rule: fm.induct) auto
 
 primrec subst:: "nat \<Rightarrow> tm \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "subst n t T = T"
-| "subst n t F = F"
-| "subst n t (Lt a) = Lt (tmsubst n t a)"
-| "subst n t (Le a) = Le (tmsubst n t a)"
-| "subst n t (Eq a) = Eq (tmsubst n t a)"
-| "subst n t (NEq a) = NEq (tmsubst n t a)"
-| "subst n t (NOT p) = NOT (subst n t p)"
-| "subst n t (And p q) = And (subst n t p) (subst n t q)"
-| "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
-| "subst n t (Imp p q) = Imp (subst n t p)  (subst n t q)"
-| "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)"
-| "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)"
-| "subst n t (A p) = A (subst (Suc n) (incrtm0 t) p)"
+  where
+    "subst n t T = T"
+  | "subst n t F = F"
+  | "subst n t (Lt a) = Lt (tmsubst n t a)"
+  | "subst n t (Le a) = Le (tmsubst n t a)"
+  | "subst n t (Eq a) = Eq (tmsubst n t a)"
+  | "subst n t (NEq a) = NEq (tmsubst n t a)"
+  | "subst n t (NOT p) = NOT (subst n t p)"
+  | "subst n t (And p q) = And (subst n t p) (subst n t q)"
+  | "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
+  | "subst n t (Imp p q) = Imp (subst n t p)  (subst n t q)"
+  | "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)"
+  | "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)"
+  | "subst n t (A p) = A (subst (Suc n) (incrtm0 t) p)"
 
 lemma subst:
   assumes nb: "boundslt (length bs) p"
@@ -860,10 +852,9 @@
 qed (auto simp add: tmsubst)
 
 lemma subst_nb:
-  assumes tnb: "tmbound m t"
+  assumes "tmbound m t"
   shows "bound m (subst m t p)"
-  using tnb tmsubst_nb incrtm0_tmbound
-  by (induct p arbitrary: m t rule: fm.induct) auto
+  using assms tmsubst_nb incrtm0_tmbound by (induct p arbitrary: m t rule: fm.induct) auto
 
 lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
   by (induct p rule: not.induct) auto
@@ -913,21 +904,20 @@
   by (induct p) simp_all
 
 fun isatom :: "fm \<Rightarrow> bool"  \<comment> \<open>test for atomicity\<close>
-where
-  "isatom T = True"
-| "isatom F = True"
-| "isatom (Lt a) = True"
-| "isatom (Le a) = True"
-| "isatom (Eq a) = True"
-| "isatom (NEq a) = True"
-| "isatom p = False"
+  where
+    "isatom T = True"
+  | "isatom F = True"
+  | "isatom (Lt a) = True"
+  | "isatom (Le a) = True"
+  | "isatom (Eq a) = True"
+  | "isatom (NEq a) = True"
+  | "isatom p = False"
 
 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
   by (induct p) simp_all
 
 definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "djf f p q \<equiv>
+  where "djf f p q \<equiv>
     (if q = T then T
      else if q = F then f p
      else (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
@@ -937,60 +927,63 @@
 
 lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
   apply (cases "q = T")
-  apply (simp add: djf_def)
+   apply (simp add: djf_def)
   apply (cases "q = F")
-  apply (simp add: djf_def)
+   apply (simp add: djf_def)
   apply (cases "f p")
-  apply (simp_all add: Let_def djf_def)
+              apply (simp_all add: Let_def djf_def)
   done
 
 lemma evaldjf_ex: "Ifm vs bs (evaldjf f ps) \<longleftrightarrow> (\<exists>p \<in> set ps. Ifm vs bs (f p))"
   by (induct ps) (simp_all add: evaldjf_def djf_Or)
 
 lemma evaldjf_bound0:
-  assumes nb: "\<forall>x\<in> set xs. bound0 (f x)"
+  assumes "\<forall>x\<in> set xs. bound0 (f x)"
   shows "bound0 (evaldjf f xs)"
-  using nb
+  using assms
   apply (induct xs)
-  apply (auto simp add: evaldjf_def djf_def Let_def)
+   apply (auto simp add: evaldjf_def djf_def Let_def)
   apply (case_tac "f a")
-  apply auto
+              apply auto
   done
 
 lemma evaldjf_qf:
-  assumes nb: "\<forall>x\<in> set xs. qfree (f x)"
+  assumes "\<forall>x\<in> set xs. qfree (f x)"
   shows "qfree (evaldjf f xs)"
-  using nb
+  using assms
   apply (induct xs)
-  apply (auto simp add: evaldjf_def djf_def Let_def)
+   apply (auto simp add: evaldjf_def djf_def Let_def)
   apply (case_tac "f a")
-  apply auto
+              apply auto
   done
 
 fun disjuncts :: "fm \<Rightarrow> fm list"
-where
-  "disjuncts (Or p q) = disjuncts p @ disjuncts q"
-| "disjuncts F = []"
-| "disjuncts p = [p]"
+  where
+    "disjuncts (Or p q) = disjuncts p @ disjuncts q"
+  | "disjuncts F = []"
+  | "disjuncts p = [p]"
 
 lemma disjuncts: "(\<exists>q \<in> set (disjuncts p). Ifm vs bs q) = Ifm vs bs p"
   by (induct p rule: disjuncts.induct) auto
 
-lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall>q \<in> set (disjuncts p). bound0 q"
+lemma disjuncts_nb:
+  assumes "bound0 p"
+  shows "\<forall>q \<in> set (disjuncts p). bound0 q"
 proof -
-  assume nb: "bound0 p"
-  then have "list_all bound0 (disjuncts p)"
-    by (induct p rule:disjuncts.induct) auto
+  from assms have "list_all bound0 (disjuncts p)"
+    by (induct p rule: disjuncts.induct) auto
   then show ?thesis
     by (simp only: list_all_iff)
 qed
 
-lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall>q \<in> set (disjuncts p). qfree q"
+lemma disjuncts_qf:
+  assumes "qfree p"
+  shows "\<forall>q \<in> set (disjuncts p). qfree q"
 proof -
-  assume qf: "qfree p"
-  then have "list_all qfree (disjuncts p)"
+  from assms have "list_all qfree (disjuncts p)"
     by (induct p rule: disjuncts.induct) auto
-  then show ?thesis by (simp only: list_all_iff)
+  then show ?thesis
+    by (simp only: list_all_iff)
 qed
 
 definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
@@ -1044,17 +1037,16 @@
 qed
 
 fun conjuncts :: "fm \<Rightarrow> fm list"
-where
-  "conjuncts (And p q) = conjuncts p @ conjuncts q"
-| "conjuncts T = []"
-| "conjuncts p = [p]"
+  where
+    "conjuncts (And p q) = conjuncts p @ conjuncts q"
+  | "conjuncts T = []"
+  | "conjuncts p = [p]"
 
 definition list_conj :: "fm list \<Rightarrow> fm"
   where "list_conj ps \<equiv> foldr conj ps T"
 
 definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
-where
-  "CJNB f p \<equiv>
+  where "CJNB f p \<equiv>
     (let cjs = conjuncts p;
       (yes, no) = partition bound0 cjs
      in conj (decr0 (list_conj yes)) (f (list_conj no)))"
@@ -1071,27 +1063,28 @@
 lemma conjuncts: "(\<forall>q\<in> set (conjuncts p). Ifm vs bs q) = Ifm vs bs p"
   by (induct p rule: conjuncts.induct) auto
 
-lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall>q\<in> set (conjuncts p). bound0 q"
+lemma conjuncts_nb:
+  assumes "bound0 p"
+  shows "\<forall>q \<in> set (conjuncts p). bound0 q"
 proof -
-  assume nb: "bound0 p"
-  then have "list_all bound0 (conjuncts p)"
+  from assms have "list_all bound0 (conjuncts p)"
     by (induct p rule:conjuncts.induct) auto
   then show ?thesis
     by (simp only: list_all_iff)
 qed
 
 fun islin :: "fm \<Rightarrow> bool"
-where
-  "islin (And p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
-| "islin (Or p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
-| "islin (Eq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
-| "islin (NEq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
-| "islin (Lt (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
-| "islin (Le (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
-| "islin (NOT p) = False"
-| "islin (Imp p q) = False"
-| "islin (Iff p q) = False"
-| "islin p = bound0 p"
+  where
+    "islin (And p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
+  | "islin (Or p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
+  | "islin (Eq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+  | "islin (NEq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+  | "islin (Lt (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+  | "islin (Le (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+  | "islin (NOT p) = False"
+  | "islin (Imp p q) = False"
+  | "islin (Iff p q) = False"
+  | "islin p = bound0 p"
 
 lemma islin_stupid:
   assumes nb: "tmbound0 p"
@@ -1109,25 +1102,25 @@
 lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
   apply (simp add: lt_def)
   apply (cases p)
-  apply simp_all
+        apply simp_all
   apply (rename_tac poly, case_tac poly)
-  apply (simp_all add: isnpoly_def)
+         apply (simp_all add: isnpoly_def)
   done
 
 lemma le: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (le p) = Ifm vs bs (Le p)"
   apply (simp add: le_def)
   apply (cases p)
-  apply simp_all
+        apply simp_all
   apply (rename_tac poly, case_tac poly)
-  apply (simp_all add: isnpoly_def)
+         apply (simp_all add: isnpoly_def)
   done
 
 lemma eq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (eq p) = Ifm vs bs (Eq p)"
   apply (simp add: eq_def)
   apply (cases p)
-  apply simp_all
+        apply simp_all
   apply (rename_tac poly, case_tac poly)
-  apply (simp_all add: isnpoly_def)
+         apply (simp_all add: isnpoly_def)
   done
 
 lemma neq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (neq p) = Ifm vs bs (NEq p)"
@@ -1136,41 +1129,41 @@
 lemma lt_lin: "tmbound0 p \<Longrightarrow> islin (lt p)"
   apply (simp add: lt_def)
   apply (cases p)
-  apply simp_all
-  apply (rename_tac poly, case_tac poly)
-  apply simp_all
+        apply simp_all
+   apply (rename_tac poly, case_tac poly)
+          apply simp_all
   apply (rename_tac nat a b, case_tac nat)
-  apply simp_all
+   apply simp_all
   done
 
 lemma le_lin: "tmbound0 p \<Longrightarrow> islin (le p)"
   apply (simp add: le_def)
   apply (cases p)
-  apply simp_all
-  apply (rename_tac poly, case_tac poly)
-  apply simp_all
+        apply simp_all
+   apply (rename_tac poly, case_tac poly)
+          apply simp_all
   apply (rename_tac nat a b, case_tac nat)
-  apply simp_all
+   apply simp_all
   done
 
 lemma eq_lin: "tmbound0 p \<Longrightarrow> islin (eq p)"
   apply (simp add: eq_def)
   apply (cases p)
-  apply simp_all
-  apply (rename_tac poly, case_tac poly)
-  apply simp_all
+        apply simp_all
+   apply (rename_tac poly, case_tac poly)
+          apply simp_all
   apply (rename_tac nat a b, case_tac nat)
-  apply simp_all
+   apply simp_all
   done
 
 lemma neq_lin: "tmbound0 p \<Longrightarrow> islin (neq p)"
   apply (simp add: neq_def eq_def)
   apply (cases p)
-  apply simp_all
-  apply (rename_tac poly, case_tac poly)
-  apply simp_all
+        apply simp_all
+   apply (rename_tac poly, case_tac poly)
+          apply simp_all
   apply (rename_tac nat a b, case_tac nat)
-  apply simp_all
+   apply simp_all
   done
 
 definition "simplt t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then lt s else Lt (CNP 0 c s))"
@@ -1178,7 +1171,7 @@
 definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))"
 definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))"
 
-lemma simplt_islin[simp]:
+lemma simplt_islin [simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "islin (simplt t)"
   unfolding simplt_def
@@ -1186,7 +1179,7 @@
   by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
 
-lemma simple_islin[simp]:
+lemma simple_islin [simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "islin (simple t)"
   unfolding simple_def
@@ -1194,7 +1187,7 @@
   by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
 
-lemma simpeq_islin[simp]:
+lemma simpeq_islin [simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "islin (simpeq t)"
   unfolding simpeq_def
@@ -1202,7 +1195,7 @@
   by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
       islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
 
-lemma simpneq_islin[simp]:
+lemma simpneq_islin [simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "islin (simpneq t)"
   unfolding simpneq_def
@@ -1215,24 +1208,24 @@
 
 lemma split0_npoly:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
-    and n: "allpolys isnpoly t"
+    and *: "allpolys isnpoly t"
   shows "isnpoly (fst (split0 t))"
     and "allpolys isnpoly (snd (split0 t))"
-  using n
+  using *
   by (induct t rule: split0.induct)
     (auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm
       polysub_norm really_stupid)
 
 lemma simplt[simp]: "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)"
 proof -
-  have n: "allpolys isnpoly (simptm t)"
+  have *: "allpolys isnpoly (simptm t)"
     by simp
   let ?t = "simptm t"
   show ?thesis
   proof (cases "fst (split0 ?t) = 0\<^sub>p")
     case True
     then show ?thesis
-      using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF n], of vs bs]
+      using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF *], of vs bs]
       by (simp add: simplt_def Let_def split_def lt)
   next
     case False
@@ -1244,14 +1237,14 @@
 
 lemma simple[simp]: "Ifm vs bs (simple t) = Ifm vs bs (Le t)"
 proof -
-  have n: "allpolys isnpoly (simptm t)"
+  have *: "allpolys isnpoly (simptm t)"
     by simp
   let ?t = "simptm t"
   show ?thesis
   proof (cases "fst (split0 ?t) = 0\<^sub>p")
     case True
     then show ?thesis
-      using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF n], of vs bs]
+      using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF *], of vs bs]
       by (simp add: simple_def Let_def split_def le)
   next
     case False
@@ -1300,44 +1293,44 @@
 lemma lt_nb: "tmbound0 t \<Longrightarrow> bound0 (lt t)"
   apply (simp add: lt_def)
   apply (cases t)
-  apply auto
+        apply auto
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma le_nb: "tmbound0 t \<Longrightarrow> bound0 (le t)"
   apply (simp add: le_def)
   apply (cases t)
-  apply auto
+        apply auto
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma eq_nb: "tmbound0 t \<Longrightarrow> bound0 (eq t)"
   apply (simp add: eq_def)
   apply (cases t)
-  apply auto
+        apply auto
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma neq_nb: "tmbound0 t \<Longrightarrow> bound0 (neq t)"
   apply (simp add: neq_def eq_def)
   apply (cases t)
-  apply auto
+        apply auto
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma simplt_nb[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
 proof (simp add: simplt_def Let_def split_def)
-  assume nb: "tmbound0 t"
-  then have nb': "tmbound0 (simptm t)"
+  assume "tmbound0 t"
+  then have *: "tmbound0 (simptm t)"
     by simp
   let ?c = "fst (split0 (simptm t))"
-  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
     by auto
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
@@ -1354,11 +1347,11 @@
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
 proof(simp add: simple_def Let_def split_def)
-  assume nb: "tmbound0 t"
-  then have nb': "tmbound0 (simptm t)"
+  assume "tmbound0 t"
+  then have *: "tmbound0 (simptm t)"
     by simp
   let ?c = "fst (split0 (simptm t))"
-  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
     by auto
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
@@ -1375,11 +1368,11 @@
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
 proof (simp add: simpeq_def Let_def split_def)
-  assume nb: "tmbound0 t"
-  then have nb': "tmbound0 (simptm t)"
+  assume "tmbound0 t"
+  then have *: "tmbound0 (simptm t)"
     by simp
   let ?c = "fst (split0 (simptm t))"
-  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
     by auto
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
@@ -1396,11 +1389,11 @@
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
 proof (simp add: simpneq_def Let_def split_def)
-  assume nb: "tmbound0 t"
-  then have nb': "tmbound0 (simptm t)"
+  assume "tmbound0 t"
+  then have *: "tmbound0 (simptm t)"
     by simp
   let ?c = "fst (split0 (simptm t))"
-  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
   have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
     by auto
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
@@ -1414,10 +1407,10 @@
 qed
 
 fun conjs :: "fm \<Rightarrow> fm list"
-where
-  "conjs (And p q) = conjs p @ conjs q"
-| "conjs T = []"
-| "conjs p = [p]"
+  where
+    "conjs (And p q) = conjs p @ conjs q"
+  | "conjs T = []"
+  | "conjs p = [p]"
 
 lemma conjs_ci: "(\<forall>q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
   by (induct p rule: conjs.induct) auto
@@ -1439,36 +1432,33 @@
 
 lemma conjs_nb: "bound n p \<Longrightarrow> \<forall>q\<in> set (conjs p). bound n q"
   apply (induct p rule: conjs.induct)
-  apply (unfold conjs.simps)
-  apply (unfold set_append)
-  apply (unfold ball_Un)
-  apply (unfold bound.simps)
-  apply auto
+              apply (unfold conjs.simps)
+              apply (unfold set_append)
+              apply (unfold ball_Un)
+              apply (unfold bound.simps)
+              apply auto
   done
 
 lemma conjs_boundslt: "boundslt n p \<Longrightarrow> \<forall>q\<in> set (conjs p). boundslt n q"
   apply (induct p rule: conjs.induct)
-  apply (unfold conjs.simps)
-  apply (unfold set_append)
-  apply (unfold ball_Un)
-  apply (unfold boundslt.simps)
-  apply blast
-  apply simp_all
+              apply (unfold conjs.simps)
+              apply (unfold set_append)
+              apply (unfold ball_Un)
+              apply (unfold boundslt.simps)
+              apply blast
+             apply simp_all
   done
 
 lemma list_conj_boundslt: " \<forall>p\<in> set ps. boundslt n p \<Longrightarrow> boundslt n (list_conj ps)"
-  unfolding list_conj_def
-  by (induct ps) auto
+  by (induct ps) (auto simp: list_conj_def)
 
 lemma list_conj_nb:
-  assumes bnd: "\<forall>p\<in> set ps. bound n p"
+  assumes "\<forall>p\<in> set ps. bound n p"
   shows "bound n (list_conj ps)"
-  using bnd
-  unfolding list_conj_def
-  by (induct ps) auto
+  using assms by (induct ps) (auto simp: list_conj_def)
 
 lemma list_conj_nb': "\<forall>p\<in>set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
-  unfolding list_conj_def by (induct ps) auto
+  by (induct ps) (auto simp: list_conj_def)
 
 lemma CJNB_qe:
   assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
@@ -1523,29 +1513,29 @@
 qed
 
 fun simpfm :: "fm \<Rightarrow> fm"
-where
-  "simpfm (Lt t) = simplt (simptm t)"
-| "simpfm (Le t) = simple (simptm t)"
-| "simpfm (Eq t) = simpeq(simptm t)"
-| "simpfm (NEq t) = simpneq(simptm t)"
-| "simpfm (And p q) = conj (simpfm p) (simpfm q)"
-| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
-| "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
-| "simpfm (Iff p q) =
-    disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
-| "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
-| "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
-| "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
-| "simpfm (NOT (Iff p q)) =
-    disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
-| "simpfm (NOT (Eq t)) = simpneq t"
-| "simpfm (NOT (NEq t)) = simpeq t"
-| "simpfm (NOT (Le t)) = simplt (Neg t)"
-| "simpfm (NOT (Lt t)) = simple (Neg t)"
-| "simpfm (NOT (NOT p)) = simpfm p"
-| "simpfm (NOT T) = F"
-| "simpfm (NOT F) = T"
-| "simpfm p = p"
+  where
+    "simpfm (Lt t) = simplt (simptm t)"
+  | "simpfm (Le t) = simple (simptm t)"
+  | "simpfm (Eq t) = simpeq(simptm t)"
+  | "simpfm (NEq t) = simpneq(simptm t)"
+  | "simpfm (And p q) = conj (simpfm p) (simpfm q)"
+  | "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
+  | "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
+  | "simpfm (Iff p q) =
+      disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
+  | "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
+  | "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
+  | "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
+  | "simpfm (NOT (Iff p q)) =
+      disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
+  | "simpfm (NOT (Eq t)) = simpneq t"
+  | "simpfm (NOT (NEq t)) = simpeq t"
+  | "simpfm (NOT (Le t)) = simplt (Neg t)"
+  | "simpfm (NOT (Lt t)) = simple (Neg t)"
+  | "simpfm (NOT (NOT p)) = simpfm p"
+  | "simpfm (NOT T) = F"
+  | "simpfm (NOT F) = T"
+  | "simpfm p = p"
 
 lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
   by (induct p arbitrary: bs rule: simpfm.induct) auto
@@ -1557,23 +1547,23 @@
 
 lemma lt_qf[simp]: "qfree (lt t)"
   apply (cases t)
-  apply (auto simp add: lt_def)
+        apply (auto simp add: lt_def)
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma le_qf[simp]: "qfree (le t)"
   apply (cases t)
-  apply (auto simp add: le_def)
+        apply (auto simp add: le_def)
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma eq_qf[simp]: "qfree (eq t)"
   apply (cases t)
-  apply (auto simp add: eq_def)
+        apply (auto simp add: eq_def)
   apply (rename_tac poly, case_tac poly)
-  apply auto
+         apply auto
   done
 
 lemma neq_qf[simp]: "qfree (neq t)"
@@ -1596,6 +1586,7 @@
 
 lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)"
   by (simp add: disj_def)
+
 lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)"
   by (simp add: conj_def)
 
@@ -1605,30 +1596,30 @@
   by (induct p rule: simpfm.induct) (simp_all add: conj_lin disj_lin)
 
 fun prep :: "fm \<Rightarrow> fm"
-where
-  "prep (E T) = T"
-| "prep (E F) = F"
-| "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
-| "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
-| "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
-| "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
-| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
-| "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
-| "prep (E p) = E (prep p)"
-| "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
-| "prep (A p) = prep (NOT (E (NOT p)))"
-| "prep (NOT (NOT p)) = prep p"
-| "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (A p)) = prep (E (NOT p))"
-| "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
-| "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
-| "prep (NOT p) = not (prep p)"
-| "prep (Or p q) = disj (prep p) (prep q)"
-| "prep (And p q) = conj (prep p) (prep q)"
-| "prep (Imp p q) = prep (Or (NOT p) q)"
-| "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
-| "prep p = p"
+  where
+    "prep (E T) = T"
+  | "prep (E F) = F"
+  | "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
+  | "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
+  | "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
+  | "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
+  | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
+  | "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+  | "prep (E p) = E (prep p)"
+  | "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
+  | "prep (A p) = prep (NOT (E (NOT p)))"
+  | "prep (NOT (NOT p)) = prep p"
+  | "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
+  | "prep (NOT (A p)) = prep (E (NOT p))"
+  | "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
+  | "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
+  | "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
+  | "prep (NOT p) = not (prep p)"
+  | "prep (Or p q) = disj (prep p) (prep q)"
+  | "prep (And p q) = conj (prep p) (prep q)"
+  | "prep (Imp p q) = prep (Or (NOT p) q)"
+  | "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+  | "prep p = p"
 
 lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
   by (induct p arbitrary: bs rule: prep.induct) auto
@@ -1636,15 +1627,15 @@
 
 text \<open>Generic quantifier elimination.\<close>
 fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
-where
-  "qelim (E p) = (\<lambda>qe. DJ (CJNB qe) (qelim p qe))"
-| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
-| "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
-| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
-| "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
-| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
-| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
-| "qelim p = (\<lambda>y. simpfm p)"
+  where
+    "qelim (E p) = (\<lambda>qe. DJ (CJNB qe) (qelim p qe))"
+  | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
+  | "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
+  | "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
+  | "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
+  | "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
+  | "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
+  | "qelim p = (\<lambda>y. simpfm p)"
 
 lemma qelim:
   assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
@@ -1655,30 +1646,30 @@
 
 subsection \<open>Core Procedure\<close>
 
-fun minusinf:: "fm \<Rightarrow> fm"  \<comment> \<open>Virtual substitution of -\<infinity>\<close>
-where
-  "minusinf (And p q) = conj (minusinf p) (minusinf q)"
-| "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
-| "minusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
-| "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
-| "minusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))"
-| "minusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))"
-| "minusinf p = p"
-
-fun plusinf:: "fm \<Rightarrow> fm"  \<comment> \<open>Virtual substitution of +\<infinity>\<close>
-where
-  "plusinf (And p q) = conj (plusinf p) (plusinf q)"
-| "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
-| "plusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
-| "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
-| "plusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))"
-| "plusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))"
-| "plusinf p = p"
+fun minusinf:: "fm \<Rightarrow> fm"  \<comment> \<open>virtual substitution of \<open>-\<infinity>\<close>\<close>
+  where
+    "minusinf (And p q) = conj (minusinf p) (minusinf q)"
+  | "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
+  | "minusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
+  | "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
+  | "minusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))"
+  | "minusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))"
+  | "minusinf p = p"
+
+fun plusinf:: "fm \<Rightarrow> fm"  \<comment> \<open>virtual substitution of \<open>+\<infinity>\<close>\<close>
+  where
+    "plusinf (And p q) = conj (plusinf p) (plusinf q)"
+  | "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
+  | "plusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
+  | "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
+  | "plusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))"
+  | "plusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))"
+  | "plusinf p = p"
 
 lemma minusinf_inf:
-  assumes lp: "islin p"
+  assumes "islin p"
   shows "\<exists>z. \<forall>x < z. Ifm vs (x#bs) (minusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
-  using lp
+  using assms
 proof (induct p rule: minusinf.induct)
   case 1
   then show ?case
@@ -1881,9 +1872,9 @@
 qed auto
 
 lemma plusinf_inf:
-  assumes lp: "islin p"
+  assumes "islin p"
   shows "\<exists>z. \<forall>x > z. Ifm vs (x#bs) (plusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
-  using lp
+  using assms
 proof (induct p rule: plusinf.induct)
   case 1
   then show ?case
@@ -2125,14 +2116,14 @@
 qed
 
 fun uset :: "fm \<Rightarrow> (poly \<times> tm) list"
-where
-  "uset (And p q) = uset p @ uset q"
-| "uset (Or p q) = uset p @ uset q"
-| "uset (Eq (CNP 0 a e)) = [(a, e)]"
-| "uset (Le (CNP 0 a e)) = [(a, e)]"
-| "uset (Lt (CNP 0 a e)) = [(a, e)]"
-| "uset (NEq (CNP 0 a e)) = [(a, e)]"
-| "uset p = []"
+  where
+    "uset (And p q) = uset p @ uset q"
+  | "uset (Or p q) = uset p @ uset q"
+  | "uset (Eq (CNP 0 a e)) = [(a, e)]"
+  | "uset (Le (CNP 0 a e)) = [(a, e)]"
+  | "uset (Lt (CNP 0 a e)) = [(a, e)]"
+  | "uset (NEq (CNP 0 a e)) = [(a, e)]"
+  | "uset p = []"
 
 lemma uset_l:
   assumes lp: "islin p"
@@ -2150,8 +2141,8 @@
       Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s"
     using lp nmi ex
     apply (induct p rule: minusinf.induct)
-    apply (auto simp add: eq le lt polyneg_norm)
-    apply (auto simp add: linorder_not_less order_le_less)
+                        apply (auto simp add: eq le lt polyneg_norm)
+      apply (auto simp add: linorder_not_less order_le_less)
     done
   then obtain c s where csU: "(c, s) \<in> set (uset p)"
     and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>
@@ -2193,18 +2184,19 @@
       Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
     using lp nmi ex
     apply (induct p rule: minusinf.induct)
-    apply (auto simp add: eq le lt polyneg_norm)
-    apply (auto simp add: linorder_not_less order_le_less)
+                        apply (auto simp add: eq le lt polyneg_norm)
+      apply (auto simp add: linorder_not_less order_le_less)
     done
-  then obtain c s where csU: "(c, s) \<in> set (uset p)"
-    and x: "Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s \<or>
-      Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
+  then obtain c s
+    where c_s: "(c, s) \<in> set (uset p)"
+      and "Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s \<or>
+        Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
     by blast
   then have "x \<le> (- Itm vs (x#bs) s) / Ipoly vs c"
     using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"]
     by (auto simp add: mult.commute)
   then show ?thesis
-    using csU by auto
+    using c_s by auto
 qed
 
 lemma plusinf_uset:
@@ -2216,12 +2208,13 @@
   from nmi have nmi': "\<not> (Ifm vs (x#bs) (plusinf p))"
     by (simp add: bound0_I[OF plusinf_nb[OF lp], where b=x and b'=a])
   from plusinf_uset0[OF lp nmi' ex]
-  obtain c s where csU: "(c,s) \<in> set (uset p)"
-    and th: "x \<le> - Itm vs (x#bs) s / Ipoly vs c"
+  obtain c s
+    where c_s: "(c,s) \<in> set (uset p)"
+      and x: "x \<le> - Itm vs (x#bs) s / Ipoly vs c"
     by blast
-  from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s"
+  from uset_l[OF lp, rule_format, OF c_s] have nb: "tmbound0 s"
     by simp
-  from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis
+  from x tmbound0_I[OF nb, of vs x bs a] c_s show ?thesis
     by auto
 qed
 
@@ -2345,7 +2338,7 @@
   next
     case N: 3
     from px neg_divide_le_eq[OF N, where a="x" and b="-?Nt x s"]
-    have px': "x >= - ?Nt x s / ?N c"
+    have px': "x \<ge> - ?Nt x s / ?N c"
       by (simp add: field_simps)
     from ycs show ?thesis
     proof
@@ -2487,9 +2480,9 @@
     have "\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
         a \<le> - ?Nt x t / ?N c \<and> a \<ge> - ?Nt x s / ?N d"
       by blast
-    then obtain "c" "t" "d" "s"
-      where ctU: "(c,t) \<in> ?U"
-        and dsU: "(d,s) \<in> ?U"
+    then obtain c t d s
+      where ctU: "(c, t) \<in> ?U"
+        and dsU: "(d, s) \<in> ?U"
         and xs1: "a \<le> - ?Nt x s / ?N d"
         and tx1: "a \<ge> - ?Nt x t / ?N c"
       by blast
@@ -3090,18 +3083,18 @@
   from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r"
     by simp_all
   note r = tmbound0_I[OF lin(3), of vs _ bs x]
-  have cd_cs: "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)"
+  have "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)"
     by auto
-  moreover
-  {
-    assume c: "?c = 0" and d: "?d = 0"
-    then have ?thesis
-      using nc nd
+  then consider "?c = 0" "?d = 0" | "?c * ?d > 0" | "?c * ?d < 0"
+    | "?c > 0" "?d = 0" | "?c < 0" "?d = 0" | "?c = 0" "?d > 0" | "?c = 0" "?d < 0"
+    by blast
+  then show ?thesis
+  proof cases
+    case 1
+    with nc nd show ?thesis
       by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)
-  }
-  moreover
-  {
-    assume dc: "?c * ?d > 0"
+  next
+    case dc: 2
     from dc have dc': "2 * ?c * ?d > 0"
       by simp
     then have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
@@ -3122,13 +3115,12 @@
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<le> 0"
       using nonzero_mult_div_cancel_left[of "2*?c*?d"] c d
       by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
-    finally  have ?thesis using dc c d  nc nd dc'
+    finally show ?thesis
+      using dc c d  nc nd dc'
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def
           Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
-  }
-  moreover
-  {
-    assume dc: "?c * ?d < 0"
+  next
+    case dc: 3
     from dc have dc': "2 * ?c * ?d < 0"
       by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos)
     then have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
@@ -3147,13 +3139,13 @@
     also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r \<le> 0"
       using nonzero_mult_div_cancel_left[of "2 * ?c * ?d"] c d
       by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
-    finally  have ?thesis using dc c d  nc nd
+    finally show ?thesis
+      using dc c d  nc nd
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def
           Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
-  }
-  moreover
-  {
-    assume c: "?c > 0" and d: "?d = 0"
+  next
+    case 4
+    have c: "?c > 0" and d: "?d = 0" by fact+
     from c have c'': "2 * ?c > 0"
       by (simp add: zero_less_mult_iff)
     from c have c': "2 * ?c \<noteq> 0"
@@ -3171,13 +3163,13 @@
     also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r \<le> 0"
       using nonzero_mult_div_cancel_left[OF c'] c
       by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
-    finally have ?thesis using c d nc nd
+    finally show ?thesis
+      using c d nc nd
       by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def
           evaldjf_ex field_simps lt polyneg_norm polymul_norm)
-  }
-  moreover
-  {
-    assume c: "?c < 0" and d: "?d = 0"
+  next
+    case 5
+    have c: "?c < 0" and d: "?d = 0" by fact+
     then have c': "2 * ?c \<noteq> 0"
       by simp
     from c have c'': "2 * ?c < 0"
@@ -3196,13 +3188,12 @@
       using nonzero_mult_div_cancel_left[OF c'] c order_less_not_sym[OF c'']
           less_imp_neq[OF c''] c''
         by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
-    finally have ?thesis using c d nc nd
+    finally show ?thesis using c d nc nd
       by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def
           evaldjf_ex field_simps lt polyneg_norm polymul_norm)
-  }
-  moreover
-  {
-    assume c: "?c = 0" and d: "?d > 0"
+  next
+    case 6
+    have c: "?c = 0" and d: "?d > 0" by fact+
     from d have d'': "2 * ?d > 0"
       by (simp add: zero_less_mult_iff)
     from d have d': "2 * ?d \<noteq> 0"
@@ -3220,13 +3211,13 @@
     also have "\<dots> \<longleftrightarrow> - ?a * ?s + 2 * ?d * ?r \<le> 0"
       using nonzero_mult_div_cancel_left[OF d'] d
       by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
-    finally have ?thesis using c d nc nd
+    finally show ?thesis
+      using c d nc nd
       by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def
           evaldjf_ex field_simps lt polyneg_norm polymul_norm)
-  }
-  moreover
-  {
-    assume c: "?c = 0" and d: "?d < 0"
+  next
+    case 7
+    have c: "?c = 0" and d: "?d < 0" by fact+
     then have d': "2 * ?d \<noteq> 0"
       by simp
     from d have d'': "2 * ?d < 0"
@@ -3243,24 +3234,24 @@
       by simp
     also have "\<dots> \<longleftrightarrow> ?a * ?s -  2 * ?d * ?r \<le> 0"
       using nonzero_mult_div_cancel_left[OF d'] d order_less_not_sym[OF d'']
-          less_imp_neq[OF d''] d''
-        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
-    finally have ?thesis using c d nc nd
+        less_imp_neq[OF d''] d''
+      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
+    finally show ?thesis
+      using c d nc nd
       by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def
           evaldjf_ex field_simps lt polyneg_norm polymul_norm)
-  }
-  ultimately show ?thesis by blast
+  qed
 qed
 
 fun msubst :: "fm \<Rightarrow> (poly \<times> tm) \<times> (poly \<times> tm) \<Rightarrow> fm"
-where
-  "msubst (And p q) ((c, t), (d, s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
-| "msubst (Or p q) ((c, t), (d, s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
-| "msubst (Eq (CNP 0 a r)) ((c, t), (d, s)) = msubsteq c t d s a r"
-| "msubst (NEq (CNP 0 a r)) ((c, t), (d, s)) = msubstneq c t d s a r"
-| "msubst (Lt (CNP 0 a r)) ((c, t), (d, s)) = msubstlt c t d s a r"
-| "msubst (Le (CNP 0 a r)) ((c, t), (d, s)) = msubstle c t d s a r"
-| "msubst p ((c, t), (d, s)) = p"
+  where
+    "msubst (And p q) ((c, t), (d, s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
+  | "msubst (Or p q) ((c, t), (d, s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
+  | "msubst (Eq (CNP 0 a r)) ((c, t), (d, s)) = msubsteq c t d s a r"
+  | "msubst (NEq (CNP 0 a r)) ((c, t), (d, s)) = msubstneq c t d s a r"
+  | "msubst (Lt (CNP 0 a r)) ((c, t), (d, s)) = msubstlt c t d s a r"
+  | "msubst (Le (CNP 0 a r)) ((c, t), (d, s)) = msubstle c t d s a r"
+  | "msubst p ((c, t), (d, s)) = p"
 
 lemma msubst_I:
   assumes lp: "islin p"
@@ -3276,12 +3267,12 @@
       msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
 
 lemma msubst_nb:
-  assumes lp: "islin p"
-    and t: "tmbound0 t"
-    and s: "tmbound0 s"
+  assumes "islin p"
+    and "tmbound0 t"
+    and "tmbound0 s"
   shows "bound0 (msubst p ((c,t),(d,s)))"
-  using lp t s
-  by (induct p rule: islin.induct, auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
+  using assms
+  by (induct p rule: islin.induct) (auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
 
 lemma fr_eq_msubst:
   assumes lp: "islin p"
@@ -3292,15 +3283,14 @@
       Ifm vs (x#bs) (msubst p ((c, t), (d, s)))))"
   (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
 proof -
-  from uset_l[OF lp]
-  have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
+  from uset_l[OF lp] have *: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
     by blast
   {
     fix c t d s
     assume ctU: "(c, t) \<in>set (uset p)"
       and dsU: "(d,s) \<in>set (uset p)"
       and pts: "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
-    from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
+    from *[rule_format, OF ctU] *[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
       by simp_all
     from msubst_I[OF lp norm, of vs x bs t s] pts
     have "Ifm vs (x # bs) (msubst p ((c, t), d, s))" ..
@@ -3311,15 +3301,15 @@
     assume ctU: "(c, t) \<in> set (uset p)"
       and dsU: "(d,s) \<in>set (uset p)"
       and pts: "Ifm vs (x # bs) (msubst p ((c, t), d, s))"
-    from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
+    from *[rule_format, OF ctU] *[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
       by simp_all
     from msubst_I[OF lp norm, of vs x bs t s] pts
     have "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p" ..
   }
-  ultimately have th': "(\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
+  ultimately have **: "(\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
       Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p) \<longleftrightarrow> ?F"
     by blast
-  from fr_eq[OF lp, of vs bs x, simplified th'] show ?thesis .
+  from fr_eq[OF lp, of vs bs x, simplified **] show ?thesis .
 qed
 
 lemma simpfm_lin:
@@ -3407,7 +3397,7 @@
     using decr0[OF th1, of vs x bs]
     apply (simp add: ferrack_def Let_def)
     apply (cases "?mp = T \<or> ?pp = T")
-    apply auto
+     apply auto
     done
   finally show ?thesis
     using thqf by blast
@@ -3425,7 +3415,7 @@
 qed
 
 
-section \<open>Second implemenation: Case splits not local\<close>
+section \<open>Second implementation: case splits not local\<close>
 
 lemma fr_eq2:
   assumes lp: "islin p"
@@ -3442,14 +3432,14 @@
   (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?Z \<or> ?U \<or> ?F)" is "?E = ?D")
 proof
   assume px: "\<exists>x. ?I x p"
-  have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
-  moreover {
-    assume "?M \<or> ?P"
-    then have "?D" by blast
-  }
-  moreover {
-    assume nmi: "\<not> ?M"
-      and npi: "\<not> ?P"
+  consider "?M \<or> ?P" | "\<not> ?M" "\<not> ?P" by blast
+  then show ?D
+  proof cases
+    case 1
+    then show ?thesis by blast
+  next
+    case 2
+    have nmi: "\<not> ?M" and npi: "\<not> ?P" by fact+
     from inf_uset[OF lp nmi npi, OF px]
     obtain c t d s where ct:
       "(c, t) \<in> set (uset p)"
@@ -3462,50 +3452,42 @@
     let ?t = "Itm vs (x # bs) t"
     have eq2: "\<And>(x::'a). x + x = 2 * x"
       by (simp add: field_simps)
-    {
-      assume "?c = 0 \<and> ?d = 0"
-      with ct have ?D by simp
-    }
-    moreover
-    {
-      assume z: "?c = 0" "?d \<noteq> 0"
-      from z have ?D using ct by auto
-    }
-    moreover
-    {
-      assume z: "?c \<noteq> 0" "?d = 0"
-      with ct have ?D by auto
-    }
-    moreover
-    {
-      assume z: "?c \<noteq> 0" "?d \<noteq> 0"
-      from z have ?F using ct
+    consider "?c = 0" "?d = 0" | "?c = 0" "?d \<noteq> 0" | "?c \<noteq> 0" "?d = 0" | "?c \<noteq> 0" "?d \<noteq> 0"
+      by auto
+    then show ?thesis
+    proof cases
+      case 1
+      with ct show ?thesis by simp
+    next
+      case 2
+      with ct show ?thesis by auto
+    next
+      case 3
+      with ct show ?thesis by auto
+    next
+      case z: 4
+      from z have ?F
+        using ct
         apply -
         apply (rule bexI[where x = "(c,t)"])
-        apply simp_all
+         apply simp_all
         apply (rule bexI[where x = "(d,s)"])
-        apply simp_all
+         apply simp_all
         done
-      then have ?D by blast
-    }
-    ultimately have ?D by auto
-  }
-  ultimately show "?D" by blast
+      then show ?thesis by blast
+    qed
+  qed
 next
-  assume "?D"
-  moreover {
-    assume m: "?M"
-    from minusinf_ex[OF lp m] have "?E" .
-  }
-  moreover {
-    assume p: "?P"
-    from plusinf_ex[OF lp p] have "?E" .
-  }
-  moreover {
-    assume f:"?F"
-    then have "?E" by blast
-  }
-  ultimately show "?E" by blast
+  assume ?D
+  then consider ?M | ?P | ?Z | ?U | ?F by blast
+  then show ?E
+  proof cases
+    case 1
+    show ?thesis by (rule minusinf_ex[OF lp \<open>?M\<close>])
+  next
+    case 2
+    show ?thesis by (rule plusinf_ex[OF lp \<open>?P\<close>])
+  qed blast+
 qed
 
 definition "msubsteq2 c t a b = Eq (Add (Mul a t) (Mul c b))"
@@ -3555,14 +3537,14 @@
   by (simp add: msubstleneg_def field_simps del: minus_add_distrib)
 
 fun msubstpos :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm"
-where
-  "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
-| "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
-| "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
-| "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
-| "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r"
-| "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r"
-| "msubstpos p c t = p"
+  where
+    "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
+  | "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
+  | "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
+  | "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+  | "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r"
+  | "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r"
+  | "msubstpos p c t = p"
 
 lemma msubstpos_I:
   assumes lp: "islin p"
@@ -3576,14 +3558,14 @@
       bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
 
 fun msubstneg :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm"
-where
-  "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
-| "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
-| "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
-| "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
-| "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r"
-| "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
-| "msubstneg p c t = p"
+  where
+    "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
+  | "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
+  | "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
+  | "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+  | "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r"
+  | "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
+  | "msubstneg p c t = p"
 
 lemma msubstneg_I:
   assumes lp: "islin p"
@@ -3595,11 +3577,9 @@
       tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x]
       bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
 
-definition
-  "msubst2 p c t =
-    disj
-      (conj (lt (CP (polyneg c))) (simpfm (msubstpos p c t)))
-      (conj (lt (CP c)) (simpfm (msubstneg p c t)))"
+definition "msubst2 p c t =
+  disj (conj (lt (CP (polyneg c))) (simpfm (msubstpos p c t)))
+    (conj (lt (CP c)) (simpfm (msubstneg p c t)))"
 
 lemma msubst2:
   assumes lp: "islin p"
@@ -3610,20 +3590,19 @@
   let ?c = "Ipoly vs c"
   from nc have anc: "allpolys isnpoly (CP c)" "allpolys isnpoly (CP (~\<^sub>p c))"
     by (simp_all add: polyneg_norm)
-  from nz have "?c > 0 \<or> ?c < 0" by arith
-  moreover
-  {
-    assume c: "?c < 0"
+  from nz consider "?c < 0" | "?c > 0" by arith
+  then show ?thesis
+  proof cases
+    case c: 1
     from c msubstneg_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
-    have ?thesis by (auto simp add: msubst2_def)
-  }
-  moreover
-  {
-    assume c: "?c > 0"
+    show ?thesis
+      by (auto simp add: msubst2_def)
+  next
+    case c: 2
     from c msubstpos_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
-    have ?thesis by (auto simp add: msubst2_def)
-  }
-  ultimately show ?thesis by blast
+    show ?thesis
+      by (auto simp add: msubst2_def)
+  qed
 qed
 
 lemma msubsteq2_nb: "tmbound0 t \<Longrightarrow> islin (Eq (CNP 0 a r)) \<Longrightarrow> bound0 (msubsteq2 c t a r)"
@@ -3664,10 +3643,12 @@
   using lp tnb
   by (simp add: msubst2_def msubstneg_nb msubstpos_nb lt_nb simpfm_bound0)
 
-lemma mult_minus2_left: "-2 * (x::'a::comm_ring_1) = - (2 * x)"
+lemma mult_minus2_left: "-2 * x = - (2 * x)"
+  for x :: "'a::comm_ring_1"
   by simp
 
-lemma mult_minus2_right: "(x::'a::comm_ring_1) * -2 = - (x * 2)"
+lemma mult_minus2_right: "x * -2 = - (x * 2)"
+  for x :: "'a::comm_ring_1"
   by simp
 
 lemma islin_qf: "islin p \<Longrightarrow> qfree p"
@@ -3685,7 +3666,7 @@
         Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))))"
   (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?Pz \<or> ?PU \<or> ?F)" is "?E = ?D")
 proof -
-  from uset_l[OF lp] have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
+  from uset_l[OF lp] have *: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
     by blast
   let ?I = "\<lambda>p. Ifm vs (x#bs) p"
   have n2: "isnpoly (C (-2,1))"
@@ -3700,7 +3681,7 @@
     {
       fix n t
       assume H: "(n, t) \<in> set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)"
-      from H(1) th have "isnpoly n"
+      from H(1) * have "isnpoly n"
         by blast
       then have nn: "isnpoly (n *\<^sub>p (C (-2,1)))"
         by (simp_all add: polymul_norm n2)
@@ -3719,7 +3700,7 @@
       assume H:
         "(n, t) \<in> set (uset p)" "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         "Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p"
-      from H(1) th have "isnpoly n"
+      from H(1) * have "isnpoly n"
         by blast
       then have nn: "isnpoly (n *\<^sub>p (C (-2,1)))" "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         using H(2) by (simp_all add: polymul_norm n2)
@@ -3740,7 +3721,7 @@
       assume H:
         "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
         "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
-      from H(1,2) th have "isnpoly c" "isnpoly d"
+      from H(1,2) * have "isnpoly c" "isnpoly d"
         by blast+
       then have nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)"
         by (simp_all add: polymul_norm n2)
@@ -3766,7 +3747,7 @@
         "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
-      from H(1,2) th have "isnpoly c" "isnpoly d"
+      from H(1,2) * have "isnpoly c" "isnpoly d"
         by blast+
       then have nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         using H(3,4) by (simp_all add: polymul_norm n2)
@@ -3780,32 +3761,31 @@
     unfolding eq0 eq1 eq2 by blast
 qed
 
-definition
-  "ferrack2 p \<equiv>
-    let
-      q = simpfm p;
-      mp = minusinf q;
-      pp = plusinf q
-    in
-      if (mp = T \<or> pp = T) then T
-      else
-       (let U = remdups (uset  q)
-        in
-          decr0
-            (list_disj
-              [mp,
-               pp,
-               simpfm (subst0 (CP 0\<^sub>p) q),
-               evaldjf (\<lambda>(c, t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U,
-               evaldjf (\<lambda>((b, a),(d, c)).
-                msubst2 q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs U)]))"
+definition "ferrack2 p \<equiv>
+  let
+    q = simpfm p;
+    mp = minusinf q;
+    pp = plusinf q
+  in
+    if (mp = T \<or> pp = T) then T
+    else
+     (let U = remdups (uset  q)
+      in
+        decr0
+          (list_disj
+            [mp,
+             pp,
+             simpfm (subst0 (CP 0\<^sub>p) q),
+             evaldjf (\<lambda>(c, t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U,
+             evaldjf (\<lambda>((b, a),(d, c)).
+              msubst2 q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs U)]))"
 
 definition "frpar2 p = simpfm (qelim (prep p) ferrack2)"
 
 lemma ferrack2:
   assumes qf: "qfree p"
   shows "qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
-  (is "_ \<and> (?rhs = ?lhs)")
+    (is "_ \<and> (?rhs = ?lhs)")
 proof -
   let ?J = "\<lambda>x p. Ifm vs (x#bs) p"
   let ?N = "\<lambda>t. Ipoly vs t"
@@ -3826,34 +3806,29 @@
     by simp
   have bnd0: "\<forall>x \<in> set ?U. bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) x)"
   proof -
-    {
-      fix c t
-      assume ct: "(c, t) \<in> set ?U"
-      then have tnb: "tmbound0 t"
-        using U_l by blast
-      from msubst2_nb[OF lq tnb]
-      have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" by simp
-    }
+    have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))"
+      if "(c, t) \<in> set ?U" for c t
+    proof -
+      from U_l that have "tmbound0 t" by blast
+      from msubst2_nb[OF lq this] show ?thesis by simp
+    qed
     then show ?thesis by auto
   qed
   have bnd1: "\<forall>x \<in> set ?Up. bound0 ((\<lambda>((b, a), (d, c)).
     msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) x)"
   proof -
-    {
-      fix b a d c
-      assume badc: "((b,a),(d,c)) \<in> set ?Up"
-      from badc U_l alluopairs_set1[of ?U]
-      have nb: "tmbound0 (Add (Mul d a) (Mul b c))"
+    have "bound0 ((\<lambda>((b, a),(d, c)).
+      msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))"
+      if  "((b,a),(d,c)) \<in> set ?Up" for b a d c
+    proof -
+      from U_l alluopairs_set1[of ?U] that have this: "tmbound0 (Add (Mul d a) (Mul b c))"
         by auto
-      from msubst2_nb[OF lq nb]
-      have "bound0 ((\<lambda>((b, a),(d, c)).
-          msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))"
+      from msubst2_nb[OF lq this] show ?thesis
         by simp
-    }
+    qed
     then show ?thesis by auto
   qed
-  have stupid: "bound0 F"
-    by simp
+  have stupid: "bound0 F" by simp
   let ?R =
     "list_disj
      [?mp,
@@ -3925,9 +3900,9 @@
   also have "\<dots> \<longleftrightarrow> ?rhs"
     unfolding ferrack2_def
     apply (cases "?mp = T")
-    apply (simp add: list_disj_def)
+     apply (simp add: list_disj_def)
     apply (cases "?pp = T")
-    apply (simp add: list_disj_def)
+     apply (simp add: list_disj_def)
     apply (simp_all add: Let_def decr0[OF nb])
     done
   finally show ?thesis using decr0_qf[OF nb]
@@ -3937,14 +3912,14 @@
 lemma frpar2: "qfree (frpar2 p) \<and> (Ifm vs bs (frpar2 p) \<longleftrightarrow> Ifm vs bs p)"
 proof -
   from ferrack2
-  have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
+  have this: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
     by blast
-  from qelim[OF th, of "prep p" bs]
-  show ?thesis
+  from qelim[OF this, of "prep p" bs] show ?thesis
     unfolding frpar2_def by (auto simp add: prep)
 qed
 
-oracle frpar_oracle = \<open>
+oracle frpar_oracle =
+\<open>
 let
 
 fun binopT T = T --> T --> T;
--- a/src/HOL/Decision_Procs/Rat_Pair.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Rat_Pair.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -5,7 +5,7 @@
 section \<open>Rational numbers as pairs\<close>
 
 theory Rat_Pair
-imports Complex_Main
+  imports Complex_Main
 begin
 
 type_synonym Num = "int \<times> int"
@@ -20,8 +20,7 @@
   where "isnormNum = (\<lambda>(a, b). if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1)"
 
 definition normNum :: "Num \<Rightarrow> Num"
-where
-  "normNum = (\<lambda>(a,b).
+  where "normNum = (\<lambda>(a,b).
     (if a = 0 \<or> b = 0 then (0, 0)
      else
       (let g = gcd a b
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -5,7 +5,7 @@
 section \<open>Implementation and verification of multivariate polynomials\<close>
 
 theory Reflected_Multivariate_Polynomial
-imports Complex_Main Rat_Pair Polynomial_List
+  imports Complex_Main Rat_Pair Polynomial_List
 begin
 
 subsection \<open>Datatype of polynomial expressions\<close>
@@ -20,95 +20,95 @@
 subsection\<open>Boundedness, substitution and all that\<close>
 
 primrec polysize:: "poly \<Rightarrow> nat"
-where
-  "polysize (C c) = 1"
-| "polysize (Bound n) = 1"
-| "polysize (Neg p) = 1 + polysize p"
-| "polysize (Add p q) = 1 + polysize p + polysize q"
-| "polysize (Sub p q) = 1 + polysize p + polysize q"
-| "polysize (Mul p q) = 1 + polysize p + polysize q"
-| "polysize (Pw p n) = 1 + polysize p"
-| "polysize (CN c n p) = 4 + polysize c + polysize p"
+  where
+    "polysize (C c) = 1"
+  | "polysize (Bound n) = 1"
+  | "polysize (Neg p) = 1 + polysize p"
+  | "polysize (Add p q) = 1 + polysize p + polysize q"
+  | "polysize (Sub p q) = 1 + polysize p + polysize q"
+  | "polysize (Mul p q) = 1 + polysize p + polysize q"
+  | "polysize (Pw p n) = 1 + polysize p"
+  | "polysize (CN c n p) = 4 + polysize c + polysize p"
 
 primrec polybound0:: "poly \<Rightarrow> bool" \<comment> \<open>a poly is INDEPENDENT of Bound 0\<close>
-where
-  "polybound0 (C c) \<longleftrightarrow> True"
-| "polybound0 (Bound n) \<longleftrightarrow> n > 0"
-| "polybound0 (Neg a) \<longleftrightarrow> polybound0 a"
-| "polybound0 (Add a b) \<longleftrightarrow> polybound0 a \<and> polybound0 b"
-| "polybound0 (Sub a b) \<longleftrightarrow> polybound0 a \<and> polybound0 b"
-| "polybound0 (Mul a b) \<longleftrightarrow> polybound0 a \<and> polybound0 b"
-| "polybound0 (Pw p n) \<longleftrightarrow> polybound0 p"
-| "polybound0 (CN c n p) \<longleftrightarrow> n \<noteq> 0 \<and> polybound0 c \<and> polybound0 p"
+  where
+    "polybound0 (C c) \<longleftrightarrow> True"
+  | "polybound0 (Bound n) \<longleftrightarrow> n > 0"
+  | "polybound0 (Neg a) \<longleftrightarrow> polybound0 a"
+  | "polybound0 (Add a b) \<longleftrightarrow> polybound0 a \<and> polybound0 b"
+  | "polybound0 (Sub a b) \<longleftrightarrow> polybound0 a \<and> polybound0 b"
+  | "polybound0 (Mul a b) \<longleftrightarrow> polybound0 a \<and> polybound0 b"
+  | "polybound0 (Pw p n) \<longleftrightarrow> polybound0 p"
+  | "polybound0 (CN c n p) \<longleftrightarrow> n \<noteq> 0 \<and> polybound0 c \<and> polybound0 p"
 
 primrec polysubst0:: "poly \<Rightarrow> poly \<Rightarrow> poly" \<comment> \<open>substitute a poly into a poly for Bound 0\<close>
-where
-  "polysubst0 t (C c) = C c"
-| "polysubst0 t (Bound n) = (if n = 0 then t else Bound n)"
-| "polysubst0 t (Neg a) = Neg (polysubst0 t a)"
-| "polysubst0 t (Add a b) = Add (polysubst0 t a) (polysubst0 t b)"
-| "polysubst0 t (Sub a b) = Sub (polysubst0 t a) (polysubst0 t b)"
-| "polysubst0 t (Mul a b) = Mul (polysubst0 t a) (polysubst0 t b)"
-| "polysubst0 t (Pw p n) = Pw (polysubst0 t p) n"
-| "polysubst0 t (CN c n p) =
-    (if n = 0 then Add (polysubst0 t c) (Mul t (polysubst0 t p))
-     else CN (polysubst0 t c) n (polysubst0 t p))"
+  where
+    "polysubst0 t (C c) = C c"
+  | "polysubst0 t (Bound n) = (if n = 0 then t else Bound n)"
+  | "polysubst0 t (Neg a) = Neg (polysubst0 t a)"
+  | "polysubst0 t (Add a b) = Add (polysubst0 t a) (polysubst0 t b)"
+  | "polysubst0 t (Sub a b) = Sub (polysubst0 t a) (polysubst0 t b)"
+  | "polysubst0 t (Mul a b) = Mul (polysubst0 t a) (polysubst0 t b)"
+  | "polysubst0 t (Pw p n) = Pw (polysubst0 t p) n"
+  | "polysubst0 t (CN c n p) =
+      (if n = 0 then Add (polysubst0 t c) (Mul t (polysubst0 t p))
+       else CN (polysubst0 t c) n (polysubst0 t p))"
 
 fun decrpoly:: "poly \<Rightarrow> poly"
-where
-  "decrpoly (Bound n) = Bound (n - 1)"
-| "decrpoly (Neg a) = Neg (decrpoly a)"
-| "decrpoly (Add a b) = Add (decrpoly a) (decrpoly b)"
-| "decrpoly (Sub a b) = Sub (decrpoly a) (decrpoly b)"
-| "decrpoly (Mul a b) = Mul (decrpoly a) (decrpoly b)"
-| "decrpoly (Pw p n) = Pw (decrpoly p) n"
-| "decrpoly (CN c n p) = CN (decrpoly c) (n - 1) (decrpoly p)"
-| "decrpoly a = a"
+  where
+    "decrpoly (Bound n) = Bound (n - 1)"
+  | "decrpoly (Neg a) = Neg (decrpoly a)"
+  | "decrpoly (Add a b) = Add (decrpoly a) (decrpoly b)"
+  | "decrpoly (Sub a b) = Sub (decrpoly a) (decrpoly b)"
+  | "decrpoly (Mul a b) = Mul (decrpoly a) (decrpoly b)"
+  | "decrpoly (Pw p n) = Pw (decrpoly p) n"
+  | "decrpoly (CN c n p) = CN (decrpoly c) (n - 1) (decrpoly p)"
+  | "decrpoly a = a"
 
 
 subsection \<open>Degrees and heads and coefficients\<close>
 
 fun degree :: "poly \<Rightarrow> nat"
-where
-  "degree (CN c 0 p) = 1 + degree p"
-| "degree p = 0"
+  where
+    "degree (CN c 0 p) = 1 + degree p"
+  | "degree p = 0"
 
 fun head :: "poly \<Rightarrow> poly"
-where
-  "head (CN c 0 p) = head p"
-| "head p = p"
+  where
+    "head (CN c 0 p) = head p"
+  | "head p = p"
 
 text \<open>More general notions of degree and head.\<close>
 
 fun degreen :: "poly \<Rightarrow> nat \<Rightarrow> nat"
-where
-  "degreen (CN c n p) = (\<lambda>m. if n = m then 1 + degreen p n else 0)"
-| "degreen p = (\<lambda>m. 0)"
+  where
+    "degreen (CN c n p) = (\<lambda>m. if n = m then 1 + degreen p n else 0)"
+  | "degreen p = (\<lambda>m. 0)"
 
 fun headn :: "poly \<Rightarrow> nat \<Rightarrow> poly"
-where
-  "headn (CN c n p) = (\<lambda>m. if n \<le> m then headn p m else CN c n p)"
-| "headn p = (\<lambda>m. p)"
+  where
+    "headn (CN c n p) = (\<lambda>m. if n \<le> m then headn p m else CN c n p)"
+  | "headn p = (\<lambda>m. p)"
 
 fun coefficients :: "poly \<Rightarrow> poly list"
-where
-  "coefficients (CN c 0 p) = c # coefficients p"
-| "coefficients p = [p]"
+  where
+    "coefficients (CN c 0 p) = c # coefficients p"
+  | "coefficients p = [p]"
 
 fun isconstant :: "poly \<Rightarrow> bool"
-where
-  "isconstant (CN c 0 p) = False"
-| "isconstant p = True"
+  where
+    "isconstant (CN c 0 p) = False"
+  | "isconstant p = True"
 
 fun behead :: "poly \<Rightarrow> poly"
-where
-  "behead (CN c 0 p) = (let p' = behead p in if p' = 0\<^sub>p then c else CN c 0 p')"
-| "behead p = 0\<^sub>p"
+  where
+    "behead (CN c 0 p) = (let p' = behead p in if p' = 0\<^sub>p then c else CN c 0 p')"
+  | "behead p = 0\<^sub>p"
 
 fun headconst :: "poly \<Rightarrow> Num"
-where
-  "headconst (CN c n p) = headconst p"
-| "headconst (C n) = n"
+  where
+    "headconst (CN c n p) = headconst p"
+  | "headconst (C n) = n"
 
 
 subsection \<open>Operations for normalization\<close>
@@ -116,70 +116,69 @@
 declare if_cong[fundef_cong del]
 declare let_cong[fundef_cong del]
 
-fun polyadd :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "+\<^sub>p" 60)
-where
-  "polyadd (C c) (C c') = C (c +\<^sub>N c')"
-| "polyadd (C c) (CN c' n' p') = CN (polyadd (C c) c') n' p'"
-| "polyadd (CN c n p) (C c') = CN (polyadd c (C c')) n p"
-| "polyadd (CN c n p) (CN c' n' p') =
-    (if n < n' then CN (polyadd c (CN c' n' p')) n p
-     else if n' < n then CN (polyadd (CN c n p) c') n' p'
-     else
-      let
-        cc' = polyadd c c';
-        pp' = polyadd p p'
-      in if pp' = 0\<^sub>p then cc' else CN cc' n pp')"
-| "polyadd a b = Add a b"
-
+fun polyadd :: "poly \<Rightarrow> poly \<Rightarrow> poly"  (infixl "+\<^sub>p" 60)
+  where
+    "polyadd (C c) (C c') = C (c +\<^sub>N c')"
+  | "polyadd (C c) (CN c' n' p') = CN (polyadd (C c) c') n' p'"
+  | "polyadd (CN c n p) (C c') = CN (polyadd c (C c')) n p"
+  | "polyadd (CN c n p) (CN c' n' p') =
+      (if n < n' then CN (polyadd c (CN c' n' p')) n p
+       else if n' < n then CN (polyadd (CN c n p) c') n' p'
+       else
+        let
+          cc' = polyadd c c';
+          pp' = polyadd p p'
+        in if pp' = 0\<^sub>p then cc' else CN cc' n pp')"
+  | "polyadd a b = Add a b"
 
 fun polyneg :: "poly \<Rightarrow> poly" ("~\<^sub>p")
-where
-  "polyneg (C c) = C (~\<^sub>N c)"
-| "polyneg (CN c n p) = CN (polyneg c) n (polyneg p)"
-| "polyneg a = Neg a"
+  where
+    "polyneg (C c) = C (~\<^sub>N c)"
+  | "polyneg (CN c n p) = CN (polyneg c) n (polyneg p)"
+  | "polyneg a = Neg a"
 
-definition polysub :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "-\<^sub>p" 60)
+definition polysub :: "poly \<Rightarrow> poly \<Rightarrow> poly"  (infixl "-\<^sub>p" 60)
   where "p -\<^sub>p q = polyadd p (polyneg q)"
 
-fun polymul :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "*\<^sub>p" 60)
-where
-  "polymul (C c) (C c') = C (c *\<^sub>N c')"
-| "polymul (C c) (CN c' n' p') =
-    (if c = 0\<^sub>N then 0\<^sub>p else CN (polymul (C c) c') n' (polymul (C c) p'))"
-| "polymul (CN c n p) (C c') =
-    (if c' = 0\<^sub>N  then 0\<^sub>p else CN (polymul c (C c')) n (polymul p (C c')))"
-| "polymul (CN c n p) (CN c' n' p') =
-    (if n < n' then CN (polymul c (CN c' n' p')) n (polymul p (CN c' n' p'))
-     else if n' < n then CN (polymul (CN c n p) c') n' (polymul (CN c n p) p')
-     else polyadd (polymul (CN c n p) c') (CN 0\<^sub>p n' (polymul (CN c n p) p')))"
-| "polymul a b = Mul a b"
+fun polymul :: "poly \<Rightarrow> poly \<Rightarrow> poly"  (infixl "*\<^sub>p" 60)
+  where
+    "polymul (C c) (C c') = C (c *\<^sub>N c')"
+  | "polymul (C c) (CN c' n' p') =
+      (if c = 0\<^sub>N then 0\<^sub>p else CN (polymul (C c) c') n' (polymul (C c) p'))"
+  | "polymul (CN c n p) (C c') =
+      (if c' = 0\<^sub>N  then 0\<^sub>p else CN (polymul c (C c')) n (polymul p (C c')))"
+  | "polymul (CN c n p) (CN c' n' p') =
+      (if n < n' then CN (polymul c (CN c' n' p')) n (polymul p (CN c' n' p'))
+       else if n' < n then CN (polymul (CN c n p) c') n' (polymul (CN c n p) p')
+       else polyadd (polymul (CN c n p) c') (CN 0\<^sub>p n' (polymul (CN c n p) p')))"
+  | "polymul a b = Mul a b"
 
 declare if_cong[fundef_cong]
 declare let_cong[fundef_cong]
 
 fun polypow :: "nat \<Rightarrow> poly \<Rightarrow> poly"
-where
-  "polypow 0 = (\<lambda>p. (1)\<^sub>p)"
-| "polypow n =
-    (\<lambda>p.
-      let
-        q = polypow (n div 2) p;
-        d = polymul q q
-      in if even n then d else polymul p d)"
+  where
+    "polypow 0 = (\<lambda>p. (1)\<^sub>p)"
+  | "polypow n =
+      (\<lambda>p.
+        let
+          q = polypow (n div 2) p;
+          d = polymul q q
+        in if even n then d else polymul p d)"
 
-abbreviation poly_pow :: "poly \<Rightarrow> nat \<Rightarrow> poly" (infixl "^\<^sub>p" 60)
+abbreviation poly_pow :: "poly \<Rightarrow> nat \<Rightarrow> poly"  (infixl "^\<^sub>p" 60)
   where "a ^\<^sub>p k \<equiv> polypow k a"
 
 function polynate :: "poly \<Rightarrow> poly"
-where
-  "polynate (Bound n) = CN 0\<^sub>p n (1)\<^sub>p"
-| "polynate (Add p q) = polynate p +\<^sub>p polynate q"
-| "polynate (Sub p q) = polynate p -\<^sub>p polynate q"
-| "polynate (Mul p q) = polynate p *\<^sub>p polynate q"
-| "polynate (Neg p) = ~\<^sub>p (polynate p)"
-| "polynate (Pw p n) = polynate p ^\<^sub>p n"
-| "polynate (CN c n p) = polynate (Add c (Mul (Bound n) p))"
-| "polynate (C c) = C (normNum c)"
+  where
+    "polynate (Bound n) = CN 0\<^sub>p n (1)\<^sub>p"
+  | "polynate (Add p q) = polynate p +\<^sub>p polynate q"
+  | "polynate (Sub p q) = polynate p -\<^sub>p polynate q"
+  | "polynate (Mul p q) = polynate p *\<^sub>p polynate q"
+  | "polynate (Neg p) = ~\<^sub>p (polynate p)"
+  | "polynate (Pw p n) = polynate p ^\<^sub>p n"
+  | "polynate (CN c n p) = polynate (Add c (Mul (Bound n) p))"
+  | "polynate (C c) = C (normNum c)"
   by pat_completeness auto
 termination by (relation "measure polysize") auto
 
@@ -190,8 +189,7 @@
 | "poly_cmul y p = C y *\<^sub>p p"
 
 definition monic :: "poly \<Rightarrow> poly \<times> bool"
-where
-  "monic p =
+  where "monic p =
     (let h = headconst p
      in if h = 0\<^sub>N then (p, False) else (C (Ninv h) *\<^sub>p p, 0>\<^sub>N h))"
 
@@ -224,28 +222,28 @@
   where "polydivide s p = polydivide_aux (head p) (degree p) p 0 s"
 
 fun poly_deriv_aux :: "nat \<Rightarrow> poly \<Rightarrow> poly"
-where
-  "poly_deriv_aux n (CN c 0 p) = CN (poly_cmul ((int n)\<^sub>N) c) 0 (poly_deriv_aux (n + 1) p)"
-| "poly_deriv_aux n p = poly_cmul ((int n)\<^sub>N) p"
+  where
+    "poly_deriv_aux n (CN c 0 p) = CN (poly_cmul ((int n)\<^sub>N) c) 0 (poly_deriv_aux (n + 1) p)"
+  | "poly_deriv_aux n p = poly_cmul ((int n)\<^sub>N) p"
 
 fun poly_deriv :: "poly \<Rightarrow> poly"
-where
-  "poly_deriv (CN c 0 p) = poly_deriv_aux 1 p"
-| "poly_deriv p = 0\<^sub>p"
+  where
+    "poly_deriv (CN c 0 p) = poly_deriv_aux 1 p"
+  | "poly_deriv p = 0\<^sub>p"
 
 
 subsection \<open>Semantics of the polynomial representation\<close>
 
 primrec Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{field_char_0,field,power}"
-where
-  "Ipoly bs (C c) = INum c"
-| "Ipoly bs (Bound n) = bs!n"
-| "Ipoly bs (Neg a) = - Ipoly bs a"
-| "Ipoly bs (Add a b) = Ipoly bs a + Ipoly bs b"
-| "Ipoly bs (Sub a b) = Ipoly bs a - Ipoly bs b"
-| "Ipoly bs (Mul a b) = Ipoly bs a * Ipoly bs b"
-| "Ipoly bs (Pw t n) = Ipoly bs t ^ n"
-| "Ipoly bs (CN c n p) = Ipoly bs c + (bs!n) * Ipoly bs p"
+  where
+    "Ipoly bs (C c) = INum c"
+  | "Ipoly bs (Bound n) = bs!n"
+  | "Ipoly bs (Neg a) = - Ipoly bs a"
+  | "Ipoly bs (Add a b) = Ipoly bs a + Ipoly bs b"
+  | "Ipoly bs (Sub a b) = Ipoly bs a - Ipoly bs b"
+  | "Ipoly bs (Mul a b) = Ipoly bs a * Ipoly bs b"
+  | "Ipoly bs (Pw t n) = Ipoly bs t ^ n"
+  | "Ipoly bs (CN c n p) = Ipoly bs c + (bs!n) * Ipoly bs p"
 
 abbreviation Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{field_char_0,field,power}"  ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
   where "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<equiv> Ipoly bs p"
@@ -262,10 +260,10 @@
 subsection \<open>Normal form and normalization\<close>
 
 fun isnpolyh:: "poly \<Rightarrow> nat \<Rightarrow> bool"
-where
-  "isnpolyh (C c) = (\<lambda>k. isnormNum c)"
-| "isnpolyh (CN c n p) = (\<lambda>k. n \<ge> k \<and> isnpolyh c (Suc n) \<and> isnpolyh p n \<and> p \<noteq> 0\<^sub>p)"
-| "isnpolyh p = (\<lambda>k. False)"
+  where
+    "isnpolyh (C c) = (\<lambda>k. isnormNum c)"
+  | "isnpolyh (CN c n p) = (\<lambda>k. n \<ge> k \<and> isnpolyh c (Suc n) \<and> isnpolyh p n \<and> p \<noteq> 0\<^sub>p)"
+  | "isnpolyh p = (\<lambda>k. False)"
 
 lemma isnpolyh_mono: "n' \<le> n \<Longrightarrow> isnpolyh p n \<Longrightarrow> isnpolyh p n'"
   by (induct p rule: isnpolyh.induct) auto
@@ -512,23 +510,22 @@
       and nn0: "n \<ge> n0"
       and nn1: "n' \<ge> n1"
       by simp_all
-    {
-      assume "n < n'"
+    consider "n < n'" | "n' < n" | "n' = n" by arith
+    then show ?case
+    proof cases
+      case 1
       with "4.hyps"(4-5)[OF np cnp', of n] and "4.hyps"(1)[OF nc cnp', of n] nn0 cnp
-      have ?case by (simp add: min_def)
-    } moreover {
-      assume "n' < n"
+      show ?thesis by (simp add: min_def)
+    next
+      case 2
       with "4.hyps"(16-17)[OF cnp np', of "n'"] and "4.hyps"(13)[OF cnp nc', of "Suc n'"] nn1 cnp'
-      have ?case by (cases "Suc n' = n") (simp_all add: min_def)
-    } moreover {
-      assume "n' = n"
+      show ?thesis by (cases "Suc n' = n") (simp_all add: min_def)
+    next
+      case 3
       with "4.hyps"(16-17)[OF cnp np', of n] and "4.hyps"(13)[OF cnp nc', of n] cnp cnp' nn1 nn0
-      have ?case
-        apply (auto intro!: polyadd_normh)
-        apply (simp_all add: min_def isnpolyh_mono[OF nn0])
-        done
-    }
-    ultimately show ?case by arith
+      show ?thesis
+        by (auto intro!: polyadd_normh) (simp_all add: min_def isnpolyh_mono[OF nn0])
+    qed
   next
     fix n0 n1 m
     assume np: "isnpolyh ?cnp n0"
@@ -538,16 +535,14 @@
     let ?d1 = "degreen ?cnp m"
     let ?d2 = "degreen ?cnp' m"
     let ?eq = "?d = (if ?cnp = 0\<^sub>p \<or> ?cnp' = 0\<^sub>p then 0  else ?d1 + ?d2)"
-    have "n' < n \<or> n < n' \<or> n' = n" by auto
-    moreover
-    {
-      assume "n' < n \<or> n < n'"
-      with "4.hyps"(3,6,18) np np' m have ?eq
-        by auto
-    }
-    moreover
-    {
-      assume nn': "n' = n"
+    consider "n' < n \<or> n < n'" | "n' = n" by linarith
+    then show ?eq
+    proof cases
+      case 1
+      with "4.hyps"(3,6,18) np np' m show ?thesis by auto
+    next
+      case 2
+      have nn': "n' = n" by fact
       then have nn: "\<not> n' < n \<and> \<not> n < n'" by arith
       from "4.hyps"(16,18)[of n n' n]
         "4.hyps"(13,14)[of n "Suc n'" n]
@@ -562,8 +557,9 @@
         "?cnp *\<^sub>p c' = 0\<^sub>p \<longleftrightarrow> c' = 0\<^sub>p"
         "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p"
         by (auto simp add: min_def)
-      {
-        assume mn: "m = n"
+      show ?thesis
+      proof (cases "m = n")
+        case mn: True
         from "4.hyps"(17,18)[OF norm(1,4), of n]
           "4.hyps"(13,15)[OF norm(1,2), of n] norm nn' mn
         have degs:
@@ -583,11 +579,9 @@
         from "4.hyps"(16-18)[OF norm(1,4), of n]
           "4.hyps"(13-15)[OF norm(1,2), of n]
           mn norm m nn' deg
-        have ?eq by simp
-      }
-      moreover
-      {
-        assume mn: "m \<noteq> n"
+        show ?thesis by simp
+      next
+        case mn: False
         then have mn': "m < n"
           using m np by auto
         from nn' m np have max1: "m \<le> max n n"
@@ -605,11 +599,10 @@
         with "4.hyps"(16-18)[OF norm(1,4) min1]
           "4.hyps"(13-15)[OF norm(1,2) min2]
           degreen_0[OF norm(3) mn']
-        have ?eq using nn' mn np np' by clarsimp
-      }
-      ultimately have ?eq by blast
-    }
-    ultimately show ?eq by blast
+          nn' mn np np'
+        show ?thesis by clarsimp
+      qed
+    qed
   }
   {
     case (2 n0 n1)
@@ -619,9 +612,9 @@
       by simp_all
     then have mn: "m \<le> n" by simp
     let ?c0p = "CN 0\<^sub>p n (?cnp *\<^sub>p p')"
-    {
-      assume C: "?cnp *\<^sub>p c' +\<^sub>p ?c0p = 0\<^sub>p" "n' = n"
-      then have nn: "\<not> n' < n \<and> \<not> n < n'"
+    have False if C: "?cnp *\<^sub>p c' +\<^sub>p ?c0p = 0\<^sub>p" "n' = n"
+    proof -
+      from C have nn: "\<not> n' < n \<and> \<not> n < n'"
         by simp
       from "4.hyps"(16-18) [of n n n]
         "4.hyps"(13-15)[of n "Suc n" n]
@@ -643,8 +636,8 @@
       have degneq: "degreen (?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
         using norm by simp
       from polyadd_eq_const_degreen[OF norm(3) cn C(1), where m="n"] degneq
-      have False by simp
-    }
+      show ?thesis by simp
+    qed
     then show ?case using "4.hyps" by clarsimp
   }
 qed auto
@@ -747,22 +740,19 @@
 
 text \<open>polypow is a power function and preserves normal forms\<close>
 
-lemma polypow[simp]:
-  "Ipoly bs (polypow n p) = (Ipoly bs p :: 'a::{field_char_0,field}) ^ n"
+lemma polypow[simp]: "Ipoly bs (polypow n p) = (Ipoly bs p :: 'a::{field_char_0,field}) ^ n"
 proof (induct n rule: polypow.induct)
   case 1
-  then show ?case
-    by simp
+  then show ?case by simp
 next
   case (2 n)
   let ?q = "polypow ((Suc n) div 2) p"
   let ?d = "polymul ?q ?q"
-  have "odd (Suc n) \<or> even (Suc n)"
-    by simp
-  moreover
-  {
-    assume odd: "odd (Suc n)"
-    have th: "(Suc (Suc (Suc 0) * (Suc n div Suc (Suc 0)))) = Suc n div 2 + Suc n div 2 + 1"
+  consider "odd (Suc n)" | "even (Suc n)" by auto
+  then show ?case
+  proof cases
+    case odd: 1
+    have *: "(Suc (Suc (Suc 0) * (Suc n div Suc (Suc 0)))) = Suc n div 2 + Suc n div 2 + 1"
       by arith
     from odd have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs (polymul p ?d)"
       by (simp add: Let_def)
@@ -771,21 +761,19 @@
     also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2 + 1)"
       by (simp only: power_add power_one_right) simp
     also have "\<dots> = (Ipoly bs p) ^ (Suc (Suc (Suc 0) * (Suc n div Suc (Suc 0))))"
-      by (simp only: th)
-    finally have ?case unfolding numeral_2_eq_2 [symmetric]
-    using odd_two_times_div_two_nat [OF odd] by simp
-  }
-  moreover
-  {
-    assume even: "even (Suc n)"
+      by (simp only: *)
+    finally show ?thesis
+      unfolding numeral_2_eq_2 [symmetric]
+      using odd_two_times_div_two_nat [OF odd] by simp
+  next
+    case even: 2
     from even have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs ?d"
       by (simp add: Let_def)
     also have "\<dots> = (Ipoly bs p) ^ (2 * (Suc n div 2))"
       using "2.hyps" by (simp only: mult_2 power_add) simp
-    finally have ?case using even_two_times_div_two [OF even]
-      by simp
-  }
-  ultimately show ?case by blast
+    finally show ?thesis
+      using even_two_times_div_two [OF even] by simp
+  qed
 qed
 
 lemma polypow_normh:
@@ -798,14 +786,13 @@
   case (2 k n)
   let ?q = "polypow (Suc k div 2) p"
   let ?d = "polymul ?q ?q"
-  from 2 have th1: "isnpolyh ?q n" and th2: "isnpolyh p n"
+  from 2 have *: "isnpolyh ?q n" and **: "isnpolyh p n"
     by blast+
-  from polymul_normh[OF th1 th1] have dn: "isnpolyh ?d n"
+  from polymul_normh[OF * *] have dn: "isnpolyh ?d n"
     by simp
-  from polymul_normh[OF th2 dn] have on: "isnpolyh (polymul p ?d) n"
+  from polymul_normh[OF ** dn] have on: "isnpolyh (polymul p ?d) n"
     by simp
   from dn on show ?case by (simp, unfold Let_def) auto
-    
 qed
 
 lemma polypow_norm:
@@ -815,8 +802,7 @@
 
 text \<open>Finally the whole normalization\<close>
 
-lemma polynate [simp]:
-  "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0,field})"
+lemma polynate [simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0,field})"
   by (induct p rule:polynate.induct) auto
 
 lemma polynate_norm[simp]:
@@ -828,7 +814,6 @@
 
 text \<open>shift1\<close>
 
-
 lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)"
   by (simp add: shift1_def)
 
@@ -909,7 +894,7 @@
 
 lemma decrpoly_normh: "isnpolyh p n0 \<Longrightarrow> polybound0 p \<Longrightarrow> isnpolyh (decrpoly p) (n0 - 1)"
   apply (induct p arbitrary: n0)
-  apply auto
+         apply auto
   apply atomize
   apply (rename_tac nat a b, erule_tac x = "Suc nat" in allE)
   apply auto
@@ -945,15 +930,15 @@
   by (induct p arbitrary: n rule: degree.induct) (auto simp add: isnpolyh_polybound0)
 
 primrec maxindex :: "poly \<Rightarrow> nat"
-where
-  "maxindex (Bound n) = n + 1"
-| "maxindex (CN c n p) = max  (n + 1) (max (maxindex c) (maxindex p))"
-| "maxindex (Add p q) = max (maxindex p) (maxindex q)"
-| "maxindex (Sub p q) = max (maxindex p) (maxindex q)"
-| "maxindex (Mul p q) = max (maxindex p) (maxindex q)"
-| "maxindex (Neg p) = maxindex p"
-| "maxindex (Pw p n) = maxindex p"
-| "maxindex (C x) = 0"
+  where
+    "maxindex (Bound n) = n + 1"
+  | "maxindex (CN c n p) = max  (n + 1) (max (maxindex c) (maxindex p))"
+  | "maxindex (Add p q) = max (maxindex p) (maxindex q)"
+  | "maxindex (Sub p q) = max (maxindex p) (maxindex q)"
+  | "maxindex (Mul p q) = max (maxindex p) (maxindex q)"
+  | "maxindex (Neg p) = maxindex p"
+  | "maxindex (Pw p n) = maxindex p"
+  | "maxindex (C x) = 0"
 
 definition wf_bs :: "'a list \<Rightarrow> poly \<Rightarrow> bool"
   where "wf_bs bs p \<longleftrightarrow> length bs \<ge> maxindex p"
@@ -964,25 +949,21 @@
   show ?case
   proof
     fix x
-    assume xc: "x \<in> set (coefficients (CN c 0 p))"
-    then have "x = c \<or> x \<in> set (coefficients p)"
-      by simp
-    moreover
-    {
-      assume "x = c"
-      then have "wf_bs bs x"
-        using "1.prems" unfolding wf_bs_def by simp
-    }
-    moreover
-    {
-      assume H: "x \<in> set (coefficients p)"
+    assume "x \<in> set (coefficients (CN c 0 p))"
+    then consider "x = c" | "x \<in> set (coefficients p)"
+      by auto
+    then show "wf_bs bs x"
+    proof cases
+      case prems: 1
+      then show ?thesis
+        using "1.prems" by (simp add: wf_bs_def)
+    next
+      case prems: 2
       from "1.prems" have "wf_bs bs p"
-        unfolding wf_bs_def by simp
-      with "1.hyps" H have "wf_bs bs x"
+        by (simp add: wf_bs_def)
+      with "1.hyps" prems show ?thesis
         by blast
-    }
-    ultimately show "wf_bs bs x"
-      by blast
+    qed
   qed
 qed simp_all
 
@@ -990,7 +971,7 @@
   by (induct p rule: coefficients.induct) auto
 
 lemma wf_bs_I: "wf_bs bs p \<Longrightarrow> Ipoly (bs @ bs') p = Ipoly bs p"
-  unfolding wf_bs_def by (induct p) (auto simp add: nth_append)
+  by (induct p) (auto simp add: nth_append wf_bs_def)
 
 lemma take_maxindex_wf:
   assumes wf: "wf_bs bs p"
@@ -1012,10 +993,10 @@
   by (induct p) auto
 
 lemma wf_bs_insensitive: "length bs = length bs' \<Longrightarrow> wf_bs bs p = wf_bs bs' p"
-  unfolding wf_bs_def by simp
+  by (simp add: wf_bs_def)
 
 lemma wf_bs_insensitive': "wf_bs (x # bs) p = wf_bs (y # bs) p"
-  unfolding wf_bs_def by simp
+  by (simp add: wf_bs_def)
 
 lemma wf_bs_coefficients': "\<forall>c \<in> set (coefficients p). wf_bs bs c \<Longrightarrow> wf_bs (x # bs) p"
   by (induct p rule: coefficients.induct) (auto simp add: wf_bs_def)
@@ -1030,31 +1011,28 @@
   unfolding wf_bs_def by (induct p rule: decrpoly.induct) auto
 
 lemma length_le_list_ex: "length xs \<le> n \<Longrightarrow> \<exists>ys. length (xs @ ys) = n"
-  apply (rule exI[where x="replicate (n - length xs) z" for z])
-  apply simp
-  done
+  by (rule exI[where x="replicate (n - length xs) z" for z]) simp
 
 lemma isnpolyh_Suc_const: "isnpolyh p (Suc n) \<Longrightarrow> isconstant p"
   apply (cases p)
-  apply auto
+         apply auto
   apply (rename_tac nat a, case_tac "nat")
-  apply simp_all
+   apply simp_all
   done
 
 lemma wf_bs_polyadd: "wf_bs bs p \<and> wf_bs bs q \<longrightarrow> wf_bs bs (p +\<^sub>p q)"
-  unfolding wf_bs_def by (induct p q rule: polyadd.induct) (auto simp add: Let_def)
+  by (induct p q rule: polyadd.induct) (auto simp add: Let_def wf_bs_def)
 
 lemma wf_bs_polyul: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p *\<^sub>p q)"
-  unfolding wf_bs_def
   apply (induct p q arbitrary: bs rule: polymul.induct)
-  apply (simp_all add: wf_bs_polyadd)
+                      apply (simp_all add: wf_bs_polyadd wf_bs_def)
   apply clarsimp
   apply (rule wf_bs_polyadd[unfolded wf_bs_def, rule_format])
   apply auto
   done
 
 lemma wf_bs_polyneg: "wf_bs bs p \<Longrightarrow> wf_bs bs (~\<^sub>p p)"
-  unfolding wf_bs_def by (induct p rule: polyneg.induct) auto
+  by (induct p rule: polyneg.induct) (auto simp: wf_bs_def)
 
 lemma wf_bs_polysub: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p -\<^sub>p q)"
   unfolding polysub_def split_def fst_conv snd_conv
@@ -1087,16 +1065,15 @@
 proof -
   let ?cf = "set (coefficients p)"
   from coefficients_normh[OF np] have cn_norm: "\<forall> q\<in> ?cf. isnpolyh q n0" .
-  {
-    fix q
-    assume q: "q \<in> ?cf"
-    from q cn_norm have th: "isnpolyh q n0"
+  have "polybound0 q" if "q \<in> ?cf" for q
+  proof -
+    from that cn_norm have *: "isnpolyh q n0"
       by blast
-    from coefficients_isconst[OF np] q have "isconstant q"
+    from coefficients_isconst[OF np] that have "isconstant q"
       by blast
-    with isconstant_polybound0[OF th] have "polybound0 q"
+    with isconstant_polybound0[OF *] show ?thesis
       by blast
-  }
+  qed
   then have "\<forall>q \<in> ?cf. polybound0 q" ..
   then have "\<forall>q \<in> ?cf. Ipoly (x # bs) q = Ipoly bs (decrpoly q)"
     using polybound0_I[where b=x and bs=bs and b'=y] decrpoly[where x=x and bs=bs]
@@ -1124,9 +1101,9 @@
   using assms
   unfolding polypoly_def
   apply (cases p)
-  apply auto
+         apply auto
   apply (rename_tac nat a, case_tac nat)
-  apply auto
+   apply auto
   done
 
 lemma head_isnpolyh: "isnpolyh p n0 \<Longrightarrow> isnpolyh (head p) n0"
@@ -1149,16 +1126,13 @@
 proof (induct "maxindex p" arbitrary: p n0 rule: less_induct)
   case less
   note np = \<open>isnpolyh p n0\<close> and zp = \<open>\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)\<close>
-  {
-    assume nz: "maxindex p = 0"
-    then obtain c where "p = C c"
-      using np by (cases p) auto
-    with zp np have "p = 0\<^sub>p"
-      unfolding wf_bs_def by simp
-  }
-  moreover
-  {
-    assume nz: "maxindex p \<noteq> 0"
+  show "p = 0\<^sub>p"
+  proof (cases "maxindex p = 0")
+    case True
+    with np obtain c where "p = C c" by (cases p) auto
+    with zp np show ?thesis by (simp add: wf_bs_def)
+  next
+    case nz: False
     let ?h = "head p"
     let ?hd = "decrpoly ?h"
     let ?ihd = "maxindex ?hd"
@@ -1173,13 +1147,13 @@
       by auto
     with decr_maxindex[OF h(2)] nz have ihd_lt_n: "?ihd < maxindex p"
       by auto
-    {
-      fix bs :: "'a list"
-      assume bs: "wf_bs bs ?hd"
+
+    have "\<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" if bs: "wf_bs bs ?hd" for bs :: "'a list"
+    proof -
       let ?ts = "take ?ihd bs"
       let ?rs = "drop ?ihd bs"
-      have ts: "wf_bs ?ts ?hd"
-        using bs unfolding wf_bs_def by simp
+      from bs have ts: "wf_bs ?ts ?hd"
+        by (simp add: wf_bs_def)
       have bs_ts_eq: "?ts @ ?rs = bs"
         by simp
       from wf_bs_decrpoly[OF ts] have tsh: " \<forall>x. wf_bs (x # ?ts) ?h"
@@ -1189,7 +1163,7 @@
       with length_le_list_ex obtain xs where xs: "length ((x # ?ts) @ xs) = maxindex p"
         by blast
       then have "\<forall>x. wf_bs ((x # ?ts) @ xs) p"
-        unfolding wf_bs_def by simp
+        by (simp add: wf_bs_def)
       with zp have "\<forall>x. Ipoly ((x # ?ts) @ xs) p = 0"
         by blast
       then have "\<forall>x. Ipoly (x # (?ts @ xs)) p = 0"
@@ -1202,24 +1176,22 @@
       then have "\<forall>c \<in> set (coefficients p). Ipoly (?ts @ xs) (decrpoly c) = 0"
         using poly_zero[where ?'a='a] by (simp add: polypoly'_def)
       with coefficients_head[of p, symmetric]
-      have th0: "Ipoly (?ts @ xs) ?hd = 0"
+      have *: "Ipoly (?ts @ xs) ?hd = 0"
         by simp
       from bs have wf'': "wf_bs ?ts ?hd"
-        unfolding wf_bs_def by simp
-      with th0 wf_bs_I[of ?ts ?hd xs] have "Ipoly ?ts ?hd = 0"
+        by (simp add: wf_bs_def)
+      with * wf_bs_I[of ?ts ?hd xs] have "Ipoly ?ts ?hd = 0"
         by simp
-      with wf'' wf_bs_I[of ?ts ?hd ?rs] bs_ts_eq have "\<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0"
+      with wf'' wf_bs_I[of ?ts ?hd ?rs] bs_ts_eq show ?thesis
         by simp
-    }
+    qed
     then have hdz: "\<forall>bs. wf_bs bs ?hd \<longrightarrow> \<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)"
       by blast
     from less(1)[OF ihd_lt_n nhd] hdz have "?hd = 0\<^sub>p"
       by blast
     then have "?h = 0\<^sub>p" by simp
-    with head_nz[OF np] have "p = 0\<^sub>p" by simp
-  }
-  ultimately show "p = 0\<^sub>p"
-    by blast
+    with head_nz[OF np] show ?thesis by simp
+  qed
 qed
 
 lemma isnpolyh_unique:
@@ -1227,7 +1199,7 @@
     and nq: "isnpolyh q n1"
   shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{field_char_0,field,power})) \<longleftrightarrow> p = q"
 proof auto
-  assume H: "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a) = \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>"
+  assume "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a) = \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>"
   then have "\<forall>bs.\<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup>= (0::'a)"
     by simp
   then have "\<forall>bs. wf_bs bs (p -\<^sub>p q) \<longrightarrow> \<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)"
@@ -1237,7 +1209,7 @@
 qed
 
 
-text \<open>consequences of unicity on the algorithms for polynomial normalization\<close>
+text \<open>Consequences of unicity on the algorithms for polynomial normalization.\<close>
 
 lemma polyadd_commute:
   assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
@@ -1311,7 +1283,7 @@
   unfolding poly_nate_polypoly' by auto
 
 
-subsection \<open>heads, degrees and all that\<close>
+subsection \<open>Heads, degrees and all that\<close>
 
 lemma degree_eq_degreen0: "degree p = degreen p 0"
   by (induct p rule: degree.induct) simp_all
@@ -1321,9 +1293,9 @@
   shows "degree (polyneg p) = degree p"
   apply (induct p rule: polyneg.induct)
   using assms
-  apply simp_all
+         apply simp_all
   apply (case_tac na)
-  apply auto
+   apply auto
   done
 
 lemma degree_polyadd:
@@ -1395,50 +1367,43 @@
     by simp
   from H have pnh: "isnpolyh p n" and p'nh: "isnpolyh p' n'"
     by auto
-  have "n = n' \<or> n < n' \<or> n > n'"
+  consider "n = n'" | "n < n'" | "n > n'"
     by arith
-  moreover
-  {
-    assume nn': "n = n'"
-    have "n = 0 \<or> n > 0" by arith
-    moreover
-    {
-      assume nz: "n = 0"
-      then have ?case using 4 nn'
+  then show ?case
+  proof cases
+    case nn': 1
+    consider "n = 0" | "n > 0" by arith
+    then show ?thesis
+    proof cases
+      case 1
+      with 4 nn' show ?thesis
         by (auto simp add: Let_def degcmc')
-    }
-    moreover
-    {
-      assume nz: "n > 0"
-      with nn' H(3) have  cc': "c = c'" and pp': "p = p'"
-        by (cases n, auto)+
-      then have ?case
+    next
+      case 2
+      with nn' H(3) have "c = c'" and "p = p'"
+        by (cases n; auto)+
+      with nn' 4 show ?thesis
         using polysub_same_0[OF p'nh, simplified polysub_def split_def fst_conv snd_conv]
         using polysub_same_0[OF c'nh, simplified polysub_def]
-        using nn' 4 by (simp add: Let_def)
-    }
-    ultimately have ?case by blast
-  }
-  moreover
-  {
-    assume nn': "n < n'"
+        by (simp add: Let_def)
+    qed
+  next
+    case nn': 2
     then have n'p: "n' > 0"
       by simp
     then have headcnp':"head (CN c' n' p') = CN c' n' p'"
       by (cases n') simp_all
-    have degcnp': "degree (CN c' n' p') = 0"
+    with 4 nn' have degcnp': "degree (CN c' n' p') = 0"
       and degcnpeq: "degree (CN c n p) = degree (CN c' n' p')"
-      using 4 nn' by (cases n', simp_all)
+      by (cases n', simp_all)
     then have "n > 0"
       by (cases n) simp_all
     then have headcnp: "head (CN c n p) = CN c n p"
       by (cases n) auto
-    from H(3) headcnp headcnp' nn' have ?case
+    from H(3) headcnp headcnp' nn' show ?thesis
       by auto
-  }
-  moreover
-  {
-    assume nn': "n > n'"
+  next
+    case nn': 3
     then have np: "n > 0" by simp
     then have headcnp:"head (CN c n p) = CN c n p"
       by (cases n) simp_all
@@ -1450,15 +1415,14 @@
       by (cases n') simp_all
     then have headcnp': "head (CN c' n' p') = CN c' n' p'"
       by (cases n') auto
-    from H(3) headcnp headcnp' nn' have ?case by auto
-  }
-  ultimately show ?case by blast
+    from H(3) headcnp headcnp' nn' show ?thesis by auto
+  qed
 qed auto
 
 lemma shift1_head : "isnpolyh p n0 \<Longrightarrow> head (shift1 p) = head p"
   by (induct p arbitrary: n0 rule: head.induct) (simp_all add: shift1_def)
 
-lemma funpow_shift1_head: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> head (funpow k shift1 p) = head p"
+lemma funpow_shift1_head: "isnpolyh p n0 \<Longrightarrow> p \<noteq> 0\<^sub>p \<Longrightarrow> head (funpow k shift1 p) = head p"
 proof (induct k arbitrary: n0 p)
   case 0
   then show ?case
@@ -1503,13 +1467,13 @@
   shows "head (p +\<^sub>p q) = (if degree p < degree q then head q else head p)"
   using np nq deg
   apply (induct p q arbitrary: n0 n1 rule: polyadd.induct)
-  apply simp_all
-  apply (case_tac n', simp, simp)
-  apply (case_tac n, simp, simp)
+                      apply simp_all
+    apply (case_tac n', simp, simp)
+   apply (case_tac n, simp, simp)
   apply (case_tac n, case_tac n', simp add: Let_def)
-  apply (auto simp add: polyadd_eq_const_degree)[2]
-  apply (metis head_nz)
-  apply (metis head_nz)
+    apply (auto simp add: polyadd_eq_const_degree)[2]
+    apply (metis head_nz)
+   apply (metis head_nz)
   apply (metis degree.simps(9) gr0_conv_Suc head.simps(1) less_Suc0 not_less_eq)
   done
 
@@ -1533,67 +1497,61 @@
   then have norm: "isnpolyh p n" "isnpolyh c (Suc n)" "isnpolyh p' n'" "isnpolyh c' (Suc n')"
     "isnpolyh (CN c n p) n" "isnpolyh (CN c' n' p') n'"
     by simp_all
-  have "n < n' \<or> n' < n \<or> n = n'" by arith
-  moreover
-  {
-    assume nn': "n < n'"
-    then have ?case
+  consider "n < n'" | "n' < n" | "n' = n" by arith
+  then show ?case
+  proof cases
+    case nn': 1
+    then show ?thesis
       using norm "4.hyps"(2)[OF norm(1,6)] "4.hyps"(1)[OF norm(2,6)]
       apply simp
       apply (cases n)
-      apply simp
+       apply simp
       apply (cases n')
-      apply simp_all
+       apply simp_all
       done
-  }
-  moreover {
-    assume nn': "n'< n"
-    then have ?case
+  next
+    case nn': 2
+    then show ?thesis
       using norm "4.hyps"(6) [OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)]
       apply simp
       apply (cases n')
-      apply simp
+       apply simp
       apply (cases n)
-      apply auto
+       apply auto
       done
-  }
-  moreover
-  {
-    assume nn': "n' = n"
+  next
+    case nn': 3
     from nn' polymul_normh[OF norm(5,4)]
-    have ncnpc':"isnpolyh (CN c n p *\<^sub>p c') n" by (simp add: min_def)
+    have ncnpc': "isnpolyh (CN c n p *\<^sub>p c') n" by (simp add: min_def)
     from nn' polymul_normh[OF norm(5,3)] norm
-    have ncnpp':"isnpolyh (CN c n p *\<^sub>p p') n" by simp
+    have ncnpp': "isnpolyh (CN c n p *\<^sub>p p') n" by simp
     from nn' ncnpp' polymul_eq0_iff[OF norm(5,3)] norm(6)
-    have ncnpp0':"isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+    have ncnpp0': "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
     from polyadd_normh[OF ncnpc' ncnpp0']
     have nth: "isnpolyh ((CN c n p *\<^sub>p c') +\<^sub>p (CN 0\<^sub>p n (CN c n p *\<^sub>p p'))) n"
       by (simp add: min_def)
-    {
-      assume np: "n > 0"
+    consider "n > 0" | "n = 0" by auto
+    then show ?thesis
+    proof cases
+      case np: 1
       with nn' head_isnpolyh_Suc'[OF np nth]
         head_isnpolyh_Suc'[OF np norm(5)] head_isnpolyh_Suc'[OF np norm(6)[simplified nn']]
-      have ?case by simp
-    }
-    moreover
-    {
-      assume nz: "n = 0"
+      show ?thesis by simp
+    next
+      case nz: 2
       from polymul_degreen[OF norm(5,4), where m="0"]
         polymul_degreen[OF norm(5,3), where m="0"] nn' nz degree_eq_degreen0
-      norm(5,6) degree_npolyhCN[OF norm(6)]
-    have dth: "degree (CN c 0 p *\<^sub>p c') < degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))"
-      by simp
-    then have dth': "degree (CN c 0 p *\<^sub>p c') \<noteq> degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))"
-      by simp
-    from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth
-    have ?case
-      using norm "4.hyps"(6)[OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)] nn' nz
-      by simp
-    }
-    ultimately have ?case
-      by (cases n) auto
-  }
-  ultimately show ?case by blast
+        norm(5,6) degree_npolyhCN[OF norm(6)]
+      have dth: "degree (CN c 0 p *\<^sub>p c') < degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))"
+        by simp
+      then have dth': "degree (CN c 0 p *\<^sub>p c') \<noteq> degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))"
+        by simp
+      from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth
+      show ?thesis
+        using norm "4.hyps"(6)[OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)] nn' nz
+        by simp
+    qed
+  qed
 qed simp_all
 
 lemma degree_coefficients: "degree p = length (coefficients p) - 1"
@@ -1663,52 +1621,50 @@
     by (simp add: isnpoly_def)
   from polypow_normh [OF head_isnpolyh[OF np0], where k="k' - k"] ap
   have nakk':"isnpolyh ?akk' 0" by blast
-  {
-    assume sz: "s = 0\<^sub>p"
-    then have ?ths
-      using np polydivide_aux.simps
-      apply clarsimp
+  show ?ths
+  proof (cases "s = 0\<^sub>p")
+    case True
+    with np show ?thesis
+      apply (clarsimp simp: polydivide_aux.simps)
       apply (rule exI[where x="0\<^sub>p"])
       apply simp
       done
-  }
-  moreover
-  {
-    assume sz: "s \<noteq> 0\<^sub>p"
-    {
-      assume dn: "degree s < n"
-      then have "?ths"
+  next
+    case sz: False
+    show ?thesis
+    proof (cases "degree s < n")
+      case True
+      then show ?thesis
         using ns ndp np polydivide_aux.simps
         apply auto
         apply (rule exI[where x="0\<^sub>p"])
         apply simp
         done
-    }
-    moreover
-    {
-      assume dn': "\<not> degree s < n"
+    next
+      case dn': False
       then have dn: "degree s \<ge> n"
         by arith
       have degsp': "degree s = degree ?p'"
         using dn ndp funpow_shift1_degree[where k = "degree s - n" and p="p"]
         by simp
-      {
-        assume ba: "?b = a"
+      show ?thesis
+      proof (cases "?b = a")
+        case ba: True
         then have headsp': "head s = head ?p'"
           using ap headp' by simp
         have nr: "isnpolyh (s -\<^sub>p ?p') 0"
           using polysub_normh[OF ns np'] by simp
         from degree_polysub_samehead[OF ns np' headsp' degsp']
-        have "degree (s -\<^sub>p ?p') < degree s \<or> s -\<^sub>p ?p' = 0\<^sub>p"
-          by simp
-        moreover
-        {
-          assume deglt:"degree (s -\<^sub>p ?p') < degree s"
+        consider "degree (s -\<^sub>p ?p') < degree s" | "s -\<^sub>p ?p' = 0\<^sub>p" by auto
+        then show ?thesis
+        proof cases
+          case deglt: 1
           from polydivide_aux.simps sz dn' ba
           have eq: "polydivide_aux a n p k s = polydivide_aux a n p k (s -\<^sub>p ?p')"
             by (simp add: Let_def)
-          {
-            assume h1: "polydivide_aux a n p k s = (k', r)"
+          have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+            if h1: "polydivide_aux a n p k s = (k', r)"
+          proof -
             from less(1)[OF deglt nr, of k k' r] trans[OF eq[symmetric] h1]
             have kk': "k \<le> k'"
               and nr: "\<exists>nr. isnpolyh r nr"
@@ -1753,54 +1709,45 @@
             have "a ^\<^sub>p (k' - k) *\<^sub>p s =
               p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 (1)\<^sub>p) +\<^sub>p q) +\<^sub>p r"
               by blast
-            then have ?qths using nq'
+            with nq' have ?qths
               apply (rule_tac x="(a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 (1)\<^sub>p) +\<^sub>p q" in exI)
               apply (rule_tac x="0" in exI)
               apply simp
               done
-            with kk' nr dr have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and>
-              (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+            with kk' nr dr show ?thesis
               by blast
-          }
-          then have ?ths by blast
-        }
-        moreover
-        {
-          assume spz:"s -\<^sub>p ?p' = 0\<^sub>p"
+          qed
+          then show ?thesis by blast
+        next
+          case spz: 2
           from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{field_char_0,field}"]
           have "\<forall>bs:: 'a::{field_char_0,field} list. Ipoly bs s = Ipoly bs ?p'"
             by simp
-          then have "\<forall>bs:: 'a::{field_char_0,field} list. Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)"
-            using np nxdn
-            apply simp
-            apply (simp only: funpow_shift1_1)
-            apply simp
-            done
+          with np nxdn have "\<forall>bs:: 'a::{field_char_0,field} list. Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)"
+            by (simp only: funpow_shift1_1) simp
           then have sp': "s = ?xdn *\<^sub>p p"
             using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]]
             by blast
-          {
-            assume h1: "polydivide_aux a n p k s = (k', r)"
-            from polydivide_aux.simps sz dn' ba
-            have eq: "polydivide_aux a n p k s = polydivide_aux a n p k (s -\<^sub>p ?p')"
-              by (simp add: Let_def)
+          have ?thesis if h1: "polydivide_aux a n p k s = (k', r)"
+          proof -
+            from sz dn' ba
+            have "polydivide_aux a n p k s = polydivide_aux a n p k (s -\<^sub>p ?p')"
+              by (simp add: Let_def polydivide_aux.simps)
             also have "\<dots> = (k,0\<^sub>p)"
-              using polydivide_aux.simps spz by simp
+              using spz by (simp add: polydivide_aux.simps)
             finally have "(k', r) = (k, 0\<^sub>p)"
-              using h1 by simp
+              by (simp add: h1)
             with sp'[symmetric] ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]]
-              polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?ths
+              polyadd_0(2)[OF polymul_normh[OF np nxdn]] show ?thesis
               apply auto
               apply (rule exI[where x="?xdn"])
               apply (auto simp add: polymul_commute[of p])
               done
-          }
-        }
-        ultimately have ?ths by blast
-      }
-      moreover
-      {
-        assume ba: "?b \<noteq> a"
+          qed
+          then show ?thesis by blast
+        qed
+      next
+        case ba: False
         from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns]
           polymul_normh[OF head_isnpolyh[OF ns] np']]
         have nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0"
@@ -1811,7 +1758,8 @@
             funpow_shift1_nz[OF pnz]
           by simp_all
         from polymul_head_polyeq[OF head_isnpolyh[OF np] ns] head_nz[OF np] sz ap head_head[OF np] pnz
-          polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz funpow_shift1_nz[OF pnz, where n="degree s - n"]
+          polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz
+          funpow_shift1_nz[OF pnz, where n="degree s - n"]
         have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')"
           using head_head[OF ns] funpow_shift1_head[OF np pnz]
             polymul_commute[OF head_isnpolyh[OF np] head_isnpolyh[OF ns]]
@@ -1823,14 +1771,22 @@
           ndp dn
         have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p')"
           by (simp add: degree_eq_degreen0[symmetric] funpow_shift1_degree)
-        {
-          assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < degree s"
+
+        consider "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < degree s" | "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p"
+          using degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns]
+            polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth]
+            polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+            head_nz[OF np] pnz sz ap[symmetric]
+          by (auto simp add: degree_eq_degreen0[symmetric])
+        then show ?thesis
+        proof cases
+          case dth: 1
           from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np] ns]
             polymul_normh[OF head_isnpolyh[OF ns]np']] ap
           have nasbp': "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0"
             by simp
-          {
-            assume h1:"polydivide_aux a n p k s = (k', r)"
+          have ?thesis if h1: "polydivide_aux a n p k s = (k', r)"
+          proof -
             from h1 polydivide_aux.simps sz dn' ba
             have eq:"polydivide_aux a n p (Suc k) ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = (k',r)"
               by (simp add: Let_def)
@@ -1843,8 +1799,10 @@
               by auto
             from kk' have kk'': "Suc (k' - Suc k) = k' - k"
               by arith
-            {
-              fix bs :: "'a::{field_char_0,field} list"
+            have "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+                Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
+              for bs :: "'a::{field_char_0,field} list"
+            proof -
               from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
               have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)"
                 by simp
@@ -1854,11 +1812,10 @@
               then have "Ipoly bs a ^ (k' - k)  * Ipoly bs s =
                 Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
                 by (simp add: kk'' funpow_shift1_1[where n="degree s - n" and p="p"])
-              then have "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
-                Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
+              then show ?thesis
                 by (simp add: field_simps)
-            }
-            then have ieq:"\<forall>bs :: 'a::{field_char_0,field} list.
+            qed
+            then have ieq: "\<forall>bs :: 'a::{field_char_0,field} list.
                 Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
                 Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)"
               by auto
@@ -1869,62 +1826,53 @@
             from ieq isnpolyh_unique[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - k"] ns, simplified ap] polyadd_normh[OF polymul_normh[OF np nqw] nr]]
             have asth: "(a ^\<^sub>p (k' - k) *\<^sub>p s) = p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r"
               by blast
-            from dr kk' nr h1 asth nqw have ?ths
+            from dr kk' nr h1 asth nqw show ?thesis
               apply simp
               apply (rule conjI)
               apply (rule exI[where x="nr"], simp)
               apply (rule exI[where x="(q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn))"], simp)
               apply (rule exI[where x="0"], simp)
               done
-          }
-          then have ?ths by blast
-        }
-        moreover
-        {
-          assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p"
-          {
+          qed
+          then show ?thesis by blast
+        next
+          case spz: 2
+          have hth: "\<forall>bs :: 'a::{field_char_0,field} list.
+            Ipoly bs (a *\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))"
+          proof
             fix bs :: "'a::{field_char_0,field} list"
             from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
             have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'"
               by simp
             then have "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p"
               by (simp add: funpow_shift1_1[where n="degree s - n" and p="p"])
-            then have "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))"
+            then show "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))"
               by simp
-          }
-          then have hth: "\<forall>bs :: 'a::{field_char_0,field} list.
-            Ipoly bs (a *\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
+          qed
           from hth have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)"
             using isnpolyh_unique[where ?'a = "'a::{field_char_0,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns]
                     polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]],
               simplified ap]
             by simp
-          {
-            assume h1: "polydivide_aux a n p k s = (k', r)"
+          have ?ths if h1: "polydivide_aux a n p k s = (k', r)"
+          proof -
             from h1 sz ba dn' spz polydivide_aux.simps polydivide_aux.simps
             have "(k', r) = (Suc k, 0\<^sub>p)"
               by (simp add: Let_def)
             with h1 np head_isnpolyh[OF np, simplified ap] ns polymul_normh[OF head_isnpolyh[OF ns] nxdn]
               polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]] asq
-            have ?ths
+            show ?thesis
               apply (clarsimp simp add: Let_def)
               apply (rule exI[where x="?b *\<^sub>p ?xdn"])
               apply simp
               apply (rule exI[where x="0"], simp)
               done
-          }
-          then have ?ths by blast
-        }
-        ultimately have ?ths
-          using degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth] polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
-            head_nz[OF np] pnz sz ap[symmetric]
-          by (auto simp add: degree_eq_degreen0[symmetric])
-      }
-      ultimately have ?ths by blast
-    }
-    ultimately have ?ths by blast
-  }
-  ultimately show ?ths by blast
+          qed
+          then show ?thesis by blast
+        qed
+      qed
+    qed
+  qed
 qed
 
 lemma polydivide_properties:
@@ -1965,14 +1913,14 @@
   shows "pnormal (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0"
 proof
   let ?p = "polypoly bs p"
-  assume H: "pnormal ?p"
-  have csz: "coefficients p \<noteq> []"
+  assume *: "pnormal ?p"
+  have "coefficients p \<noteq> []"
     using assms by (cases p) auto
-  from coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] pnormal_last_nonzero[OF H]
+  from coefficients_head[of p] last_map[OF this, of "Ipoly bs"] pnormal_last_nonzero[OF *]
   show "Ipoly bs (head p) \<noteq> 0"
     by (simp add: polypoly_def)
 next
-  assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+  assume *: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
   let ?p = "polypoly bs p"
   have csz: "coefficients p \<noteq> []"
     using assms by (cases p) auto
@@ -1980,7 +1928,7 @@
     by (simp add: polypoly_def)
   then have lg: "length ?p > 0"
     by simp
-  from h coefficients_head[of p] last_map[OF csz, of "Ipoly bs"]
+  from * coefficients_head[of p] last_map[OF csz, of "Ipoly bs"]
   have lz: "last ?p \<noteq> 0"
     by (simp add: polypoly_def)
   from pnormal_last_length[OF lg lz] show "pnormal ?p" .
@@ -1999,73 +1947,71 @@
   shows "nonconstant (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0"
 proof
   let ?p = "polypoly bs p"
-  assume nc: "nonconstant ?p"
-  from isnonconstant_pnormal_iff[OF assms, of bs] nc
-  show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+  assume "nonconstant ?p"
+  with isnonconstant_pnormal_iff[OF assms, of bs] show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
     unfolding nonconstant_def by blast
 next
   let ?p = "polypoly bs p"
-  assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
-  from isnonconstant_pnormal_iff[OF assms, of bs] h
-  have pn: "pnormal ?p"
+  assume "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+  with isnonconstant_pnormal_iff[OF assms, of bs] have pn: "pnormal ?p"
     by blast
-  {
-    fix x
-    assume H: "?p = [x]"
+  have False if H: "?p = [x]" for x
+  proof -
     from H have "length (coefficients p) = 1"
-      unfolding polypoly_def by auto
+      by (auto simp: polypoly_def)
     with isnonconstant_coefficients_length[OF assms]
-    have False by arith
-  }
+    show ?thesis by arith
+  qed
   then show "nonconstant ?p"
     using pn unfolding nonconstant_def by blast
 qed
 
 lemma pnormal_length: "p \<noteq> [] \<Longrightarrow> pnormal p \<longleftrightarrow> length (pnormalize p) = length p"
   apply (induct p)
-  apply (simp_all add: pnormal_def)
+   apply (simp_all add: pnormal_def)
   apply (case_tac "p = []")
-  apply simp_all
+   apply simp_all
   done
 
 lemma degree_degree:
   assumes "isnonconstant p"
   shows "degree p = Polynomial_List.degree (polypoly bs p) \<longleftrightarrow> \<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+    (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   let ?p = "polypoly bs p"
-  assume H: "degree p = Polynomial_List.degree ?p"
-  from isnonconstant_coefficients_length[OF assms] have pz: "?p \<noteq> []"
-    unfolding polypoly_def by auto
-  from H degree_coefficients[of p] isnonconstant_coefficients_length[OF assms]
-  have lg: "length (pnormalize ?p) = length ?p"
-    unfolding Polynomial_List.degree_def polypoly_def by simp
-  then have "pnormal ?p"
-    using pnormal_length[OF pz] by blast
-  with isnonconstant_pnormal_iff[OF assms] show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
-    by blast
-next
-  let ?p = "polypoly bs p"
-  assume H: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
-  with isnonconstant_pnormal_iff[OF assms] have "pnormal ?p"
-    by blast
-  with degree_coefficients[of p] isnonconstant_coefficients_length[OF assms]
-  show "degree p = Polynomial_List.degree ?p"
-    unfolding polypoly_def pnormal_def Polynomial_List.degree_def by auto
+  {
+    assume ?lhs
+    from isnonconstant_coefficients_length[OF assms] have "?p \<noteq> []"
+      by (auto simp: polypoly_def)
+    from \<open>?lhs\<close> degree_coefficients[of p] isnonconstant_coefficients_length[OF assms]
+    have "length (pnormalize ?p) = length ?p"
+      by (simp add: Polynomial_List.degree_def polypoly_def)
+    with pnormal_length[OF \<open>?p \<noteq> []\<close>] have "pnormal ?p"
+      by blast
+    with isnonconstant_pnormal_iff[OF assms] show ?rhs
+      by blast
+  next
+    assume ?rhs
+    with isnonconstant_pnormal_iff[OF assms] have "pnormal ?p"
+      by blast
+    with degree_coefficients[of p] isnonconstant_coefficients_length[OF assms] show ?lhs
+      by (auto simp: polypoly_def pnormal_def Polynomial_List.degree_def)
+  }
 qed
 
 
-section \<open>Swaps ; Division by a certain variable\<close>
+section \<open>Swaps -- division by a certain variable\<close>
 
 primrec swap :: "nat \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> poly"
-where
-  "swap n m (C x) = C x"
-| "swap n m (Bound k) = Bound (if k = n then m else if k = m then n else k)"
-| "swap n m (Neg t) = Neg (swap n m t)"
-| "swap n m (Add s t) = Add (swap n m s) (swap n m t)"
-| "swap n m (Sub s t) = Sub (swap n m s) (swap n m t)"
-| "swap n m (Mul s t) = Mul (swap n m s) (swap n m t)"
-| "swap n m (Pw t k) = Pw (swap n m t) k"
-| "swap n m (CN c k p) = CN (swap n m c) (if k = n then m else if k=m then n else k) (swap n m p)"
+  where
+    "swap n m (C x) = C x"
+  | "swap n m (Bound k) = Bound (if k = n then m else if k = m then n else k)"
+  | "swap n m (Neg t) = Neg (swap n m t)"
+  | "swap n m (Add s t) = Add (swap n m s) (swap n m t)"
+  | "swap n m (Sub s t) = Sub (swap n m s) (swap n m t)"
+  | "swap n m (Mul s t) = Mul (swap n m s) (swap n m t)"
+  | "swap n m (Pw t k) = Pw (swap n m t) k"
+  | "swap n m (CN c k p) = CN (swap n m c) (if k = n then m else if k=m then n else k) (swap n m p)"
 
 lemma swap:
   assumes "n < length bs"
@@ -2116,10 +2062,10 @@
   by (induct p) simp_all
 
 fun isweaknpoly :: "poly \<Rightarrow> bool"
-where
-  "isweaknpoly (C c) = True"
-| "isweaknpoly (CN c n p) \<longleftrightarrow> isweaknpoly c \<and> isweaknpoly p"
-| "isweaknpoly p = False"
+  where
+    "isweaknpoly (C c) = True"
+  | "isweaknpoly (CN c n p) \<longleftrightarrow> isweaknpoly c \<and> isweaknpoly p"
+  | "isweaknpoly p = False"
 
 lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \<Longrightarrow> isweaknpoly p"
   by (induct p arbitrary: n0) auto
--- a/src/HOL/Decision_Procs/Reflective_Field.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/Reflective_Field.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -5,7 +5,7 @@
 *)
 
 theory Reflective_Field
-imports "~~/src/HOL/Decision_Procs/Commutative_Ring"
+  imports Commutative_Ring
 begin
 
 datatype fexpr =
@@ -18,29 +18,28 @@
   | FDiv fexpr fexpr
   | FPow fexpr nat
 
-fun (in field) nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" where
-  "nth_el [] n = \<zero>"
-| "nth_el (x # xs) 0 = x"
-| "nth_el (x # xs) (Suc n) = nth_el xs n"
+fun (in field) nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a"
+  where
+    "nth_el [] n = \<zero>"
+  | "nth_el (x # xs) 0 = x"
+  | "nth_el (x # xs) (Suc n) = nth_el xs n"
 
-lemma (in field) nth_el_Cons:
-  "nth_el (x # xs) n = (if n = 0 then x else nth_el xs (n - 1))"
+lemma (in field) nth_el_Cons: "nth_el (x # xs) n = (if n = 0 then x else nth_el xs (n - 1))"
   by (cases n) simp_all
 
-lemma (in field) nth_el_closed [simp]:
-  "in_carrier xs \<Longrightarrow> nth_el xs n \<in> carrier R"
+lemma (in field) nth_el_closed [simp]: "in_carrier xs \<Longrightarrow> nth_el xs n \<in> carrier R"
   by (induct xs n rule: nth_el.induct) (simp_all add: in_carrier_def)
 
 primrec (in field) feval :: "'a list \<Rightarrow> fexpr \<Rightarrow> 'a"
-where
-  "feval xs (FCnst c) = \<guillemotleft>c\<guillemotright>"
-| "feval xs (FVar n) = nth_el xs n"
-| "feval xs (FAdd a b) = feval xs a \<oplus> feval xs b"
-| "feval xs (FSub a b) = feval xs a \<ominus> feval xs b"
-| "feval xs (FMul a b) = feval xs a \<otimes> feval xs b"
-| "feval xs (FNeg a) = \<ominus> feval xs a"
-| "feval xs (FDiv a b) = feval xs a \<oslash> feval xs b"
-| "feval xs (FPow a n) = feval xs a (^) n"
+  where
+    "feval xs (FCnst c) = \<guillemotleft>c\<guillemotright>"
+  | "feval xs (FVar n) = nth_el xs n"
+  | "feval xs (FAdd a b) = feval xs a \<oplus> feval xs b"
+  | "feval xs (FSub a b) = feval xs a \<ominus> feval xs b"
+  | "feval xs (FMul a b) = feval xs a \<otimes> feval xs b"
+  | "feval xs (FNeg a) = \<ominus> feval xs a"
+  | "feval xs (FDiv a b) = feval xs a \<oslash> feval xs b"
+  | "feval xs (FPow a n) = feval xs a (^) n"
 
 lemma (in field) feval_Cnst:
   "feval xs (FCnst 0) = \<zero>"
@@ -75,29 +74,29 @@
   case (PExpr1 e')
   then show ?thesis
     apply (cases e')
-    apply simp_all
-    apply (erule assms)+
+        apply simp_all
+        apply (erule assms)+
     done
 next
   case (PExpr2 e')
   then show ?thesis
     apply (cases e')
-    apply simp_all
-    apply (erule assms)+
+     apply simp_all
+     apply (erule assms)+
     done
 qed
 
 lemmas pexpr_cases2 = pexpr_cases [case_product pexpr_cases]
 
 fun (in field) peval :: "'a list \<Rightarrow> pexpr \<Rightarrow> 'a"
-where
-  "peval xs (PExpr1 (PCnst c)) = \<guillemotleft>c\<guillemotright>"
-| "peval xs (PExpr1 (PVar n)) = nth_el xs n"
-| "peval xs (PExpr1 (PAdd a b)) = peval xs a \<oplus> peval xs b"
-| "peval xs (PExpr1 (PSub a b)) = peval xs a \<ominus> peval xs b"
-| "peval xs (PExpr1 (PNeg a)) = \<ominus> peval xs a"
-| "peval xs (PExpr2 (PMul a b)) = peval xs a \<otimes> peval xs b"
-| "peval xs (PExpr2 (PPow a n)) = peval xs a (^) n"
+  where
+    "peval xs (PExpr1 (PCnst c)) = \<guillemotleft>c\<guillemotright>"
+  | "peval xs (PExpr1 (PVar n)) = nth_el xs n"
+  | "peval xs (PExpr1 (PAdd a b)) = peval xs a \<oplus> peval xs b"
+  | "peval xs (PExpr1 (PSub a b)) = peval xs a \<ominus> peval xs b"
+  | "peval xs (PExpr1 (PNeg a)) = \<ominus> peval xs a"
+  | "peval xs (PExpr2 (PMul a b)) = peval xs a \<otimes> peval xs b"
+  | "peval xs (PExpr2 (PPow a n)) = peval xs a (^) n"
 
 lemma (in field) peval_Cnst:
   "peval xs (PExpr1 (PCnst 0)) = \<zero>"
@@ -113,44 +112,44 @@
   by (induct e and e1 and e2) simp_all
 
 definition npepow :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr"
-where
-  "npepow e n =
-     (if n = 0 then PExpr1 (PCnst 1)
-      else if n = 1 then e
-      else (case e of
-          PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c ^ n))
-        | _ \<Rightarrow> PExpr2 (PPow e n)))"
+  where "npepow e n =
+   (if n = 0 then PExpr1 (PCnst 1)
+    else if n = 1 then e
+    else
+      (case e of
+        PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c ^ n))
+      | _ \<Rightarrow> PExpr2 (PPow e n)))"
 
 lemma (in field) npepow_correct:
   "in_carrier xs \<Longrightarrow> peval xs (npepow e n) = peval xs (PExpr2 (PPow e n))"
-  by (cases e rule: pexpr_cases)
-    (simp_all add: npepow_def)
+  by (cases e rule: pexpr_cases) (simp_all add: npepow_def)
 
 fun npemul :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
-where
-  "npemul x y = (case x of
-       PExpr1 (PCnst c) \<Rightarrow>
-         if c = 0 then x
-         else if c = 1 then y else
-           (case y of
-              PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c * d))
-            | _ \<Rightarrow> PExpr2 (PMul x y))
-     | PExpr2 (PPow e1 n) \<Rightarrow>
-         (case y of
-            PExpr2 (PPow e2 m) \<Rightarrow>
-              if n = m then npepow (npemul e1 e2) n
-              else PExpr2 (PMul x y)
-          | PExpr1 (PCnst d) \<Rightarrow>
-              if d = 0 then y
-              else if d = 1 then x
-              else PExpr2 (PMul x y)
+  where "npemul x y =
+    (case x of
+      PExpr1 (PCnst c) \<Rightarrow>
+        if c = 0 then x
+        else if c = 1 then y else
+          (case y of
+            PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c * d))
           | _ \<Rightarrow> PExpr2 (PMul x y))
-     | _ \<Rightarrow> (case y of
-         PExpr1 (PCnst d) \<Rightarrow>
-           if d = 0 then y
-           else if d = 1 then x
-           else PExpr2 (PMul x y)
-       | _ \<Rightarrow> PExpr2 (PMul x y)))"
+    | PExpr2 (PPow e1 n) \<Rightarrow>
+        (case y of
+          PExpr2 (PPow e2 m) \<Rightarrow>
+            if n = m then npepow (npemul e1 e2) n
+            else PExpr2 (PMul x y)
+        | PExpr1 (PCnst d) \<Rightarrow>
+            if d = 0 then y
+            else if d = 1 then x
+            else PExpr2 (PMul x y)
+        | _ \<Rightarrow> PExpr2 (PMul x y))
+    | _ \<Rightarrow>
+      (case y of
+        PExpr1 (PCnst d) \<Rightarrow>
+          if d = 0 then y
+          else if d = 1 then x
+          else PExpr2 (PMul x y)
+      | _ \<Rightarrow> PExpr2 (PMul x y)))"
 
 lemma (in field) npemul_correct:
   "in_carrier xs \<Longrightarrow> peval xs (npemul e1 e2) = peval xs (PExpr2 (PMul e1 e2))"
@@ -160,104 +159,101 @@
   proof (cases x y rule: pexpr_cases2)
     case (PPow_PPow e n e' m)
     then show ?thesis
-    by (simp add: 1 npepow_correct nat_pow_distr
-      npemul.simps [of "PExpr2 (PPow e n)" "PExpr2 (PPow e' m)"]
-      del: npemul.simps)
+      by (simp del: npemul.simps add: 1 npepow_correct nat_pow_distr
+          npemul.simps [of "PExpr2 (PPow e n)" "PExpr2 (PPow e' m)"])
   qed simp_all
 qed
 
 declare npemul.simps [simp del]
 
 definition npeadd :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
-where
-  "npeadd x y = (case x of
-       PExpr1 (PCnst c) \<Rightarrow>
-         if c = 0 then y else
-           (case y of
-              PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c + d))
-            | _ \<Rightarrow> PExpr1 (PAdd x y))
-     | _ \<Rightarrow> (case y of
-         PExpr1 (PCnst d) \<Rightarrow>
-           if d = 0 then x
-           else PExpr1 (PAdd x y)
-       | _ \<Rightarrow> PExpr1 (PAdd x y)))"
+  where "npeadd x y =
+    (case x of
+      PExpr1 (PCnst c) \<Rightarrow>
+        if c = 0 then y
+        else
+          (case y of
+            PExpr1 (PCnst d) \<Rightarrow> PExpr1 (PCnst (c + d))
+          | _ \<Rightarrow> PExpr1 (PAdd x y))
+    | _ \<Rightarrow>
+      (case y of
+        PExpr1 (PCnst d) \<Rightarrow>
+          if d = 0 then x
+          else PExpr1 (PAdd x y)
+      | _ \<Rightarrow> PExpr1 (PAdd x y)))"
 
 lemma (in field) npeadd_correct:
   "in_carrier xs \<Longrightarrow> peval xs (npeadd e1 e2) = peval xs (PExpr1 (PAdd e1 e2))"
   by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npeadd_def)
 
 definition npesub :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
-where
-  "npesub x y = (case y of
-       PExpr1 (PCnst d) \<Rightarrow>
-         if d = 0 then x else
-           (case x of
-              PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c - d))
-            | _ \<Rightarrow> PExpr1 (PSub x y))
-     | _ \<Rightarrow> (case x of
-         PExpr1 (PCnst c) \<Rightarrow>
-           if c = 0 then PExpr1 (PNeg y)
-           else PExpr1 (PSub x y)
-       | _ \<Rightarrow> PExpr1 (PSub x y)))"
+  where "npesub x y =
+    (case y of
+      PExpr1 (PCnst d) \<Rightarrow>
+        if d = 0 then x
+        else
+          (case x of
+            PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (c - d))
+          | _ \<Rightarrow> PExpr1 (PSub x y))
+    | _ \<Rightarrow>
+      (case x of
+        PExpr1 (PCnst c) \<Rightarrow>
+          if c = 0 then PExpr1 (PNeg y)
+          else PExpr1 (PSub x y)
+      | _ \<Rightarrow> PExpr1 (PSub x y)))"
 
 lemma (in field) npesub_correct:
   "in_carrier xs \<Longrightarrow> peval xs (npesub e1 e2) = peval xs (PExpr1 (PSub e1 e2))"
   by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npesub_def)
 
 definition npeneg :: "pexpr \<Rightarrow> pexpr"
-where
-  "npeneg e = (case e of
-       PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (- c))
-     | _ \<Rightarrow> PExpr1 (PNeg e))"
+  where "npeneg e =
+    (case e of
+      PExpr1 (PCnst c) \<Rightarrow> PExpr1 (PCnst (- c))
+    | _ \<Rightarrow> PExpr1 (PNeg e))"
 
-lemma (in field) npeneg_correct:
-  "peval xs (npeneg e) = peval xs (PExpr1 (PNeg e))"
+lemma (in field) npeneg_correct: "peval xs (npeneg e) = peval xs (PExpr1 (PNeg e))"
   by (cases e rule: pexpr_cases) (simp_all add: npeneg_def)
 
 lemma option_pair_cases [case_names None Some]:
-  assumes
-    "x = None \<Longrightarrow> P"
-    "\<And>p q. x = Some (p, q) \<Longrightarrow> P"
-  shows P
+  obtains (None) "x = None" | (Some) p q where "x = Some (p, q)"
 proof (cases x)
   case None
-  then show ?thesis by (rule assms)
+  then show ?thesis by (rule that)
 next
   case (Some r)
   then show ?thesis
-    apply (cases r)
-    apply simp
-    by (rule assms)
+    by (cases r, simp) (rule that)
 qed
 
-fun isin :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> nat \<Rightarrow> (nat * pexpr) option"
-where
-  "isin e n (PExpr2 (PMul e1 e2)) m =
-     (case isin e n e1 m of
+fun isin :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> nat \<Rightarrow> (nat \<times> pexpr) option"
+  where
+    "isin e n (PExpr2 (PMul e1 e2)) m =
+      (case isin e n e1 m of
         Some (k, e3) \<Rightarrow>
           if k = 0 then Some (0, npemul e3 (npepow e2 m))
-          else (case isin e k e2 m of
+          else
+            (case isin e k e2 m of
               Some (l, e4) \<Rightarrow> Some (l, npemul e3 e4)
             | None \<Rightarrow> Some (k, npemul e3 (npepow e2 m)))
-      | None \<Rightarrow> (case isin e n e2 m of
-          Some (k, e3) \<Rightarrow> Some (k, npemul (npepow e1 m) e3)
-        | None \<Rightarrow> None))"
-| "isin e n (PExpr2 (PPow e' k)) m =
-     (if k = 0 then None else isin e n e' (k * m))"
-| "isin (PExpr1 e) n (PExpr1 e') m =
-     (if e = e' then
-        if n >= m then Some (n - m, PExpr1 (PCnst 1))
+      | None \<Rightarrow>
+          (case isin e n e2 m of
+            Some (k, e3) \<Rightarrow> Some (k, npemul (npepow e1 m) e3)
+          | None \<Rightarrow> None))"
+  | "isin e n (PExpr2 (PPow e' k)) m =
+      (if k = 0 then None else isin e n e' (k * m))"
+  | "isin (PExpr1 e) n (PExpr1 e') m =
+      (if e = e' then
+        if n \<ge> m then Some (n - m, PExpr1 (PCnst 1))
         else Some (0, npepow (PExpr1 e) (m - n))
-      else None)"
-| "isin (PExpr2 e) n (PExpr1 e') m = None"
+       else None)"
+  | "isin (PExpr2 e) n (PExpr1 e') m = None"
 
 lemma (in field) isin_correct:
   assumes "in_carrier xs"
-  and "isin e n e' m = Some (p, e'')"
-  shows
-    "peval xs (PExpr2 (PPow e' m)) =
-     peval xs (PExpr2 (PMul (PExpr2 (PPow e (n - p))) e''))"
-    "p \<le> n"
+    and "isin e n e' m = Some (p, e'')"
+  shows "peval xs (PExpr2 (PPow e' m)) = peval xs (PExpr2 (PMul (PExpr2 (PPow e (n - p))) e''))"
+    and "p \<le> n"
   using assms
   by (induct e n e' m arbitrary: p e'' rule: isin.induct)
     (force
@@ -268,38 +264,37 @@
 
 lemma (in field) isin_correct':
   "in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow>
-   peval xs e' = peval xs e (^) (n - p) \<otimes> peval xs e''"
+    peval xs e' = peval xs e (^) (n - p) \<otimes> peval xs e''"
   "in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow> p \<le> n"
-  using isin_correct [where m=1]
-  by simp_all
+  using isin_correct [where m = 1] by simp_all
 
 fun split_aux :: "pexpr \<Rightarrow> nat \<Rightarrow> pexpr \<Rightarrow> pexpr \<times> pexpr \<times> pexpr"
-where
-  "split_aux (PExpr2 (PMul e1 e2)) n e =
-     (let
-        (left1, common1, right1) = split_aux e1 n e;
-        (left2, common2, right2) = split_aux e2 n right1
-      in (npemul left1 left2, npemul common1 common2, right2))"
-| "split_aux (PExpr2 (PPow e' m)) n e =
-     (if m = 0 then (PExpr1 (PCnst 1), PExpr1 (PCnst 1), e)
-      else split_aux e' (m * n) e)"
-| "split_aux (PExpr1 e') n e =
-     (case isin (PExpr1 e') n e 1 of
-        Some (m, e'') \<Rightarrow>
-          (if m = 0 then (PExpr1 (PCnst 1), npepow (PExpr1 e') n, e'')
-           else (npepow (PExpr1 e') m, npepow (PExpr1 e') (n - m), e''))
-      | None \<Rightarrow> (npepow (PExpr1 e') n, PExpr1 (PCnst 1), e))"
+  where
+    "split_aux (PExpr2 (PMul e1 e2)) n e =
+       (let
+          (left1, common1, right1) = split_aux e1 n e;
+          (left2, common2, right2) = split_aux e2 n right1
+        in (npemul left1 left2, npemul common1 common2, right2))"
+  | "split_aux (PExpr2 (PPow e' m)) n e =
+       (if m = 0 then (PExpr1 (PCnst 1), PExpr1 (PCnst 1), e)
+        else split_aux e' (m * n) e)"
+  | "split_aux (PExpr1 e') n e =
+       (case isin (PExpr1 e') n e 1 of
+          Some (m, e'') \<Rightarrow>
+            (if m = 0 then (PExpr1 (PCnst 1), npepow (PExpr1 e') n, e'')
+             else (npepow (PExpr1 e') m, npepow (PExpr1 e') (n - m), e''))
+        | None \<Rightarrow> (npepow (PExpr1 e') n, PExpr1 (PCnst 1), e))"
 
-hide_const Left Right
+hide_const Left Right  (* FIXME !? *)
 
-abbreviation Left :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr" where
-  "Left e1 e2 \<equiv> fst (split_aux e1 (Suc 0) e2)"
+abbreviation Left :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
+  where "Left e1 e2 \<equiv> fst (split_aux e1 (Suc 0) e2)"
 
-abbreviation Common :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr" where
-  "Common e1 e2 \<equiv> fst (snd (split_aux e1 (Suc 0) e2))"
+abbreviation Common :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
+  where "Common e1 e2 \<equiv> fst (snd (split_aux e1 (Suc 0) e2))"
 
-abbreviation Right :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr" where
-  "Right e1 e2 \<equiv> snd (snd (split_aux e1 (Suc 0) e2))"
+abbreviation Right :: "pexpr \<Rightarrow> pexpr \<Rightarrow> pexpr"
+  where "Right e1 e2 \<equiv> snd (snd (split_aux e1 (Suc 0) e2))"
 
 lemma split_aux_induct [case_names 1 2 3]:
   assumes I1: "\<And>e1 e2 n e. P e1 n e \<Longrightarrow> P e2 n (snd (snd (split_aux e1 n e))) \<Longrightarrow>
@@ -321,11 +316,11 @@
 
 lemma (in field) split_aux_correct:
   "in_carrier xs \<Longrightarrow>
-   peval xs (PExpr2 (PPow e\<^sub>1 n)) =
-   peval xs (PExpr2 (PMul (fst (split_aux e\<^sub>1 n e\<^sub>2)) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"
+    peval xs (PExpr2 (PPow e\<^sub>1 n)) =
+    peval xs (PExpr2 (PMul (fst (split_aux e\<^sub>1 n e\<^sub>2)) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"
   "in_carrier xs \<Longrightarrow>
-   peval xs e\<^sub>2 =
-   peval xs (PExpr2 (PMul (snd (snd (split_aux e\<^sub>1 n e\<^sub>2))) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"
+    peval xs e\<^sub>2 =
+    peval xs (PExpr2 (PMul (snd (snd (split_aux e\<^sub>1 n e\<^sub>2))) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))"
   by (induct e\<^sub>1 n e\<^sub>2 rule: split_aux_induct)
     (auto simp add: split_beta
        nat_pow_distr nat_pow_pow nat_pow_mult m_ac
@@ -334,81 +329,78 @@
 
 lemma (in field) split_aux_correct':
   "in_carrier xs \<Longrightarrow>
-   peval xs e\<^sub>1 = peval xs (Left e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"
+    peval xs e\<^sub>1 = peval xs (Left e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"
   "in_carrier xs \<Longrightarrow>
-   peval xs e\<^sub>2 = peval xs (Right e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"
-  using split_aux_correct [where n=1]
-  by simp_all
+    peval xs e\<^sub>2 = peval xs (Right e\<^sub>1 e\<^sub>2) \<otimes> peval xs (Common e\<^sub>1 e\<^sub>2)"
+  using split_aux_correct [where n = 1] by simp_all
 
 fun fnorm :: "fexpr \<Rightarrow> pexpr \<times> pexpr \<times> pexpr list"
-where
-  "fnorm (FCnst c) = (PExpr1 (PCnst c), PExpr1 (PCnst 1), [])"
-| "fnorm (FVar n) = (PExpr1 (PVar n), PExpr1 (PCnst 1), [])"
-| "fnorm (FAdd e1 e2) =
-     (let
-        (xn, xd, xc) = fnorm e1;
-        (yn, yd, yc) = fnorm e2;
-        (left, common, right) = split_aux xd 1 yd
-      in
-        (npeadd (npemul xn right) (npemul yn left),
-         npemul left (npemul right common),
-         List.union xc yc))"
-| "fnorm (FSub e1 e2) =
-     (let
-        (xn, xd, xc) = fnorm e1;
-        (yn, yd, yc) = fnorm e2;
-        (left, common, right) = split_aux xd 1 yd
-      in
-        (npesub (npemul xn right) (npemul yn left),
-         npemul left (npemul right common),
-         List.union xc yc))"
-| "fnorm (FMul e1 e2) =
-     (let
-        (xn, xd, xc) = fnorm e1;
-        (yn, yd, yc) = fnorm e2;
-        (left1, common1, right1) = split_aux xn 1 yd;
-        (left2, common2, right2) = split_aux yn 1 xd
-      in
-        (npemul left1 left2,
-         npemul right2 right1,
-         List.union xc yc))"
-| "fnorm (FNeg e) =
-     (let (n, d, c) = fnorm e
-      in (npeneg n, d, c))"
-| "fnorm (FDiv e1 e2) =
-     (let
-        (xn, xd, xc) = fnorm e1;
-        (yn, yd, yc) = fnorm e2;
-        (left1, common1, right1) = split_aux xn 1 yn;
-        (left2, common2, right2) = split_aux xd 1 yd
-      in
-        (npemul left1 right2,
-         npemul left2 right1,
-         List.insert yn (List.union xc yc)))"
-| "fnorm (FPow e m) =
-     (let (n, d, c) = fnorm e
-      in (npepow n m, npepow d m, c))"
+  where
+    "fnorm (FCnst c) = (PExpr1 (PCnst c), PExpr1 (PCnst 1), [])"
+  | "fnorm (FVar n) = (PExpr1 (PVar n), PExpr1 (PCnst 1), [])"
+  | "fnorm (FAdd e1 e2) =
+       (let
+          (xn, xd, xc) = fnorm e1;
+          (yn, yd, yc) = fnorm e2;
+          (left, common, right) = split_aux xd 1 yd
+        in
+          (npeadd (npemul xn right) (npemul yn left),
+           npemul left (npemul right common),
+           List.union xc yc))"
+  | "fnorm (FSub e1 e2) =
+       (let
+          (xn, xd, xc) = fnorm e1;
+          (yn, yd, yc) = fnorm e2;
+          (left, common, right) = split_aux xd 1 yd
+        in
+          (npesub (npemul xn right) (npemul yn left),
+           npemul left (npemul right common),
+           List.union xc yc))"
+  | "fnorm (FMul e1 e2) =
+       (let
+          (xn, xd, xc) = fnorm e1;
+          (yn, yd, yc) = fnorm e2;
+          (left1, common1, right1) = split_aux xn 1 yd;
+          (left2, common2, right2) = split_aux yn 1 xd
+        in
+          (npemul left1 left2,
+           npemul right2 right1,
+           List.union xc yc))"
+  | "fnorm (FNeg e) =
+       (let (n, d, c) = fnorm e
+        in (npeneg n, d, c))"
+  | "fnorm (FDiv e1 e2) =
+       (let
+          (xn, xd, xc) = fnorm e1;
+          (yn, yd, yc) = fnorm e2;
+          (left1, common1, right1) = split_aux xn 1 yn;
+          (left2, common2, right2) = split_aux xd 1 yd
+        in
+          (npemul left1 right2,
+           npemul left2 right1,
+           List.insert yn (List.union xc yc)))"
+  | "fnorm (FPow e m) =
+       (let (n, d, c) = fnorm e
+        in (npepow n m, npepow d m, c))"
 
-abbreviation Num :: "fexpr \<Rightarrow> pexpr" where
-  "Num e \<equiv> fst (fnorm e)"
+abbreviation Num :: "fexpr \<Rightarrow> pexpr"
+  where "Num e \<equiv> fst (fnorm e)"
 
-abbreviation Denom :: "fexpr \<Rightarrow> pexpr" where
-  "Denom e \<equiv> fst (snd (fnorm e))"
+abbreviation Denom :: "fexpr \<Rightarrow> pexpr"
+  where "Denom e \<equiv> fst (snd (fnorm e))"
 
-abbreviation Cond :: "fexpr \<Rightarrow> pexpr list" where
-  "Cond e \<equiv> snd (snd (fnorm e))"
+abbreviation Cond :: "fexpr \<Rightarrow> pexpr list"
+  where "Cond e \<equiv> snd (snd (fnorm e))"
 
 primrec (in field) nonzero :: "'a list \<Rightarrow> pexpr list \<Rightarrow> bool"
-where
-  "nonzero xs [] = True"
-| "nonzero xs (p # ps) = (peval xs p \<noteq> \<zero> \<and> nonzero xs ps)"
+  where
+    "nonzero xs [] \<longleftrightarrow> True"
+  | "nonzero xs (p # ps) \<longleftrightarrow> peval xs p \<noteq> \<zero> \<and> nonzero xs ps"
 
-lemma (in field) nonzero_singleton:
-  "nonzero xs [p] = (peval xs p \<noteq> \<zero>)"
+lemma (in field) nonzero_singleton: "nonzero xs [p] = (peval xs p \<noteq> \<zero>)"
   by simp
 
-lemma (in field) nonzero_append:
-  "nonzero xs (ps @ qs) = (nonzero xs ps \<and> nonzero xs qs)"
+lemma (in field) nonzero_append: "nonzero xs (ps @ qs) = (nonzero xs ps \<and> nonzero xs qs)"
   by (induct ps) simp_all
 
 lemma (in field) nonzero_idempotent:
@@ -426,12 +418,13 @@
 
 lemma (in field) fnorm_correct:
   assumes "in_carrier xs"
-  and "nonzero xs (Cond e)"
+    and "nonzero xs (Cond e)"
   shows "feval xs e = peval xs (Num e) \<oslash> peval xs (Denom e)"
-  and "peval xs (Denom e) \<noteq> \<zero>"
+    and "peval xs (Denom e) \<noteq> \<zero>"
   using assms
 proof (induct e)
-  case (FCnst c) {
+  case (FCnst c)
+  {
     case 1
     show ?case by simp
   next
@@ -439,7 +432,8 @@
     show ?case by simp
   }
 next
-  case (FVar n) {
+  case (FVar n)
+  {
     case 1
     then show ?case by simp
   next
@@ -448,21 +442,17 @@
   }
 next
   case (FAdd e1 e2)
-  note split = split_aux_correct' [where xs=xs and
-    e\<^sub>1="Denom e1" and e\<^sub>2="Denom e2"]
+  note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]
   {
     case 1
     let ?left = "peval xs (Left (Denom e1) (Denom e2))"
     let ?common = "peval xs (Common (Denom e1) (Denom e2))"
     let ?right = "peval xs (Right (Denom e1) (Denom e2))"
-    from 1 FAdd
-    have "feval xs (FAdd e1 e2) =
+    from 1 FAdd have "feval xs (FAdd e1 e2) =
       (?common \<otimes> (peval xs (Num e1) \<otimes> ?right \<oplus> peval xs (Num e2) \<otimes> ?left)) \<oslash>
       (?common \<otimes> (?left \<otimes> (?right \<otimes> ?common)))"
-      by (simp add: split_beta split nonzero_union
-        add_frac_eq r_distr m_ac)
-    also from 1 FAdd have "\<dots> =
-      peval xs (Num (FAdd e1 e2)) \<oslash> peval xs (Denom (FAdd e1 e2))"
+      by (simp add: split_beta split nonzero_union add_frac_eq r_distr m_ac)
+    also from 1 FAdd have "\<dots> = peval xs (Num (FAdd e1 e2)) \<oslash> peval xs (Denom (FAdd e1 e2))"
       by (simp add: split_beta split nonzero_union npeadd_correct npemul_correct integral_iff)
     finally show ?case .
   next
@@ -472,8 +462,7 @@
   }
 next
   case (FSub e1 e2)
-  note split = split_aux_correct' [where xs=xs and
-    e\<^sub>1="Denom e1" and e\<^sub>2="Denom e2"]
+  note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]
   {
     case 1
     let ?left = "peval xs (Left (Denom e1) (Denom e2))"
@@ -483,10 +472,8 @@
     have "feval xs (FSub e1 e2) =
       (?common \<otimes> (peval xs (Num e1) \<otimes> ?right \<ominus> peval xs (Num e2) \<otimes> ?left)) \<oslash>
       (?common \<otimes> (?left \<otimes> (?right \<otimes> ?common)))"
-      by (simp add: split_beta split nonzero_union
-        diff_frac_eq r_diff_distr m_ac)
-    also from 1 FSub have "\<dots> =
-      peval xs (Num (FSub e1 e2)) \<oslash> peval xs (Denom (FSub e1 e2))"
+      by (simp add: split_beta split nonzero_union diff_frac_eq r_diff_distr m_ac)
+    also from 1 FSub have "\<dots> = peval xs (Num (FSub e1 e2)) \<oslash> peval xs (Denom (FSub e1 e2))"
       by (simp add: split_beta split nonzero_union npesub_correct npemul_correct integral_iff)
     finally show ?case .
   next
@@ -497,10 +484,8 @@
 next
   case (FMul e1 e2)
   note split =
-    split_aux_correct' [where xs=xs and
-      e\<^sub>1="Num e1" and e\<^sub>2="Denom e2"]
-    split_aux_correct' [where xs=xs and
-      e\<^sub>1="Num e2" and e\<^sub>2="Denom e1"]
+    split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Denom e2"]
+    split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e2" and e\<^sub>2 = "Denom e1"]
   {
     case 1
     let ?left\<^sub>1 = "peval xs (Left (Num e1) (Denom e2))"
@@ -509,14 +494,12 @@
     let ?left\<^sub>2 = "peval xs (Left (Num e2) (Denom e1))"
     let ?common\<^sub>2 = "peval xs (Common (Num e2) (Denom e1))"
     let ?right\<^sub>2 = "peval xs (Right (Num e2) (Denom e1))"
-    from 1 FMul
-    have "feval xs (FMul e1 e2) =
+    from 1 FMul have "feval xs (FMul e1 e2) =
       ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>1 \<otimes> ?left\<^sub>2)) \<oslash>
       ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?right\<^sub>2 \<otimes> ?right\<^sub>1))"
       by (simp add: split_beta split nonzero_union
         nonzero_divide_divide_eq_left m_ac)
-    also from 1 FMul have "\<dots> =
-      peval xs (Num (FMul e1 e2)) \<oslash> peval xs (Denom (FMul e1 e2))"
+    also from 1 FMul have "\<dots> = peval xs (Num (FMul e1 e2)) \<oslash> peval xs (Denom (FMul e1 e2))"
       by (simp add: split_beta split nonzero_union npemul_correct integral_iff)
     finally show ?case .
   next
@@ -538,10 +521,8 @@
 next
   case (FDiv e1 e2)
   note split =
-    split_aux_correct' [where xs=xs and
-      e\<^sub>1="Num e1" and e\<^sub>2="Num e2"]
-    split_aux_correct' [where xs=xs and
-      e\<^sub>1="Denom e1" and e\<^sub>2="Denom e2"]
+    split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Num e2"]
+    split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"]
   {
     case 1
     let ?left\<^sub>1 = "peval xs (Left (Num e1) (Num e2))"
@@ -556,8 +537,7 @@
       ((?common\<^sub>1 \<otimes> ?common\<^sub>2) \<otimes> (?left\<^sub>2 \<otimes> ?right\<^sub>1))"
       by (simp add: split_beta split nonzero_union nonzero_insert
         nonzero_divide_divide_eq m_ac)
-    also from 1 FDiv have "\<dots> =
-      peval xs (Num (FDiv e1 e2)) \<oslash> peval xs (Denom (FDiv e1 e2))"
+    also from 1 FDiv have "\<dots> = peval xs (Num (FDiv e1 e2)) \<oslash> peval xs (Denom (FDiv e1 e2))"
       by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff)
     finally show ?case .
   next
@@ -580,16 +560,15 @@
 
 lemma (in field) feval_eq0:
   assumes "in_carrier xs"
-  and "fnorm e = (n, d, c)"
-  and "nonzero xs c"
-  and "peval xs n = \<zero>"
+    and "fnorm e = (n, d, c)"
+    and "nonzero xs c"
+    and "peval xs n = \<zero>"
   shows "feval xs e = \<zero>"
-  using assms fnorm_correct [of xs e]
-  by simp
+  using assms fnorm_correct [of xs e] by simp
 
 lemma (in field) fexpr_in_carrier:
   assumes "in_carrier xs"
-  and "nonzero xs (Cond e)"
+    and "nonzero xs (Cond e)"
   shows "feval xs e \<in> carrier R"
   using assms
 proof (induct e)
@@ -610,8 +589,8 @@
 
 lemma (in field) feval_eq:
   assumes "in_carrier xs"
-  and "fnorm (FSub e e') = (n, d, c)"
-  and "nonzero xs c"
+    and "fnorm (FSub e e') = (n, d, c)"
+    and "nonzero xs c"
   shows "(feval xs e = feval xs e') = (peval xs n = \<zero>)"
 proof -
   from assms have "nonzero xs (Cond e)" "nonzero xs (Cond e')"
@@ -940,7 +919,8 @@
 end
 \<close>
 
-context field begin
+context field
+begin
 
 local_setup \<open>
 Local_Theory.declaration {syntax = false, pervasive = false}
--- a/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy	Sun Dec 03 19:09:42 2017 +0100
+++ b/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy	Sun Dec 03 22:28:19 2017 +0100
@@ -5,7 +5,7 @@
 section \<open>Some examples demonstrating the ring and field methods\<close>
 
 theory Commutative_Ring_Ex
-imports "../Reflective_Field"
+  imports "../Reflective_Field"
 begin
 
 lemma "4 * (x::int) ^ 5 * y ^ 3 * x ^ 2 * 3 + x * z + 3 ^ 5 = 12 * x ^ 7 * y ^ 3 + z * x + 243"