misc tuning;
authorwenzelm
Fri, 21 Sep 2012 20:54:48 +0200
changeset 49519 4d2c93e1d9c8
parent 49518 b377da40244b
child 49520 506f2a634473
misc tuning;
src/HOL/Decision_Procs/DP_Library.thy
src/HOL/Decision_Procs/Dense_Linear_Order.thy
--- a/src/HOL/Decision_Procs/DP_Library.thy	Fri Sep 21 19:17:49 2012 +0200
+++ b/src/HOL/Decision_Procs/DP_Library.thy	Fri Sep 21 20:54:48 2012 +0200
@@ -2,35 +2,38 @@
 imports Main
 begin
 
-primrec alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list" where
+primrec alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list"
+where
   "alluopairs [] = []"
 | "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)"
 
 lemma alluopairs_set1: "set (alluopairs xs) \<le> {(x,y). x\<in> set xs \<and> y\<in> set xs}"
-by (induct xs, auto)
+  by (induct xs) auto
 
 lemma alluopairs_set:
   "\<lbrakk>x\<in> set xs ; y \<in> set xs\<rbrakk> \<Longrightarrow> (x,y) \<in> set (alluopairs xs) \<or> (y,x) \<in> set (alluopairs xs) "
-by (induct xs, auto)
+  by (induct xs) auto
 
 lemma alluopairs_bex:
   assumes Pc: "\<forall> x \<in> set xs. \<forall>y\<in> set xs. P x y = P y x"
   shows "(\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)"
 proof
   assume "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y"
-  then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y"  by blast
+  then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y"
+    by blast
   from alluopairs_set[OF x y] P Pc x y show"\<exists>(x, y)\<in>set (alluopairs xs). P x y" 
     by auto
 next
   assume "\<exists>(x, y)\<in>set (alluopairs xs). P x y"
-  then obtain "x" and "y"  where xy:"(x,y) \<in> set (alluopairs xs)" and P: "P x y" by blast+
+  then obtain "x" and "y" where xy: "(x,y) \<in> set (alluopairs xs)" and P: "P x y"
+    by blast+
   from xy have "x \<in> set xs \<and> y\<in> set xs" using alluopairs_set1 by blast
   with P show "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" by blast
 qed
 
 lemma alluopairs_ex:
-  "\<forall> x y. P x y = P y x \<Longrightarrow>
-  (\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)"
-by(blast intro!: alluopairs_bex)
+  "\<forall>x y. P x y = P y x \<Longrightarrow>
+    (\<exists>x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists>(x,y) \<in> set (alluopairs xs). P x y)"
+  by (blast intro!: alluopairs_bex)
 
 end
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Fri Sep 21 19:17:49 2012 +0200
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy	Fri Sep 21 20:54:48 2012 +0200
@@ -17,17 +17,21 @@
 context linorder
 begin
 
-lemma less_not_permute[no_atp]: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
+lemma less_not_permute[no_atp]: "\<not> (x < y \<and> y < x)"
+  by (simp add: not_less linear)
 
 lemma gather_simps[no_atp]: 
