src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy
changeset 67123 3fe40ff1b921
parent 62390 842917225d56
child 68442 477b3f7067c9
--- 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