tuned specifications and proofs;
authorwenzelm
Tue, 25 Feb 2014 23:12:48 +0100
changeset 55754 d14072d53c1e
parent 55752 43d0e2a34c9d
child 55755 e5128a9c4311
tuned specifications and proofs;
src/HOL/Decision_Procs/Commutative_Ring.thy
src/HOL/Decision_Procs/Commutative_Ring_Complete.thy
src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy	Tue Feb 25 22:13:57 2014 +0100
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy	Tue Feb 25 23:12:48 2014 +0100
@@ -6,7 +6,7 @@
 header {* Proving equalities in commutative rings *}
 
 theory Commutative_Ring
-imports Main Parity
+imports Parity
 begin
 
 text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *}
@@ -26,15 +26,13 @@
 
 text {* Interpretation functions for the shadow syntax. *}
 
-primrec
-  Ipol :: "'a::{comm_ring_1} list \<Rightarrow> 'a pol \<Rightarrow> 'a"
+primrec Ipol :: "'a::{comm_ring_1} list \<Rightarrow> 'a pol \<Rightarrow> 'a"
 where
     "Ipol l (Pc c) = c"
   | "Ipol l (Pinj i P) = Ipol (drop i l) P"
   | "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q"
 
-primrec
-  Ipolex :: "'a::{comm_ring_1} list \<Rightarrow> 'a polex \<Rightarrow> 'a"
+primrec Ipolex :: "'a::{comm_ring_1} list \<Rightarrow> 'a polex \<Rightarrow> 'a"
 where
     "Ipolex l (Pol P) = Ipol l P"
   | "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q"
@@ -45,111 +43,106 @@
 
 text {* Create polynomial normalized polynomials given normalized inputs. *}
 
-definition
-  mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
+definition mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
+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 :: "'a::{comm_ring} pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" 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 |
-    PX P2 i2 Q2 \<Rightarrow> (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )"
+definition mkPX :: "'a::comm_ring pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
+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
+    | PX P2 i2 Q2 \<Rightarrow> if Q2 = Pc 0 then PX P2 (i + i2) Q else PX P i Q)"
 
 text {* Defining the basic ring operations on normalized polynomials *}
 
-function
-  add :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<oplus>" 65)
+function add :: "'a::comm_ring pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"  (infixl "\<oplus>" 65)
 where
-    "Pc a \<oplus> Pc b = Pc (a + b)"
-  | "Pc c \<oplus> Pinj i P = Pinj i (P \<oplus> Pc c)"
-  | "Pinj i P \<oplus> Pc c = Pinj i (P \<oplus> Pc c)"
-  | "Pc c \<oplus> PX P i Q = PX P i (Q \<oplus> Pc c)"
-  | "PX P i Q \<oplus> Pc c = PX P i (Q \<oplus> Pc c)"
-  | "Pinj x P \<oplus> Pinj y Q =
-      (if x = y then mkPinj x (P \<oplus> Q)
-       else (if x > y then mkPinj y (Pinj (x - y) P \<oplus> Q)
-         else mkPinj x (Pinj (y - x) Q \<oplus> P)))"
-  | "Pinj x P \<oplus> PX Q y R =
-      (if x = 0 then P \<oplus> PX Q y R
-       else (if x = 1 then PX Q y (R \<oplus> P)
-         else PX Q y (R \<oplus> Pinj (x - 1) P)))"
-  | "PX P x R \<oplus> Pinj y Q =
-      (if y = 0 then PX P x R \<oplus> Q
-       else (if y = 1 then PX P x (R \<oplus> Q)
-         else PX P x (R \<oplus> Pinj (y - 1) Q)))"
-  | "PX P1 x P2 \<oplus> PX Q1 y Q2 =
-      (if x = y then mkPX (P1 \<oplus> Q1) x (P2 \<oplus> Q2)
-       else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<oplus> Q1) y (P2 \<oplus> Q2)
-         else mkPX (PX Q1 (y-x) (Pc 0) \<oplus> P1) x (P2 \<oplus> Q2)))"
+  "Pc a \<oplus> Pc b = Pc (a + b)"
+| "Pc c \<oplus> Pinj i P = Pinj i (P \<oplus> Pc c)"
+| "Pinj i P \<oplus> Pc c = Pinj i (P \<oplus> Pc c)"
+| "Pc c \<oplus> PX P i Q = PX P i (Q \<oplus> Pc c)"
+| "PX P i Q \<oplus> Pc c = PX P i (Q \<oplus> Pc c)"
+| "Pinj x P \<oplus> Pinj y Q =
+    (if x = y then mkPinj x (P \<oplus> Q)
+     else (if x > y then mkPinj y (Pinj (x - y) P \<oplus> Q)
+       else mkPinj x (Pinj (y - x) Q \<oplus> P)))"
+| "Pinj x P \<oplus> PX Q y R =
+    (if x = 0 then P \<oplus> PX Q y R
+     else (if x = 1 then PX Q y (R \<oplus> P)
+       else PX Q y (R \<oplus> Pinj (x - 1) P)))"
+| "PX P x R \<oplus> Pinj y Q =
+    (if y = 0 then PX P x R \<oplus> Q
+     else (if y = 1 then PX P x (R \<oplus> Q)
+       else PX P x (R \<oplus> Pinj (y - 1) Q)))"
+| "PX P1 x P2 \<oplus> PX Q1 y Q2 =
+    (if x = y then mkPX (P1 \<oplus> Q1) x (P2 \<oplus> Q2)
+     else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<oplus> Q1) y (P2 \<oplus> Q2)
+       else mkPX (PX Q1 (y-x) (Pc 0) \<oplus> P1) x (P2 \<oplus> Q2)))"
 by pat_completeness auto
 termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto
 
-function
-  mul :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<otimes>" 70)
+function mul :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"  (infixl "\<otimes>" 70)
 where
-    "Pc a \<otimes> Pc b = Pc (a * b)"
-  | "Pc c \<otimes> Pinj i P =
-      (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
-  | "Pinj i P \<otimes> Pc c =
-      (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
-  | "Pc c \<otimes> PX P i Q =
-      (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
-  | "PX P i Q \<otimes> Pc c =
-      (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
-  | "Pinj x P \<otimes> Pinj y Q =
-      (if x = y then mkPinj x (P \<otimes> Q) else
-         (if x > y then mkPinj y (Pinj (x-y) P \<otimes> Q)
-           else mkPinj x (Pinj (y - x) Q \<otimes> P)))"
-  | "Pinj x P \<otimes> PX Q y R =
-      (if x = 0 then P \<otimes> PX Q y R else
-         (if x = 1 then mkPX (Pinj x P \<otimes> Q) y (R \<otimes> P)
-           else mkPX (Pinj x P \<otimes> Q) y (R \<otimes> Pinj (x - 1) P)))"
-  | "PX P x R \<otimes> Pinj y Q =
-      (if y = 0 then PX P x R \<otimes> Q else
-         (if y = 1 then mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Q)
-           else mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Pinj (y - 1) Q)))"
-  | "PX P1 x P2 \<otimes> PX Q1 y Q2 =
-      mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2) \<oplus>
-        (mkPX (P1 \<otimes> mkPinj 1 Q2) x (Pc 0) \<oplus>
-          (mkPX (Q1 \<otimes> mkPinj 1 P2) y (Pc 0)))"
+  "Pc a \<otimes> Pc b = Pc (a * b)"
+| "Pc c \<otimes> Pinj i P =
+    (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
+| "Pinj i P \<otimes> Pc c =
+    (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
+| "Pc c \<otimes> PX P i Q =
+    (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
+| "PX P i Q \<otimes> Pc c =
+    (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
+| "Pinj x P \<otimes> Pinj y Q =
+    (if x = y then mkPinj x (P \<otimes> Q) else
+       (if x > y then mkPinj y (Pinj (x-y) P \<otimes> Q)
+         else mkPinj x (Pinj (y - x) Q \<otimes> P)))"
+| "Pinj x P \<otimes> PX Q y R =
+    (if x = 0 then P \<otimes> PX Q y R else
+       (if x = 1 then mkPX (Pinj x P \<otimes> Q) y (R \<otimes> P)
+         else mkPX (Pinj x P \<otimes> Q) y (R \<otimes> Pinj (x - 1) P)))"
+| "PX P x R \<otimes> Pinj y Q =
+    (if y = 0 then PX P x R \<otimes> Q else
+       (if y = 1 then mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Q)
+         else mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Pinj (y - 1) Q)))"
+| "PX P1 x P2 \<otimes> PX Q1 y Q2 =
+    mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2) \<oplus>
+      (mkPX (P1 \<otimes> mkPinj 1 Q2) x (Pc 0) \<oplus>
+        (mkPX (Q1 \<otimes> mkPinj 1 P2) y (Pc 0)))"
 by pat_completeness auto
 termination by (relation "measure (\<lambda>(x, y). size x + size y)")
   (auto simp add: mkPinj_def split: pol.split)
 
 text {* Negation*}
-primrec
-  neg :: "'a::{comm_ring} pol \<Rightarrow> 'a pol"
+primrec neg :: "'a::{comm_ring} pol \<Rightarrow> 'a pol"
 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)"
+  "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 {* Substraction *}
-definition
-  sub :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<ominus>" 65)
-where
-  "sub P Q = P \<oplus> neg Q"
+definition sub :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"  (infixl "\<ominus>" 65)
+  where "sub P Q = P \<oplus> neg Q"
 
 text {* Square for Fast Exponentation *}
-primrec
-  sqr :: "'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
+primrec sqr :: "'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
 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) \<oplus>
-      mkPX (Pc (1 + 1) \<otimes> A \<otimes> mkPinj 1 B) x (Pc 0)"
+  "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) \<oplus> mkPX (Pc (1 + 1) \<otimes> A \<otimes> mkPinj 1 B) x (Pc 0)"
 
 text {* Fast Exponentation *}
-fun
-  pow :: "nat \<Rightarrow> 'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
+fun pow :: "nat \<Rightarrow> 'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
 where
-    "pow 0 P = Pc 1"
-  | "pow n P = (if even n then pow (n div 2) (sqr P)
-       else P \<otimes> pow (n div 2) (sqr P))"
+  "pow 0 P = Pc 1"
+| "pow n P =
+    (if even n then pow (n div 2) (sqr P)
+     else P \<otimes> pow (n div 2) (sqr P))"
   
 lemma pow_if:
   "pow n P =
@@ -160,15 +153,14 @@
 
 text {* Normalization of polynomial expressions *}
 
-primrec
-  norm :: "'a::{comm_ring_1} polex \<Rightarrow> 'a pol"
+primrec norm :: "'a::{comm_ring_1} polex \<Rightarrow> 'a pol"
 where
-    "norm (Pol P) = P"
-  | "norm (Add P Q) = norm P \<oplus> norm Q"
-  | "norm (Sub P Q) = norm P \<ominus> norm Q"
-  | "norm (Mul P Q) = norm P \<otimes> norm Q"
-  | "norm (Pow P n) = pow n (norm P)"
-  | "norm (Neg P) = neg (norm P)"
+  "norm (Pol P) = P"
+| "norm (Add P Q) = norm P \<oplus> norm Q"
+| "norm (Sub P Q) = norm P \<ominus> norm Q"
+| "norm (Mul P Q) = norm P \<otimes> norm Q"
+| "norm (Pow P n) = pow n (norm P)"
+| "norm (Neg P) = neg (norm P)"
 
 text {* mkPinj preserve semantics *}
 lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)"
--- a/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Tue Feb 25 22:13:57 2014 +0100
+++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy	Tue Feb 25 23:12:48 2014 +0100
@@ -14,15 +14,15 @@
 text {* Formalization of normal form *}
 fun isnorm :: "'a::comm_ring 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"
+  "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 *)
 lemma norm_Pinj_0_False: "isnorm (Pinj 0 P) = False"
@@ -35,12 +35,24 @@
   by (cases i) (simp add: norm_Pinj_0_False norm_PX_0_False, cases Q, auto)
 
 lemma norm_PX2: "isnorm (PX P i Q) \<Longrightarrow> isnorm Q"
-  by (cases i) (auto, cases P, auto, case_tac pol2, auto)
+  apply (cases i)
+  apply auto
+  apply (cases P)
+  apply auto
+  apply (case_tac pol2)
+  apply auto
+  done
 
 lemma norm_PX1: "isnorm (PX P i Q) \<Longrightarrow> isnorm P"
-  by (cases i) (auto, cases P, auto, case_tac pol2, auto)
+  apply (cases i)
+  apply auto
+  apply (cases P)
+  apply auto
+  apply (case_tac pol2)
+  apply auto
+  done
 
-lemma mkPinj_cn: "y ~= 0 \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (mkPinj y Q)"
+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 (case_tac nat, auto simp add: norm_Pinj_0_False)
   apply (case_tac pol, auto)
@@ -48,11 +60,17 @@
   done
 
 lemma norm_PXtrans: 
-  assumes A: "isnorm (PX P x Q)" and "isnorm Q2" 
+  assumes A: "isnorm (PX P x Q)"
+    and "isnorm Q2" 
   shows "isnorm (PX P x Q2)"
 proof (cases P)
   case (PX p1 y p2)
-  with assms show ?thesis by (cases x) (auto, cases p2, auto)
+  with assms show ?thesis
+    apply (cases x)
+    apply auto
+    apply (cases p2)
+    apply auto
+    done
 next
   case Pc
   with assms show ?thesis by (cases x) auto
@@ -63,10 +81,15 @@
  
 lemma norm_PXtrans2:
   assumes "isnorm (PX P x Q)" and "isnorm Q2"
-  shows "isnorm (PX P (Suc (n+x)) Q2)"
+  shows "isnorm (PX P (Suc (n + x)) Q2)"
 proof (cases P)
   case (PX p1 y p2)
-  with assms show ?thesis by (cases x) (auto, cases p2, auto)
+  with assms show ?thesis
+    apply (cases x)
+    apply auto
+    apply (cases p2)
+    apply auto
+    done
 next
   case Pc
   with assms show ?thesis by (cases x) auto
@@ -77,14 +100,18 @@
 
 text {* mkPX conserves normalizedness (@{text "_cn"}) *}
 lemma mkPX_cn: 
-  assumes "x \<noteq> 0" and "isnorm P" and "isnorm Q" 
+  assumes "x \<noteq> 0"
+    and "isnorm P"
+    and "isnorm Q" 
   shows "isnorm (mkPX P x Q)"
-proof(cases P)
+proof (cases P)
   case (Pc c)
-  with assms show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def)
+  with assms show ?thesis
+    by (cases x) (auto simp add: mkPinj_cn mkPX_def)
 next
   case (Pinj i Q)
-  with assms show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def)
+  with assms show ?thesis
+    by (cases x) (auto simp add: mkPinj_cn mkPX_def)
 next
   case (PX P1 y P2)
   with assms have Y0: "y > 0" by (cases y) auto
@@ -98,10 +125,10 @@
 lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<oplus> Q)"
 proof (induct P Q rule: add.induct)
   case (2 c i P2)