-  shows 
-  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
-  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
-  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
-  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"  by auto
+  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow>
+    (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
+  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow>
+    (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
+  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow>
+    (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
+  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow>
+    (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"
+  by auto
 
-lemma 
-  gather_start[no_atp]: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
+lemma gather_start [no_atp]: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
   by simp
 
 text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
@@ -43,8 +47,8 @@
 lemma minf_P[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
 
 text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
-lemma pinf_gt[no_atp]:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
-lemma pinf_lt[no_atp]: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
+lemma pinf_gt[no_atp]:  "\<exists>z. \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
+lemma pinf_lt[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
   by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
 
 lemma pinf_ge[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
@@ -81,60 +85,84 @@
 lemma  npi_disj[no_atp]: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
   \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
 
-lemma lin_dense_lt[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
+lemma lin_dense_lt[no_atp]:
+  "t \<in> U \<Longrightarrow>
+    \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
 proof(clarsimp)
-  fix x l u y  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
+  fix x l u y
+  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
     and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
   from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "t < y"
+  { assume H: "t < y"
     from less_trans[OF lx px] less_trans[OF H yu]
-    have "l < t \<and> t < u"  by simp
-    with tU noU have "False" by auto}
-  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
-  thus "y < t" using tny by (simp add: less_le)
+    have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto }
+  then have "\<not> t < y" by auto
+  then have "y \<le> t" by (simp add: not_less)
+  then show "y < t" using tny by (simp add: less_le)
 qed
 
-lemma lin_dense_gt[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
+lemma lin_dense_gt[no_atp]:
+  "t \<in> U \<Longrightarrow>
+    \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
+proof(clarsimp)
+  fix x l u y
+  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
+    and px: "t < x" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  { assume H: "y< t"
+    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto }
+  then have "\<not> y<t" by auto
+  then have "t \<le> y" by (auto simp add: not_less)
+  then show "t < y" using tny by (simp add: less_le)
+qed
+
+lemma lin_dense_le[no_atp]:
+  "t \<in> U \<Longrightarrow>
+    \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
 proof(clarsimp)
   fix x l u y
   assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
-  and px: "t < x" and ly: "l<y" and yu:"y < u"
+    and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
   from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "y< t"
-    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
-  thus "t < y" using tny by (simp add:less_le)
+  { assume H: "t < y"
+    from less_le_trans[OF lx px] less_trans[OF H yu]
+    have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto }
+  then have "\<not> t < y" by auto
+  then show "y \<le> t" by (simp add: not_less)
 qed
 
-lemma lin_dense_le[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
+lemma lin_dense_ge[no_atp]:
+  "t \<in> U \<Longrightarrow>
+    \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
 proof(clarsimp)
   fix x l u y
   assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
-  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
+    and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
   from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "t < y"
-    from less_le_trans[OF lx px] less_trans[OF H yu]
+  { assume H: "y< t"
+    from less_trans[OF ly H] le_less_trans[OF px xu]
     have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
+    with tU noU have "False" by auto }
+  then have "\<not> y<t" by auto
+  then show "t \<le> y" by (simp add: not_less)
 qed
 
-lemma lin_dense_ge[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
-proof(clarsimp)
-  fix x l u y
-  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
-  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "y< t"
-    from less_trans[OF ly H] le_less_trans[OF px xu]
-    have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
-qed
-lemma lin_dense_eq[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)"  by auto
-lemma lin_dense_neq[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)"  by auto
-lemma lin_dense_P[no_atp]: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)"  by auto
+lemma lin_dense_eq[no_atp]:
+  "t \<in> U \<Longrightarrow>
+    \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)"
+  by auto
+
+lemma lin_dense_neq[no_atp]:
+  "t \<in> U \<Longrightarrow>
+    \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)"
+  by auto
+
+lemma lin_dense_P[no_atp]:
+  "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)"
+  by auto
 
 lemma lin_dense_conj[no_atp]:
   "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
@@ -155,13 +183,13 @@
 
 lemma npmibnd[no_atp]: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
   \<Longrightarrow> \<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<le> x \<and> x \<le> u')"
-by auto
+  by auto
 
 lemma finite_set_intervals[no_atp]:
   assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
-  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
+    and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
   shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
-proof-
+proof -
   let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
   let ?xM = "{y. y\<in> S \<and> x \<le> y}"
   let ?a = "Max ?Mx"
@@ -180,37 +208,45 @@
   have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
   have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
   proof(clarsimp)
-    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
+    fix y
+    assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
     from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
-    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
-    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
-    ultimately show "False" by blast
+    moreover {
+      assume "y \<in> ?Mx"
+      hence "y \<le> ?a" using Mxne fMx by auto
+      with ay have "False" by (simp add: not_le[symmetric])
+    }
+    moreover {
+      assume "y \<in> ?xM"
+      hence "?b \<le> y" using xMne fxM by auto
+      with yb have "False" by (simp add: not_le[symmetric])
+    }
+    ultimately show False by blast
   qed
   from ainS binS noy ax xb px show ?thesis by blast
 qed
 
 lemma finite_set_intervals2[no_atp]:
   assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
-  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
+    and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
   shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
 proof-
   from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
-  obtain a and b where
-    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
-    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
+  obtain a and b where as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
+    and axb: "a \<le> x \<and> x \<le> b \<and> P x" by auto
   from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
   thus ?thesis using px as bs noS by blast
 qed
 
 end
 
+
 section {* The classical QE after Langford for dense linear orders *}
 
 context dense_linorder
 begin
 
-lemma interval_empty_iff:
-  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
+lemma interval_empty_iff: "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
   by (auto dest: dense)
 
 lemma dlo_qe_bnds[no_atp]: 
@@ -219,11 +255,11 @@
 proof (simp only: atomize_eq, rule iffI)
   assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
   then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
-  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
+  { fix l u assume l: "l \<in> L" and u: "u \<in> U"
     have "l < x" using xL l by blast
     also have "x < u" using xU u by blast
-    finally (less_trans) have "l < u" .}
-  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
+    finally (less_trans) have "l < u" . }
+  then show "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
 next
   assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
   let ?ML = "Max L"
@@ -300,6 +336,7 @@
 
 locale linorder_stupid_syntax = linorder
 begin
+
 notation
   less_eq  ("op \<sqsubseteq>") and
   less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
@@ -311,6 +348,7 @@
 locale linorder_no_ub = linorder_stupid_syntax +
   assumes gt_ex: "\<exists>y. less x y"
 begin
+
 lemma ge_ex[no_atp]: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
 
 text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
@@ -323,7 +361,7 @@
      and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
   from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
-  {fix x assume H: "z \<sqsubset> x"
+  { fix x assume H: "z \<sqsubset> x"
     from less_trans[OF zz1 H] less_trans[OF zz2 H]
     have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
   }
@@ -332,14 +370,14 @@
 
 lemma pinf_disj[no_atp]:
   assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+    and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
 proof-
   from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
      and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
   from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
-  {fix x assume H: "z \<sqsubset> x"
+  { fix x assume H: "z \<sqsubset> x"
     from less_trans[OF zz1 H] less_trans[OF zz2 H]
     have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
   }
@@ -347,7 +385,7 @@
 qed
 
 lemma pinf_ex[no_atp]: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
-proof-
+proof -
   from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
   from gt_ex obtain x where x: "z \<sqsubset> x" by blast
   from z x p1 show ?thesis by blast
@@ -360,19 +398,20 @@
 locale linorder_no_lb = linorder_stupid_syntax +
   assumes lt_ex: "\<exists>y. less y x"
 begin
+
 lemma le_ex[no_atp]: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
 
 
 text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
 lemma minf_conj[no_atp]:
   assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+    and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
 proof-
   from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
   from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
-  {fix x assume H: "x \<sqsubset> z"
+  { fix x assume H: "x \<sqsubset> z"
     from less_trans[OF H zz1] less_trans[OF H zz2]
     have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
   }
@@ -381,21 +420,25 @@
 
 lemma minf_disj[no_atp]:
   assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+    and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+proof -
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+    and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
   from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
-  {fix x assume H: "x \<sqsubset> z"
+  { fix x assume H: "x \<sqsubset> z"
     from less_trans[OF H zz1] less_trans[OF H zz2]
     have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
   }
   thus ?thesis by blast
 qed
 
-lemma minf_ex[no_atp]: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
-proof-
+lemma minf_ex[no_atp]:
+  assumes ex: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)"
+    and p1: P1
+  shows "\<exists> x. P x"
+proof -
   from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
   from lt_ex obtain x where x: "x \<sqsubset> z" by blast
   from z x p1 show ?thesis by blast
@@ -407,29 +450,31 @@
 locale constr_dense_linorder = linorder_no_lb + linorder_no_ub +
   fixes between
   assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
-     and  between_same: "between x x = x"
+    and between_same: "between x x = x"
 
 sublocale  constr_dense_linorder < dense_linorder 
   apply unfold_locales
   using gt_ex lt_ex between_less
-    by (auto, rule_tac x="between x y" in exI, simp)
+  apply auto
+  apply (rule_tac x="between x y" in exI)
+  apply simp
+  done
 
-context  constr_dense_linorder
+context constr_dense_linorder
 begin
 
 lemma rinf_U[no_atp]:
   assumes fU: "finite U"
-  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
-  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
-  and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
-  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
+    and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
+      \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
+    and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
+    and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
   shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
-proof-
+proof -
   from ex obtain x where px: "P x" by blast
   from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
   then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto
   from uU have Une: "U \<noteq> {}" by auto
-  term "linorder.Min less_eq"
   let ?l = "linorder.Min less_eq U"
   let ?u = "linorder.Max less_eq U"
   have linM: "?l \<in> U" using fU Une by simp
@@ -443,49 +488,51 @@
   from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
   have "(\<exists> s\<in> U. P s) \<or>
       (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" .
-  moreover { fix u assume um: "u\<in>U" and pu: "P u"
+  moreover {
+    fix u assume um: "u\<in>U" and pu: "P u"
     have "between u u = u" by (simp add: between_same)
     with um pu have "P (between u u)" by simp
-    with um have ?thesis by blast}
-  moreover{
+    with um have ?thesis by blast }
+  moreover {
     assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x"
-      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
-        and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x"
-        by blast
-      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
-      let ?u = "between t1 t2"
-      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
-      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
-      with t1M t2M have ?thesis by blast}
-    ultimately show ?thesis by blast
-  qed
+    then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
+      and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U"
+      and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x" by blast
+    from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
+    let ?u = "between t1 t2"
+    from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
+    from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
+    with t1M t2M have ?thesis by blast
+  }
+  ultimately show ?thesis by blast
+qed
 
 theorem fr_eq[no_atp]:
   assumes fU: "finite U"
-  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
-   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
-  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
-  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
-  and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)"  and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
+    and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
+     \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
+    and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
+    and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
+    and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)"  and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
   shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
   (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
-proof-
- {
-   assume px: "\<exists> x. P x"
-   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
-   moreover {assume "MP \<or> PP" hence "?D" by blast}
-   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
-     from npmibnd[OF nmibnd npibnd]
-     have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
-     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
-   ultimately have "?D" by blast}
- moreover
- { assume "?D"
-   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
-   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
-   moreover {assume f:"?F" hence "?E" by blast}
-   ultimately have "?E" by blast}
- ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
+proof -
+  { assume px: "\<exists> x. P x"
+    have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
+    moreover { assume "MP \<or> PP" hence "?D" by blast }
+    moreover {
+      assume nmi: "\<not> MP" and npi: "\<not> PP"
+      from npmibnd[OF nmibnd npibnd]
+      have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
+      from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast }
+    ultimately have "?D" by blast }
+  moreover
+  { assume "?D"
+    moreover { assume m:"MP" from minf_ex[OF mi m] have "?E" . }
+    moreover { assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
+    moreover { assume f:"?F" hence "?E" by blast }
+    ultimately have "?E" by blast }
+  ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
 qed
 
 lemmas minf_thms[no_atp] = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
@@ -497,6 +544,7 @@
 
 lemma ferrack_axiom[no_atp]: "constr_dense_linorder less_eq less between"
   by (rule constr_dense_linorder_axioms)
+
 lemma atoms[no_atp]:
   shows "TERM (less :: 'a \<Rightarrow> _)"
     and "TERM (less_eq :: 'a \<Rightarrow> _)"
@@ -508,31 +556,33 @@
 
 declaration {*
 let
-fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
-fun generic_whatis phi =
- let
-  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
-  fun h x t =
-   case term_of t of
-     Const(@{const_name HOL.eq}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
-                            else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
-                            else Ferrante_Rackoff_Data.Nox
-   | b$y$z => if Term.could_unify (b, lt) then
-                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
-                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+  fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
+  fun generic_whatis phi =
+    let
+      val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
+      fun h x t =
+        case term_of t of
+          Const(@{const_name HOL.eq}, _)$y$z =>
+            if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+            else Ferrante_Rackoff_Data.Nox
+       | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
+            if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+            else Ferrante_Rackoff_Data.Nox
+       | b$y$z => if Term.could_unify (b, lt) then
+                     if term_of x aconv y then Ferrante_Rackoff_Data.Lt
+                     else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+                     else Ferrante_Rackoff_Data.Nox
+                 else if Term.could_unify (b, le) then
+                     if term_of x aconv y then Ferrante_Rackoff_Data.Le
+                     else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
+                     else Ferrante_Rackoff_Data.Nox
                  else Ferrante_Rackoff_Data.Nox
-             else if Term.could_unify (b, le) then
-                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
-                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
-                 else Ferrante_Rackoff_Data.Nox
-             else Ferrante_Rackoff_Data.Nox
-   | _ => Ferrante_Rackoff_Data.Nox
- in h end
- fun ss phi = HOL_ss addsimps (simps phi)
+       | _ => Ferrante_Rackoff_Data.Nox
+  in h end
+  fun ss phi = HOL_ss addsimps (simps phi)
 in
- Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
-  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
+  Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
+    {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
 end
 *}
 
@@ -544,38 +594,45 @@
   Scan.succeed (SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
 *} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
 
+
 subsection {* Ferrante and Rackoff algorithm over ordered fields *}
 
 lemma neg_prod_lt:"(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
-proof-
+proof -
   assume H: "c < 0"
-  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] algebra_simps)
+  have "c*x < 0 = (0/c < x)"
+    by (simp only: neg_divide_less_eq[OF H] algebra_simps)
   also have "\<dots> = (0 < x)" by simp
   finally show  "(c*x < 0) == (x > 0)" by simp
 qed
 
 lemma pos_prod_lt:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
-proof-
+proof -
   assume H: "c > 0"
-  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] algebra_simps)
+  then have "c*x < 0 = (0/c > x)"
+    by (simp only: pos_less_divide_eq[OF H] algebra_simps)
   also have "\<dots> = (0 > x)" by simp
   finally show  "(c*x < 0) == (x < 0)" by simp
 qed
 
 lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
-proof-
+proof -
   assume H: "c < 0"
-  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] algebra_simps)
+  have "c*x + t< 0 = (c*x < -t)"
+    by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c < x)"
+    by (simp only: neg_divide_less_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t < x)" by simp
-  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
+  finally show "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
 qed
 
 lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
-proof-
+proof -
   assume H: "c > 0"
-  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] algebra_simps)
+  have "c*x + t< 0 = (c*x < -t)"
+    by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c > x)"
+    by (simp only: pos_less_divide_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t > x)" by simp
   finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
 qed
@@ -584,50 +641,60 @@
   using less_diff_eq[where a= x and b=t and c=0] by simp
 
 lemma neg_prod_le:"(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
-proof-
+proof -
   assume H: "c < 0"
-  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] algebra_simps)
+  have "c*x <= 0 = (0/c <= x)"
+    by (simp only: neg_divide_le_eq[OF H] algebra_simps)
   also have "\<dots> = (0 <= x)" by simp
   finally show  "(c*x <= 0) == (x >= 0)" by simp
 qed
 
 lemma pos_prod_le:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