-  thus ?case by (cases P2) (simp_all, cases i, simp_all)
+  then show ?case by (cases P2) (simp_all, cases i, simp_all)
 next
   case (3 i P2 c)
-  thus ?case by (cases P2) (simp_all, cases i, simp_all)
+  then show ?case by (cases P2) (simp_all, cases i, simp_all)
 next
   case (4 c P2 i Q2)
   then have "isnorm P2" "isnorm Q2"
@@ -120,7 +147,8 @@
   with 6 have X0: "x>0" by (cases x) (auto simp add: norm_Pinj_0_False)
   have "x < y \<or> x = y \<or> x > y" by arith
   moreover
-  { assume "x<y" hence "EX d. y =d + x" by arith
+  { assume "x < y"
+    then have "\<exists>d. y = d + x" by arith
     then obtain d where y: "y = d + x" ..
     moreover
     note 6 X0
@@ -132,7 +160,7 @@
       by (cases d, simp, cases Q2, auto)
     ultimately have ?case by (simp add: mkPinj_cn) }
   moreover
-  { assume "x=y"
+  { assume "x = y"
     moreover
     from 6 have "isnorm P2" "isnorm Q2"
       by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
@@ -140,7 +168,7 @@
     note 6 Y0
     ultimately have ?case by (simp add: mkPinj_cn) }
   moreover
-  { assume "x>y" hence "EX d. x = d + y" by arith
+  { assume "x > y" then have "\<exists>d. x = d + y" by arith
     then obtain d where x: "x = d + y"..
     moreover
     note 6 Y0
@@ -166,7 +194,7 @@
     with 7 `x = 1` have ?case
       by (simp add: norm_PXtrans[of Q2 y _]) }
   moreover
-  { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
+  { assume "x > 1" then have "\<exists>d. x=Suc (Suc d)" by arith
     then obtain d where X: "x=Suc (Suc d)" ..
     with 7 have NR: "isnorm R" "isnorm P2"
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
@@ -186,7 +214,7 @@
     with 8 `x = 1` have "isnorm (R \<oplus> P2)" by simp
     with 8 `x = 1` have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
   moreover
-  { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
+  { assume "x > 1" then have "\<exists>d. x=Suc (Suc d)" by arith
     then obtain d where X: "x = Suc (Suc d)" ..
     with 8 have NR: "isnorm R" "isnorm P2"
       by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
@@ -204,7 +232,7 @@
     by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2])
   have "y < x \<or> x = y \<or> x < y" by arith
   moreover
-  { assume sm1: "y < x" hence "EX d. x = d + y" by arith
+  { assume sm1: "y < x" then have "\<exists>d. x = d + y" by arith
     then obtain d where sm2: "x = d + y" ..
     note 9 NQ1 NP1 NP2 NQ2 sm1 sm2
     moreover
@@ -224,7 +252,7 @@
     with 9 NP1 NP2 NQ1 NQ2 have "isnorm (P2 \<oplus> Q2)" "isnorm (P1 \<oplus> Q1)" by auto
     with `x = y` Y0 have ?case by (simp add: mkPX_cn) }
   moreover
-  { assume sm1: "x < y" hence "EX d. y = d + x" by arith
+  { assume sm1: "x < y" then have "\<exists>d. y = d + x" by arith
     then obtain d where sm2: "y = d + x" ..
     note 9 NQ1 NP1 NP2 NQ2 sm1 sm2
     moreover
@@ -245,10 +273,10 @@
 text {* mul concerves normalizedness *}
 lemma mul_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<otimes> Q)"
 proof (induct P Q rule: mul.induct)
-  case (2 c i P2) thus ?case 
+  case (2 c i P2) then show ?case 
     by (cases P2) (simp_all, cases i, simp_all add: mkPinj_cn)
 next
-  case (3 i P2 c) thus ?case 
+  case (3 i P2 c) then show ?case 
     by (cases P2) (simp_all, cases i, simp_all add: mkPinj_cn)
 next
   case (4 c P2 i Q2)
@@ -266,7 +294,7 @@
   case (6 x P2 y Q2)
   have "x < y \<or> x = y \<or> x > y" by arith
   moreover
-  { assume "x < y" hence "EX d. y = d + x" by arith
+  { assume "x < y" then have "\<exists>d. y = d + x" by arith
     then obtain d where y: "y = d + x" ..
     moreover
     note 6
@@ -287,7 +315,7 @@
     note 6
     ultimately have ?case by (simp add: mkPinj_cn) }
   moreover
-  { assume "x > y" hence "EX d. x = d + y" by arith
+  { assume "x > y" then have "\<exists>d. x = d + y" by arith
     then obtain d where x: "x = d + y" ..
     moreover
     note 6
@@ -311,7 +339,7 @@
     with 7 `x = 1` have "isnorm (R \<otimes> P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
     with 7 `x = 1` Y0 have ?case by (simp add: mkPX_cn) }
   moreover
-  { assume "x > 1" hence "EX d. x = Suc (Suc d)" by arith
+  { assume "x > 1" then have "\<exists>d. x = Suc (Suc d)" by arith
     then obtain d where X: "x = Suc (Suc d)" ..
     from 7 have NR: "isnorm R" "isnorm Q2"
       by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R])
@@ -326,7 +354,7 @@
   ultimately show ?case by blast
 next
   case (8 Q2 y R x P2)
-  then have Y0: "y>0" by (cases y) auto
+  then have Y0: "y > 0" by (cases y) auto
   have "x = 0 \<or> x = 1 \<or> x > 1" by arith
   moreover
   { assume "x = 0" with 8 have ?case by (auto simp add: norm_Pinj_0_False) }
@@ -336,7 +364,7 @@
     with 8 `x = 1` have "isnorm (R \<otimes> P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
     with 8 `x = 1` Y0 have ?case by (simp add: mkPX_cn) }
   moreover
-  { assume "x > 1" hence "EX d. x = Suc (Suc d)" by arith
+  { assume "x > 1" then have "\<exists>d. x = Suc (Suc d)" by arith
     then obtain d where X: "x = Suc (Suc d)" ..
     from 8 have NR: "isnorm R" "isnorm Q2"
       by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R])
@@ -366,7 +394,7 @@
     "isnorm (mkPX (P1 \<otimes> mkPinj (Suc 0) Q2) x (Pc 0))"  
     "isnorm (mkPX (Q1 \<otimes> mkPinj (Suc 0) P2) y (Pc 0))" 
     by (auto simp add: mkPX_cn)
-  thus ?case by (simp add: add_cn)
+  then show ?case by (simp add: add_cn)
 qed simp
 
 text {* neg conserves normalizedness *}
@@ -404,7 +432,7 @@
     by (cases Q) (auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_cn)
 next 
   case (PX P1 x P2)
-  then have "x + x ~= 0" "isnorm P2" "isnorm P1"
+  then have "x + x \<noteq> 0" "isnorm P2" "isnorm P1"
     by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2])
   with PX have "isnorm (mkPX (Pc (1 + 1) \<otimes> P1 \<otimes> mkPinj (Suc 0) P2) x (Pc 0))"
       and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))"
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Tue Feb 25 22:13:57 2014 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Tue Feb 25 23:12:48 2014 +0100
@@ -5,17 +5,22 @@
 header{* A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008 *}
 
 theory Parametric_Ferrante_Rackoff
-imports Reflected_Multivariate_Polynomial Dense_Linear_Order DP_Library
-  "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
+imports
+  Reflected_Multivariate_Polynomial
+  Dense_Linear_Order
+  DP_Library
+  "~~/src/HOL/Library/Code_Target_Numeral"
+  "~~/src/HOL/Library/Old_Recdef"
 begin
 
 subsection {* Terms *}
 
-datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm 
+datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm
   | Neg tm | Sub tm tm | CNP nat poly tm
-  (* A size for poly to make inductive proofs simpler*)
 
-primrec tmsize :: "tm \<Rightarrow> nat" where
+(* A size for poly to make inductive proofs simpler*)
+primrec tmsize :: "tm \<Rightarrow> nat"
+where
   "tmsize (CP c) = polysize c"
 | "tmsize (Bound n) = 1"
 | "tmsize (Neg a) = 1 + tmsize a"
@@ -24,18 +29,19 @@
 | "tmsize (Mul c a) = 1 + polysize c + tmsize a"
 | "tmsize (CNP n c a) = 3 + polysize c + tmsize a "
 
-  (* Semantics of terms tm *)
-primrec Itm :: "'a::{field_char_0, field_inverse_zero} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a" where
+(* Semantics of terms tm *)
+primrec Itm :: "'a::{field_char_0, field_inverse_zero} 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"   
+| "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
+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)"
@@ -44,46 +50,55 @@
 | "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
 | "allpolys P p = True"
 
-primrec tmboundslt:: "nat \<Rightarrow> tm \<Rightarrow> bool" where
+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 (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)"
 | "tmboundslt n (Mul i a) = tmboundslt n a"
 
-primrec tmbound0:: "tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound 0 *) where
+primrec tmbound0 :: "tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound 0 *)
+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 (Sub a b) = (tmbound0 a \<and> tmbound0 b)"
 | "tmbound0 (Mul i a) = tmbound0 a"
+
 lemma tmbound0_I:
   assumes nb: "tmbound0 a"
   shows "Itm vs (b#bs) a = Itm vs (b'#bs) a"
-using nb
-by (induct a rule: tm.induct,auto)
+  using nb
+  by (induct a rule: tm.induct,auto)
 
-primrec tmbound:: "nat \<Rightarrow> tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound n *) where
+primrec tmbound :: "nat \<Rightarrow> tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound n *)
+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 (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)
+
+lemma tmbound0_tmbound_iff: "tmbound 0 t = tmbound0 t"
+  by (induct t) auto
 
-lemma tmbound_I: 
-  assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound n t" and le: "n \<le> length bs"
+lemma tmbound_I:
+  assumes bnd: "tmboundslt (length bs) t"
+    and nb: "tmbound n t"
+    and le: "n \<le> length bs"
   shows "Itm vs (bs[n:=x]) t = Itm vs bs t"
   using nb le bnd
-  by (induct t rule: tm.induct , auto)
+  by (induct t rule: tm.induct) auto
 
-fun decrtm0:: "tm \<Rightarrow> tm" where
+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)"
@@ -92,7 +107,8 @@
 | "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
 | "decrtm0 a = a"
 
-fun incrtm0:: "tm \<Rightarrow> tm" where
+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)"
@@ -101,14 +117,16 @@
 | "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
 | "incrtm0 a = a"
 
-lemma decrtm0: assumes nb: "tmbound0 t"
-  shows "Itm vs (x#bs) t = Itm vs bs (decrtm0 t)"
-  using nb by (induct t rule: decrtm0.induct, simp_all)
+lemma decrtm0:
+  assumes nb: "tmbound0 t"
+  shows "Itm vs (x # bs) t = Itm vs bs (decrtm0 t)"
+  using nb by (induct t rule: decrtm0.induct) simp_all
 
 lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t"
-  by (induct t rule: decrtm0.induct, simp_all)
+  by (induct t rule: decrtm0.induct) simp_all
 
-primrec decrtm:: "nat \<Rightarrow> tm \<Rightarrow> tm" where
+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)"
@@ -117,39 +135,47 @@
 | "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
+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)))"
 
 lemma removen_same: "n \<ge> length xs \<Longrightarrow> removen n xs = xs"
-  by (induct xs arbitrary: n, auto)
+  by (induct xs arbitrary: n) auto
 
 lemma nth_length_exceeds: "n \<ge> length xs \<Longrightarrow> xs!n = []!(n - length xs)"
+  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)
 
-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)
-lemma removen_nth: "(removen n xs)!m = (if n \<ge> length xs then xs!m 
-  else if m < n then xs!m else if m \<le> length xs then xs!(Suc m) else []!(m - (length xs - 1)))"
-proof(induct xs arbitrary: n m)
-  case Nil thus ?case by simp
+lemma removen_nth:
+  "(removen n xs)!m =
+    (if n \<ge> length xs then xs!m
+     else if m < n then xs!m
+     else if m \<le> length xs then xs!(Suc m)
+     else []!(m - (length xs - 1)))"
+proof (induct xs arbitrary: n m)
+  case Nil
+  thus ?case by simp
 next
   case (Cons x xs n m)
   {assume nxs: "n \<ge> length (x#xs)" hence ?case using removen_same[OF nxs] by simp}
   moreover
-  {assume nxs: "\<not> (n \<ge> length (x#xs))" 
+  {assume nxs: "\<not> (n \<ge> length (x#xs))"
     {assume mln: "m < n" hence ?case using Cons by (cases m, auto)}
     moreover
-    {assume mln: "\<not> (m < n)" 
+    {assume mln: "\<not> (m < n)"
       {assume mxs: "m \<le> length (x#xs)" hence ?case using Cons by (cases m, auto)}
       moreover
-      {assume mxs: "\<not> (m \<le> length (x#xs))" 
-        have th: "length (removen n (x#xs)) = length xs" 
+      {assume mxs: "\<not> (m \<le> length (x#xs))"
+        have th: "length (removen n (x#xs)) = length xs"
           using removen_length[where n="n" and xs="x#xs"] nxs by simp
         with mxs have mxs':"m \<ge> length (removen n (x#xs))" by auto
-        hence "(removen n (x#xs))!m = [] ! (m - length xs)" 
+        hence "(removen n (x#xs))!m = [] ! (m - length xs)"
           using th nth_length_exceeds[OF mxs'] by auto
-        hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))" 
+        hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))"
           by auto
         hence ?case using nxs mln mxs by auto }
       ultimately have ?case by blast
@@ -158,196 +184,225 @@
   } ultimately show ?case by blast
 qed
 
-lemma decrtm: assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound m t" 
-  and nle: "m \<le> length bs" 
+lemma decrtm:
+  assumes bnd: "tmboundslt (length bs) t"
+    and nb: "tmbound m t"
+    and nle: "m \<le> length bs"
   shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t"
   using bnd nb nle by (induct t rule: tm.induct) (auto simp add: removen_nth)
 
-primrec tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm" where
+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 (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)"
 | "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
-lemma tmsubst0:
-  shows "Itm vs (x#bs) (tmsubst0 t a) = Itm vs ((Itm vs (x#bs) t)#bs) 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
 
 lemma tmsubst0_nb: "tmbound0 t \<Longrightarrow> tmbound0 (tmsubst0 t a)"
   by (induct a rule: tm.induct) auto
 
-primrec tmsubst:: "nat \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm" where
+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 (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 (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" and nlt: "n \<le> length bs"
+lemma tmsubst:
+  assumes nb: "tmboundslt (length bs) a"
+    and nlt: "n \<le> length bs"
   shows "Itm vs bs (tmsubst n t a) = Itm vs (bs[n:= Itm vs bs t]) a"
-using nb nlt
-by (induct a rule: tm.induct,auto)
+  using nb nlt
+  by (induct a rule: tm.induct) auto
 
-lemma tmsubst_nb0: assumes tnb: "tmbound0 t"
-shows "tmbound0 (tmsubst 0 t a)"
-using tnb
-by (induct a rule: tm.induct, auto)
+lemma tmsubst_nb0:
+  assumes tnb: "tmbound0 t"
+  shows "tmbound0 (tmsubst 0 t a)"
+  using tnb
+  by (induct a rule: tm.induct) auto
 
-lemma tmsubst_nb: assumes tnb: "tmbound m t"
-shows "tmbound m (tmsubst m t a)"
-using tnb
-by (induct a rule: tm.induct, auto)
+lemma tmsubst_nb:
+  assumes tnb: "tmbound m t"
+  shows "tmbound m (tmsubst m t a)"
+  using tnb
+  by (induct a rule: tm.induct) auto
+
 lemma incrtm0_tmbound: "tmbound n t \<Longrightarrow> tmbound (Suc n) (incrtm0 t)"
-  by (induct t, auto)
-  (* Simplification *)
+  by (induct t) auto
 
-consts
-  tmadd:: "tm \<times> tm \<Rightarrow> tm"
+(* Simplification *)
+
+consts tmadd:: "tm \<times> tm \<Rightarrow> tm"
 recdef tmadd "measure (\<lambda> (t,s). size t + size s)"
   "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))" 
+    (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"
+  "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, simp_all add: Let_def)
-apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \<le> n2", simp_all)
-apply (case_tac "n1 = n2", simp_all add: field_simps)
-apply (simp only: distrib_left[symmetric]) 
-by (auto simp del: polyadd simp add: polyadd[symmetric])
+lemma tmadd[simp]: "Itm vs bs (tmadd (t, s)) = Itm vs bs (Add t s)"
+  apply (induct t s rule: tmadd.induct, simp_all add: Let_def)
+  apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \<le> n2", simp_all)
+  apply (case_tac "n1 = n2", simp_all add: field_simps)
+  apply (simp only: distrib_left[symmetric])
+  apply (auto simp del: polyadd simp add: polyadd[symmetric])
+  done
+
+lemma tmadd_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmadd (t, s))"
+  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
 
-lemma tmadd_nb0[simp]: "\<lbrakk> tmbound0 t ; tmbound0 s\<rbrakk> \<Longrightarrow> tmbound0 (tmadd (t,s))"
-by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+lemma tmadd_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmadd (t, s))"
+  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
+
+lemma tmadd_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmadd (t, s))"
+  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
 
-lemma tmadd_nb[simp]: "\<lbrakk> tmbound n t ; tmbound n s\<rbrakk> \<Longrightarrow> tmbound n (tmadd (t,s))"
-by (induct t s rule: tmadd.induct, auto simp add: Let_def)
-lemma tmadd_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmadd (t,s))"
-by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+lemma tmadd_allpolys_npoly[simp]:
+  "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd(t, s))"
+  by (induct t s rule: tmadd.induct) (simp_all add: Let_def polyadd_norm)
 
-lemma tmadd_allpolys_npoly[simp]: "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd(t,s))" by (induct t s rule: tmadd.induct, simp_all add: Let_def polyadd_norm)
-
-fun tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm" where
+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)"
 
 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)
+  by (induct t arbitrary: i rule: tmmul.induct) (simp_all add: field_simps)
 
 lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
-by (induct t arbitrary: i rule: tmmul.induct, auto )
+  by (induct t arbitrary: i rule: tmmul.induct) auto
 
 lemma tmmul_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmmul t i)"
-by (induct t arbitrary: n rule: tmmul.induct, auto )
+  by (induct t arbitrary: n rule: tmmul.induct) auto
+
 lemma tmmul_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmmul t i)"
-by (induct t arbitrary: i rule: tmmul.induct, auto simp add: Let_def)
+  by (induct t arbitrary: i rule: tmmul.induct) (auto simp add: Let_def)
 
-lemma tmmul_allpolys_npoly[simp]: 
+lemma tmmul_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
+  shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)"
+  by (induct t rule: tmmul.induct) (simp_all add: Let_def polymul_norm)
 
-definition tmneg :: "tm \<Rightarrow> tm" where
-  "tmneg t \<equiv> tmmul t (C (- 1,1))"
+definition tmneg :: "tm \<Rightarrow> tm"
+  where "tmneg t \<equiv> tmmul t (C (- 1,1))"
 
-definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm" where
-  "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))"
+definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
+  where "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd (s, tmneg t))"
 
 lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
-using tmneg_def[of t] 
-apply simp
-done
+  using tmneg_def[of t] by simp
 
 lemma tmneg_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmneg t)"
-using tmneg_def by simp
+  using tmneg_def by simp
 
 lemma tmneg_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmneg t)"
-using tmneg_def by simp
+  using tmneg_def by simp
+
 lemma tmneg_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmneg t)"
-using tmneg_def by simp
-lemma [simp]: "isnpoly (C (-1,1))" unfolding isnpoly_def by simp
-lemma tmneg_allpolys_npoly[simp]: 
+  using tmneg_def by simp
+
+lemma [simp]: "isnpoly (C (-1, 1))"
+  unfolding isnpoly_def by simp
+
+lemma tmneg_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)" 
+  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)"
   unfolding tmneg_def by auto
 
 lemma tmsub[simp]: "Itm vs bs (tmsub a b) = Itm vs bs (Sub a b)"
-using tmsub_def by simp
+  using tmsub_def by simp
+
+lemma tmsub_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmsub t s)"
+  using tmsub_def by simp
 
-lemma tmsub_nb0[simp]: "\<lbrakk> tmbound0 t ; tmbound0 s\<rbrakk> \<Longrightarrow> tmbound0 (tmsub t s)"
-using tmsub_def by simp
-lemma tmsub_nb[simp]: "\<lbrakk> tmbound n t ; tmbound n s\<rbrakk> \<Longrightarrow> tmbound n (tmsub t s)"
-using tmsub_def by simp
-lemma tmsub_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmsub t s )"
-using tmsub_def by simp
-lemma tmsub_allpolys_npoly[simp]: 
+lemma tmsub_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmsub t s)"
+  using tmsub_def by simp
+
+lemma tmsub_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmsub t s)"
+  using tmsub_def by simp
+
+lemma tmsub_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)" 
+  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)"
   unfolding tmsub_def by (simp add: isnpoly_def)
 
-fun simptm:: "tm \<Rightarrow> tm" where
+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))"
+| "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: 
+lemma polynate_stupid:
   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a)"
-apply (subst polynate[symmetric])
-apply simp
-done
+  apply (subst polynate[symmetric])
+  apply simp
+  done
 
 lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t"
-by (induct t rule: simptm.induct, auto simp add: tmneg tmadd tmsub tmmul Let_def polynate_stupid) 
+  by (induct t rule: simptm.induct)
+    (auto simp add: tmneg tmadd tmsub tmmul Let_def polynate_stupid)
 
-lemma simptm_tmbound0[simp]: 
-  "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
-by (induct t rule: simptm.induct, auto simp add: Let_def)
+lemma simptm_tmbound0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
+  by (induct t rule: simptm.induct) (auto simp add: Let_def)
 
 lemma simptm_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (simptm t)"
-by (induct t rule: simptm.induct, auto simp add: Let_def)
+  by (induct t rule: simptm.induct) (auto simp add: Let_def)
+
 lemma simptm_nlt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (simptm t)"
-by (induct t rule: simptm.induct, auto simp add: Let_def)
+  by (induct t rule: simptm.induct) (auto simp add: Let_def)
 
-lemma [simp]: "isnpoly 0\<^sub>p" and [simp]: "isnpoly (C(1,1))" 
+lemma [simp]: "isnpoly 0\<^sub>p"
+  and [simp]: "isnpoly (C(1,1))"
   by (simp_all add: isnpoly_def)
-lemma simptm_allpolys_npoly[simp]: 
+
+lemma simptm_allpolys_npoly[simp]:
   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "allpolys isnpoly (simptm p)"
-  by (induct p rule: simptm.induct, auto simp add: Let_def)
+  by (induct p rule: simptm.induct) (auto simp add: Let_def)
 
 declare let_cong[fundef_cong del]
 
-fun split0 :: "tm \<Rightarrow> (poly \<times> tm)" where
+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 (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]
 
-lemma split0_stupid[simp]: "\<exists>x y. (x,y) = split0 p"
+lemma split0_stupid[simp]: "\<exists>x y. (x, y) = split0 p"
   apply (rule exI[where x="fst (split0 p)"])
   apply (rule exI[where x="snd (split0 p)"])
-  by simp
+  apply simp
+  done
 
 lemma split0:
   "tmbound 0 (snd (split0 t)) \<and> (Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t)"
@@ -364,50 +419,61 @@
   done
 
 lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
-proof-
+proof -
   fix c' t'
-  assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "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')" by simp
+  assume "split0 t = (c', t')"
+  hence "c' = fst (split0 t)" and "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')"
+    by simp
 qed
 
-lemma split0_nb0: 
+lemma split0_nb0:
   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "split0 t = (c',t') \<Longrightarrow>  tmbound 0 t'"
-proof-
+proof -
   fix c' t'
-  assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto
-  with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'" by simp
+  assume "split0 t = (c', t')"
+  hence "c' = fst (split0 t)" and "t' = snd (split0 t)"
+    by auto
+  with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'"
+    by simp
 qed
 
-lemma split0_nb0'[simp]:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+lemma split0_nb0'[simp]:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "tmbound0 (snd (split0 t))"
-  using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"] by (simp add: tmbound0_tmbound_iff)
+  using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"]
+  by (simp add: tmbound0_tmbound_iff)
 
-
-lemma split0_nb: assumes nb:"tmbound n t" shows "tmbound n (snd (split0 t))"
-  using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+lemma split0_nb:
+  assumes nb: "tmbound n t"
+  shows "tmbound n (snd (split0 t))"
+  using nb by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
 
-lemma split0_blt: assumes nb:"tmboundslt n t" shows "tmboundslt n (snd (split0 t))"
-  using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
-
-lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0"
- by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+lemma split0_blt:
+  assumes nb: "tmboundslt n t"
+  shows "tmboundslt n (snd (split0 t))"
+  using nb by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
 
-lemma tmboundslt_split0: "tmboundslt n t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0 \<or> n > 0"
-by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
+  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
 
-lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0"
- by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+lemma tmboundslt_split0: "tmboundslt n t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0 \<or> n > 0"
+  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
+
+lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
+  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
 
 lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
-by (induct p rule: split0.induct, auto simp  add: isnpoly_def Let_def split_def split0_stupid)
+  by (induct p rule: split0.induct) (auto simp  add: isnpoly_def Let_def split_def)
 
-lemma isnpoly_fst_split0:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows 
-  "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
-  by (induct p rule: split0.induct, 
-    auto simp  add: polyadd_norm polysub_norm polyneg_norm polymul_norm 
-    Let_def split_def split0_stupid)
+lemma isnpoly_fst_split0:
+  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
+  shows "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
+  by (induct p rule: split0.induct)
+    (auto simp  add: polyadd_norm polysub_norm polyneg_norm polymul_norm Let_def split_def)
+
 
 subsection{* Formulae *}
 
@@ -415,8 +481,9 @@
   NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
 
 
-  (* A size for fm *)
-fun fmsize :: "fm \<Rightarrow> nat" where
+(* A size for fm *)
+fun fmsize :: "fm \<Rightarrow> nat"
+where
   "fmsize (NOT p) = 1 + fmsize p"
 | "fmsize (And p q) = 1 + fmsize p + fmsize q"
 | "fmsize (Or p q) = 1 + fmsize p + fmsize q"
@@ -425,9 +492,10 @@
 | "fmsize (E p) = 1 + fmsize p"
 | "fmsize (A p) = 4+ fmsize p"
 | "fmsize p = 1"
-  (* several lemmas about fmsize *)
-lemma fmsize_pos[termination_simp]: "fmsize p > 0"        
-by (induct p rule: fmsize.induct) simp_all
+
+(* several lemmas about fmsize *)
+lemma fmsize_pos[termination_simp]: "fmsize p > 0"
+  by (induct p rule: fmsize.induct) simp_all
 
   (* Semantics of formulae (fm) *)
 primrec Ifm ::"'a::{linordered_field_inverse_zero} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool" where
@@ -442,8 +510,8 @@
 | "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)"
+| "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"
@@ -455,49 +523,74 @@
 | "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
+  by (induct p rule: not.induct) auto
 
-definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" 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 else 
-   if p = q then p else And p q)"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+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
+     else if p = q then p
+     else And p q)"
+
 lemma conj[simp]: "Ifm vs bs (conj p q) = Ifm vs bs (And p q)"
-by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
+  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> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
-       else if p=q then p else Or p q)"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+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
+     else if p = q then p
+     else Or p q)"
 
 lemma disj[simp]: "Ifm vs bs (disj p q) = Ifm vs bs (Or p q)"
-by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
+  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> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p 
-    else Imp p q)"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+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
+     else Imp p q)"
+
 lemma imp[simp]: "Ifm vs bs (imp p q) = Ifm vs bs (Imp p q)"
-by (cases "p=F \<or> q=T",simp_all add: imp_def) 
+  by (cases "p=F \<or> q=T") (simp_all add: imp_def)
 
-definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" 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 else if q=F then not p else if p=T then q else if q=T then p else 
-  Iff p q)"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+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
+    else if q = F then not p
+    else if p = T then q
+    else if q = T then p
+    else Iff p q)"
+
 lemma iff[simp]: "Ifm vs bs (iff p q) = Ifm vs bs (Iff p q)"
   by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", auto)
 
-  (* Quantifier freeness *)
-fun qfree:: "fm \<Rightarrow> bool" where
+(* Quantifier freeness *)
+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 (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"
 
-  (* Boundedness and substitution *)
-
-primrec boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool" where
+(* Boundedness and substitution *)
+primrec boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool"
+where
   "boundslt n T = True"
 | "boundslt n F = True"
 | "boundslt n (Lt t) = (tmboundslt n t)"
@@ -512,7 +605,8 @@
 | "boundslt n (E p) = boundslt (Suc n) p"
 | "boundslt n (A p) = boundslt (Suc n) p"
 
-fun bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *) where
+fun bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
+where
   "bound0 T = True"
 | "bound0 F = True"
 | "bound0 (Lt a) = tmbound0 a"
@@ -525,13 +619,15 @@
 | "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"
   shows "Ifm vs (b#bs) p = Ifm vs (b'#bs) p"
-using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"]
-by (induct p rule: bound0.induct,auto)
+  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" (* A Formula is independent of Bound n *) where
+primrec bound:: "nat \<Rightarrow> fm \<Rightarrow> bool" (* A Formula is independent of Bound n *)
+where
   "bound m T = True"
 | "bound m F = True"
 | "bound m (Lt t) = tmbound m t"
@@ -547,22 +643,31 @@
 | "bound m (A p) = bound (Suc m) p"
 
 lemma bound_I:
-  assumes bnd: "boundslt (length bs) p" and nb: "bound n p" and le: "n \<le> length bs"
+  assumes bnd: "boundslt (length bs) p"
+    and nb: "bound n p"
+    and le: "n \<le> length bs"
   shows "Ifm vs (bs[n:=x]) p = Ifm vs bs p"
   using bnd nb le tmbound_I[where bs=bs and vs = vs]
-proof(induct p arbitrary: bs n rule: fm.induct)
-  case (E p bs n) 
-  {fix y
-    from E have bnd: "boundslt (length (y#bs)) p" 
+proof (induct p arbitrary: bs n rule: fm.induct)
+  case (E p bs n)
+  {
+    fix y
+    from E have bnd: "boundslt (length (y#bs)) p"
       and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
-    from E.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" .   }
-  thus ?case by simp 
+    from E.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" .
+  }
+  thus ?case by simp
 next
-  case (A p bs n) {fix y
-    from A have bnd: "boundslt (length (y#bs)) p" 
-      and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
-    from A.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" .   }
-  thus ?case by simp 
+  case (A p bs n)
+  {
+    fix y
+    from A have bnd: "boundslt (length (y#bs)) p"
+      and nb: "bound (Suc n) p"
+      and le: "Suc n \<le> length (y#bs)"
+      by simp_all
+    from A.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" .
+  }
+  thus ?case by simp
 qed auto
 
 fun decr0 :: "fm \<Rightarrow> fm" where
@@ -570,26 +675,28 @@
 | "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 (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"
+lemma decr0:
+  assumes nb: "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 nb
+  by (induct p rule: decr0.induct) (simp_all add: decrtm0)
 
-primrec decr :: "nat \<Rightarrow> fm \<Rightarrow> fm" where
+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 (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)"
@@ -597,27 +704,40 @@
 | "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" and nb: "bound m p" 
-  and nle: "m < length bs" 
+lemma decr:
+  assumes bnd: "boundslt (length bs) p"
+    and nb: "bound m p"
+    and nle: "m < length bs"
   shows "Ifm vs (removen m bs) (decr m p) = Ifm vs bs p"
   using bnd nb nle
-proof(induct p arbitrary: bs m rule: fm.induct)
-  case (E p bs m) 
-  {fix x
-    from E have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" 
-  and nle: "Suc m < length (x#bs)" by auto
-    from E(1)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p".
-  } thus ?case by auto 
+proof (induct p arbitrary: bs m rule: fm.induct)
+  case (E p bs m)
+  { fix x
+    from E
+    have bnd: "boundslt (length (x#bs)) p"
+      and nb: "bound (Suc m) p"
+      and nle: "Suc m < length (x#bs)"
+      by auto
+    from E(1)[OF bnd nb nle]
+    have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p" .
+  }
+  thus ?case by auto
 next
-  case (A p bs m)  
-  {fix x
-    from A have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" 
-  and nle: "Suc m < length (x#bs)" by auto
-    from A(1)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p".
-  } thus ?case by auto
+  case (A p bs m)
+  { fix x
+    from A
+    have bnd: "boundslt (length (x#bs)) p"
+      and nb: "bound (Suc m) p"
+      and nle: "Suc m < length (x#bs)"
+      by auto
+    from A(1)[OF bnd nb nle]
+    have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p" .
+  }
+  thus ?case by auto
 qed (auto simp add: decrtm removen_nth)
 
-primrec subst0:: "tm \<Rightarrow> fm \<Rightarrow> fm" where
+primrec subst0 :: "tm \<Rightarrow> fm \<Rightarrow> fm"
+where
   "subst0 t T = T"
 | "subst0 t F = F"
 | "subst0 t (Lt a) = Lt (tmsubst0 t a)"
@@ -632,18 +752,21 @@
 | "subst0 t (E p) = E p"
 | "subst0 t (A p) = A p"
 
-lemma subst0: assumes qf: "qfree p"
-  shows "Ifm vs (x#bs) (subst0 t p) = Ifm vs ((Itm vs (x#bs) t)#bs) p"
-using qf tmsubst0[where x="x" and bs="bs" and t="t"]
-by (induct p rule: fm.induct, auto)
+lemma subst0:
+  assumes qf: "qfree p"
+  shows "Ifm vs (x # bs) (subst0 t p) = Ifm vs ((Itm vs (x # bs) t) # bs) p"
+  using qf tmsubst0[where x="x" and bs="bs" and t="t"]
+  by (induct p rule: fm.induct) auto
 
 lemma subst0_nb:
-  assumes bp: "tmbound0 t" and qf: "qfree p"
+  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
+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)"
@@ -658,82 +781,98 @@
 | "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" and nlm: "n \<le> length bs"
+lemma subst:
+  assumes nb: "boundslt (length bs) p"
+    and nlm: "n \<le> length bs"
   shows "Ifm vs bs (subst n t p) = Ifm vs (bs[n:= Itm vs bs t]) p"
   using nb nlm
 proof (induct p arbitrary: bs n t rule: fm.induct)
-  case (E p bs n) 
-  {fix x 
-    from E have bn: "boundslt (length (x#bs)) p" by simp 
-    from E have nlm: "Suc n \<le> length (x#bs)" by simp
-    from E(1)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp 
+  case (E p bs n)
+  {
+    fix x
+    from E have bn: "boundslt (length (x#bs)) p"
+      by simp
+    from E have nlm: "Suc n \<le> length (x#bs)"
+      by simp
+    from E(1)[OF bn nlm]
+    have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p"
+      by simp
     hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p"
-    by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) }  
-thus ?case by simp 
+      by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
+  }
+  thus ?case by simp
 next
-  case (A p bs n)   
-  {fix x 
-    from A have bn: "boundslt (length (x#bs)) p" by simp 
-    from A have nlm: "Suc n \<le> length (x#bs)" by simp
-    from A(1)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp 
+  case (A p bs n)
+  {
+    fix x
+    from A have bn: "boundslt (length (x#bs)) p"
+      by simp
+    from A have nlm: "Suc n \<le> length (x#bs)"
+      by simp
+    from A(1)[OF bn nlm]
+    have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p"
+      by simp
     hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p"
-    by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) }  
-thus ?case by simp 
-qed(auto simp add: tmsubst)
+      by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
+  }
+  thus ?case by simp
+qed (auto simp add: tmsubst)
 
-lemma subst_nb: assumes tnb: "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)
+lemma subst_nb:
+  assumes tnb: "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
 
 lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
-by (induct p rule: not.induct, auto)
+  by (induct p rule: not.induct) auto
 lemma not_bn0[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
-by (induct p rule: not.induct, auto)
+  by (induct p rule: not.induct) auto
 lemma not_nb[simp]: "bound n p \<Longrightarrow> bound n (not p)"
-by (induct p rule: not.induct, auto)
+  by (induct p rule: not.induct) auto
 lemma not_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n (not p)"
- by (induct p rule: not.induct, auto)
+  by (induct p rule: not.induct) auto
 
-lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
-using conj_def by auto 
-lemma conj_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
-using conj_def by auto 
-lemma conj_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (conj p q)"
-using conj_def by auto 
+lemma conj_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (conj p q)"
+  using conj_def by auto
+lemma conj_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (conj p q)"
+  using conj_def by auto
+lemma conj_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (conj p q)"
+  using conj_def by auto
 lemma conj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
-using conj_def by auto 
+  using conj_def by auto
 
-lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
-using disj_def by auto 
-lemma disj_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
-using disj_def by auto 
-lemma disj_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (disj p q)"
-using disj_def by auto 
+lemma disj_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (disj p q)"
+  using disj_def by auto
+lemma disj_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (disj p q)"
+  using disj_def by auto
+lemma disj_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (disj p q)"
+  using disj_def by auto
 lemma disj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (disj p q)"
-using disj_def by auto 
+  using disj_def by auto
 
-lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
-using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
-lemma imp_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
-using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
-lemma imp_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (imp p q)"
-using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
+lemma imp_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (imp p q)"
+  using imp_def by (cases "p=F \<or> q=T") (simp_all add: imp_def)
+lemma imp_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (imp p q)"
+  using imp_def by (cases "p=F \<or> q=T \<or> p=q") (simp_all add: imp_def)
+lemma imp_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (imp p q)"
+  using imp_def by (cases "p=F \<or> q=T \<or> p=q") (simp_all add: imp_def)
 lemma imp_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (imp p q)"
-using imp_def by auto 
+  using imp_def by auto
 
-lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
-  by (unfold iff_def,cases "p=q", auto)
-lemma iff_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)"
-using iff_def by (unfold iff_def,cases "p=q", auto)
-lemma iff_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (iff p q)"
-using iff_def by (unfold iff_def,cases "p=q", auto)
+lemma iff_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (iff p q)"
+  unfolding iff_def by (cases "p = q") auto
+lemma iff_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (iff p q)"
+  using iff_def unfolding iff_def by (cases "p = q") auto
+lemma iff_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (iff p q)"
+  using iff_def unfolding iff_def by (cases "p = q") auto
 lemma iff_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (iff p q)"
-using iff_def by auto 
+  using iff_def by auto
 lemma decr0_qf: "bound0 p \<Longrightarrow> qfree (decr0 p)"
-by (induct p, simp_all)
+  by (induct p) simp_all
 
-fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *) where
+fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
+where
   "isatom T = True"
 | "isatom F = True"
 | "isatom (Lt a) = True"
@@ -743,47 +882,53 @@
 | "isatom p = False"
 
 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
-by (induct p, simp_all)
+  by (induct p) simp_all
 
-definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" 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))"
-definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
-  "evaldjf f ps \<equiv> foldr (djf f) ps F"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+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))"
+
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+  where "evaldjf f ps \<equiv> foldr (djf f) ps F"
 
 lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
-by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) 
-(cases "f p", simp_all add: Let_def djf_def) 
+  by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def)
+    (cases "f p", simp_all add: Let_def djf_def)
 
-lemma evaldjf_ex: "Ifm vs bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm vs bs (f p))"
-  by(induct ps, simp_all add: evaldjf_def djf_Or)
+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)"
+lemma evaldjf_bound0:
+  assumes nb: "\<forall>x\<in> set xs. bound0 (f x)"
   shows "bound0 (evaldjf f xs)"
-  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
+  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
 
-lemma evaldjf_qf: 
-  assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
+lemma evaldjf_qf:
+  assumes nb: "\<forall>x\<in> set xs. qfree (f x)"
   shows "qfree (evaldjf f xs)"
-  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
+  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)"
+fun disjuncts :: "fm \<Rightarrow> fm list"
+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: "(\<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"
-proof-
+lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall>q \<in> set (disjuncts p). bound0 q"
+proof -
   assume nb: "bound0 p"
-  hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto)
+  hence "list_all bound0 (disjuncts p)"
+    by (induct p rule:disjuncts.induct,auto)
   thus ?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: "qfree p \<Longrightarrow> \<forall>q\<in> set (disjuncts p). qfree q"
 proof-
   assume qf: "qfree p"
   hence "list_all qfree (disjuncts p)"
@@ -794,38 +939,38 @@
 definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "DJ f p \<equiv> evaldjf f (disjuncts p)"
 
-lemma DJ: assumes fdj: "\<forall> p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))"
+lemma DJ: assumes fdj: "\<forall>p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))"
   and fF: "f F = F"
   shows "Ifm vs bs (DJ f p) = Ifm vs bs (f p)"
 proof-
-  have "Ifm vs bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm vs bs (f q))"
-    by (simp add: DJ_def evaldjf_ex) 
+  have "Ifm vs bs (DJ f p) = (\<exists>q \<in> set (disjuncts p). Ifm vs bs (f q))"
+    by (simp add: DJ_def evaldjf_ex)
   also have "\<dots> = Ifm vs bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
   finally show ?thesis .
 qed
 
-lemma DJ_qf: assumes 
-  fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
+lemma DJ_qf: assumes
+  fqf: "\<forall>p. qfree p \<longrightarrow> qfree (f p)"
   shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
 proof(clarify)
   fix  p assume qf: "qfree p"
   have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def)
-  from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" .
-  with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
-  
+  from disjuncts_qf[OF qf] have "\<forall>q\<in> set (disjuncts p). qfree q" .
+  with fqf have th':"\<forall>q\<in> set (disjuncts p). qfree (f q)" by blast
+
   from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp
 qed
 
-lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
-  shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm vs bs ((DJ qe p)) = Ifm vs bs (E p))"
+lemma DJ_qe: assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+  shows "\<forall>bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm vs bs ((DJ qe p)) = Ifm vs bs (E p))"
 proof(clarify)
   fix p::fm and bs
   assume qf: "qfree p"
-  from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast
+  from qe have qth: "\<forall>p. qfree p \<longrightarrow> qfree (qe p)" by blast
   from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto
-  have "Ifm vs bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm vs bs (qe q))"
+  have "Ifm vs bs (DJ qe p) = (\<exists>q\<in> set (disjuncts p). Ifm vs bs (qe q))"
     by (simp add: DJ_def evaldjf_ex)
-  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm vs bs (E q))" using qe disjuncts_qf[OF qf] by auto
+  also have "\<dots> = (\<exists>q \<in> set(disjuncts p). Ifm vs bs (E q))" using qe disjuncts_qf[OF qf] by auto
   also have "\<dots> = Ifm vs bs (E p)" by (induct p rule: disjuncts.induct, auto)
   finally show "qfree (DJ qe p) \<and> Ifm vs bs (DJ qe p) = Ifm vs bs (E p)" using qfth by blast
 qed
@@ -842,7 +987,7 @@
   "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs
                    in conj (decr0 (list_conj yes)) (f (list_conj no)))"
 
-lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
+lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall>q\<in> set (conjuncts p). qfree q"
 proof-
   assume qf: "qfree p"
   hence "list_all qfree (conjuncts p)"
@@ -850,10 +995,10 @@
   thus ?thesis by (simp only: list_all_iff)
 qed
 
-lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm vs bs q) = Ifm vs bs p"
+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: "bound0 p \<Longrightarrow> \<forall>q\<in> set (conjuncts p). bound0 q"
 proof-
   assume nb: "bound0 p"
   hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto)
@@ -937,24 +1082,24 @@
 
 lemma simplt_islin[simp]:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "islin (simplt t)"
-  unfolding simplt_def 
+  unfolding simplt_def
   using split0_nb0'
 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]:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "islin (simple t)"
-  unfolding simple_def 
+  unfolding simple_def
   using split0_nb0'
 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]:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "islin (simpeq t)"
-  unfolding simpeq_def 
+  unfolding simpeq_def
   using split0_nb0'
 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]:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "islin (simpneq t)"
-  unfolding simpneq_def 
+  unfolding simpneq_def
   using split0_nb0'
 by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] neq_lin)
 
@@ -1061,7 +1206,7 @@
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
   have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
   from iffD1[OF isnpolyh_unique[OF ths] th]