-proof-
+proof -
   assume H: "c > 0"
-  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] algebra_simps)
+  hence "c*x <= 0 = (0/c >= x)"
+    by (simp only: pos_le_divide_eq[OF H] algebra_simps)
   also have "\<dots> = (0 >= x)" by simp
   finally show  "(c*x <= 0) == (x <= 0)" by simp
 qed
 
 lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
-proof-
+proof -
   assume H: "c < 0"
-  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] algebra_simps)
+  have "c*x + t <= 0 = (c*x <= -t)"
+    by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c <= x)"
+    by (simp only: neg_divide_le_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t <= x)" by simp
-  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
+  finally show "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
 qed
 
 lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
-proof-
+proof -
   assume H: "c > 0"
-  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] algebra_simps)
+  have "c*x + t <= 0 = (c*x <= -t)"
+    by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c >= x)"
+    by (simp only: pos_le_divide_eq[OF H] algebra_simps)
   also have "\<dots> = ((- 1/c)*t >= x)" by simp
-  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
+  finally show "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
 qed
 
 lemma sum_le:"((x::'a::ordered_ab_group_add) + t <= 0) == (x <= - t)"
   using le_diff_eq[where a= x and b=t and c=0] by simp
 
 lemma nz_prod_eq:"(c\<Colon>'a\<Colon>linordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