-  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  have "fst (split0 (simptm t)) = 0\<^sub>p" .
   thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (lt (snd (split0 (simptm t))))) \<and>
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def lt_nb)
 qed
@@ -1078,7 +1223,7 @@
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
   have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
   from iffD1[OF isnpolyh_unique[OF ths] th]
-  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  have "fst (split0 (simptm t)) = 0\<^sub>p" .
   thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (le (snd (split0 (simptm t))))) \<and>
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def le_nb)
 qed
@@ -1095,7 +1240,7 @@
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
   have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
   from iffD1[OF isnpolyh_unique[OF ths] th]
-  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  have "fst (split0 (simptm t)) = 0\<^sub>p" .
   thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (eq (snd (split0 (simptm t))))) \<and>
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpeq_def Let_def split_def eq_nb)
 qed
@@ -1112,7 +1257,7 @@
   from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
   have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
   from iffD1[OF isnpolyh_unique[OF ths] th]
-  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  have "fst (split0 (simptm t)) = 0\<^sub>p" .
   thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (neq (snd (split0 (simptm t))))) \<and>
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpneq_def Let_def split_def neq_nb)
 qed
@@ -1121,7 +1266,7 @@
   "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"
+lemma conjs_ci: "(\<forall>q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
 by (induct p rule: conjs.induct, auto)
 definition list_disj :: "fm list \<Rightarrow> fm" where
   "list_disj ps \<equiv> foldr disj ps F"
@@ -1137,7 +1282,7 @@
   unfolding conj_def by auto
 
 lemma conjs_nb: "bound n p \<Longrightarrow> \<forall>q\<in> set (conjs p). bound n q"
-  apply (induct p rule: conjs.induct) 
+  apply (induct p rule: conjs.induct)
   apply (unfold conjs.simps)
   apply (unfold set_append)
   apply (unfold ball_Un)
@@ -1146,7 +1291,7 @@
   done
 
 lemma conjs_boundslt: "boundslt n p \<Longrightarrow> \<forall>q\<in> set (conjs p). boundslt n q"
-  apply (induct p rule: conjs.induct) 
+  apply (induct p rule: conjs.induct)
   apply (unfold conjs.simps)
   apply (unfold set_append)
   apply (unfold ball_Un)
@@ -1167,9 +1312,9 @@
 lemma list_conj_nb': "\<forall>p\<in>set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
 unfolding list_conj_def by (induct ps , auto)
 
-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))"
-  shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm vs bs ((CJNB qe p)) = Ifm vs bs (E p))"
+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))"
+  shows "\<forall>bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm vs bs ((CJNB qe p)) = Ifm vs bs (E p))"
 proof(clarify)
   fix bs p
   assume qfp: "qfree p"
@@ -1179,15 +1324,15 @@
   let ?cno = "list_conj ?no"
   let ?cyes = "list_conj ?yes"
   have part: "partition bound0 ?cjs = (?yes,?no)" by simp
-  from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast 
-  hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb') 
+  from partition_P[OF part] have "\<forall>q\<in> set ?yes. bound0 q" by blast
+  hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb')
   hence yes_qf: "qfree (decr0 ?cyes )" by (simp add: decr0_qf)
-  from conjuncts_qf[OF qfp] partition_set[OF part] 
+  from conjuncts_qf[OF qfp] partition_set[OF part]
   have " \<forall>q\<in> set ?no. qfree q" by auto
   hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
-  with qe have cno_qf:"qfree (qe ?cno )" 
+  with qe have cno_qf:"qfree (qe ?cno )"
     and noE: "Ifm vs bs (qe ?cno) = Ifm vs bs (E ?cno)" by blast+
-  from cno_qf yes_qf have qf: "qfree (CJNB qe p)" 
+  from cno_qf yes_qf have qf: "qfree (CJNB qe p)"
     by (simp add: CJNB_def Let_def conj_qf split_def)
   {fix bs
     from conjuncts have "Ifm vs bs p = (\<forall>q\<in> set ?cjs. Ifm vs bs q)" by blast
@@ -1201,7 +1346,7 @@
     by (auto simp add: decr0[OF yes_nb] simp del: partition_filter_conv)
   also have "\<dots> = (Ifm vs bs (conj (decr0 ?cyes) (qe ?cno)))"
     using qe[rule_format, OF no_qf] by auto
-  finally have "Ifm vs bs (E p) = Ifm vs bs (CJNB qe p)" 
+  finally have "Ifm vs bs (E p) = Ifm vs bs (CJNB qe p)"
     by (simp add: Let_def CJNB_def split_def)
   with qf show "qfree (CJNB qe p) \<and> Ifm vs bs (CJNB qe p) = Ifm vs bs (E p)" by blast
 qed
@@ -1262,7 +1407,7 @@
 lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)" by (simp add: conj_def)
 
 lemma   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
-  shows "qfree p \<Longrightarrow> islin (simpfm p)" 
+  shows "qfree p \<Longrightarrow> islin (simpfm p)"
   apply (induct p rule: simpfm.induct)
   apply (simp_all add: conj_lin disj_lin)
   done
@@ -1273,7 +1418,7 @@
   "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 (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)))"
@@ -1303,8 +1448,8 @@
   "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 (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)"
@@ -1312,7 +1457,7 @@
 termination by (relation "measure fmsize") auto
 
 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))"
+  assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
   shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm vs bs (qelim p qe) = Ifm vs bs p)"
 using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
 by (induct p rule: qelim.induct) auto
@@ -1320,8 +1465,8 @@
 subsection{* Core Procedure *}
 
 fun minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*) where
-  "minusinf (And p q) = conj (minusinf p) (minusinf q)" 
-| "minusinf (Or p q) = disj (minusinf p) (minusinf q)" 
+  "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)))"
@@ -1329,8 +1474,8 @@
 | "minusinf p = p"
 
 fun plusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of +\<infinity>*) where
-  "plusinf (And p q) = conj (plusinf p) (plusinf q)" 
-| "plusinf (Or p q) = disj (plusinf p) (plusinf q)" 
+  "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))"
@@ -1351,7 +1496,7 @@
   let ?c = "Ipoly vs c"
   let ?e = "Itm vs (y#bs) e"
   have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
-  moreover {assume "?c = 0" hence ?case 
+  moreover {assume "?c = 0" hence ?case
       using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
   moreover {assume cp: "?c > 0"
     {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
@@ -1447,10 +1592,10 @@
   let ?c = "Ipoly vs c"
   let ?e = "Itm vs (y#bs) e"
   have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
-  moreover {assume "?c = 0" hence ?case 
+  moreover {assume "?c = 0" hence ?case
       using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
   moreover {assume cp: "?c > 0"
-    {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e" 
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
         using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
       hence "?c * x + ?e > 0" by simp
       hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
@@ -1529,17 +1674,17 @@
   ultimately show ?case by blast
 qed (auto)
 
-lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)" 
+lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)"
   by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb)
-lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf p)" 
+lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf p)"
   by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb)
 
 lemma minusinf_ex: assumes lp: "islin p" and ex: "Ifm vs (x#bs) (minusinf p)"
   shows "\<exists>x. Ifm vs (x#bs) p"
 proof-
   from bound0_I [OF minusinf_nb[OF lp], where b="a" and bs ="bs"] ex
-  have th: "\<forall> x. Ifm vs (x#bs) (minusinf p)" by auto
-  from minusinf_inf[OF lp, where bs="bs"] 
+  have th: "\<forall>x. Ifm vs (x#bs) (minusinf p)" by auto
+  from minusinf_inf[OF lp, where bs="bs"]
   obtain z where z_def: "\<forall>x<z. Ifm vs (x # bs) (minusinf p) = Ifm vs (x # bs) p" by blast
   from th have "Ifm vs ((z - 1)#bs) (minusinf p)" by simp
   moreover have "z - 1 < z" by simp
@@ -1550,8 +1695,8 @@
   shows "\<exists>x. Ifm vs (x#bs) p"
 proof-
   from bound0_I [OF plusinf_nb[OF lp], where b="a" and bs ="bs"] ex
-  have th: "\<forall> x. Ifm vs (x#bs) (plusinf p)" by auto
-  from plusinf_inf[OF lp, where bs="bs"] 
+  have th: "\<forall>x. Ifm vs (x#bs) (plusinf p)" by auto
+  from plusinf_inf[OF lp, where bs="bs"]
   obtain z where z_def: "\<forall>x>z. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p" by blast
   from th have "Ifm vs ((z + 1)#bs) (plusinf p)" by simp
   moreover have "z + 1 > z" by simp
@@ -1569,20 +1714,20 @@
 
 lemma uset_l:
   assumes lp: "islin p"
-  shows "\<forall> (c,s) \<in> set (uset p). isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+  shows "\<forall>(c,s) \<in> set (uset p). isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
 using lp by(induct p rule: uset.induct,auto)
 
 lemma minusinf_uset0:
   assumes lp: "islin p"
   and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))"
   and ex: "Ifm vs (x#bs) p" (is "?I x p")
-  shows "\<exists> (c,s) \<in> set (uset p). x \<ge> - Itm vs (x#bs) s / Ipoly vs c" 
+  shows "\<exists>(c,s) \<in> set (uset p). x \<ge> - Itm vs (x#bs) s / Ipoly vs c"
 proof-
-  have "\<exists> (c,s) \<in> set (uset p). (Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>  (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" 
+  have "\<exists>(c,s) \<in> set (uset p). (Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>  (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, auto simp add: eq le lt polyneg_norm)
     apply (auto simp add: linorder_not_less order_le_less)
-    done 
+    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>  (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" by blast
   hence "x \<ge> (- Itm vs (x#bs) s) / Ipoly vs c"
     using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x]
@@ -1594,11 +1739,11 @@
   assumes lp: "islin p"
   and nmi: "\<not> (Ifm vs (a#bs) (minusinf p))"
   and ex: "Ifm vs (x#bs) p" (is "?I x p")
-  shows "\<exists> (c,s) \<in> set (uset p). x \<ge> - Itm vs (a#bs) s / Ipoly vs c" 
+  shows "\<exists>(c,s) \<in> set (uset p). x \<ge> - Itm vs (a#bs) s / Ipoly vs c"
 proof-
-  from nmi have nmi': "\<not> (Ifm vs (x#bs) (minusinf p))" 
+  from nmi have nmi': "\<not> (Ifm vs (x#bs) (minusinf p))"
     by (simp add: bound0_I[OF minusinf_nb[OF lp], where b=x and b'=a])
-  from minusinf_uset0[OF lp nmi' ex] 
+  from minusinf_uset0[OF lp nmi' ex]
   obtain c s where csU: "(c,s) \<in> set (uset p)" and th: "x \<ge> - Itm vs (x#bs) s / Ipoly vs c" by blast
   from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp
   from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto
@@ -1609,13 +1754,13 @@
   assumes lp: "islin p"
   and nmi: "\<not> (Ifm vs (x#bs) (plusinf p))"
   and ex: "Ifm vs (x#bs) p" (is "?I x p")
-  shows "\<exists> (c,s) \<in> set (uset p). x \<le> - Itm vs (x#bs) s / Ipoly vs c" 
+  shows "\<exists>(c,s) \<in> set (uset p). x \<le> - Itm vs (x#bs) s / Ipoly vs c"
 proof-
-  have "\<exists> (c,s) \<in> set (uset p). (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)" 
+  have "\<exists>(c,s) \<in> set (uset p). (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)"
     using lp nmi ex
     apply (induct p rule: minusinf.induct, auto simp add: eq le lt polyneg_norm)
     apply (auto simp add: linorder_not_less order_le_less)
-    done 
+    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)" by blast
   hence "x \<le> (- Itm vs (x#bs) s) / Ipoly vs c"
     using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"]
@@ -1627,27 +1772,27 @@
   assumes lp: "islin p"
   and nmi: "\<not> (Ifm vs (a#bs) (plusinf p))"
   and ex: "Ifm vs (x#bs) p" (is "?I x p")
-  shows "\<exists> (c,s) \<in> set (uset p). x \<le> - Itm vs (a#bs) s / Ipoly vs c" 
+  shows "\<exists>(c,s) \<in> set (uset p). x \<le> - Itm vs (a#bs) s / Ipoly vs c"
 proof-
-  from nmi have nmi': "\<not> (Ifm vs (x#bs) (plusinf p))" 
+  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] 
+  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" by blast
   from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp
   from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto
 qed
 
-lemma lin_dense: 
+lemma lin_dense:
   assumes lp: "islin p"
-  and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (c,t). - Itm vs (x#bs) t / Ipoly vs c) ` set (uset p)" 
-  (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (c,t). - ?Nt x t / ?N c) ` ?U p")
+  and noS: "\<forall>t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (c,t). - Itm vs (x#bs) t / Ipoly vs c) ` set (uset p)"
+  (is "\<forall>t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (c,t). - ?Nt x t / ?N c) ` ?U p")
   and lx: "l < x" and xu:"x < u" and px:" Ifm vs (x#bs) p"
   and ly: "l < y" and yu: "y < u"
   shows "Ifm vs (y#bs) p"
 using lp px noS
-proof (induct p rule: islin.induct) 
+proof (induct p rule: islin.induct)
   case (5 c s)
-  from "5.prems" 
+  from "5.prems"
   have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
     and px: "Ifm vs (x # bs) (Lt (CNP 0 c s))"
     and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
@@ -1658,16 +1803,16 @@
   {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
   moreover
   {assume c: "?N c > 0"
-      from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]  
-      have px': "x < - ?Nt x s / ?N c" 
-        by (auto simp add: not_less field_simps) 
-    {assume y: "y < - ?Nt x s / ?N c" 
+      from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]
+      have px': "x < - ?Nt x s / ?N c"
+        by (auto simp add: not_less field_simps)
+    {assume y: "y < - ?Nt x s / ?N c"
       hence "y * ?N c < - ?Nt x s"
         by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
       hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
-    {assume y: "y > -?Nt x s / ?N c" 
+    {assume y: "y > -?Nt x s / ?N c"
       with yu have eu: "u > - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
       with lx px' have "False" by simp  hence ?case by simp }
@@ -1675,16 +1820,16 @@
   }
   moreover
   {assume c: "?N c < 0"
-      from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]  
-      have px': "x > - ?Nt x s / ?N c" 
-        by (auto simp add: not_less field_simps) 
-    {assume y: "y > - ?Nt x s / ?N c" 
+      from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]
+      have px': "x > - ?Nt x s / ?N c"
+        by (auto simp add: not_less field_simps)
+    {assume y: "y > - ?Nt x s / ?N c"
       hence "y * ?N c < - ?Nt x s"
         by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
       hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
-    {assume y: "y < -?Nt x s / ?N c" 
+    {assume y: "y < -?Nt x s / ?N c"
       with ly have eu: "l < - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
       with xu px' have "False" by simp  hence ?case by simp }
@@ -1693,7 +1838,7 @@
   ultimately show ?case by blast
 next
   case (6 c s)
-  from "6.prems" 
+  from "6.prems"
   have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
     and px: "Ifm vs (x # bs) (Le (CNP 0 c s))"
     and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
@@ -1704,15 +1849,15 @@
   {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
   moreover
   {assume c: "?N c > 0"
-      from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"]  
-      have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less field_simps) 
-    {assume y: "y < - ?Nt x s / ?N c" 
+      from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"]
+      have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less field_simps)
+    {assume y: "y < - ?Nt x s / ?N c"
       hence "y * ?N c < - ?Nt x s"
         by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
       hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
-    {assume y: "y > -?Nt x s / ?N c" 
+    {assume y: "y > -?Nt x s / ?N c"
       with yu have eu: "u > - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
       with lx px' have "False" by simp  hence ?case by simp }
@@ -1720,15 +1865,15 @@
   }
   moreover
   {assume c: "?N c < 0"
-      from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"]  
-      have px': "x >= - ?Nt x s / ?N c" by (simp add: field_simps) 
-    {assume y: "y > - ?Nt x s / ?N c" 
+      from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"]
+      have px': "x >= - ?Nt x s / ?N c" by (simp add: field_simps)
+    {assume y: "y > - ?Nt x s / ?N c"
       hence "y * ?N c < - ?Nt x s"
         by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
       hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
-    {assume y: "y < -?Nt x s / ?N c" 
+    {assume y: "y < -?Nt x s / ?N c"
       with ly have eu: "l < - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
       with xu px' have "False" by simp  hence ?case by simp }
@@ -1737,7 +1882,7 @@
   ultimately show ?case by blast
 next
     case (3 c s)
-  from "3.prems" 
+  from "3.prems"
   have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
     and px: "Ifm vs (x # bs) (Eq (CNP 0 c s))"
     and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
@@ -1750,12 +1895,12 @@
   {assume c: "?N c > 0" hence cnz: "?N c \<noteq> 0" by simp
     from px eq_divide_eq[of "x" "-?Nt x s" "?N c"]  cnz
     have px': "x = - ?Nt x s / ?N c" by (simp add: field_simps)
-    {assume y: "y < -?Nt x s / ?N c" 
+    {assume y: "y < -?Nt x s / ?N c"
       with ly have eu: "l < - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
       with xu px' have "False" by simp  hence ?case by simp }
     moreover
-    {assume y: "y > -?Nt x s / ?N c" 
+    {assume y: "y > -?Nt x s / ?N c"
       with yu have eu: "u > - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
       with lx px' have "False" by simp  hence ?case by simp }
@@ -1765,12 +1910,12 @@
   {assume c: "?N c < 0" hence cnz: "?N c \<noteq> 0" by simp
     from px eq_divide_eq[of "x" "-?Nt x s" "?N c"]  cnz
     have px': "x = - ?Nt x s / ?N c" by (simp add: field_simps)
-    {assume y: "y < -?Nt x s / ?N c" 
+    {assume y: "y < -?Nt x s / ?N c"
       with ly have eu: "l < - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
       with xu px' have "False" by simp  hence ?case by simp }
     moreover
-    {assume y: "y > -?Nt x s / ?N c" 
+    {assume y: "y > -?Nt x s / ?N c"
       with yu have eu: "u > - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
       with lx px' have "False" by simp  hence ?case by simp }
@@ -1779,7 +1924,7 @@
   ultimately show ?case by blast
 next
     case (4 c s)
-  from "4.prems" 
+  from "4.prems"
   have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
     and px: "Ifm vs (x # bs) (NEq (CNP 0 c s))"
     and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
@@ -1799,8 +1944,8 @@
   assumes lp: "islin p"
   and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))" (is "\<not> (Ifm vs (x#bs) (?M p))")
   and npi: "\<not> (Ifm vs (x#bs) (plusinf p))" (is "\<not> (Ifm vs (x#bs) (?P p))")
-  and ex: "\<exists> x.  Ifm vs (x#bs) p" (is "\<exists> x. ?I x p")
-  shows "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). ?I ((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / 2) p" 
+  and ex: "\<exists>x.  Ifm vs (x#bs) p" (is "\<exists>x. ?I x p")
+  shows "\<exists>(c,t) \<in> set (uset p). \<exists>(d,s) \<in> set (uset p). ?I ((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / 2) p"
 proof-
   let ?Nt = "\<lambda> x t. Itm vs (x#bs) t"
   let ?N = "Ipoly vs"
@@ -1810,16 +1955,16 @@
   have nmi': "\<not> (?I a (?M p))" by simp
   from bound0_I[OF plusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] npi
   have npi': "\<not> (?I a (?P p))" by simp
-  have "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). ?I ((- ?Nt a t/?N c + - ?Nt a s /?N d) / 2) p"
+  have "\<exists>(c,t) \<in> set (uset p). \<exists>(d,s) \<in> set (uset p). ?I ((- ?Nt a t/?N c + - ?Nt a s /?N d) / 2) p"
   proof-
     let ?M = "(\<lambda> (c,t). - ?Nt a t / ?N c) ` ?U"
     have fM: "finite ?M" by auto
-    from minusinf_uset[OF lp nmi pa] plusinf_uset[OF lp npi pa] 
-    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" 
+    from minusinf_uset[OF lp nmi pa] plusinf_uset[OF lp npi pa]
+    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"
       and xs1: "a \<le> - ?Nt x s / ?N d" and tx1: "a \<ge> - ?Nt x t / ?N c" by blast
-    from uset_l[OF lp] ctU dsU tmbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 
+    from uset_l[OF lp] ctU dsU tmbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1
     have xs: "a \<le> - ?Nt a s / ?N d" and tx: "a \<ge> - ?Nt a t / ?N c" by auto
     from ctU have Mne: "?M \<noteq> {}" by auto
     hence Une: "?U \<noteq> {}" by simp
@@ -1828,28 +1973,28 @@
     have linM: "?l \<in> ?M" using fM Mne by simp
     have uinM: "?u \<in> ?M" using fM Mne by simp
     have ctM: "- ?Nt a t / ?N c \<in> ?M" using ctU by auto
-    have dsM: "- ?Nt a s / ?N d \<in> ?M" using dsU by auto 
-    have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
-    have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
+    have dsM: "- ?Nt a s / ?N d \<in> ?M" using dsU by auto
+    have lM: "\<forall>t\<in> ?M. ?l \<le> t" using Mne fM by auto
+    have Mu: "\<forall>t\<in> ?M. t \<le> ?u" using Mne fM by auto
     have "?l \<le> - ?Nt a t / ?N c" using ctM Mne by simp hence lx: "?l \<le> a" using tx by simp
     have "- ?Nt a s / ?N d \<le> ?u" using dsM Mne by simp hence xu: "a \<le> ?u" using xs by simp
     from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
-    have "(\<exists> s\<in> ?M. ?I s p) \<or> 
-      (\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p)" .
+    have "(\<exists>s\<in> ?M. ?I s p) \<or>
+      (\<exists>t1\<in> ?M. \<exists>t2 \<in> ?M. (\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p)" .
     moreover {fix u assume um: "u\<in> ?M" and pu: "?I u p"
-      hence "\<exists> (nu,tu) \<in> ?U. u = - ?Nt a tu / ?N nu" by auto
+      hence "\<exists>(nu,tu) \<in> ?U. u = - ?Nt a tu / ?N nu" by auto
       then obtain "tu" "nu" where tuU: "(nu,tu) \<in> ?U" and tuu:"u= - ?Nt a tu / ?N nu" by blast
       from pu tuu
       have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / 2) p" by simp
       with tuU have ?thesis by blast}
     moreover{
-      assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
-      then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
-        and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
+      assume "\<exists>t1\<in> ?M. \<exists>t2 \<in> ?M. (\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
+      then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M"
+        and noM: "\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
         by blast
-      from t1M have "\<exists> (t1n,t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n" by auto
+      from t1M have "\<exists>(t1n,t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n" by auto
       then obtain "t1u" "t1n" where t1uU: "(t1n,t1u) \<in> ?U" and t1u: "t1 = - ?Nt a t1u / ?N t1n" by blast
-      from t2M have "\<exists> (t2n,t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n" by auto
+      from t2M have "\<exists>(t2n,t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n" by auto
       then obtain "t2u" "t2n" where t2uU: "(t2n,t2u) \<in> ?U" and t2u: "t2 = - ?Nt a t2u / ?N t2n" by blast
       from t1x xt2 have t1t2: "t1 < t2" by simp
       let ?u = "(t1 + t2) / 2"
@@ -1858,10 +2003,10 @@
       with t1uU t2uU t1u t2u have ?thesis by blast}
     ultimately show ?thesis by blast
   qed
-  then obtain "l" "n" "s"  "m" where lnU: "(n,l) \<in> ?U" and smU:"(m,s) \<in> ?U" 
+  then obtain "l" "n" "s"  "m" where lnU: "(n,l) \<in> ?U" and smU:"(m,s) \<in> ?U"
     and pu: "?I ((- ?Nt a l / ?N n + - ?Nt a s / ?N m) / 2) p" by blast
   from lnU smU uset_l[OF lp] have nbl: "tmbound0 l" and nbs: "tmbound0 s" by auto
-  from tmbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] 
+  from tmbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"]
     tmbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
   have "?I ((- ?Nt x l / ?N n + - ?Nt x s / ?N m) / 2) p" by simp
   with lnU smU
@@ -1870,19 +2015,19 @@
 
     (* The Ferrante - Rackoff Theorem *)
 
-theorem fr_eq: 
+theorem fr_eq:
   assumes lp: "islin p"
-  shows "(\<exists> x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists> (n,t) \<in> set (uset p). \<exists> (m,s) \<in> set (uset p). Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) / 2)#bs) p))"
-  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
+  shows "(\<exists>x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists>(n,t) \<in> set (uset p). \<exists>(m,s) \<in> set (uset p). Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) / 2)#bs) p))"
+  (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
 proof
-  assume px: "\<exists> x. ?I x p"
+  assume px: "\<exists>x. ?I x p"
   have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
   moreover {assume "?M \<or> ?P" hence "?D" by blast}
   moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
     from inf_uset[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
   ultimately show "?D" by blast
 next
-  assume "?D" 
+  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" hence "?E" by blast}
@@ -1890,8 +2035,8 @@
 qed
 
 section{* First implementation : Naive by encoding all case splits locally *}
-definition "msubsteq c t d s a r = 
-  evaldjf (split conj) 
+definition "msubsteq c t d s a r =
+  evaldjf (split conj)
   [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
    (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
    (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
@@ -1925,7 +2070,7 @@
   moreover
   {assume c: "?c = 0" and d: "?d=0"
     hence ?thesis  by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex)}
-  moreover 
+  moreover
   {assume c: "?c = 0" and d: "?d\<noteq>0"
     from c have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2*?d)" by simp
     have "?rhs = Ifm vs (-?s / (2*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
@@ -1934,9 +2079,9 @@
       using d mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * (2*?d / (2*?d)) + 2*?d*?r= 0"
       by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left)
-    
-    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r = 0" using d by simp 
-    finally have ?thesis using c d 
+
+    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r = 0" using d by simp
+    finally have ?thesis using c d
       by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex)
   }
   moreover
@@ -1944,36 +2089,36 @@
     from d have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2*?c)" by simp
     have "?rhs = Ifm vs (-?t / (2*?c) # bs) (Eq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (-?t / (2*?c)) + ?r = 0" by (simp add: r[of "- (?t/ (2 * ?c))"])
-    also have "\<dots> \<longleftrightarrow> 2*?c * (?a * (-?t / (2*?c)) + ?r) = 0" 
+    also have "\<dots> \<longleftrightarrow> 2*?c * (?a * (-?t / (2*?c)) + ?r) = 0"
       using c mult_cancel_left[of "2*?c" "(?a * (-?t / (2*?c)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (?a * -?t)* (2*?c) / (2*?c) + 2*?c*?r= 0"
       by (simp add: field_simps distrib_left[of "2*?c"] del: distrib_left)
-    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2*?c*?r = 0" using c by simp 
-    finally have ?thesis using c d 
+    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2*?c*?r = 0" using c by simp
+    finally have ?thesis using c d
       by (simp add: r[of "- (?t/ (2*?c))"] msubsteq_def Let_def evaldjf_ex)
   }
   moreover
   {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *2 \<noteq> 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
-    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)" 
+    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)"
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0"
       by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) =0 "
       using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r =0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r =0"
       using nonzero_mult_divide_cancel_left [OF dc] c d
       by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
-    finally  have ?thesis using c d 
+    finally  have ?thesis using c d
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex field_simps)
   }
   ultimately show ?thesis by blast
 qed
 
 
-definition "msubstneq c t d s a r = 
-  evaldjf (split conj) 
+definition "msubstneq c t d s a r =
+  evaldjf (split conj)
   [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
    (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
    (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
@@ -1982,7 +2127,7 @@
 lemma msubstneq_nb: assumes lp: "islin (NEq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
   shows "bound0 (msubstneq c t d s a r)"
 proof-
-  have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))), 
+  have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
     (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
     (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
     (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]. bound0 (split conj x)"
@@ -2007,18 +2152,18 @@
   moreover
   {assume c: "?c = 0" and d: "?d=0"
     hence ?thesis  by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex)}
-  moreover 
+  moreover
   {assume c: "?c = 0" and d: "?d\<noteq>0"
     from c have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2*?d)" by simp
     have "?rhs = Ifm vs (-?s / (2*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (-?s / (2*?d)) + ?r \<noteq> 0" by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
-    also have "\<dots> \<longleftrightarrow> 2*?d * (?a * (-?s / (2*?d)) + ?r) \<noteq> 0" 
+    also have "\<dots> \<longleftrightarrow> 2*?d * (?a * (-?s / (2*?d)) + ?r) \<noteq> 0"
       using d mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * (2*?d / (2*?d)) + 2*?d*?r\<noteq> 0"
       by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left)
-    
-    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r \<noteq> 0" using d by simp 
-    finally have ?thesis using c d 
+
+    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r \<noteq> 0" using d by simp
+    finally have ?thesis using c d
       by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex)
   }
   moreover
@@ -2026,35 +2171,35 @@
     from d have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2*?c)" by simp
     have "?rhs = Ifm vs (-?t / (2*?c) # bs) (NEq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (-?t / (2*?c)) + ?r \<noteq> 0" by (simp add: r[of "- (?t/ (2 * ?c))"])
-    also have "\<dots> \<longleftrightarrow> 2*?c * (?a * (-?t / (2*?c)) + ?r) \<noteq> 0" 
+    also have "\<dots> \<longleftrightarrow> 2*?c * (?a * (-?t / (2*?c)) + ?r) \<noteq> 0"
       using c mult_cancel_left[of "2*?c" "(?a * (-?t / (2*?c)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (?a * -?t)* (2*?c) / (2*?c) + 2*?c*?r \<noteq> 0"
       by (simp add: field_simps distrib_left[of "2*?c"] del: distrib_left)
-    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2*?c*?r \<noteq> 0" using c by simp 
-    finally have ?thesis using c d 
+    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2*?c*?r \<noteq> 0" using c by simp
+    finally have ?thesis using c d
       by (simp add: r[of "- (?t/ (2*?c))"] msubstneq_def Let_def evaldjf_ex)
   }
   moreover
   {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *2 \<noteq> 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
-    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)" 
+    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)"
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<noteq> 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<noteq> 0"
       by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<noteq> 0 "
       using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<noteq> 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<noteq> 0"
       using nonzero_mult_divide_cancel_left[OF dc] c d
       by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
-    finally  have ?thesis using c d 
+    finally  have ?thesis using c d
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex field_simps)
   }
   ultimately show ?thesis by blast
 qed
 
-definition "msubstlt c t d s a r = 
-  evaldjf (split conj) 
+definition "msubstlt c t d s a r =
+  evaldjf (split conj)
   [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
   (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
    (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
@@ -2078,8 +2223,8 @@
 qed
 
 
-lemma msubstlt: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Lt (CNP 0 a r))" 
-  shows "Ifm vs (x#bs) (msubstlt c t d s a r) \<longleftrightarrow> 
+lemma msubstlt: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Lt (CNP 0 a r))"
+  shows "Ifm vs (x#bs) (msubstlt c t d s a r) \<longleftrightarrow>
   Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) (Lt (CNP 0 a r))" (is "?lhs = ?rhs")
 proof-
   let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
@@ -2097,49 +2242,49 @@
   {assume c: "?c=0" and d: "?d=0"
     hence ?thesis  using nc nd by (simp add: polyneg_norm lt r[of 0] msubstlt_def Let_def evaldjf_ex)}
   moreover
-  {assume dc: "?c*?d > 0" 
+  {assume dc: "?c*?d > 0"
     from dc have dc': "2*?c *?d > 0" by simp
     hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
     from dc' have dc'': "\<not> 2*?c *?d < 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
-    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)" 
+    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)"
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0"
       by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) < 0"
-      
+
       using dc' dc'' mult_less_cancel_left_disj[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r < 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r < 0"
       using nonzero_mult_divide_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'
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
   }
   moreover
-  {assume dc: "?c*?d < 0" 
+  {assume dc: "?c*?d < 0"
 
     from dc have dc': "2*?c *?d < 0"
-      by (simp add: mult_less_0_iff field_simps) 
+      by (simp add: mult_less_0_iff field_simps)
     hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
-    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)" 
+    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)"
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0"
       by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
 
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) > 0"
-      
+
       using dc' order_less_not_sym[OF dc'] mult_less_cancel_left_disj[of "2 * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r"] by simp
-    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2*?c*?d*?r < 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2*?c*?d*?r < 0"
       using nonzero_mult_divide_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
-      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
   }
   moreover
-  {assume c: "?c > 0" and d: "?d=0"  
+  {assume c: "?c > 0" and d: "?d=0"
     from c have c'': "2*?c > 0" by (simp add: zero_less_mult_iff)
     from c have c': "2*?c \<noteq> 0" by simp
     from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)"  by (simp add: field_simps)
@@ -2147,10 +2292,10 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / (2*?c))+ ?r < 0" by (simp add: r[of "- (?t / (2*?c))"])
     also have "\<dots> \<longleftrightarrow> 2*?c * (?a* (- ?t / (2*?c))+ ?r) < 0"
       using c mult_less_cancel_left_disj[of "2 * ?c" "?a* (- ?t / (2*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
-    also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r < 0" 
+    also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r < 0"
       using nonzero_mult_divide_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 have ?thesis using c d nc nd
       by(simp add: r[of "- (?t / (2*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
   }
   moreover
@@ -2161,14 +2306,14 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / (2*?c))+ ?r < 0" by (simp add: r[of "- (?t / (2*?c))"])
     also have "\<dots> \<longleftrightarrow> 2*?c * (?a* (- ?t / (2*?c))+ ?r) > 0"
       using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"] by simp
-    also have "\<dots> \<longleftrightarrow> ?a*?t -  2*?c *?r < 0" 
+    also have "\<dots> \<longleftrightarrow> ?a*?t -  2*?c *?r < 0"
       using nonzero_mult_divide_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 have ?thesis using c d nc nd
       by(simp add: r[of "- (?t / (2*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
   }
   moreover
-  {assume c: "?c = 0" and d: "?d>0"  
+  {assume c: "?c = 0" and d: "?d>0"
     from d have d'': "2*?d > 0" by (simp add: zero_less_mult_iff)
     from d have d': "2*?d \<noteq> 0" by simp
     from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)"  by (simp add: field_simps)
@@ -2176,10 +2321,10 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / (2*?d))+ ?r < 0" by (simp add: r[of "- (?s / (2*?d))"])
     also have "\<dots> \<longleftrightarrow> 2*?d * (?a* (- ?s / (2*?d))+ ?r) < 0"
       using d mult_less_cancel_left_disj[of "2 * ?d" "?a* (- ?s / (2*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
-    also have "\<dots> \<longleftrightarrow> - ?a*?s+  2*?d *?r < 0" 
+    also have "\<dots> \<longleftrightarrow> - ?a*?s+  2*?d *?r < 0"
       using nonzero_mult_divide_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 have ?thesis using c d nc nd
       by(simp add: r[of "- (?s / (2*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
   }
   moreover
@@ -2190,17 +2335,17 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / (2*?d))+ ?r < 0" by (simp add: r[of "- (?s / (2*?d))"])
     also have "\<dots> \<longleftrightarrow> 2*?d * (?a* (- ?s / (2*?d))+ ?r) > 0"
       using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"] by simp
-    also have "\<dots> \<longleftrightarrow> ?a*?s -  2*?d *?r < 0" 
+    also have "\<dots> \<longleftrightarrow> ?a*?s -  2*?d *?r < 0"
       using nonzero_mult_divide_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 
+    finally have ?thesis using c d nc nd
       by(simp add: r[of "- (?s / (2*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
   }
 ultimately show ?thesis by blast
 qed
 
-definition "msubstle c t d s a r = 
-  evaldjf (split conj) 
+definition "msubstle c t d s a r =
+  evaldjf (split conj)
   [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
   (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
    (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
@@ -2223,8 +2368,8 @@
   from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstle_def)
 qed
 
-lemma msubstle: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Le (CNP 0 a r))" 
-  shows "Ifm vs (x#bs) (msubstle c t d s a r) \<longleftrightarrow> 
+lemma msubstle: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Le (CNP 0 a r))"
+  shows "Ifm vs (x#bs) (msubstle c t d s a r) \<longleftrightarrow>
   Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) (Le (CNP 0 a r))" (is "?lhs = ?rhs")
 proof-
   let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
@@ -2242,49 +2387,49 @@
   {assume c: "?c=0" and d: "?d=0"
     hence ?thesis  using nc nd by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)}
   moreover
-  {assume dc: "?c*?d > 0" 
+  {assume dc: "?c*?d > 0"
     from dc have dc': "2*?c *?d > 0" by simp
     hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
     from dc' have dc'': "\<not> 2*?c *?d < 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
-    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)" 
+    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)"
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0"
       by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) <= 0"
-      
+
       using dc' dc'' mult_le_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r <= 0"
       using nonzero_mult_divide_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'
-      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) 
+      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" 
+  {assume dc: "?c*?d < 0"
 
     from dc have dc': "2*?c *?d < 0"
       by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos)
     hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
-    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)" 
+    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2*?c*?d)"
       by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
-    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0"
       by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
 
     also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) >= 0"
-      
+
       using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "2 * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r"] by simp
-    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2*?c*?d*?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2*?c*?d*?r <= 0"
       using nonzero_mult_divide_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
-      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) 
+      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"  
+  {assume c: "?c > 0" and d: "?d=0"
     from c have c'': "2*?c > 0" by (simp add: zero_less_mult_iff)
     from c have c': "2*?c \<noteq> 0" by simp
     from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)"  by (simp add: field_simps)
@@ -2292,10 +2437,10 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / (2*?c))+ ?r <= 0" by (simp add: r[of "- (?t / (2*?c))"])
     also have "\<dots> \<longleftrightarrow> 2*?c * (?a* (- ?t / (2*?c))+ ?r) <= 0"
       using c mult_le_cancel_left[of "2 * ?c" "?a* (- ?t / (2*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
-    also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r <= 0"
       using nonzero_mult_divide_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 have ?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
@@ -2306,14 +2451,14 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / (2*?c))+ ?r <= 0" by (simp add: r[of "- (?t / (2*?c))"])
     also have "\<dots> \<longleftrightarrow> 2*?c * (?a* (- ?t / (2*?c))+ ?r) >= 0"
       using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"] by simp
-    also have "\<dots> \<longleftrightarrow> ?a*?t -  2*?c *?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> ?a*?t -  2*?c *?r <= 0"
       using nonzero_mult_divide_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 have ?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"  
+  {assume c: "?c = 0" and d: "?d>0"
     from d have d'': "2*?d > 0" by (simp add: zero_less_mult_iff)
     from d have d': "2*?d \<noteq> 0" by simp
     from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)"  by (simp add: field_simps)
@@ -2321,10 +2466,10 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / (2*?d))+ ?r <= 0" by (simp add: r[of "- (?s / (2*?d))"])
     also have "\<dots> \<longleftrightarrow> 2*?d * (?a* (- ?s / (2*?d))+ ?r) <= 0"
       using d mult_le_cancel_left[of "2 * ?d" "?a* (- ?s / (2*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
-    also have "\<dots> \<longleftrightarrow> - ?a*?s+  2*?d *?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> - ?a*?s+  2*?d *?r <= 0"
       using nonzero_mult_divide_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 have ?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
@@ -2335,10 +2480,10 @@
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / (2*?d))+ ?r <= 0" by (simp add: r[of "- (?s / (2*?d))"])
     also have "\<dots> \<longleftrightarrow> 2*?d * (?a* (- ?s / (2*?d))+ ?r) >= 0"
       using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"] by simp
-    also have "\<dots> \<longleftrightarrow> ?a*?s -  2*?d *?r <= 0" 
+    also have "\<dots> \<longleftrightarrow> ?a*?s -  2*?d *?r <= 0"
       using nonzero_mult_divide_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 
+    finally have ?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
@@ -2367,38 +2512,38 @@
   using lp t s
   by (induct p rule: islin.induct, auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
 
-lemma fr_eq_msubst: 
+lemma fr_eq_msubst:
   assumes lp: "islin p"
-  shows "(\<exists> x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs (x#bs) (msubst p ((c,t),(d,s)))))"
-  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
+  shows "(\<exists>x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists>(c,t) \<in> set (uset p). \<exists>(d,s) \<in> set (uset p). 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" by blast
-{fix c t d s assume ctU: "(c,t) \<in>set (uset p)" and dsU: "(d,s) \<in>set (uset p)" 
+{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" 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))" ..}
 moreover
-{fix c t d s assume ctU: "(c,t) \<in>set (uset p)" and dsU: "(d,s) \<in>set (uset p)" 
+{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 (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" 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). 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
+ultimately have th': "(\<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 .
-qed 
+qed
 
 lemma simpfm_lin:   assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
   shows "qfree p \<Longrightarrow> islin (simpfm p)"
   by (induct p rule: simpfm.induct, auto simp add: conj_lin disj_lin)
 
-definition 
+definition
   "ferrack p \<equiv> let q = simpfm p ; mp = minusinf q ; pp = plusinf q
-  in if (mp = T \<or> pp = T) then T 
+  in if (mp = T \<or> pp = T) then T
      else (let U = alluopairs (remdups (uset  q))
            in decr0 (disj mp (disj pp (evaldjf (simpfm o (msubst q)) U ))))"
 
-lemma ferrack: 
+lemma ferrack:
   assumes qf: "qfree p"
   shows "qfree (ferrack p) \<and> ((Ifm vs bs (ferrack p)) = (Ifm vs bs (E p)))"
   (is "_ \<and> (?rhs = ?lhs)")
@@ -2406,7 +2551,7 @@
   let ?I = "\<lambda> x p. Ifm vs (x#bs) p"
   let ?N = "\<lambda> t. Ipoly vs t"
   let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
-  let ?q = "simpfm p" 
+  let ?q = "simpfm p"
   let ?U = "remdups(uset ?q)"
   let ?Up = "alluopairs ?U"
   let ?mp = "minusinf ?q"
@@ -2422,22 +2567,22 @@
     from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t]
     have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: field_simps)}
   hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (msubst ?q (x, y)) \<longleftrightarrow> ?I (msubst ?q (y, x))" by clarsimp
-  {fix x assume xUp: "x \<in> set ?Up" 
-    then  obtain c t d s where ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U" 
-      and x: "x = ((c,t),(d,s))" using alluopairs_set1[of ?U] by auto  
-    from U_l[rule_format, OF ctU] U_l[rule_format, OF dsU] 
+  {fix x assume xUp: "x \<in> set ?Up"
+    then  obtain c t d s where ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
+      and x: "x = ((c,t),(d,s))" using alluopairs_set1[of ?U] by auto
+    from U_l[rule_format, OF ctU] U_l[rule_format, OF dsU]
     have nbs: "tmbound0 t" "tmbound0 s" by simp_all
-    from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]] 
+    from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]]
     have "bound0 ((simpfm o (msubst (simpfm p))) x)" using x by simp}
   with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
   have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
-  with mp_nb pp_nb 
+  with mp_nb pp_nb
   have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by simp
   from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def)
   have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
   also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp
-  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists> (x,y) \<in> set ?Up. ?I ((simpfm o (msubst ?q)) (x,y)))" using alluopairs_bex[OF th0] by simp
-  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (evaldjf (simpfm o (msubst ?q)) ?Up)" 
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(x,y) \<in> set ?Up. ?I ((simpfm o (msubst ?q)) (x,y)))" using alluopairs_bex[OF th0] by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (evaldjf (simpfm o (msubst ?q)) ?Up)"
     by (simp add: evaldjf_ex)
   also have "\<dots> \<longleftrightarrow> ?I (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up)))" by simp
   also have "\<dots> \<longleftrightarrow> ?rhs" using decr0[OF th1, of vs x bs]