+
 lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>linordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
-proof-
+proof -
   assume H: "c \<noteq> 0"
-  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] algebra_simps)
+  have "c*x + t = 0 = (c*x = -t)"
+    by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (x = -t/c)"
+    by (simp only: nonzero_eq_divide_eq[OF H] algebra_simps)
   finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
 qed
+
 lemma sum_eq:"((x::'a::ordered_ab_group_add) + t = 0) == (x = - t)"
   using eq_diff_eq[where a= x and b=t and c=0] by simp
 
@@ -635,23 +702,24 @@
 interpretation class_dense_linordered_field: constr_dense_linorder
  "op <=" "op <"
    "\<lambda> x y. 1/2 * ((x::'a::{linordered_field}) + y)"
-by (unfold_locales, dlo, dlo, auto)
+  by unfold_locales (dlo, dlo, auto)
 
 declaration{*
 let
-fun earlier [] x y = false
-        | earlier (h::t) x y =
-    if h aconvc y then false else if h aconvc x then true else earlier t x y;
+  fun earlier [] x y = false
+    | earlier (h::t) x y =
+        if h aconvc y then false else if h aconvc x then true else earlier t x y;
 
-fun dest_frac ct = case term_of ct of
-   Const (@{const_name Fields.divide},_) $ a $ b=>
-    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
- | Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
- | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
+fun dest_frac ct =
+  case term_of ct of
+    Const (@{const_name Fields.divide},_) $ a $ b=>
+      Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
+  | Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
+  | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
 
 fun mk_frac phi cT x =
- let val (a, b) = Rat.quotient_of_rat x
- in if b = 1 then Numeral.mk_cnumber cT a
+  let val (a, b) = Rat.quotient_of_rat x
+  in if b = 1 then Numeral.mk_cnumber cT a
     else Thm.apply
          (Thm.apply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
                      (Numeral.mk_cnumber cT a))