@@ -2457,18 +2602,18 @@
 section{* Second implemenation: Case splits not local *}
 
 lemma fr_eq2:  assumes lp: "islin p"
-  shows "(\<exists> x. Ifm vs (x#bs) p) \<longleftrightarrow> 
-   ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> 
-    (Ifm vs (0#bs) p) \<or> 
-    (\<exists> (n,t) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ifm vs ((- Itm vs (x#bs) t /  (Ipoly vs n * 2))#bs) p) \<or> 
-    (\<exists> (n,t) \<in> set (uset p). \<exists> (m,s) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ipoly vs m \<noteq> 0 \<and> Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /2)#bs) p))"
-  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?Z \<or> ?U \<or> ?F)" is "?E = ?D")
+  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
+   ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or>
+    (Ifm vs (0#bs) p) \<or>
+    (\<exists>(n,t) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ifm vs ((- Itm vs (x#bs) t /  (Ipoly vs n * 2))#bs) p) \<or>
+    (\<exists>(n,t) \<in> set (uset p). \<exists>(m,s) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ipoly vs m \<noteq> 0 \<and> Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /2)#bs) p))"
+  (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"
+  assume px: "\<exists>x. ?I x p"
   have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
   moreover {assume "?M \<or> ?P" hence "?D" by blast}
   moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
-    from inf_uset[OF lp nmi npi, OF px] 
+    from inf_uset[OF lp nmi npi, OF px]
     obtain c t d s where ct: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" "?I ((- 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>) / ((1\<Colon>'a) + (1\<Colon>'a))) p"
       by auto
     let ?c = "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
@@ -2494,7 +2639,7 @@
     ultimately have ?D by auto}
   ultimately show "?D" by blast
 next
-  assume "?D" 
+  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" hence "?E" by blast}
@@ -2507,31 +2652,31 @@
 definition "msubstltneg c t a b = Lt (Neg (Add (Mul a t) (Mul c b)))"
 definition "msubstleneg c t a b = Le (Neg (Add (Mul a t) (Mul c b)))"
 
-lemma msubsteq2: 
+lemma msubsteq2:
   assumes nz: "Ipoly vs c \<noteq> 0" and l: "islin (Eq (CNP 0 a b))"
   shows "Ifm vs (x#bs) (msubsteq2 c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Eq (CNP 0 a b))" (is "?lhs = ?rhs")
   using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
   by (simp add: msubsteq2_def field_simps)
 
-lemma msubstltpos: 
+lemma msubstltpos:
   assumes nz: "Ipoly vs c > 0" and l: "islin (Lt (CNP 0 a b))"
   shows "Ifm vs (x#bs) (msubstltpos c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs")
   using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
   by (simp add: msubstltpos_def field_simps)
 
-lemma msubstlepos: 
+lemma msubstlepos:
   assumes nz: "Ipoly vs c > 0" and l: "islin (Le (CNP 0 a b))"
   shows "Ifm vs (x#bs) (msubstlepos c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs")
   using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
   by (simp add: msubstlepos_def field_simps)
 
-lemma msubstltneg: 
+lemma msubstltneg:
   assumes nz: "Ipoly vs c < 0" and l: "islin (Lt (CNP 0 a b))"
   shows "Ifm vs (x#bs) (msubstltneg c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs")
   using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
   by (simp add: msubstltneg_def field_simps del: minus_add_distrib)
 
-lemma msubstleneg: 
+lemma msubstleneg:
   assumes nz: "Ipoly vs c < 0" and l: "islin (Le (CNP 0 a b))"
   shows "Ifm vs (x#bs) (msubstleneg c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs")
   using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
@@ -2545,8 +2690,8 @@
 | "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: 
+
+lemma msubstpos_I:
   assumes lp: "islin p" and pos: "Ipoly vs c > 0"
   shows "Ifm vs (x#bs) (msubstpos p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
   using lp pos
@@ -2561,7 +2706,7 @@
 | "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
 | "msubstneg p c t = p"
 
-lemma msubstneg_I: 
+lemma msubstneg_I:
   assumes lp: "islin p" and pos: "Ipoly vs c < 0"
   shows "Ifm vs (x#bs) (msubstneg p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
   using lp pos
@@ -2574,7 +2719,7 @@
   shows "Ifm vs (x#bs) (msubst2 p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
 proof-
   let ?c = "Ipoly vs c"
-  from nc have anc: "allpolys isnpoly (CP c)" "allpolys isnpoly (CP (~\<^sub>p 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
@@ -2625,16 +2770,16 @@
 
 lemma islin_qf: "islin p \<Longrightarrow> qfree p"
   by (induct p rule: islin.induct, auto simp add: bound0_qf)
-lemma fr_eq_msubst2: 
+lemma fr_eq_msubst2:
   assumes lp: "islin p"
-  shows "(\<exists> x. Ifm vs (x#bs) p) \<longleftrightarrow> ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \<or> (\<exists>(n, t)\<in>set (uset p). Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<or> (\<exists> (c,t) \<in> set (uset p). \<exists> (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)))))"
-  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?Pz \<or> ?PU \<or> ?F)" is "?E = ?D")
+  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow> ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \<or> (\<exists>(n, t)\<in>set (uset p). Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<or> (\<exists>(c,t) \<in> set (uset p). \<exists>(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)))))"
+  (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" by blast
   let ?I = "\<lambda>p. Ifm vs (x#bs) p"
   have n2: "isnpoly (C (-2,1))" by (simp add: isnpoly_def)
   note eq0 = subst0[OF islin_qf[OF lp], of vs x bs "CP 0\<^sub>p", simplified]
-  
+
   have eq1: "(\<exists>(n, t)\<in>set (uset p). ?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p)"
   proof-
     {fix n t assume H: "(n, t)\<in>set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)"
@@ -2642,7 +2787,7 @@
       hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" by (simp_all add: polymul_norm n2)
       have nn': "allpolys isnpoly (CP (~\<^sub>p (n *\<^sub>p C (-2, 1))))"
         by (simp add: polyneg_norm nn)
-      hence nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" using H(2) nn' nn 
+      hence nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" using H(2) nn' nn
         by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
       from msubst2[OF lp nn nn2(1), of x bs t]
       have "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p"
@@ -2655,13 +2800,13 @@
       from msubst2[OF lp nn, of x bs t] have "?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)" using H(2,3) by (simp add: mult_minus2_right)}
     ultimately show ?thesis by blast
   qed
-  have eq2: "(\<exists> (c,t) \<in> set (uset p). \<exists> (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)))) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p).
-     \<exists>(m, s)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p)" 
+  have eq2: "(\<exists>(c,t) \<in> set (uset p). \<exists>(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)))) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p).
+     \<exists>(m, s)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p)"
   proof-
-    {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" 
+    {fix c t d s 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" by blast+
-      hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" 
+      hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)"
         by (simp_all add: polymul_norm n2)
       have stupid: "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))" "allpolys isnpoly (CP ((C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
         by (simp_all add: polyneg_norm nn)
@@ -2672,26 +2817,26 @@
         by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult_commute)
     }
     moreover
-    {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" 
+    {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
       "\<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" by blast+
       hence 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)
-      from msubst2[OF lp nn, of x bs ] H(3,4,5) 
+      from msubst2[OF lp nn, of x bs ] H(3,4,5)
       have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
         by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left mult_commute)
     }
     ultimately show ?thesis by blast
   qed
   from fr_eq2[OF lp, of vs bs x] show ?thesis
-    unfolding eq0 eq1 eq2 by blast  
+    unfolding eq0 eq1 eq2 by blast
 qed
 
-definition 
+definition
 "ferrack2 p \<equiv> let q = simpfm p ; mp = minusinf q ; pp = plusinf q
- in if (mp = T \<or> pp = T) then T 
+ 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, 
+    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)"
@@ -2703,7 +2848,7 @@
   let ?J = "\<lambda> x p. Ifm vs (x#bs) p"
   let ?N = "\<lambda> t. Ipoly vs t"
   let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
-  let ?q = "simpfm p" 
+  let ?q = "simpfm p"
   let ?qz = "subst0 (CP 0\<^sub>p) ?q"
   let ?U = "remdups(uset ?q)"
   let ?Up = "alluopairs ?U"
@@ -2715,7 +2860,7 @@
   from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" .
   from uset_l[OF lq] have U_l: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
     by simp
-  have bnd0: "\<forall>x \<in> set ?U. bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) x)" 
+  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"
       hence tnb: "tmbound0 t" using U_l by blast
@@ -2723,16 +2868,16 @@
       have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" by simp}
     thus ?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)" 
+  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] 
+      from badc U_l alluopairs_set1[of ?U]
       have nb: "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)))" by simp}
     thus ?thesis by auto
   qed
   have stupid: "bound0 F" by simp
-  let ?R = "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, 
+  let ?R = "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)]"
   from subst0_nb[of "CP 0\<^sub>p" ?q] q_qf evaldjf_bound0[OF bnd1] evaldjf_bound0[OF bnd0] mp_nb pp_nb stupid
   have nb: "bound0 ?R "
@@ -2740,7 +2885,7 @@
   let ?s = "\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))"
 
   {fix b a d c assume baU: "(b,a) \<in> set ?U" and dcU: "(d,c) \<in> set ?U"
-    from U_l baU dcU have norm: "isnpoly b" "isnpoly d" "isnpoly (C (-2, 1))" 
+    from U_l baU dcU have norm: "isnpoly b" "isnpoly d" "isnpoly (C (-2, 1))"
       by auto (simp add: isnpoly_def)
     have norm2: "isnpoly (C (-2, 1) *\<^sub>p b*\<^sub>p d)" "isnpoly (C (-2, 1) *\<^sub>p d*\<^sub>p b)"
       using norm by (simp_all add: polymul_norm)
@@ -2749,17 +2894,17 @@
     have "?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) = ?I (msubst2 ?q (C (-2, 1) *\<^sub>p d*\<^sub>p b) (Add (Mul b c) (Mul d a)))" (is "?lhs \<longleftrightarrow> ?rhs")
     proof
       assume H: ?lhs
-      hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" 
+      hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
-      from msubst2[OF lq norm2(1) z(1), of x bs] 
-        msubst2[OF lq norm2(2) z(2), of x bs] H 
+      from msubst2[OF lq norm2(1) z(1), of x bs]
+        msubst2[OF lq norm2(2) z(2), of x bs] H
       show ?rhs by (simp add: field_simps)
     next
       assume H: ?rhs
-      hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" 
+      hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
         by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
-      from msubst2[OF lq norm2(1) z(1), of x bs] 
-        msubst2[OF lq norm2(2) z(2), of x bs] H 
+      from msubst2[OF lq norm2(1) z(1), of x bs]
+        msubst2[OF lq norm2(2) z(2), of x bs] H
       show ?lhs by (simp add: field_simps)
     qed}
   hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
@@ -2768,28 +2913,30 @@
   have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
   also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists>(b, a)\<in>set ?U. \<exists>(d, c)\<in>set ?U. ?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))))"
     using fr_eq_msubst2[OF lq, of vs bs x] by simp
-  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists> x\<in>set ?U. \<exists> y \<in>set ?U. ?I (?s (x,y)))"
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists>x\<in>set ?U. \<exists>y \<in>set ?U. ?I (?s (x,y)))"
     by (simp add: split_def)
-  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists> (x,y) \<in> set ?Up. ?I (?s (x,y)))"
-    using alluopairs_bex[OF th0] by simp 
-  also have "\<dots> \<longleftrightarrow> ?I ?R" 
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists>(x,y) \<in> set ?Up. ?I (?s (x,y)))"
+    using alluopairs_bex[OF th0] by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?R"
     by (simp add: list_disj_def evaldjf_ex split_def)
   also have "\<dots> \<longleftrightarrow> ?rhs"
     unfolding ferrack2_def
-    apply (cases "?mp = T") 
+    apply (cases "?mp = T")
     apply (simp add: list_disj_def)
-    apply (cases "?pp = T") 
+    apply (cases "?pp = T")
     apply (simp add: list_disj_def)
     by (simp_all add: Let_def decr0[OF nb])
-  finally show ?thesis using decr0_qf[OF nb]  
+  finally show ?thesis using decr0_qf[OF nb]
     by (simp  add: ferrack2_def Let_def)
 qed
 
 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)" by blast
-  from qelim[OF th, of "prep p" bs] 
-show ?thesis  unfolding frpar2_def by (auto simp add: prep)
+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)"
+    by blast
+  from qelim[OF th, of "prep p" bs]
+  show ?thesis unfolding frpar2_def by (auto simp add: prep)
 qed
 
 oracle frpar_oracle = {*
@@ -2829,7 +2976,7 @@
   | tm_of_term fs ps (Const(@{const_name Groups.plus}, _) $ a $ b) = @{code Add} (tm_of_term fs ps a, tm_of_term fs ps b)
   | tm_of_term fs ps (Const(@{const_name Groups.minus}, _) $ a $ b) = @{code Sub} (tm_of_term fs ps a, tm_of_term fs ps b)
   | tm_of_term fs ps (Const(@{const_name Groups.times}, _) $ a $ b) = @{code Mul} (num_of_term ps a, tm_of_term fs ps b)
-  | tm_of_term fs ps t = (@{code CP} (num_of_term ps t) 
+  | tm_of_term fs ps t = (@{code CP} (num_of_term ps t)
       handle TERM _ => mk_Bound (the_index fs t)
            | General.Subscript => mk_Bound (the_index fs t));
 
@@ -2856,7 +3003,7 @@
       in @{code A} (fm_of_term (Free (xn', xT) :: fs) ps p') end
   | fm_of_term fs ps _ = error "fm_of_term";
 
-fun term_of_num T ps (@{code poly.C} (a, b)) = 
+fun term_of_num T ps (@{code poly.C} (a, b)) =
       let
         val (c, d) = pairself (@{code integer_of_int}) (a, b)
       in
@@ -2897,7 +3044,7 @@
   | term_of_fm T fs ps _ = error "term_of_fm: quantifiers";
 
 fun frpar_procedure alternative T ps fm =
-  let 
+  let
     val frpar = if alternative then @{code frpar2} else @{code frpar};
     val fs = subtract (op aconv) (map Free (Term.add_frees fm [])) ps;
     val eval = term_of_fm T fs ps o frpar o fm_of_term fs ps;
@@ -2908,18 +3055,18 @@
 
 in
 
-  fn (ctxt, alternative, ty, ps, ct) => 
+  fn (ctxt, alternative, ty, ps, ct) =>
     cterm_of (Proof_Context.theory_of ctxt)
       (frpar_procedure alternative ty ps (term_of ct))
 
 end
 *}
 
-ML {* 
-structure Parametric_Ferrante_Rackoff = 
+ML {*
+structure Parametric_Ferrante_Rackoff =
 struct
 
-fun tactic ctxt alternative T ps = 
+fun tactic ctxt alternative T ps =
   Object_Logic.full_atomize_tac ctxt
   THEN' CSUBGOAL (fn (g, i) =>
     let
@@ -2955,14 +3102,16 @@
   apply (simp add: field_simps)
   apply (rule spec[where x=y])
   apply (frpar type: "'a::{linordered_field_inverse_zero}" pars: "z::'a::{linordered_field_inverse_zero}")
-  by simp
+  apply simp
+  done
 
 lemma "\<exists>(x::'a::{linordered_field_inverse_zero}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
   apply (frpar2 type: "'a::{linordered_field_inverse_zero}" pars: "y::'a::{linordered_field_inverse_zero}")
   apply (simp add: field_simps)
   apply (rule spec[where x=y])
   apply (frpar2 type: "'a::{linordered_field_inverse_zero}" pars: "z::'a::{linordered_field_inverse_zero}")
-  by simp
+  apply simp
+  done
 
 text{* Collins/Jones Problem *}
 (*