merged
authorpaulson
Mon, 29 Apr 2019 16:50:34 +0100
changeset 70213 ff8386fc191d
parent 70211 2388e0d2827b (diff)
parent 70212 e886b58bf698 (current diff)
child 70214 58191e01f0b1
merged
--- a/etc/symbols	Mon Apr 29 16:50:20 2019 +0100
+++ b/etc/symbols	Mon Apr 29 16:50:34 2019 +0100
@@ -394,6 +394,7 @@
 \<^class>               argument: cartouche
 \<^class_syntax>        argument: cartouche
 \<^command_keyword>     argument: cartouche
+\<^const>               argument: cartouche
 \<^const_abbrev>        argument: cartouche
 \<^const_name>          argument: cartouche
 \<^const_syntax>        argument: cartouche
--- a/src/HOL/Nonstandard_Analysis/HTranscendental.thy	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/HOL/Nonstandard_Analysis/HTranscendental.thy	Mon Apr 29 16:50:34 2019 +0100
@@ -12,602 +12,524 @@
 begin
 
 definition
-  exphr :: "real => hypreal" where
+  exphr :: "real \<Rightarrow> hypreal" where
     \<comment> \<open>define exponential function using standard part\<close>
-  "exphr x =  st(sumhr (0, whn, %n. inverse (fact n) * (x ^ n)))"
+  "exphr x \<equiv>  st(sumhr (0, whn, \<lambda>n. inverse (fact n) * (x ^ n)))"
 
 definition
-  sinhr :: "real => hypreal" where
-  "sinhr x = st(sumhr (0, whn, %n. sin_coeff n * x ^ n))"
+  sinhr :: "real \<Rightarrow> hypreal" where
+  "sinhr x \<equiv> st(sumhr (0, whn, \<lambda>n. sin_coeff n * x ^ n))"
   
 definition
-  coshr :: "real => hypreal" where
-  "coshr x = st(sumhr (0, whn, %n. cos_coeff n * x ^ n))"
+  coshr :: "real \<Rightarrow> hypreal" where
+  "coshr x \<equiv> st(sumhr (0, whn, \<lambda>n. cos_coeff n * x ^ n))"
 
 
 subsection\<open>Nonstandard Extension of Square Root Function\<close>
 
 lemma STAR_sqrt_zero [simp]: "( *f* sqrt) 0 = 0"
-by (simp add: starfun star_n_zero_num)
+  by (simp add: starfun star_n_zero_num)
 
 lemma STAR_sqrt_one [simp]: "( *f* sqrt) 1 = 1"
-by (simp add: starfun star_n_one_num)
+  by (simp add: starfun star_n_one_num)
 
 lemma hypreal_sqrt_pow2_iff: "(( *f* sqrt)(x) ^ 2 = x) = (0 \<le> x)"
-apply (cases x)
-apply (auto simp add: star_n_le star_n_zero_num starfun hrealpow star_n_eq_iff
-            simp del: hpowr_Suc power_Suc)
-done
-
-lemma hypreal_sqrt_gt_zero_pow2: "!!x. 0 < x ==> ( *f* sqrt) (x) ^ 2 = x"
-by (transfer, simp)
+proof (cases x)
+  case (star_n X)
+  then show ?thesis
+    by (simp add: star_n_le star_n_zero_num starfun hrealpow star_n_eq_iff del: hpowr_Suc power_Suc)
+qed
 
-lemma hypreal_sqrt_pow2_gt_zero: "0 < x ==> 0 < ( *f* sqrt) (x) ^ 2"
-by (frule hypreal_sqrt_gt_zero_pow2, auto)
+lemma hypreal_sqrt_gt_zero_pow2: "\<And>x. 0 < x \<Longrightarrow> ( *f* sqrt) (x) ^ 2 = x"
+  by transfer simp
 
-lemma hypreal_sqrt_not_zero: "0 < x ==> ( *f* sqrt) (x) \<noteq> 0"
-apply (frule hypreal_sqrt_pow2_gt_zero)
-apply (auto simp add: numeral_2_eq_2)
-done
+lemma hypreal_sqrt_pow2_gt_zero: "0 < x \<Longrightarrow> 0 < ( *f* sqrt) (x) ^ 2"
+  by (frule hypreal_sqrt_gt_zero_pow2, auto)
+
+lemma hypreal_sqrt_not_zero: "0 < x \<Longrightarrow> ( *f* sqrt) (x) \<noteq> 0"
+  using hypreal_sqrt_gt_zero_pow2 by fastforce
 
 lemma hypreal_inverse_sqrt_pow2:
-     "0 < x ==> inverse (( *f* sqrt)(x)) ^ 2 = inverse x"
-apply (cut_tac n = 2 and a = "( *f* sqrt) x" in power_inverse [symmetric])
-apply (auto dest: hypreal_sqrt_gt_zero_pow2)
-done
+     "0 < x \<Longrightarrow> inverse (( *f* sqrt)(x)) ^ 2 = inverse x"
+  by (simp add: hypreal_sqrt_gt_zero_pow2 power_inverse)
 
 lemma hypreal_sqrt_mult_distrib: 
-    "!!x y. [|0 < x; 0 <y |] ==>
+    "\<And>x y. \<lbrakk>0 < x; 0 <y\<rbrakk> \<Longrightarrow>
       ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)"
-apply transfer
-apply (auto intro: real_sqrt_mult) 
-done
+  by transfer (auto intro: real_sqrt_mult) 
 
 lemma hypreal_sqrt_mult_distrib2:
-     "[|0\<le>x; 0\<le>y |] ==>  
-     ( *f* sqrt)(x*y) =  ( *f* sqrt)(x) * ( *f* sqrt)(y)"
+     "\<lbrakk>0\<le>x; 0\<le>y\<rbrakk> \<Longrightarrow>  ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)"
 by (auto intro: hypreal_sqrt_mult_distrib simp add: order_le_less)
 
 lemma hypreal_sqrt_approx_zero [simp]:
-     "0 < x ==> (( *f* sqrt)(x) \<approx> 0) = (x \<approx> 0)"
-apply (auto simp add: mem_infmal_iff [symmetric])
-apply (rule hypreal_sqrt_gt_zero_pow2 [THEN subst])
-apply (auto intro: Infinitesimal_mult 
-            dest!: hypreal_sqrt_gt_zero_pow2 [THEN ssubst] 
-            simp add: numeral_2_eq_2)
-done
+  assumes "0 < x"
+  shows "(( *f* sqrt) x \<approx> 0) \<longleftrightarrow> (x \<approx> 0)"
+proof -
+  have "( *f* sqrt) x \<in> Infinitesimal \<longleftrightarrow> ((*f* sqrt) x)\<^sup>2 \<in> Infinitesimal"
+    by (metis Infinitesimal_hrealpow pos2 power2_eq_square Infinitesimal_square_iff)
+  also have "... \<longleftrightarrow> x \<in> Infinitesimal"
+    by (simp add: assms hypreal_sqrt_gt_zero_pow2)
+  finally show ?thesis
+    using mem_infmal_iff by blast
+qed
 
 lemma hypreal_sqrt_approx_zero2 [simp]:
-     "0 \<le> x ==> (( *f* sqrt)(x) \<approx> 0) = (x \<approx> 0)"
-by (auto simp add: order_le_less)
+  "0 \<le> x \<Longrightarrow> (( *f* sqrt)(x) \<approx> 0) = (x \<approx> 0)"
+  by (auto simp add: order_le_less)
 
-lemma hypreal_sqrt_sum_squares [simp]:
-     "(( *f* sqrt)(x*x + y*y + z*z) \<approx> 0) = (x*x + y*y + z*z \<approx> 0)"
-apply (rule hypreal_sqrt_approx_zero2)
-apply (rule add_nonneg_nonneg)+
-apply (auto)
-done
+lemma hypreal_sqrt_gt_zero: "\<And>x. 0 < x \<Longrightarrow> 0 < ( *f* sqrt)(x)"
+  by transfer (simp add: real_sqrt_gt_zero)
 
-lemma hypreal_sqrt_sum_squares2 [simp]:
-     "(( *f* sqrt)(x*x + y*y) \<approx> 0) = (x*x + y*y \<approx> 0)"
-apply (rule hypreal_sqrt_approx_zero2)
-apply (rule add_nonneg_nonneg)
-apply (auto)
-done
+lemma hypreal_sqrt_ge_zero: "0 \<le> x \<Longrightarrow> 0 \<le> ( *f* sqrt)(x)"
+  by (auto intro: hypreal_sqrt_gt_zero simp add: order_le_less)
 
-lemma hypreal_sqrt_gt_zero: "!!x. 0 < x ==> 0 < ( *f* sqrt)(x)"
-apply transfer
-apply (auto intro: real_sqrt_gt_zero)
-done
+lemma hypreal_sqrt_hrabs [simp]: "\<And>x. ( *f* sqrt)(x\<^sup>2) = \<bar>x\<bar>"
+  by transfer simp
 
-lemma hypreal_sqrt_ge_zero: "0 \<le> x ==> 0 \<le> ( *f* sqrt)(x)"
-by (auto intro: hypreal_sqrt_gt_zero simp add: order_le_less)
-
-lemma hypreal_sqrt_hrabs [simp]: "!!x. ( *f* sqrt)(x\<^sup>2) = \<bar>x\<bar>"
-by (transfer, simp)
-
-lemma hypreal_sqrt_hrabs2 [simp]: "!!x. ( *f* sqrt)(x*x) = \<bar>x\<bar>"
-by (transfer, simp)
+lemma hypreal_sqrt_hrabs2 [simp]: "\<And>x. ( *f* sqrt)(x*x) = \<bar>x\<bar>"
+  by transfer simp
 
 lemma hypreal_sqrt_hyperpow_hrabs [simp]:
-     "!!x. ( *f* sqrt)(x pow (hypnat_of_nat 2)) = \<bar>x\<bar>"
-by (transfer, simp)
+  "\<And>x. ( *f* sqrt)(x pow (hypnat_of_nat 2)) = \<bar>x\<bar>"
+  by transfer simp
 
 lemma star_sqrt_HFinite: "\<lbrakk>x \<in> HFinite; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HFinite"
-apply (rule HFinite_square_iff [THEN iffD1])
-apply (simp only: hypreal_sqrt_mult_distrib2 [symmetric], simp) 
-done
+  by (metis HFinite_square_iff hypreal_sqrt_pow2_iff power2_eq_square)
 
 lemma st_hypreal_sqrt:
-     "[| x \<in> HFinite; 0 \<le> x |] ==> st(( *f* sqrt) x) = ( *f* sqrt)(st x)"
-apply (rule power_inject_base [where n=1])
-apply (auto intro!: st_zero_le hypreal_sqrt_ge_zero)
-apply (rule st_mult [THEN subst])
-apply (rule_tac [3] hypreal_sqrt_mult_distrib2 [THEN subst])
-apply (rule_tac [5] hypreal_sqrt_mult_distrib2 [THEN subst])
-apply (auto simp add: st_hrabs st_zero_le star_sqrt_HFinite)
-done
+  assumes "x \<in> HFinite" "0 \<le> x"
+  shows "st(( *f* sqrt) x) = ( *f* sqrt)(st x)"
+proof (rule power_inject_base)
+  show "st ((*f* sqrt) x) ^ Suc 1 = (*f* sqrt) (st x) ^ Suc 1"
+    using assms hypreal_sqrt_pow2_iff [of x]
+    by (metis HFinite_square_iff hypreal_sqrt_hrabs2 power2_eq_square st_hrabs st_mult)
+  show "0 \<le> st ((*f* sqrt) x)"
+    by (simp add: assms hypreal_sqrt_ge_zero st_zero_le star_sqrt_HFinite)
+  show "0 \<le> (*f* sqrt) (st x)"
+    by (simp add: assms hypreal_sqrt_ge_zero st_zero_le)
+qed
 
-lemma hypreal_sqrt_sum_squares_ge1 [simp]: "!!x y. x \<le> ( *f* sqrt)(x\<^sup>2 + y\<^sup>2)"
-by transfer (rule real_sqrt_sum_squares_ge1)
-
-lemma HFinite_hypreal_sqrt:
-     "[| 0 \<le> x; x \<in> HFinite |] ==> ( *f* sqrt) x \<in> HFinite"
-apply (auto simp add: order_le_less)
-apply (rule HFinite_square_iff [THEN iffD1])
-apply (drule hypreal_sqrt_gt_zero_pow2)
-apply (simp add: numeral_2_eq_2)
-done
+lemma hypreal_sqrt_sum_squares_ge1 [simp]: "\<And>x y. x \<le> ( *f* sqrt)(x\<^sup>2 + y\<^sup>2)"
+  by transfer (rule real_sqrt_sum_squares_ge1)
 
 lemma HFinite_hypreal_sqrt_imp_HFinite:
-     "[| 0 \<le> x; ( *f* sqrt) x \<in> HFinite |] ==> x \<in> HFinite"
-apply (auto simp add: order_le_less)
-apply (drule HFinite_square_iff [THEN iffD2])
-apply (drule hypreal_sqrt_gt_zero_pow2)
-apply (simp add: numeral_2_eq_2 del: HFinite_square_iff)
-done
+  "\<lbrakk>0 \<le> x; ( *f* sqrt) x \<in> HFinite\<rbrakk> \<Longrightarrow> x \<in> HFinite"
+  by (metis HFinite_mult hrealpow_two hypreal_sqrt_pow2_iff numeral_2_eq_2)
 
 lemma HFinite_hypreal_sqrt_iff [simp]:
-     "0 \<le> x ==> (( *f* sqrt) x \<in> HFinite) = (x \<in> HFinite)"
-by (blast intro: HFinite_hypreal_sqrt HFinite_hypreal_sqrt_imp_HFinite)
-
-lemma HFinite_sqrt_sum_squares [simp]:
-     "(( *f* sqrt)(x*x + y*y) \<in> HFinite) = (x*x + y*y \<in> HFinite)"
-apply (rule HFinite_hypreal_sqrt_iff)
-apply (rule add_nonneg_nonneg)
-apply (auto)
-done
+  "0 \<le> x \<Longrightarrow> (( *f* sqrt) x \<in> HFinite) = (x \<in> HFinite)"
+  by (blast intro: star_sqrt_HFinite HFinite_hypreal_sqrt_imp_HFinite)
 
 lemma Infinitesimal_hypreal_sqrt:
-     "[| 0 \<le> x; x \<in> Infinitesimal |] ==> ( *f* sqrt) x \<in> Infinitesimal"
-apply (auto simp add: order_le_less)
-apply (rule Infinitesimal_square_iff [THEN iffD2])
-apply (drule hypreal_sqrt_gt_zero_pow2)
-apply (simp add: numeral_2_eq_2)
-done
+     "\<lbrakk>0 \<le> x; x \<in> Infinitesimal\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> Infinitesimal"
+  by (simp add: mem_infmal_iff)
 
 lemma Infinitesimal_hypreal_sqrt_imp_Infinitesimal:
-     "[| 0 \<le> x; ( *f* sqrt) x \<in> Infinitesimal |] ==> x \<in> Infinitesimal"
-apply (auto simp add: order_le_less)
-apply (drule Infinitesimal_square_iff [THEN iffD1])
-apply (drule hypreal_sqrt_gt_zero_pow2)
-apply (simp add: numeral_2_eq_2 del: Infinitesimal_square_iff [symmetric])
-done
+     "\<lbrakk>0 \<le> x; ( *f* sqrt) x \<in> Infinitesimal\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal"
+  using hypreal_sqrt_approx_zero2 mem_infmal_iff by blast
 
 lemma Infinitesimal_hypreal_sqrt_iff [simp]:
-     "0 \<le> x ==> (( *f* sqrt) x \<in> Infinitesimal) = (x \<in> Infinitesimal)"
+     "0 \<le> x \<Longrightarrow> (( *f* sqrt) x \<in> Infinitesimal) = (x \<in> Infinitesimal)"
 by (blast intro: Infinitesimal_hypreal_sqrt_imp_Infinitesimal Infinitesimal_hypreal_sqrt)
 
-lemma Infinitesimal_sqrt_sum_squares [simp]:
-     "(( *f* sqrt)(x*x + y*y) \<in> Infinitesimal) = (x*x + y*y \<in> Infinitesimal)"
-apply (rule Infinitesimal_hypreal_sqrt_iff)
-apply (rule add_nonneg_nonneg)
-apply (auto)
-done
-
 lemma HInfinite_hypreal_sqrt:
-     "[| 0 \<le> x; x \<in> HInfinite |] ==> ( *f* sqrt) x \<in> HInfinite"
-apply (auto simp add: order_le_less)
-apply (rule HInfinite_square_iff [THEN iffD1])
-apply (drule hypreal_sqrt_gt_zero_pow2)
-apply (simp add: numeral_2_eq_2)
-done
+     "\<lbrakk>0 \<le> x; x \<in> HInfinite\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HInfinite"
+  by (simp add: HInfinite_HFinite_iff)
 
 lemma HInfinite_hypreal_sqrt_imp_HInfinite:
-     "[| 0 \<le> x; ( *f* sqrt) x \<in> HInfinite |] ==> x \<in> HInfinite"
-apply (auto simp add: order_le_less)
-apply (drule HInfinite_square_iff [THEN iffD2])
-apply (drule hypreal_sqrt_gt_zero_pow2)
-apply (simp add: numeral_2_eq_2 del: HInfinite_square_iff)
-done
+     "\<lbrakk>0 \<le> x; ( *f* sqrt) x \<in> HInfinite\<rbrakk> \<Longrightarrow> x \<in> HInfinite"
+  using HFinite_hypreal_sqrt_iff HInfinite_HFinite_iff by blast
 
 lemma HInfinite_hypreal_sqrt_iff [simp]:
-     "0 \<le> x ==> (( *f* sqrt) x \<in> HInfinite) = (x \<in> HInfinite)"
+     "0 \<le> x \<Longrightarrow> (( *f* sqrt) x \<in> HInfinite) = (x \<in> HInfinite)"
 by (blast intro: HInfinite_hypreal_sqrt HInfinite_hypreal_sqrt_imp_HInfinite)
 
-lemma HInfinite_sqrt_sum_squares [simp]:
-     "(( *f* sqrt)(x*x + y*y) \<in> HInfinite) = (x*x + y*y \<in> HInfinite)"
-apply (rule HInfinite_hypreal_sqrt_iff)
-apply (rule add_nonneg_nonneg)
-apply (auto)
-done
-
 lemma HFinite_exp [simp]:
-     "sumhr (0, whn, %n. inverse (fact n) * x ^ n) \<in> HFinite"
-unfolding sumhr_app
-apply (simp only: star_zero_def starfun2_star_of atLeast0LessThan)
-apply (rule NSBseqD2)
-apply (rule NSconvergent_NSBseq)
-apply (rule convergent_NSconvergent_iff [THEN iffD1])
-apply (rule summable_iff_convergent [THEN iffD1])
-apply (rule summable_exp)
-done
+  "sumhr (0, whn, \<lambda>n. inverse (fact n) * x ^ n) \<in> HFinite"
+  unfolding sumhr_app star_zero_def starfun2_star_of atLeast0LessThan
+  by (metis NSBseqD2 NSconvergent_NSBseq convergent_NSconvergent_iff summable_iff_convergent summable_exp)
 
 lemma exphr_zero [simp]: "exphr 0 = 1"
-apply (simp add: exphr_def sumhr_split_add [OF hypnat_one_less_hypnat_omega, symmetric])
-apply (rule st_unique, simp)
-apply (rule subst [where P="\<lambda>x. 1 \<approx> x", OF _ approx_refl])
-apply (rule rev_mp [OF hypnat_one_less_hypnat_omega])
-apply (rule_tac x="whn" in spec)
-apply (unfold sumhr_app, transfer, simp add: power_0_left)
-done
+proof -
+  have "\<forall>x>1. 1 = sumhr (0, 1, \<lambda>n. inverse (fact n) * 0 ^ n) + sumhr (1, x, \<lambda>n. inverse (fact n) * 0 ^ n)"
+    unfolding sumhr_app by transfer (simp add: power_0_left)
+  then have "sumhr (0, 1, \<lambda>n. inverse (fact n) * 0 ^ n) + sumhr (1, whn, \<lambda>n. inverse (fact n) * 0 ^ n) \<approx> 1"
+    by auto
+  then show ?thesis
+    unfolding exphr_def
+    using sumhr_split_add [OF hypnat_one_less_hypnat_omega] st_unique by auto
+qed
 
 lemma coshr_zero [simp]: "coshr 0 = 1"
-apply (simp add: coshr_def sumhr_split_add
-                   [OF hypnat_one_less_hypnat_omega, symmetric]) 
-apply (rule st_unique, simp)
-apply (rule subst [where P="\<lambda>x. 1 \<approx> x", OF _ approx_refl])
-apply (rule rev_mp [OF hypnat_one_less_hypnat_omega])
-apply (rule_tac x="whn" in spec)
-apply (unfold sumhr_app, transfer, simp add: cos_coeff_def power_0_left)
-done
+  proof -
+  have "\<forall>x>1. 1 = sumhr (0, 1, \<lambda>n. cos_coeff n * 0 ^ n) + sumhr (1, x, \<lambda>n. cos_coeff n * 0 ^ n)"
+    unfolding sumhr_app by transfer (simp add: power_0_left)
+  then have "sumhr (0, 1, \<lambda>n. cos_coeff n * 0 ^ n) + sumhr (1, whn, \<lambda>n. cos_coeff n * 0 ^ n) \<approx> 1"
+    by auto
+  then show ?thesis
+    unfolding coshr_def
+    using sumhr_split_add [OF hypnat_one_less_hypnat_omega] st_unique by auto
+qed
 
 lemma STAR_exp_zero_approx_one [simp]: "( *f* exp) (0::hypreal) \<approx> 1"
-apply (subgoal_tac "( *f* exp) (0::hypreal) = 1", simp)
-apply (transfer, simp)
-done
+  proof -
+  have "(*f* exp) (0::real star) = 1"
+    by transfer simp
+  then show ?thesis
+    by auto
+qed
 
-lemma STAR_exp_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* exp) (x::hypreal) \<approx> 1"
-apply (case_tac "x = 0")
-apply (cut_tac [2] x = 0 in DERIV_exp)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
-apply (drule_tac x = x in bspec, auto)
-apply (drule_tac c = x in approx_mult1)
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] 
-            simp add: mult.assoc)
-apply (rule approx_add_right_cancel [where d="-1"])
-apply (rule approx_sym [THEN [2] approx_trans2])
-apply (auto simp add: mem_infmal_iff)
-done
+lemma STAR_exp_Infinitesimal: 
+  assumes "x \<in> Infinitesimal" shows "( *f* exp) (x::hypreal) \<approx> 1"
+proof (cases "x = 0")
+  case False
+  have "NSDERIV exp 0 :> 1"
+    by (metis DERIV_exp NSDERIV_DERIV_iff exp_zero)
+  then have "((*f* exp) x - 1) / x \<approx> 1"
+    using nsderiv_def False NSDERIVD2 assms by fastforce
+  then have "(*f* exp) x - 1 \<approx> x"
+    using NSDERIVD4 \<open>NSDERIV exp 0 :> 1\<close> assms by fastforce
+  then show ?thesis
+    by (meson Infinitesimal_approx approx_minus_iff approx_trans2 assms not_Infinitesimal_not_zero)
+qed auto
 
 lemma STAR_exp_epsilon [simp]: "( *f* exp) \<epsilon> \<approx> 1"
-by (auto intro: STAR_exp_Infinitesimal)
+  by (auto intro: STAR_exp_Infinitesimal)
 
 lemma STAR_exp_add:
-  "!!(x::'a:: {banach,real_normed_field} star) y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
-by transfer (rule exp_add)
+  "\<And>(x::'a:: {banach,real_normed_field} star) y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
+  by transfer (rule exp_add)
 
 lemma exphr_hypreal_of_real_exp_eq: "exphr x = hypreal_of_real (exp x)"
-apply (simp add: exphr_def)
-apply (rule st_unique, simp)
-apply (subst starfunNat_sumr [symmetric])
-unfolding atLeast0LessThan
-apply (rule NSLIMSEQ_D [THEN approx_sym])
-apply (rule LIMSEQ_NSLIMSEQ)
-apply (subst sums_def [symmetric])
-apply (cut_tac exp_converges [where x=x], simp)
-apply (rule HNatInfinite_whn)
-done
+proof -
+  have "(\<lambda>n. inverse (fact n) * x ^ n) sums exp x"
+    using exp_converges [of x] by simp
+  then have "(\<lambda>n. \<Sum>n<n. inverse (fact n) * x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S exp x"
+    using NSsums_def sums_NSsums_iff by blast
+  then have "hypreal_of_real (exp x) \<approx> sumhr (0, whn, \<lambda>n. inverse (fact n) * x ^ n)"
+    unfolding starfunNat_sumr [symmetric] atLeast0LessThan
+    using HNatInfinite_whn NSLIMSEQ_iff approx_sym by blast
+  then show ?thesis
+    unfolding exphr_def using st_eq_approx_iff by auto
+qed
 
-lemma starfun_exp_ge_add_one_self [simp]: "!!x::hypreal. 0 \<le> x ==> (1 + x) \<le> ( *f* exp) x"
-by transfer (rule exp_ge_add_one_self_aux)
+lemma starfun_exp_ge_add_one_self [simp]: "\<And>x::hypreal. 0 \<le> x \<Longrightarrow> (1 + x) \<le> ( *f* exp) x"
+  by transfer (rule exp_ge_add_one_self_aux)
 
-(* exp (oo) is infinite *)
+text\<open>exp maps infinities to infinities\<close>
 lemma starfun_exp_HInfinite:
-     "[| x \<in> HInfinite; 0 \<le> x |] ==> ( *f* exp) (x::hypreal) \<in> HInfinite"
-apply (frule starfun_exp_ge_add_one_self)
-apply (rule HInfinite_ge_HInfinite, assumption)
-apply (rule order_trans [of _ "1+x"], auto) 
-done
+  fixes x :: hypreal
+  assumes "x \<in> HInfinite" "0 \<le> x"
+  shows "( *f* exp) x \<in> HInfinite"
+proof -
+  have "x \<le> 1 + x"
+    by simp
+  also have "\<dots> \<le> (*f* exp) x"
+    by (simp add: \<open>0 \<le> x\<close>)
+  finally show ?thesis
+    using HInfinite_ge_HInfinite assms by blast
+qed
 
 lemma starfun_exp_minus:
-  "!!x::'a:: {banach,real_normed_field} star. ( *f* exp) (-x) = inverse(( *f* exp) x)"
-by transfer (rule exp_minus)
+  "\<And>x::'a:: {banach,real_normed_field} star. ( *f* exp) (-x) = inverse(( *f* exp) x)"
+  by transfer (rule exp_minus)
 
-(* exp (-oo) is infinitesimal *)
+text\<open>exp maps infinitesimals to infinitesimals\<close>
 lemma starfun_exp_Infinitesimal:
-     "[| x \<in> HInfinite; x \<le> 0 |] ==> ( *f* exp) (x::hypreal) \<in> Infinitesimal"
-apply (subgoal_tac "\<exists>y. x = - y")
-apply (rule_tac [2] x = "- x" in exI)
-apply (auto intro!: HInfinite_inverse_Infinitesimal starfun_exp_HInfinite
-            simp add: starfun_exp_minus HInfinite_minus_iff)
-done
+  fixes x :: hypreal
+  assumes "x \<in> HInfinite" "x \<le> 0"
+  shows "( *f* exp) x \<in> Infinitesimal"
+proof -
+  obtain y where "x = -y" "y \<ge> 0"
+    by (metis abs_of_nonpos assms(2) eq_abs_iff')
+  then have "( *f* exp) y \<in> HInfinite"
+    using HInfinite_minus_iff assms(1) starfun_exp_HInfinite by blast
+  then show ?thesis
+    by (simp add: HInfinite_inverse_Infinitesimal \<open>x = - y\<close> starfun_exp_minus)
+qed
 
-lemma starfun_exp_gt_one [simp]: "!!x::hypreal. 0 < x ==> 1 < ( *f* exp) x"
-by transfer (rule exp_gt_one)
+lemma starfun_exp_gt_one [simp]: "\<And>x::hypreal. 0 < x \<Longrightarrow> 1 < ( *f* exp) x"
+  by transfer (rule exp_gt_one)
 
 abbreviation real_ln :: "real \<Rightarrow> real" where 
   "real_ln \<equiv> ln"
 
-lemma starfun_ln_exp [simp]: "!!x. ( *f* real_ln) (( *f* exp) x) = x"
-by transfer (rule ln_exp)
+lemma starfun_ln_exp [simp]: "\<And>x. ( *f* real_ln) (( *f* exp) x) = x"
+  by transfer (rule ln_exp)
 
-lemma starfun_exp_ln_iff [simp]: "!!x. (( *f* exp)(( *f* real_ln) x) = x) = (0 < x)"
-by transfer (rule exp_ln_iff)
+lemma starfun_exp_ln_iff [simp]: "\<And>x. (( *f* exp)(( *f* real_ln) x) = x) = (0 < x)"
+  by transfer (rule exp_ln_iff)
 
-lemma starfun_exp_ln_eq: "!!u x. ( *f* exp) u = x ==> ( *f* real_ln) x = u"
-by transfer (rule ln_unique)
+lemma starfun_exp_ln_eq: "\<And>u x. ( *f* exp) u = x \<Longrightarrow> ( *f* real_ln) x = u"
+  by transfer (rule ln_unique)
 
-lemma starfun_ln_less_self [simp]: "!!x. 0 < x ==> ( *f* real_ln) x < x"
-by transfer (rule ln_less_self)
-
-lemma starfun_ln_ge_zero [simp]: "!!x. 1 \<le> x ==> 0 \<le> ( *f* real_ln) x"
-by transfer (rule ln_ge_zero)
+lemma starfun_ln_less_self [simp]: "\<And>x. 0 < x \<Longrightarrow> ( *f* real_ln) x < x"
+  by transfer (rule ln_less_self)
 
-lemma starfun_ln_gt_zero [simp]: "!!x .1 < x ==> 0 < ( *f* real_ln) x"
-by transfer (rule ln_gt_zero)
+lemma starfun_ln_ge_zero [simp]: "\<And>x. 1 \<le> x \<Longrightarrow> 0 \<le> ( *f* real_ln) x"
+  by transfer (rule ln_ge_zero)
 
-lemma starfun_ln_not_eq_zero [simp]: "!!x. [| 0 < x; x \<noteq> 1 |] ==> ( *f* real_ln) x \<noteq> 0"
-by transfer simp
+lemma starfun_ln_gt_zero [simp]: "\<And>x .1 < x \<Longrightarrow> 0 < ( *f* real_ln) x"
+  by transfer (rule ln_gt_zero)
 
-lemma starfun_ln_HFinite: "[| x \<in> HFinite; 1 \<le> x |] ==> ( *f* real_ln) x \<in> HFinite"
-apply (rule HFinite_bounded)
-apply assumption 
-apply (simp_all add: starfun_ln_less_self order_less_imp_le)
-done
+lemma starfun_ln_not_eq_zero [simp]: "\<And>x. \<lbrakk>0 < x; x \<noteq> 1\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<noteq> 0"
+  by transfer simp
 
-lemma starfun_ln_inverse: "!!x. 0 < x ==> ( *f* real_ln) (inverse x) = -( *f* ln) x"
-by transfer (rule ln_inverse)
+lemma starfun_ln_HFinite: "\<lbrakk>x \<in> HFinite; 1 \<le> x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<in> HFinite"
+  by (metis HFinite_HInfinite_iff less_le_trans starfun_exp_HInfinite starfun_exp_ln_iff starfun_ln_ge_zero zero_less_one)
+
+lemma starfun_ln_inverse: "\<And>x. 0 < x \<Longrightarrow> ( *f* real_ln) (inverse x) = -( *f* ln) x"
+  by transfer (rule ln_inverse)
 
 lemma starfun_abs_exp_cancel: "\<And>x. \<bar>( *f* exp) (x::hypreal)\<bar> = ( *f* exp) x"
-by transfer (rule abs_exp_cancel)
+  by transfer (rule abs_exp_cancel)
 
 lemma starfun_exp_less_mono: "\<And>x y::hypreal. x < y \<Longrightarrow> ( *f* exp) x < ( *f* exp) y"
-by transfer (rule exp_less_mono)
+  by transfer (rule exp_less_mono)
 
-lemma starfun_exp_HFinite: "x \<in> HFinite ==> ( *f* exp) (x::hypreal) \<in> HFinite"
-apply (auto simp add: HFinite_def, rename_tac u)
-apply (rule_tac x="( *f* exp) u" in rev_bexI)
-apply (simp add: Reals_eq_Standard)
-apply (simp add: starfun_abs_exp_cancel)
-apply (simp add: starfun_exp_less_mono)
-done
+lemma starfun_exp_HFinite: 
+  fixes x :: hypreal
+  assumes "x \<in> HFinite"
+  shows "( *f* exp) x \<in> HFinite"
+proof -
+  obtain u where "u \<in> \<real>" "\<bar>x\<bar> < u"
+    using HFiniteD assms by force
+  with assms have "\<bar>(*f* exp) x\<bar> < (*f* exp) u" 
+    using starfun_abs_exp_cancel starfun_exp_less_mono by auto
+  with \<open>u \<in> \<real>\<close> show ?thesis
+    by (force simp: HFinite_def Reals_eq_Standard)
+qed
 
 lemma starfun_exp_add_HFinite_Infinitesimal_approx:
-     "[|x \<in> Infinitesimal; z \<in> HFinite |] ==> ( *f* exp) (z + x::hypreal) \<approx> ( *f* exp) z"
-apply (simp add: STAR_exp_add)
-apply (frule STAR_exp_Infinitesimal)
-apply (drule approx_mult2)
-apply (auto intro: starfun_exp_HFinite)
-done
+  fixes x :: hypreal
+  shows "\<lbrakk>x \<in> Infinitesimal; z \<in> HFinite\<rbrakk> \<Longrightarrow> ( *f* exp) (z + x::hypreal) \<approx> ( *f* exp) z"
+  using STAR_exp_Infinitesimal approx_mult2 starfun_exp_HFinite by (fastforce simp add: STAR_exp_add)
 
-(* using previous result to get to result *)
 lemma starfun_ln_HInfinite:
-     "[| x \<in> HInfinite; 0 < x |] ==> ( *f* real_ln) x \<in> HInfinite"
-apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
-apply (drule starfun_exp_HFinite)
-apply (simp add: starfun_exp_ln_iff [THEN iffD2] HFinite_HInfinite_iff)
-done
+  "\<lbrakk>x \<in> HInfinite; 0 < x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<in> HInfinite"
+  by (metis HInfinite_HFinite_iff starfun_exp_HFinite starfun_exp_ln_iff)
 
 lemma starfun_exp_HInfinite_Infinitesimal_disj:
- "x \<in> HInfinite \<Longrightarrow> ( *f* exp) x \<in> HInfinite \<or> ( *f* exp) (x::hypreal) \<in> Infinitesimal"
-apply (insert linorder_linear [of x 0]) 
-apply (auto intro: starfun_exp_HInfinite starfun_exp_Infinitesimal)
-done
+  fixes x :: hypreal
+  shows "x \<in> HInfinite \<Longrightarrow> ( *f* exp) x \<in> HInfinite \<or> ( *f* exp) (x::hypreal) \<in> Infinitesimal"
+  by (meson linear starfun_exp_HInfinite starfun_exp_Infinitesimal)
 
-(* check out this proof!!! *)
 lemma starfun_ln_HFinite_not_Infinitesimal:
-     "[| x \<in> HFinite - Infinitesimal; 0 < x |] ==> ( *f* real_ln) x \<in> HFinite"
-apply (rule ccontr, drule HInfinite_HFinite_iff [THEN iffD2])
-apply (drule starfun_exp_HInfinite_Infinitesimal_disj)
-apply (simp add: starfun_exp_ln_iff [symmetric] HInfinite_HFinite_iff
-            del: starfun_exp_ln_iff)
-done
+     "\<lbrakk>x \<in> HFinite - Infinitesimal; 0 < x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<in> HFinite"
+  by (metis DiffD1 DiffD2 HInfinite_HFinite_iff starfun_exp_HInfinite_Infinitesimal_disj starfun_exp_ln_iff)
 
 (* we do proof by considering ln of 1/x *)
 lemma starfun_ln_Infinitesimal_HInfinite:
-     "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* real_ln) x \<in> HInfinite"
-apply (drule Infinitesimal_inverse_HInfinite)
-apply (frule positive_imp_inverse_positive)
-apply (drule_tac [2] starfun_ln_HInfinite)
-apply (auto simp add: starfun_ln_inverse HInfinite_minus_iff)
-done
+  assumes "x \<in> Infinitesimal" "0 < x"
+  shows "( *f* real_ln) x \<in> HInfinite"
+proof -
+  have "inverse x \<in> HInfinite"
+    using Infinitesimal_inverse_HInfinite assms by blast
+  then show ?thesis
+    using HInfinite_minus_iff assms(2) starfun_ln_HInfinite starfun_ln_inverse by fastforce
+qed
 
-lemma starfun_ln_less_zero: "!!x. [| 0 < x; x < 1 |] ==> ( *f* real_ln) x < 0"
-by transfer (rule ln_less_zero)
+lemma starfun_ln_less_zero: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> ( *f* real_ln) x < 0"
+  by transfer (rule ln_less_zero)
 
 lemma starfun_ln_Infinitesimal_less_zero:
-     "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* real_ln) x < 0"
-by (auto intro!: starfun_ln_less_zero simp add: Infinitesimal_def)
+  "\<lbrakk>x \<in> Infinitesimal; 0 < x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x < 0"
+  by (auto intro!: starfun_ln_less_zero simp add: Infinitesimal_def)
 
 lemma starfun_ln_HInfinite_gt_zero:
-     "[| x \<in> HInfinite; 0 < x |] ==> 0 < ( *f* real_ln) x"
-by (auto intro!: starfun_ln_gt_zero simp add: HInfinite_def)
+  "\<lbrakk>x \<in> HInfinite; 0 < x\<rbrakk> \<Longrightarrow> 0 < ( *f* real_ln) x"
+  by (auto intro!: starfun_ln_gt_zero simp add: HInfinite_def)
 
 
-(*
-Goalw [NSLIM_def] "(%h. ((x powr h) - 1) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S ln x"
-*)
-
-lemma HFinite_sin [simp]: "sumhr (0, whn, %n. sin_coeff n * x ^ n) \<in> HFinite"
-unfolding sumhr_app
-apply (simp only: star_zero_def starfun2_star_of atLeast0LessThan)
-apply (rule NSBseqD2)
-apply (rule NSconvergent_NSBseq)
-apply (rule convergent_NSconvergent_iff [THEN iffD1])
-apply (rule summable_iff_convergent [THEN iffD1])
-using summable_norm_sin [of x]
-apply (simp add: summable_rabs_cancel)
-done
+lemma HFinite_sin [simp]: "sumhr (0, whn, \<lambda>n. sin_coeff n * x ^ n) \<in> HFinite"
+proof -
+  have "summable (\<lambda>i. sin_coeff i * x ^ i)"
+    using summable_norm_sin [of x] by (simp add: summable_rabs_cancel)
+  then have "(*f* (\<lambda>n. \<Sum>n<n. sin_coeff n * x ^ n)) whn \<in> HFinite"
+    unfolding summable_sums_iff sums_NSsums_iff NSsums_def NSLIMSEQ_iff
+    using HFinite_star_of HNatInfinite_whn approx_HFinite approx_sym by blast
+  then show ?thesis
+    unfolding sumhr_app
+    by (simp only: star_zero_def starfun2_star_of atLeast0LessThan)
+qed
 
 lemma STAR_sin_zero [simp]: "( *f* sin) 0 = 0"
-by transfer (rule sin_zero)
+  by transfer (rule sin_zero)
 
 lemma STAR_sin_Infinitesimal [simp]:
   fixes x :: "'a::{real_normed_field,banach} star"
-  shows "x \<in> Infinitesimal ==> ( *f* sin) x \<approx> x"
-apply (case_tac "x = 0")
-apply (cut_tac [2] x = 0 in DERIV_sin)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
-apply (drule bspec [where x = x], auto)
-apply (drule approx_mult1 [where c = x])
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-           simp add: mult.assoc)
-done
+  assumes "x \<in> Infinitesimal"
+  shows "( *f* sin) x \<approx> x"
+proof (cases "x = 0")
+  case False
+  have "NSDERIV sin 0 :> 1"
+    by (metis DERIV_sin NSDERIV_DERIV_iff cos_zero)
+  then have "(*f* sin) x / x \<approx> 1"
+    using False NSDERIVD2 assms by fastforce
+  with assms show ?thesis
+    unfolding star_one_def
+    by (metis False Infinitesimal_approx Infinitesimal_ratio approx_star_of_HFinite)
+qed auto
 
-lemma HFinite_cos [simp]: "sumhr (0, whn, %n. cos_coeff n * x ^ n) \<in> HFinite"
-unfolding sumhr_app
-apply (simp only: star_zero_def starfun2_star_of atLeast0LessThan)
-apply (rule NSBseqD2)
-apply (rule NSconvergent_NSBseq)
-apply (rule convergent_NSconvergent_iff [THEN iffD1])
-apply (rule summable_iff_convergent [THEN iffD1])
-using summable_norm_cos [of x]
-apply (simp add: summable_rabs_cancel)
-done
+lemma HFinite_cos [simp]: "sumhr (0, whn, \<lambda>n. cos_coeff n * x ^ n) \<in> HFinite"
+proof -
+  have "summable (\<lambda>i. cos_coeff i * x ^ i)"
+    using summable_norm_cos [of x] by (simp add: summable_rabs_cancel)
+  then have "(*f* (\<lambda>n. \<Sum>n<n. cos_coeff n * x ^ n)) whn \<in> HFinite"
+    unfolding summable_sums_iff sums_NSsums_iff NSsums_def NSLIMSEQ_iff
+    using HFinite_star_of HNatInfinite_whn approx_HFinite approx_sym by blast
+  then show ?thesis
+    unfolding sumhr_app
+    by (simp only: star_zero_def starfun2_star_of atLeast0LessThan)
+qed
 
 lemma STAR_cos_zero [simp]: "( *f* cos) 0 = 1"
-by transfer (rule cos_zero)
+  by transfer (rule cos_zero)
 
 lemma STAR_cos_Infinitesimal [simp]:
   fixes x :: "'a::{real_normed_field,banach} star"
-  shows "x \<in> Infinitesimal ==> ( *f* cos) x \<approx> 1"
-apply (case_tac "x = 0")
-apply (cut_tac [2] x = 0 in DERIV_cos)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
-apply (drule bspec [where x = x])
-apply auto
-apply (drule approx_mult1 [where c = x])
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-            simp add: mult.assoc)
-apply (rule approx_add_right_cancel [where d = "-1"])
-apply simp
-done
+  assumes "x \<in> Infinitesimal"
+  shows "( *f* cos) x \<approx> 1"
+proof (cases "x = 0")
+  case False
+  have "NSDERIV cos 0 :> 0"
+    by (metis DERIV_cos NSDERIV_DERIV_iff add.inverse_neutral sin_zero)
+  then have "(*f* cos) x - 1 \<approx> 0"
+    using NSDERIV_approx assms by fastforce
+  with assms show ?thesis
+    using approx_minus_iff by blast
+qed auto
 
 lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
-by transfer (rule tan_zero)
+  by transfer (rule tan_zero)
 
-lemma STAR_tan_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* tan) x \<approx> x"
-apply (case_tac "x = 0")
-apply (cut_tac [2] x = 0 in DERIV_tan)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
-apply (drule bspec [where x = x], auto)
-apply (drule approx_mult1 [where c = x])
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-             simp add: mult.assoc)
-done
+lemma STAR_tan_Infinitesimal [simp]:
+  assumes "x \<in> Infinitesimal"
+  shows "( *f* tan) x \<approx> x"
+proof (cases "x = 0")
+  case False
+  have "NSDERIV tan 0 :> 1"
+    using DERIV_tan [of 0] by (simp add: NSDERIV_DERIV_iff)
+  then have "(*f* tan) x / x \<approx> 1"
+    using False NSDERIVD2 assms by fastforce
+  with assms show ?thesis
+    unfolding star_one_def
+    by (metis False Infinitesimal_approx Infinitesimal_ratio approx_star_of_HFinite)
+qed auto
 
 lemma STAR_sin_cos_Infinitesimal_mult:
   fixes x :: "'a::{real_normed_field,banach} star"
-  shows "x \<in> Infinitesimal ==> ( *f* sin) x * ( *f* cos) x \<approx> x"
-using approx_mult_HFinite [of "( *f* sin) x" _ "( *f* cos) x" 1] 
-by (simp add: Infinitesimal_subset_HFinite [THEN subsetD])
+  shows "x \<in> Infinitesimal \<Longrightarrow> ( *f* sin) x * ( *f* cos) x \<approx> x"
+  using approx_mult_HFinite [of "( *f* sin) x" _ "( *f* cos) x" 1] 
+  by (simp add: Infinitesimal_subset_HFinite [THEN subsetD])
 
 lemma HFinite_pi: "hypreal_of_real pi \<in> HFinite"
-by simp
-
-(* lemmas *)
+  by simp
 
-lemma lemma_split_hypreal_of_real:
-     "N \<in> HNatInfinite  
-      ==> hypreal_of_real a =  
-          hypreal_of_hypnat N * (inverse(hypreal_of_hypnat N) * hypreal_of_real a)"
-by (simp add: mult.assoc [symmetric] zero_less_HNatInfinite)
 
 lemma STAR_sin_Infinitesimal_divide:
   fixes x :: "'a::{real_normed_field,banach} star"
-  shows "[|x \<in> Infinitesimal; x \<noteq> 0 |] ==> ( *f* sin) x/x \<approx> 1"
-using DERIV_sin [of "0::'a"]
-by (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
+  shows "\<lbrakk>x \<in> Infinitesimal; x \<noteq> 0\<rbrakk> \<Longrightarrow> ( *f* sin) x/x \<approx> 1"
+  using DERIV_sin [of "0::'a"]
+  by (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
 
-(*------------------------------------------------------------------------*) 
-(* sin* (1/n) * 1/(1/n) \<approx> 1 for n = oo                                   *)
-(*------------------------------------------------------------------------*)
+subsection \<open>Proving $\sin* (1/n) \times 1/(1/n) \approx 1$ for $n = \infty$ \<close> 
 
 lemma lemma_sin_pi:
-     "n \<in> HNatInfinite  
-      ==> ( *f* sin) (inverse (hypreal_of_hypnat n))/(inverse (hypreal_of_hypnat n)) \<approx> 1"
-apply (rule STAR_sin_Infinitesimal_divide)
-apply (auto simp add: zero_less_HNatInfinite)
-done
+  "n \<in> HNatInfinite  
+      \<Longrightarrow> ( *f* sin) (inverse (hypreal_of_hypnat n))/(inverse (hypreal_of_hypnat n)) \<approx> 1"
+  by (simp add: STAR_sin_Infinitesimal_divide zero_less_HNatInfinite)
 
 lemma STAR_sin_inverse_HNatInfinite:
      "n \<in> HNatInfinite  
-      ==> ( *f* sin) (inverse (hypreal_of_hypnat n)) * hypreal_of_hypnat n \<approx> 1"
-apply (frule lemma_sin_pi)
-apply (simp add: divide_inverse)
-done
+      \<Longrightarrow> ( *f* sin) (inverse (hypreal_of_hypnat n)) * hypreal_of_hypnat n \<approx> 1"
+  by (metis field_class.field_divide_inverse inverse_inverse_eq lemma_sin_pi)
 
 lemma Infinitesimal_pi_divide_HNatInfinite: 
      "N \<in> HNatInfinite  
-      ==> hypreal_of_real pi/(hypreal_of_hypnat N) \<in> Infinitesimal"
-apply (simp add: divide_inverse)
-apply (auto intro: Infinitesimal_HFinite_mult2)
-done
+      \<Longrightarrow> hypreal_of_real pi/(hypreal_of_hypnat N) \<in> Infinitesimal"
+  by (simp add: Infinitesimal_HFinite_mult2 field_class.field_divide_inverse)
 
 lemma pi_divide_HNatInfinite_not_zero [simp]:
-     "N \<in> HNatInfinite ==> hypreal_of_real pi/(hypreal_of_hypnat N) \<noteq> 0"
-by (simp add: zero_less_HNatInfinite)
+  "N \<in> HNatInfinite \<Longrightarrow> hypreal_of_real pi/(hypreal_of_hypnat N) \<noteq> 0"
+  by (simp add: zero_less_HNatInfinite)
 
 lemma STAR_sin_pi_divide_HNatInfinite_approx_pi:
-     "n \<in> HNatInfinite  
-      ==> ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n)) * hypreal_of_hypnat n  
-          \<approx> hypreal_of_real pi"
-apply (frule STAR_sin_Infinitesimal_divide
-               [OF Infinitesimal_pi_divide_HNatInfinite 
-                   pi_divide_HNatInfinite_not_zero])
-apply (auto)
-apply (rule approx_SReal_mult_cancel [of "inverse (hypreal_of_real pi)"])
-apply (auto intro: Reals_inverse simp add: divide_inverse ac_simps)
-done
+  assumes "n \<in> HNatInfinite"
+  shows "(*f* sin) (hypreal_of_real pi / hypreal_of_hypnat n) * hypreal_of_hypnat n \<approx>
+         hypreal_of_real pi"
+proof -
+  have "(*f* sin) (hypreal_of_real pi / hypreal_of_hypnat n) / (hypreal_of_real pi / hypreal_of_hypnat n) \<approx> 1"
+    using Infinitesimal_pi_divide_HNatInfinite STAR_sin_Infinitesimal_divide assms pi_divide_HNatInfinite_not_zero by blast
+  then have "hypreal_of_hypnat n * star_of sin \<star> (hypreal_of_real pi / hypreal_of_hypnat n) / hypreal_of_real pi \<approx> 1"
+    by (simp add: mult.commute starfun_def)
+  then show ?thesis
+    apply (simp add: starfun_def field_simps)
+    by (metis (no_types, lifting) approx_mult_subst_star_of approx_refl mult_cancel_right1 nonzero_eq_divide_eq pi_neq_zero star_of_eq_0)
+qed
 
 lemma STAR_sin_pi_divide_HNatInfinite_approx_pi2:
      "n \<in> HNatInfinite  
-      ==> hypreal_of_hypnat n *  
-          ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n))  
-          \<approx> hypreal_of_real pi"
-apply (rule mult.commute [THEN subst])
-apply (erule STAR_sin_pi_divide_HNatInfinite_approx_pi)
-done
+      \<Longrightarrow> hypreal_of_hypnat n * ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n)) \<approx> hypreal_of_real pi"
+  by (metis STAR_sin_pi_divide_HNatInfinite_approx_pi mult.commute)
 
 lemma starfunNat_pi_divide_n_Infinitesimal: 
-     "N \<in> HNatInfinite ==> ( *f* (%x. pi / real x)) N \<in> Infinitesimal"
-by (auto intro!: Infinitesimal_HFinite_mult2 
-         simp add: starfun_mult [symmetric] divide_inverse
-                   starfun_inverse [symmetric] starfunNat_real_of_nat)
+     "N \<in> HNatInfinite \<Longrightarrow> ( *f* (\<lambda>x. pi / real x)) N \<in> Infinitesimal"
+  by (simp add: Infinitesimal_HFinite_mult2 divide_inverse starfunNat_real_of_nat)
 
 lemma STAR_sin_pi_divide_n_approx:
-     "N \<in> HNatInfinite ==>  
-      ( *f* sin) (( *f* (%x. pi / real x)) N) \<approx>  
-      hypreal_of_real pi/(hypreal_of_hypnat N)"
-apply (simp add: starfunNat_real_of_nat [symmetric])
-apply (rule STAR_sin_Infinitesimal)
-apply (simp add: divide_inverse)
-apply (rule Infinitesimal_HFinite_mult2)
-apply (subst starfun_inverse)
-apply (erule starfunNat_inverse_real_of_nat_Infinitesimal)
-apply simp
-done
+  assumes "N \<in> HNatInfinite"
+  shows "( *f* sin) (( *f* (\<lambda>x. pi / real x)) N) \<approx> hypreal_of_real pi/(hypreal_of_hypnat N)"
+proof -
+  have "\<exists>s. (*f* sin) ((*f* (\<lambda>n. pi / real n)) N) \<approx> s \<and> hypreal_of_real pi / hypreal_of_hypnat N \<approx> s"
+    by (metis (lifting) Infinitesimal_approx Infinitesimal_pi_divide_HNatInfinite STAR_sin_Infinitesimal assms starfunNat_pi_divide_n_Infinitesimal)
+  then show ?thesis
+    by (meson approx_trans2)
+qed
 
-lemma NSLIMSEQ_sin_pi: "(%n. real n * sin (pi / real n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S pi"
-apply (auto simp add: NSLIMSEQ_def starfun_mult [symmetric] starfunNat_real_of_nat)
-apply (rule_tac f1 = sin in starfun_o2 [THEN subst])
-apply (auto simp add: starfun_mult [symmetric] starfunNat_real_of_nat divide_inverse)
-apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
-apply (auto dest: STAR_sin_pi_divide_HNatInfinite_approx_pi 
-            simp add: starfunNat_real_of_nat mult.commute divide_inverse)
-done
+lemma NSLIMSEQ_sin_pi: "(\<lambda>n. real n * sin (pi / real n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S pi"
+proof -
+  have *: "hypreal_of_hypnat N * (*f* sin) ((*f* (\<lambda>x. pi / real x)) N) \<approx> hypreal_of_real pi"
+    if "N \<in> HNatInfinite"
+    for N :: "nat star"
+    using that
+    by simp (metis STAR_sin_pi_divide_HNatInfinite_approx_pi2 starfunNat_real_of_nat)
+  show ?thesis
+    by (simp add: NSLIMSEQ_def starfunNat_real_of_nat) (metis * starfun_o2)
+qed
 
-lemma NSLIMSEQ_cos_one: "(%n. cos (pi / real n))\<longlonglongrightarrow>\<^sub>N\<^sub>S 1"
-apply (simp add: NSLIMSEQ_def, auto)
-apply (rule_tac f1 = cos in starfun_o2 [THEN subst])
-apply (rule STAR_cos_Infinitesimal)
-apply (auto intro!: Infinitesimal_HFinite_mult2 
-            simp add: starfun_mult [symmetric] divide_inverse
-                      starfun_inverse [symmetric] starfunNat_real_of_nat)
-done
+lemma NSLIMSEQ_cos_one: "(\<lambda>n. cos (pi / real n))\<longlonglongrightarrow>\<^sub>N\<^sub>S 1"
+proof -
+  have "(*f* cos) ((*f* (\<lambda>x. pi / real x)) N) \<approx> 1"
+    if "N \<in> HNatInfinite" for N 
+    using that STAR_cos_Infinitesimal starfunNat_pi_divide_n_Infinitesimal by blast
+  then show ?thesis
+    by (simp add: NSLIMSEQ_def) (metis STAR_cos_Infinitesimal starfunNat_pi_divide_n_Infinitesimal starfun_o2)
+qed
 
 lemma NSLIMSEQ_sin_cos_pi:
-     "(%n. real n * sin (pi / real n) * cos (pi / real n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S pi"
-by (insert NSLIMSEQ_mult [OF NSLIMSEQ_sin_pi NSLIMSEQ_cos_one], simp)
+  "(\<lambda>n. real n * sin (pi / real n) * cos (pi / real n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S pi"
+  using NSLIMSEQ_cos_one NSLIMSEQ_mult NSLIMSEQ_sin_pi by force
 
 
 text\<open>A familiar approximation to \<^term>\<open>cos x\<close> when \<^term>\<open>x\<close> is small\<close>
 
 lemma STAR_cos_Infinitesimal_approx:
   fixes x :: "'a::{real_normed_field,banach} star"
-  shows "x \<in> Infinitesimal ==> ( *f* cos) x \<approx> 1 - x\<^sup>2"
-apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
-apply (auto simp add: Infinitesimal_approx_minus [symmetric] 
-            add.assoc [symmetric] numeral_2_eq_2)
-done
+  shows "x \<in> Infinitesimal \<Longrightarrow> ( *f* cos) x \<approx> 1 - x\<^sup>2"
+  by (metis Infinitesimal_square_iff STAR_cos_Infinitesimal approx_diff approx_sym diff_zero mem_infmal_iff power2_eq_square)
 
 lemma STAR_cos_Infinitesimal_approx2:
   fixes x :: hypreal  \<comment> \<open>perhaps could be generalised, like many other hypreal results\<close>
-  shows "x \<in> Infinitesimal ==> ( *f* cos) x \<approx> 1 - (x\<^sup>2)/2"
-apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
-apply (auto intro: Infinitesimal_SReal_divide Infinitesimal_mult
-            simp add: Infinitesimal_approx_minus [symmetric] numeral_2_eq_2)
-done
+  assumes "x \<in> Infinitesimal"
+  shows "( *f* cos) x \<approx> 1 - (x\<^sup>2)/2"
+proof -
+  have "1 \<approx> 1 - x\<^sup>2 / 2"
+    using assms
+    by (auto intro: Infinitesimal_SReal_divide simp add: Infinitesimal_approx_minus [symmetric] numeral_2_eq_2)
+  then show ?thesis
+    using STAR_cos_Infinitesimal approx_trans assms by blast
+qed
 
 end
--- a/src/HOL/ex/Dedekind_Real.thy	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/HOL/ex/Dedekind_Real.thy	Mon Apr 29 16:50:34 2019 +0100
@@ -1,42 +1,36 @@
+section \<open>The Reals as Dedekind Sections of Positive Rationals\<close>
+
+text \<open>Fundamentals of Abstract Analysis [Gleason- p. 121] provides some of the definitions.\<close>
+
 (*  Title:      HOL/ex/Dedekind_Real.thy
     Author:     Jacques D. Fleuriot, University of Cambridge
-    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
-
-The positive reals as Dedekind sections of positive
-rationals. Fundamentals of Abstract Analysis [Gleason- p. 121]
-provides some of the definitions.
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4; 2019
 *)
 
 theory Dedekind_Real
-imports Complex_Main
+imports Complex_Main 
 begin
 
-section \<open>Positive real numbers\<close>
+text\<open>Could be moved to \<open>Groups\<close>\<close>
+lemma add_eq_exists: "\<exists>x. a+x = (b::'a::ab_group_add)"
+  by (rule_tac x="b-a" in exI, simp)
 
-text\<open>Could be generalized and moved to \<open>Groups\<close>\<close>
-lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
-by (rule_tac x="b-a" in exI, simp)
+subsection \<open>Dedekind cuts or sections\<close>
 
 definition
-  cut :: "rat set => bool" where
-  "cut A = ({} \<subset> A &
-            A < {r. 0 < r} &
-            (\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u))))"
-
-lemma interval_empty_iff:
-  "{y. (x::'a::unbounded_dense_linorder) < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
-  by (auto dest: dense)
-
+  cut :: "rat set \<Rightarrow> bool" where
+  "cut A \<equiv> {} \<subset> A \<and> A \<subset> {0<..} \<and>
+            (\<forall>y \<in> A. ((\<forall>z. 0<z \<and> z < y \<longrightarrow> z \<in> A) \<and> (\<exists>u \<in> A. y < u)))"
 
 lemma cut_of_rat: 
-  assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" (is "cut ?A")
+  assumes q: "0 < q" shows "cut {r::rat. 0 < r \<and> r < q}" (is "cut ?A")
 proof -
-  from q have pos: "?A < {r. 0 < r}" by force
+  from q have pos: "?A \<subset> {0<..}" by force
   have nonempty: "{} \<subset> ?A"
   proof
     show "{} \<subseteq> ?A" by simp
     show "{} \<noteq> ?A"
-      by (force simp only: q eq_commute [of "{}"] interval_empty_iff)
+      using field_lbound_gt_zero q by auto
   qed
   show ?thesis
     by (simp add: cut_def pos nonempty,
@@ -51,62 +45,61 @@
   "(\<And>x. cut x \<Longrightarrow> P (Abs_preal x)) \<Longrightarrow> P x"
   using Abs_preal_induct [of P x] by simp
 
-lemma Rep_preal:
-  "cut (Rep_preal x)"
+lemma cut_Rep_preal [simp]: "cut (Rep_preal x)"
   using Rep_preal [of x] by simp
 
 definition
-  psup :: "preal set => preal" where
+  psup :: "preal set \<Rightarrow> preal" where
   "psup P = Abs_preal (\<Union>X \<in> P. Rep_preal X)"
 
 definition
-  add_set :: "[rat set,rat set] => rat set" where
+  add_set :: "[rat set,rat set] \<Rightarrow> rat set" where
   "add_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
 
 definition
-  diff_set :: "[rat set,rat set] => rat set" where
-  "diff_set A B = {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}"
+  diff_set :: "[rat set,rat set] \<Rightarrow> rat set" where
+  "diff_set A B = {w. \<exists>x. 0 < w \<and> 0 < x \<and> x \<notin> B \<and> x + w \<in> A}"
 
 definition
-  mult_set :: "[rat set,rat set] => rat set" where
+  mult_set :: "[rat set,rat set] \<Rightarrow> rat set" where
   "mult_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
 
 definition
-  inverse_set :: "rat set => rat set" where
-  "inverse_set A = {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}"
+  inverse_set :: "rat set \<Rightarrow> rat set" where
+  "inverse_set A \<equiv> {x. \<exists>y. 0 < x \<and> x < y \<and> inverse y \<notin> A}"
 
 instantiation preal :: "{ord, plus, minus, times, inverse, one}"
 begin
 
 definition
   preal_less_def:
-    "R < S == Rep_preal R < Rep_preal S"
+    "r < s \<equiv> Rep_preal r < Rep_preal s"
 
 definition
   preal_le_def:
-    "R \<le> S == Rep_preal R \<subseteq> Rep_preal S"
+    "r \<le> s \<equiv> Rep_preal r \<subseteq> Rep_preal s"
 
 definition
   preal_add_def:
-    "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))"
+    "r + s \<equiv> Abs_preal (add_set (Rep_preal r) (Rep_preal s))"
 
 definition
   preal_diff_def:
-    "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))"
+    "r - s \<equiv> Abs_preal (diff_set (Rep_preal r) (Rep_preal s))"
 
 definition
   preal_mult_def:
-    "R * S == Abs_preal (mult_set (Rep_preal R) (Rep_preal S))"
+    "r * s \<equiv> Abs_preal (mult_set (Rep_preal r) (Rep_preal s))"
 
 definition
   preal_inverse_def:
-    "inverse R == Abs_preal (inverse_set (Rep_preal R))"
+    "inverse r \<equiv> Abs_preal (inverse_set (Rep_preal r))"
 
-definition "R div S = R * inverse (S::preal)"
+definition "r div s = r * inverse (s::preal)"
 
 definition
   preal_one_def:
-    "1 == Abs_preal {x. 0 < x & x < 1}"
+    "1 \<equiv> Abs_preal {x. 0 < x \<and> x < 1}"
 
 instance ..
 
@@ -117,37 +110,27 @@
 declare Abs_preal_inject [simp]
 declare Abs_preal_inverse [simp]
 
-lemma rat_mem_preal: "0 < q ==> cut {r::rat. 0 < r & r < q}"
+lemma rat_mem_preal: "0 < q \<Longrightarrow> cut {r::rat. 0 < r \<and> r < q}"
 by (simp add: cut_of_rat)
 
-lemma preal_nonempty: "cut A ==> \<exists>x\<in>A. 0 < x"
+lemma preal_nonempty: "cut A \<Longrightarrow> \<exists>x\<in>A. 0 < x"
   unfolding cut_def [abs_def] by blast
 
 lemma preal_Ex_mem: "cut A \<Longrightarrow> \<exists>x. x \<in> A"
-  apply (drule preal_nonempty)
-  apply fast
-  done
-
-lemma preal_imp_psubset_positives: "cut A ==> A < {r. 0 < r}"
-  by (force simp add: cut_def)
+  using preal_nonempty by blast
 
-lemma preal_exists_bound: "cut A ==> \<exists>x. 0 < x & x \<notin> A"
-  apply (drule preal_imp_psubset_positives)
-  apply auto
-  done
+lemma preal_exists_bound: "cut A \<Longrightarrow> \<exists>x. 0 < x \<and> x \<notin> A"
+  using Dedekind_Real.cut_def by fastforce
 
-lemma preal_exists_greater: "[| cut A; y \<in> A |] ==> \<exists>u \<in> A. y < u"
+lemma preal_exists_greater: "\<lbrakk>cut A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>u \<in> A. y < u"
   unfolding cut_def [abs_def] by blast
 
-lemma preal_downwards_closed: "[| cut A; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
+lemma preal_downwards_closed: "\<lbrakk>cut A; y \<in> A; 0 < z; z < y\<rbrakk> \<Longrightarrow> z \<in> A"
   unfolding cut_def [abs_def] by blast
 
 text\<open>Relaxing the final premise\<close>
-lemma preal_downwards_closed':
-     "[| cut A; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
-apply (simp add: order_le_less)
-apply (blast intro: preal_downwards_closed)
-done
+lemma preal_downwards_closed': "\<lbrakk>cut A; y \<in> A; 0 < z; z \<le> y\<rbrakk> \<Longrightarrow> z \<in> A"
+  using less_eq_rat_def preal_downwards_closed by blast
 
 text\<open>A positive fraction not in a positive real is an upper bound.
  Gleason p. 122 - Remark (1)\<close>
@@ -174,12 +157,12 @@
 
 lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
 thm preal_Ex_mem
-by (rule preal_Ex_mem [OF Rep_preal])
+by (rule preal_Ex_mem [OF cut_Rep_preal])
 
 lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X"
-by (rule preal_exists_bound [OF Rep_preal])
+by (rule preal_exists_bound [OF cut_Rep_preal])
 
-lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal]
+lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF cut_Rep_preal]
 
 
 subsection\<open>Properties of Ordering\<close>
@@ -199,21 +182,18 @@
 next
   fix z w :: preal
   show "z < w \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
-  by (auto simp add: preal_le_def preal_less_def Rep_preal_inject)
+  by (auto simp: preal_le_def preal_less_def Rep_preal_inject)
 qed  
 
-lemma preal_imp_pos: "[|cut A; r \<in> A|] ==> 0 < r"
-by (insert preal_imp_psubset_positives, blast)
+lemma preal_imp_pos: "\<lbrakk>cut A; r \<in> A\<rbrakk> \<Longrightarrow> 0 < r"
+  by (auto simp: cut_def)
 
 instance preal :: linorder
 proof
   fix x y :: preal
-  show "x <= y | y <= x"
-    apply (auto simp add: preal_le_def)
-    apply (rule ccontr)
-    apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal]
-             elim: order_less_asym)
-    done
+  show "x \<le> y \<or> y \<le> x"
+    unfolding preal_le_def
+    by (meson cut_Rep_preal not_in_preal_ub preal_downwards_closed preal_imp_pos subsetI)
 qed
 
 instantiation preal :: distrib_lattice
@@ -227,114 +207,85 @@
 
 instance
   by intro_classes
-    (auto simp add: inf_preal_def sup_preal_def max_min_distrib2)
+    (auto simp: inf_preal_def sup_preal_def max_min_distrib2)
 
 end
 
 subsection\<open>Properties of Addition\<close>
 
 lemma preal_add_commute: "(x::preal) + y = y + x"
-apply (unfold preal_add_def add_set_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (force simp add: add.commute)
-done
+  unfolding preal_add_def add_set_def
+  by (metis (no_types, hide_lams) add.commute)
 
 text\<open>Lemmas for proving that addition of two positive reals gives
  a positive real\<close>
 
-text\<open>Part 1 of Dedekind sections definition\<close>
-lemma add_set_not_empty:
-     "[|cut A; cut B|] ==> {} \<subset> add_set A B"
-apply (drule preal_nonempty)+
-apply (auto simp add: add_set_def)
-done
-
-text\<open>Part 2 of Dedekind sections definition.  A structured version of
-this proof is \<open>preal_not_mem_mult_set_Ex\<close> below.\<close>
-lemma preal_not_mem_add_set_Ex:
-     "[|cut A; cut B|] ==> \<exists>q>0. q \<notin> add_set A B"
-apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) 
-apply (rule_tac x = "x+xa" in exI)
-apply (simp add: add_set_def, clarify)
-apply (drule (3) not_in_preal_ub)+
-apply (force dest: add_strict_mono)
-done
-
-lemma add_set_not_rat_set:
-   assumes A: "cut A" 
-       and B: "cut B"
-     shows "add_set A B < {r. 0 < r}"
-proof
-  from preal_imp_pos [OF A] preal_imp_pos [OF B]
-  show "add_set A B \<subseteq> {r. 0 < r}" by (force simp add: add_set_def) 
-next
-  show "add_set A B \<noteq> {r. 0 < r}"
-    by (insert preal_not_mem_add_set_Ex [OF A B], blast) 
+lemma mem_add_set:
+  assumes "cut A" "cut B"
+  shows "cut (add_set A B)"
+proof -
+  have "{} \<subset> add_set A B"
+    using assms by (force simp: add_set_def dest: preal_nonempty)
+  moreover
+  obtain q where "q > 0" "q \<notin> add_set A B"
+  proof -
+    obtain a b where "a > 0" "a \<notin> A" "b > 0" "b \<notin> B" "\<And>x. x \<in> A \<Longrightarrow> x < a" "\<And>y. y \<in> B \<Longrightarrow> y < b"
+      by (meson assms preal_exists_bound not_in_preal_ub)
+    with assms have "a+b \<notin> add_set A B"
+      by (fastforce simp add: add_set_def)
+    then show thesis
+      using \<open>0 < a\<close> \<open>0 < b\<close> add_pos_pos that by blast
+  qed
+  then have "add_set A B \<subset> {0<..}"
+    unfolding add_set_def
+    using preal_imp_pos [OF \<open>cut A\<close>] preal_imp_pos [OF \<open>cut B\<close>]  by fastforce
+  moreover have "z \<in> add_set A B" 
+    if u: "u \<in> add_set A B" and "0 < z" "z < u" for u z
+    using u unfolding add_set_def
+  proof (clarify)
+    fix x::rat and y::rat
+    assume ueq: "u = x + y" and x: "x \<in> A" and y:"y \<in> B"
+    have xpos [simp]: "x > 0" and ypos [simp]: "y > 0"
+      using assms preal_imp_pos x y by blast+
+    have xypos [simp]: "x+y > 0" by (simp add: pos_add_strict)
+    let ?f = "z/(x+y)"
+    have fless: "?f < 1"
+      using divide_less_eq_1_pos \<open>z < u\<close> ueq xypos by blast
+    show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"
+    proof (intro bexI)
+      show "z = x*?f + y*?f"
+        by (simp add: distrib_right [symmetric] divide_inverse ac_simps order_less_imp_not_eq2)
+    next
+      show "y * ?f \<in> B"
+      proof (rule preal_downwards_closed [OF \<open>cut B\<close> y])
+        show "0 < y * ?f"
+          by (simp add: \<open>0 < z\<close>)
+      next
+        show "y * ?f < y"
+          by (insert mult_strict_left_mono [OF fless ypos], simp)
+      qed
+    next
+      show "x * ?f \<in> A"
+      proof (rule preal_downwards_closed [OF \<open>cut A\<close> x])
+        show "0 < x * ?f"
+          by (simp add: \<open>0 < z\<close>)
+      next
+        show "x * ?f < x"
+          by (insert mult_strict_left_mono [OF fless xpos], simp)
+      qed
+    qed
+  qed
+  moreover
+  have "\<And>y. y \<in> add_set A B \<Longrightarrow> \<exists>u \<in> add_set A B. y < u"
+    unfolding add_set_def using preal_exists_greater assms by fastforce
+  ultimately show ?thesis
+    by (simp add: Dedekind_Real.cut_def)
 qed
 
-text\<open>Part 3 of Dedekind sections definition\<close>
-lemma add_set_lemma3:
-     "[|cut A; cut B; u \<in> add_set A B; 0 < z; z < u|] 
-      ==> z \<in> add_set A B"
-proof (unfold add_set_def, clarify)
-  fix x::rat and y::rat
-  assume A: "cut A" 
-    and B: "cut B"
-    and [simp]: "0 < z"
-    and zless: "z < x + y"
-    and x:  "x \<in> A"
-    and y:  "y \<in> B"
-  have xpos [simp]: "0<x" by (rule preal_imp_pos [OF A x])
-  have ypos [simp]: "0<y" by (rule preal_imp_pos [OF B y])
-  have xypos [simp]: "0 < x+y" by (simp add: pos_add_strict)
-  let ?f = "z/(x+y)"
-  have fless: "?f < 1" by (simp add: zless pos_divide_less_eq)
-  show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"
-  proof (intro bexI)
-    show "z = x*?f + y*?f"
-      by (simp add: distrib_right [symmetric] divide_inverse ac_simps
-          order_less_imp_not_eq2)
-  next
-    show "y * ?f \<in> B"
-    proof (rule preal_downwards_closed [OF B y])
-      show "0 < y * ?f"
-        by (simp add: divide_inverse zero_less_mult_iff)
-    next
-      show "y * ?f < y"
-        by (insert mult_strict_left_mono [OF fless ypos], simp)
-    qed
-  next
-    show "x * ?f \<in> A"
-    proof (rule preal_downwards_closed [OF A x])
-      show "0 < x * ?f"
-        by (simp add: divide_inverse zero_less_mult_iff)
-    next
-      show "x * ?f < x"
-        by (insert mult_strict_left_mono [OF fless xpos], simp)
-    qed
-  qed
-qed
-
-text\<open>Part 4 of Dedekind sections definition\<close>
-lemma add_set_lemma4:
-     "[|cut A; cut B; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
-apply (auto simp add: add_set_def)
-apply (frule preal_exists_greater [of A], auto) 
-apply (rule_tac x="u + ya" in exI)
-apply (auto intro: add_strict_left_mono)
-done
-
-lemma mem_add_set:
-     "[|cut A; cut B|] ==> cut (add_set A B)"
-apply (simp (no_asm_simp) add: cut_def)
-apply (blast intro!: add_set_not_empty add_set_not_rat_set
-                     add_set_lemma3 add_set_lemma4)
-done
-
 lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"
-apply (simp add: preal_add_def mem_add_set Rep_preal)
-apply (force simp add: add_set_def ac_simps)
-done
+  apply (simp add: preal_add_def mem_add_set)
+  apply (force simp: add_set_def ac_simps)
+  done
 
 instance preal :: ab_semigroup_add
 proof
@@ -349,120 +300,80 @@
 text\<open>Proofs essentially same as for addition\<close>
 
 lemma preal_mult_commute: "(x::preal) * y = y * x"
-apply (unfold preal_mult_def mult_set_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (force simp add: mult.commute)
-done
+  unfolding preal_mult_def mult_set_def
+  by (metis (no_types, hide_lams) mult.commute)
 
 text\<open>Multiplication of two positive reals gives a positive real.\<close>
 
-text\<open>Lemmas for proving positive reals multiplication set in \<^typ>\<open>preal\<close>\<close>
-
-text\<open>Part 1 of Dedekind sections definition\<close>
-lemma mult_set_not_empty:
-     "[|cut A; cut B|] ==> {} \<subset> mult_set A B"
-apply (insert preal_nonempty [of A] preal_nonempty [of B]) 
-apply (auto simp add: mult_set_def)
-done
-
-text\<open>Part 2 of Dedekind sections definition\<close>
-lemma preal_not_mem_mult_set_Ex:
-  assumes A: "cut A" 
-    and B: "cut B"
-  shows "\<exists>q. 0 < q & q \<notin> mult_set A B"
+lemma mem_mult_set:
+  assumes "cut A" "cut B"
+  shows "cut (mult_set A B)"
 proof -
-  from preal_exists_bound [OF A] obtain x where 1 [simp]: "0 < x" "x \<notin> A" by blast
-  from preal_exists_bound [OF B] obtain y where 2 [simp]: "0 < y" "y \<notin> B" by blast
-  show ?thesis
-  proof (intro exI conjI)
-    show "0 < x*y" by simp
-    show "x * y \<notin> mult_set A B"
+  have "{} \<subset> mult_set A B"
+    using assms
+      by (force simp: mult_set_def dest: preal_nonempty)
+    moreover
+    obtain q where "q > 0" "q \<notin> mult_set A B"
     proof -
-      {
-        fix u::rat and v::rat
-        assume u: "u \<in> A" and v: "v \<in> B" and xy: "x*y = u*v"
-        moreover from A B 1 2 u v have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+
-        moreover
-        from A B 1 2 u v have "0\<le>v"
-          by (blast intro: preal_imp_pos [OF B] order_less_imp_le)
-        moreover
-        from A B 1 \<open>u < x\<close> \<open>v < y\<close> \<open>0 \<le> v\<close>
-        have "u*v < x*y" by (blast intro: mult_strict_mono)
-        ultimately have False by force
-      }
-      thus ?thesis by (auto simp add: mult_set_def)
+      obtain x y where x [simp]: "0 < x" "x \<notin> A" and y [simp]: "0 < y" "y \<notin> B"
+        using preal_exists_bound assms by blast
+      show thesis
+      proof 
+        show "0 < x*y" by simp
+        show "x * y \<notin> mult_set A B"
+        proof -
+          {
+            fix u::rat and v::rat
+            assume u: "u \<in> A" and v: "v \<in> B" and xy: "x*y = u*v"
+            moreover have "u<x" and "v<y" using assms x y u v by (blast dest: not_in_preal_ub)+
+            moreover have "0\<le>v"
+              using less_imp_le preal_imp_pos assms x y u v by blast
+            moreover have "u*v < x*y"
+                using assms x \<open>u < x\<close> \<open>v < y\<close> \<open>0 \<le> v\<close> by (blast intro: mult_strict_mono)
+            ultimately have False by force
+          }
+          thus ?thesis by (auto simp: mult_set_def)
+        qed
+      qed
+    qed
+  then have "mult_set A B \<subset> {0<..}"
+    unfolding mult_set_def
+    using preal_imp_pos [OF \<open>cut A\<close>] preal_imp_pos [OF \<open>cut B\<close>]  by fastforce
+  moreover have "z \<in> mult_set A B"
+    if u: "u \<in> mult_set A B" and "0 < z" "z < u" for u z
+    using u unfolding mult_set_def
+  proof (clarify)
+    fix x::rat and y::rat
+    assume ueq: "u = x * y" and x: "x \<in> A" and y: "y \<in> B"  
+    have [simp]: "y > 0"
+      using \<open>cut B\<close> preal_imp_pos y by blast
+    show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"
+    proof
+      have "z = (z/y)*y"
+          by (simp add: divide_inverse mult.commute [of y] mult.assoc order_less_imp_not_eq2)
+      then show "\<exists>y'\<in>B. z = (z/y) * y'"
+        using y by blast
+    next
+      show "z/y \<in> A"
+      proof (rule preal_downwards_closed [OF \<open>cut A\<close> x])
+        show "0 < z/y"
+          by (simp add: \<open>0 < z\<close>)
+        show "z/y < x"
+          using \<open>0 < y\<close> pos_divide_less_eq \<open>z < u\<close> ueq by blast  
+      qed
     qed
   qed
-qed
-
-lemma mult_set_not_rat_set:
-  assumes A: "cut A" 
-    and B: "cut B"
-  shows "mult_set A B < {r. 0 < r}"
-proof
-  show "mult_set A B \<subseteq> {r. 0 < r}"
-    by (force simp add: mult_set_def
-      intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos_pos)
-  show "mult_set A B \<noteq> {r. 0 < r}"
-    using preal_not_mem_mult_set_Ex [OF A B] by blast
+  moreover have "\<And>y. y \<in> mult_set A B \<Longrightarrow> \<exists>u \<in> mult_set A B. y < u"
+    apply (simp add: mult_set_def)
+    by (metis preal_exists_greater mult_strict_right_mono preal_imp_pos assms)
+  ultimately show ?thesis
+    by (simp add: Dedekind_Real.cut_def)
 qed
 
-
-
-text\<open>Part 3 of Dedekind sections definition\<close>
-lemma mult_set_lemma3:
-     "[|cut A; cut B; u \<in> mult_set A B; 0 < z; z < u|] 
-      ==> z \<in> mult_set A B"
-proof (unfold mult_set_def, clarify)
-  fix x::rat and y::rat
-  assume A: "cut A" 
-    and B: "cut B"
-    and [simp]: "0 < z"
-    and zless: "z < x * y"
-    and x:  "x \<in> A"
-    and y:  "y \<in> B"
-  have [simp]: "0<y" by (rule preal_imp_pos [OF B y])
-  show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"
-  proof
-    show "\<exists>y'\<in>B. z = (z/y) * y'"
-    proof
-      show "z = (z/y)*y"
-        by (simp add: divide_inverse mult.commute [of y] mult.assoc
-                      order_less_imp_not_eq2)
-      show "y \<in> B" by fact
-    qed
-  next
-    show "z/y \<in> A"
-    proof (rule preal_downwards_closed [OF A x])
-      show "0 < z/y"
-        by (simp add: zero_less_divide_iff)
-      show "z/y < x" by (simp add: pos_divide_less_eq zless)
-    qed
-  qed
-qed
-
-text\<open>Part 4 of Dedekind sections definition\<close>
-lemma mult_set_lemma4:
-     "[|cut A; cut B; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
-apply (auto simp add: mult_set_def)
-apply (frule preal_exists_greater [of A], auto) 
-apply (rule_tac x="u * ya" in exI)
-apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B] 
-                   mult_strict_right_mono)
-done
-
-
-lemma mem_mult_set:
-     "[|cut A; cut B|] ==> cut (mult_set A B)"
-apply (simp (no_asm_simp) add: cut_def)
-apply (blast intro!: mult_set_not_empty mult_set_not_rat_set
-                     mult_set_lemma3 mult_set_lemma4)
-done
-
 lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"
-apply (simp add: preal_mult_def mem_mult_set Rep_preal)
-apply (force simp add: mult_set_def ac_simps)
-done
+  apply (simp add: preal_mult_def mem_mult_set Rep_preal)
+  apply (force simp: mult_set_def ac_simps)
+  done
 
 instance preal :: ab_semigroup_mult
 proof
@@ -478,7 +389,7 @@
 proof (induct z)
   fix A :: "rat set"
   assume A: "cut A"
-  have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
+  have "{w. \<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
   proof
     show "?lhs \<subseteq> A"
     proof clarify
@@ -487,8 +398,7 @@
       have vpos: "0<v" by (rule preal_imp_pos [OF A v])
       hence "u*v < 1*v" by (simp only: mult_strict_right_mono upos \<open>u < 1\<close> v)
       thus "u * v \<in> A"
-        by (force intro: preal_downwards_closed [OF A v] mult_pos_pos 
-          upos vpos)
+        by (force intro: preal_downwards_closed [OF A v] mult_pos_pos  upos vpos)
     qed
   next
     show "A \<subseteq> ?lhs"
@@ -505,44 +415,38 @@
           by (simp add: zero_less_divide_iff xpos vpos)
         show "x / v < 1"
           by (simp add: pos_divide_less_eq vpos xlessv)
-        show "\<exists>v'\<in>A. x = (x / v) * v'"
-        proof
-          show "x = (x/v)*v"
-            by (simp add: divide_inverse mult.assoc vpos
-                          order_less_imp_not_eq2)
-          show "v \<in> A" by fact
-        qed
+        have "x = (x/v)*v"
+            by (simp add: divide_inverse mult.assoc vpos order_less_imp_not_eq2)
+        then show "\<exists>v'\<in>A. x = (x / v) * v'"
+          using v by blast
       qed
     qed
   qed
   thus "1 * Abs_preal A = Abs_preal A"
-    by (simp add: preal_one_def preal_mult_def mult_set_def 
-                  rat_mem_preal A)
+    by (simp add: preal_one_def preal_mult_def mult_set_def rat_mem_preal A)
 qed
 
 instance preal :: comm_monoid_mult
-by intro_classes (rule preal_mult_1)
+  by intro_classes (rule preal_mult_1)
 
 
 subsection\<open>Distribution of Multiplication across Addition\<close>
 
 lemma mem_Rep_preal_add_iff:
-      "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)"
-apply (simp add: preal_add_def mem_add_set Rep_preal)
-apply (simp add: add_set_def) 
-done
+  "(z \<in> Rep_preal(r+s)) = (\<exists>x \<in> Rep_preal r. \<exists>y \<in> Rep_preal s. z = x + y)"
+  apply (simp add: preal_add_def mem_add_set Rep_preal)
+  apply (simp add: add_set_def) 
+  done
 
 lemma mem_Rep_preal_mult_iff:
-      "(z \<in> Rep_preal(R*S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x * y)"
-apply (simp add: preal_mult_def mem_mult_set Rep_preal)
-apply (simp add: mult_set_def) 
-done
+  "(z \<in> Rep_preal(r*s)) = (\<exists>x \<in> Rep_preal r. \<exists>y \<in> Rep_preal s. z = x * y)"
+  apply (simp add: preal_mult_def mem_mult_set Rep_preal)
+  apply (simp add: mult_set_def) 
+  done
 
 lemma distrib_subset1:
-     "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"
-apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
-apply (force simp add: distrib_left)
-done
+  "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"
+  by (force simp: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff distrib_left)
 
 lemma preal_add_mult_distrib_mean:
   assumes a: "a \<in> Rep_preal w"
@@ -553,7 +457,7 @@
 proof
   let ?c = "(a*d + b*e)/(d+e)"
   have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e"
-    by (blast intro: preal_imp_pos [OF Rep_preal] a b d e pos_add_strict)+
+    by (blast intro: preal_imp_pos [OF cut_Rep_preal] a b d e pos_add_strict)+
   have cpos: "0 < ?c"
     by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict)
   show "a * d + b * e = ?c * (d + e)"
@@ -564,116 +468,83 @@
     hence "?c \<le> b"
       by (simp add: pos_divide_le_eq distrib_left mult_right_mono
                     order_less_imp_le)
-    thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos])
+    thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal b cpos])
   next
     assume "b \<le> a"
     hence "?c \<le> a"
       by (simp add: pos_divide_le_eq distrib_left mult_right_mono
                     order_less_imp_le)
-    thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos])
+    thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal a cpos])
   qed
 qed
 
 lemma distrib_subset2:
-     "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"
-apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
-apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto)
-done
+  "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"
+  apply (clarsimp simp: mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
+  using mem_Rep_preal_add_iff preal_add_mult_distrib_mean by blast
 
 lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"
-apply (rule Rep_preal_inject [THEN iffD1])
-apply (rule equalityI [OF distrib_subset1 distrib_subset2])
-done
+  by (metis Rep_preal_inverse distrib_subset1 distrib_subset2 subset_antisym)
 
 lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)"
-by (simp add: preal_mult_commute preal_add_mult_distrib2)
+  by (simp add: preal_mult_commute preal_add_mult_distrib2)
 
 instance preal :: comm_semiring
-by intro_classes (rule preal_add_mult_distrib)
+  by intro_classes (rule preal_add_mult_distrib)
 
 
 subsection\<open>Existence of Inverse, a Positive Real\<close>
 
-lemma mem_inv_set_ex:
-  assumes A: "cut A" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
+lemma mem_inverse_set:
+  assumes "cut A" shows "cut (inverse_set A)"
 proof -
-  from preal_exists_bound [OF A]
-  obtain x where [simp]: "0<x" "x \<notin> A" by blast
-  show ?thesis
-  proof (intro exI conjI)
-    show "0 < inverse (x+1)"
-      by (simp add: order_less_trans [OF _ less_add_one]) 
-    show "inverse(x+1) < inverse x"
-      by (simp add: less_imp_inverse_less less_add_one)
-    show "inverse (inverse x) \<notin> A"
-      by (simp add: order_less_imp_not_eq2)
-  qed
-qed
-
-text\<open>Part 1 of Dedekind sections definition\<close>
-lemma inverse_set_not_empty:
-     "cut A ==> {} \<subset> inverse_set A"
-apply (insert mem_inv_set_ex [of A])
-apply (auto simp add: inverse_set_def)
-done
-
-text\<open>Part 2 of Dedekind sections definition\<close>
-
-lemma preal_not_mem_inverse_set_Ex:
-   assumes A: "cut A"  shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
-proof -
-  from preal_nonempty [OF A]
-  obtain x where x: "x \<in> A" and  xpos [simp]: "0<x" ..
-  show ?thesis
-  proof (intro exI conjI)
-    show "0 < inverse x" by simp
-    show "inverse x \<notin> inverse_set A"
-    proof -
-      { fix y::rat 
-        assume ygt: "inverse x < y"
-        have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
-        have iyless: "inverse y < x" 
-          by (simp add: inverse_less_imp_less [of x] ygt)
-        have "inverse y \<in> A"
-          by (simp add: preal_downwards_closed [OF A x] iyless)}
-     thus ?thesis by (auto simp add: inverse_set_def)
+  have "\<exists>x y. 0 < x \<and> x < y \<and> inverse y \<notin> A"
+  proof -
+    from preal_exists_bound [OF \<open>cut A\<close>]
+    obtain x where [simp]: "0<x" "x \<notin> A" by blast
+    show ?thesis
+    proof (intro exI conjI)
+      show "0 < inverse (x+1)"
+        by (simp add: order_less_trans [OF _ less_add_one]) 
+      show "inverse(x+1) < inverse x"
+        by (simp add: less_imp_inverse_less less_add_one)
+      show "inverse (inverse x) \<notin> A"
+        by (simp add: order_less_imp_not_eq2)
     qed
   qed
-qed
-
-lemma inverse_set_not_rat_set:
-   assumes A: "cut A"  shows "inverse_set A < {r. 0 < r}"
-proof
-  show "inverse_set A \<subseteq> {r. 0 < r}"  by (force simp add: inverse_set_def)
-next
-  show "inverse_set A \<noteq> {r. 0 < r}"
-    by (insert preal_not_mem_inverse_set_Ex [OF A], blast)
+  then have "{} \<subset> inverse_set A"
+    using inverse_set_def by fastforce
+  moreover obtain q where "q > 0" "q \<notin> inverse_set A"
+  proof -
+    from preal_nonempty [OF \<open>cut A\<close>]
+    obtain x where x: "x \<in> A" and  xpos [simp]: "0<x" ..
+    show ?thesis
+    proof 
+      show "0 < inverse x" by simp
+      show "inverse x \<notin> inverse_set A"
+      proof -
+        { fix y::rat 
+          assume ygt: "inverse x < y"
+          have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
+          have iyless: "inverse y < x" 
+            by (simp add: inverse_less_imp_less [of x] ygt)
+          have "inverse y \<in> A"
+            by (simp add: preal_downwards_closed [OF \<open>cut A\<close> x] iyless)}
+        thus ?thesis by (auto simp: inverse_set_def)
+      qed
+    qed
+  qed
+  moreover have "inverse_set A \<subset> {0<..}"
+    using calculation inverse_set_def by blast
+  moreover have "z \<in> inverse_set A"
+    if u: "u \<in> inverse_set A" and "0 < z" "z < u" for u z
+    using u that less_trans unfolding inverse_set_def by auto
+  moreover have "\<And>y. y \<in> inverse_set A \<Longrightarrow> \<exists>u \<in> inverse_set A. y < u"
+    by (simp add: inverse_set_def) (meson dense less_trans)
+  ultimately show ?thesis
+    by (simp add: Dedekind_Real.cut_def)
 qed
 
-text\<open>Part 3 of Dedekind sections definition\<close>
-lemma inverse_set_lemma3:
-     "[|cut A; u \<in> inverse_set A; 0 < z; z < u|] 
-      ==> z \<in> inverse_set A"
-apply (auto simp add: inverse_set_def)
-apply (auto intro: order_less_trans)
-done
-
-text\<open>Part 4 of Dedekind sections definition\<close>
-lemma inverse_set_lemma4:
-     "[|cut A; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
-apply (auto simp add: inverse_set_def)
-apply (drule dense [of y]) 
-apply (blast intro: order_less_trans)
-done
-
-
-lemma mem_inverse_set:
-     "cut A ==> cut (inverse_set A)"
-apply (simp (no_asm_simp) add: cut_def)
-apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set
-                     inverse_set_lemma3 inverse_set_lemma4)
-done
-
 
 subsection\<open>Gleason's Lemma 9-3.4, page 122\<close>
 
@@ -693,7 +564,7 @@
     case (Suc k)
     then obtain b where b: "b \<in> A" "b + of_nat k * u \<in> A" ..
     hence "b + of_int (int k)*u + u \<in> A" by (simp add: assms)
-    thus ?case by (force simp add: algebra_simps b)
+    thus ?case by (force simp: algebra_simps b)
   qed
 next
   case (neg n)
@@ -702,7 +573,7 @@
 
 lemma Gleason9_34_contra:
   assumes A: "cut A"
-    shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False"
+    shows "\<lbrakk>\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A\<rbrakk> \<Longrightarrow> False"
 proof (induct u, induct y)
   fix a::int and b::int
   fix c::int and d::int
@@ -741,16 +612,9 @@
 
 
 lemma Gleason9_34:
-  assumes A: "cut A"
-    and upos: "0 < u"
+  assumes "cut A" "0 < u"
   shows "\<exists>r \<in> A. r + u \<notin> A"
-proof (rule ccontr, simp)
-  assume closed: "\<forall>r\<in>A. r + u \<in> A"
-  from preal_exists_bound [OF A]
-  obtain y where y: "y \<notin> A" and ypos: "0 < y" by blast
-  show False
-    by (rule Gleason9_34_contra [OF A closed upos ypos y])
-qed
+  using assms Gleason9_34_contra preal_exists_bound by blast
 
 
 
@@ -788,9 +652,9 @@
     have "r + ?d < r*x"
     proof -
       have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
-      also from ypos have "... = (r/y) * (y + ?d)"
+      also from ypos have "\<dots> = (r/y) * (y + ?d)"
         by (simp only: algebra_simps divide_inverse, simp)
-      also have "... = r*x" using ypos
+      also have "\<dots> = r*x" using ypos
         by simp
       finally show "r + ?d < r*x" .
     qed
@@ -802,11 +666,10 @@
 subsection\<open>Existence of Inverse: Part 2\<close>
 
 lemma mem_Rep_preal_inverse_iff:
-      "(z \<in> Rep_preal(inverse R)) = 
-       (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal R))"
-apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)
-apply (simp add: inverse_set_def) 
-done
+  "(z \<in> Rep_preal(inverse r)) \<longleftrightarrow> (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal r))"
+  apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)
+  apply (simp add: inverse_set_def) 
+  done
 
 lemma Rep_preal_one:
      "Rep_preal 1 = {x. 0 < x \<and> x < 1}"
@@ -814,24 +677,24 @@
 
 lemma subset_inverse_mult_lemma:
   assumes xpos: "0 < x" and xless: "x < 1"
-  shows "\<exists>r u y. 0 < r & r < y & inverse y \<notin> Rep_preal R & 
-    u \<in> Rep_preal R & x = r * u"
+  shows "\<exists>v u y. 0 < v \<and> v < y \<and> inverse y \<notin> Rep_preal R \<and> 
+    u \<in> Rep_preal R \<and> x = v * u"
 proof -
   from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)
-  from lemma_gleason9_36 [OF Rep_preal this]
-  obtain r where r: "r \<in> Rep_preal R" 
-             and notin: "r * (inverse x) \<notin> Rep_preal R" ..
-  have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
-  from preal_exists_greater [OF Rep_preal r]
-  obtain u where u: "u \<in> Rep_preal R" and rless: "r < u" ..
-  have upos: "0<u" by (rule preal_imp_pos [OF Rep_preal u])
+  from lemma_gleason9_36 [OF cut_Rep_preal this]
+  obtain t where t: "t \<in> Rep_preal R" 
+             and notin: "t * (inverse x) \<notin> Rep_preal R" ..
+  have rpos: "0<t" by (rule preal_imp_pos [OF cut_Rep_preal t])
+  from preal_exists_greater [OF cut_Rep_preal t]
+  obtain u where u: "u \<in> Rep_preal R" and rless: "t < u" ..
+  have upos: "0<u" by (rule preal_imp_pos [OF cut_Rep_preal u])
   show ?thesis
   proof (intro exI conjI)
     show "0 < x/u" using xpos upos
       by (simp add: zero_less_divide_iff)  
-    show "x/u < x/r" using xpos upos rpos
+    show "x/u < x/t" using xpos upos rpos
       by (simp add: divide_inverse mult_less_cancel_left rless) 
-    show "inverse (x / r) \<notin> Rep_preal R" using notin
+    show "inverse (x / t) \<notin> Rep_preal R" using notin
       by (simp add: divide_inverse mult.commute) 
     show "u \<in> Rep_preal R" by (rule u) 
     show "x = x / u * u" using upos 
@@ -840,263 +703,192 @@
 qed
 
 lemma subset_inverse_mult: 
-     "Rep_preal 1 \<subseteq> Rep_preal(inverse R * R)"
-apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff 
-                      mem_Rep_preal_mult_iff)
-apply (blast dest: subset_inverse_mult_lemma) 
-done
+     "Rep_preal 1 \<subseteq> Rep_preal(inverse r * r)"
+  by (force simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff dest: subset_inverse_mult_lemma)
 
-lemma inverse_mult_subset_lemma:
-  assumes rpos: "0 < r" 
-    and rless: "r < y"
-    and notin: "inverse y \<notin> Rep_preal R"
-    and q: "q \<in> Rep_preal R"
-  shows "r*q < 1"
-proof -
-  have "q < inverse y" using rpos rless
-    by (simp add: not_in_preal_ub [OF Rep_preal notin] q)
-  hence "r * q < r/y" using rpos
-    by (simp add: divide_inverse mult_less_cancel_left)
-  also have "... \<le> 1" using rpos rless
-    by (simp add: pos_divide_le_eq)
-  finally show ?thesis .
-qed
+lemma inverse_mult_subset: "Rep_preal(inverse r * r) \<subseteq> Rep_preal 1"
+  proof -
+  have "0 < u * v" if "v \<in> Rep_preal r" "0 < u" "u < t" for u v t :: rat
+    using that by (simp add: zero_less_mult_iff preal_imp_pos [OF cut_Rep_preal]) 
+  moreover have "t * q < 1"
+    if "q \<in> Rep_preal r" "0 < t" "t < y" "inverse y \<notin> Rep_preal r"
+    for t q y :: rat
+  proof -
+    have "q < inverse y"
+      using not_in_Rep_preal_ub that by auto 
+    hence "t * q < t/y" 
+      using that by (simp add: divide_inverse mult_less_cancel_left)
+    also have "\<dots> \<le> 1" 
+      using that by (simp add: pos_divide_le_eq)
+    finally show ?thesis .
+  qed
+  ultimately show ?thesis
+    by (auto simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff)
+qed 
 
-lemma inverse_mult_subset:
-     "Rep_preal(inverse R * R) \<subseteq> Rep_preal 1"
-apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff
-                      mem_Rep_preal_mult_iff)
-apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal]) 
-apply (blast intro: inverse_mult_subset_lemma) 
-done
+lemma preal_mult_inverse: "inverse r * r = (1::preal)"
+  by (meson Rep_preal_inject inverse_mult_subset subset_antisym subset_inverse_mult)
 
-lemma preal_mult_inverse: "inverse R * R = (1::preal)"
-apply (rule Rep_preal_inject [THEN iffD1])
-apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult]) 
-done
-
-lemma preal_mult_inverse_right: "R * inverse R = (1::preal)"
-apply (rule preal_mult_commute [THEN subst])
-apply (rule preal_mult_inverse)
-done
+lemma preal_mult_inverse_right: "r * inverse r = (1::preal)"
+  using preal_mult_commute preal_mult_inverse by auto
 
 
 text\<open>Theorems needing \<open>Gleason9_34\<close>\<close>
 
-lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)"
+lemma Rep_preal_self_subset: "Rep_preal (r) \<subseteq> Rep_preal(r + s)"
 proof 
-  fix r
-  assume r: "r \<in> Rep_preal R"
-  have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
-  from mem_Rep_preal_Ex 
-  obtain y where y: "y \<in> Rep_preal S" ..
-  have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
-  have ry: "r+y \<in> Rep_preal(R + S)" using r y
-    by (auto simp add: mem_Rep_preal_add_iff)
-  show "r \<in> Rep_preal(R + S)" using r ypos rpos 
-    by (simp add:  preal_downwards_closed [OF Rep_preal ry]) 
+  fix x
+  assume x: "x \<in> Rep_preal r"
+  obtain y where y: "y \<in> Rep_preal s" and "y > 0"
+    using Rep_preal preal_nonempty by blast
+  have ry: "x+y \<in> Rep_preal(r + s)" using x y
+    by (auto simp: mem_Rep_preal_add_iff)
+  then show "x \<in> Rep_preal(r + s)"
+    by (meson \<open>0 < y\<close> add_less_same_cancel1 not_in_Rep_preal_ub order.asym preal_imp_pos [OF cut_Rep_preal x])
 qed
 
-lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \<subseteq> Rep_preal(R)"
+lemma Rep_preal_sum_not_subset: "~ Rep_preal (r + s) \<subseteq> Rep_preal(r)"
 proof -
-  from mem_Rep_preal_Ex 
-  obtain y where y: "y \<in> Rep_preal S" ..
-  have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
-  from  Gleason9_34 [OF Rep_preal ypos]
-  obtain r where r: "r \<in> Rep_preal R" and notin: "r + y \<notin> Rep_preal R" ..
-  have "r + y \<in> Rep_preal (R + S)" using r y
-    by (auto simp add: mem_Rep_preal_add_iff)
+  obtain y where y: "y \<in> Rep_preal s" and "y > 0"
+    using Rep_preal preal_nonempty by blast
+  obtain x where "x \<in> Rep_preal r" and notin: "x + y \<notin> Rep_preal r"
+    using Dedekind_Real.Rep_preal Gleason9_34 \<open>0 < y\<close> by blast 
+  then have "x + y \<in> Rep_preal (r + s)" using y
+    by (auto simp: mem_Rep_preal_add_iff)
   thus ?thesis using notin by blast
 qed
 
-lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)"
-by (insert Rep_preal_sum_not_subset, blast)
-
 text\<open>at last, Gleason prop. 9-3.5(iii) page 123\<close>
-lemma preal_self_less_add_left: "(R::preal) < R + S"
-apply (unfold preal_less_def less_le)
-apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym])
-done
+proposition preal_self_less_add_left: "(r::preal) < r + s"
+  by (meson Rep_preal_sum_not_subset not_less preal_le_def)
 
 
 subsection\<open>Subtraction for Positive Reals\<close>
 
-text\<open>Gleason prop. 9-3.5(iv), page 123: proving \<^prop>\<open>A < B ==> \<exists>D. A + D =
-B\<close>. We define the claimed \<^term>\<open>D\<close> and show that it is a positive real\<close>
-
-text\<open>Part 1 of Dedekind sections definition\<close>
-lemma diff_set_not_empty:
-     "R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)"
-apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) 
-apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater])
-apply (drule preal_imp_pos [OF Rep_preal], clarify)
-apply (cut_tac a=x and b=u in add_eq_exists, force) 
-done
-
-text\<open>Part 2 of Dedekind sections definition\<close>
-lemma diff_set_nonempty:
-     "\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)"
-apply (cut_tac X = S in Rep_preal_exists_bound)
-apply (erule exE)
-apply (rule_tac x = x in exI, auto)
-apply (simp add: diff_set_def) 
-apply (auto dest: Rep_preal [THEN preal_downwards_closed])
-done
-
-lemma diff_set_not_rat_set:
-  "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs")
-proof
-  show "?lhs \<subseteq> ?rhs" by (auto simp add: diff_set_def) 
-  show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast
-qed
-
-text\<open>Part 3 of Dedekind sections definition\<close>
-lemma diff_set_lemma3:
-     "[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] 
-      ==> z \<in> diff_set (Rep_preal S) (Rep_preal R)"
-apply (auto simp add: diff_set_def) 
-apply (rule_tac x=x in exI) 
-apply (drule Rep_preal [THEN preal_downwards_closed], auto)
-done
-
-text\<open>Part 4 of Dedekind sections definition\<close>
-lemma diff_set_lemma4:
-     "[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|] 
-      ==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u"
-apply (auto simp add: diff_set_def) 
-apply (drule Rep_preal [THEN preal_exists_greater], clarify) 
-apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify)  
-apply (rule_tac x="y+xa" in exI) 
-apply (auto simp add: ac_simps)
-done
+text\<open>gleason prop. 9-3.5(iv), page 123: proving \<^prop>\<open>a < b \<Longrightarrow> \<exists>d. a + d = b\<close>. 
+We define the claimed \<^term>\<open>D\<close> and show that it is a positive real\<close>
 
 lemma mem_diff_set:
-     "R < S ==> cut (diff_set (Rep_preal S) (Rep_preal R))"
-apply (unfold cut_def)
-apply (blast intro!: diff_set_not_empty diff_set_not_rat_set
-                     diff_set_lemma3 diff_set_lemma4)
-done
-
-lemma mem_Rep_preal_diff_iff:
-      "R < S ==>
-       (z \<in> Rep_preal(S-R)) = 
-       (\<exists>x. 0 < x & 0 < z & x \<notin> Rep_preal R & x + z \<in> Rep_preal S)"
-apply (simp add: preal_diff_def mem_diff_set Rep_preal)
-apply (force simp add: diff_set_def) 
-done
-
-
-text\<open>proving that \<^term>\<open>R + D \<le> S\<close>\<close>
-
-lemma less_add_left_lemma:
-  assumes Rless: "R < S"
-    and a: "a \<in> Rep_preal R"
-    and cb: "c + b \<in> Rep_preal S"
-    and "c \<notin> Rep_preal R"
-    and "0 < b"
-    and "0 < c"
-  shows "a + b \<in> Rep_preal S"
+  assumes "r < s"
+  shows "cut (diff_set (Rep_preal s) (Rep_preal r))"
 proof -
-  have "0<a" by (rule preal_imp_pos [OF Rep_preal a])
+  obtain p where "Rep_preal r \<subseteq> Rep_preal s" "p \<in> Rep_preal s" "p \<notin> Rep_preal r"
+    using assms unfolding preal_less_def by auto
+  then have "{} \<subset> diff_set (Rep_preal s) (Rep_preal r)"
+    apply (simp add: diff_set_def psubset_eq)
+    by (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater preal_imp_pos)
   moreover
-  have "a < c" using assms by (blast intro: not_in_Rep_preal_ub ) 
+  obtain q where "q > 0" "q \<notin> Rep_preal s"
+    using Rep_preal_exists_bound by blast
+  then have qnot: "q \<notin> diff_set (Rep_preal s) (Rep_preal r)"
+    by (auto simp: diff_set_def dest: cut_Rep_preal [THEN preal_downwards_closed])
+  moreover have "diff_set (Rep_preal s) (Rep_preal r) \<subset> {0<..}" (is "?lhs < ?rhs")
+    using \<open>0 < q\<close> diff_set_def qnot by blast
+  moreover have "z \<in> diff_set (Rep_preal s) (Rep_preal r)"
+    if u: "u \<in> diff_set (Rep_preal s) (Rep_preal r)" and "0 < z" "z < u" for u z
+    using u that less_trans Rep_preal unfolding diff_set_def Dedekind_Real.cut_def by auto
+  moreover have "\<exists>u \<in> diff_set (Rep_preal s) (Rep_preal r). y < u"
+    if y: "y \<in> diff_set (Rep_preal s) (Rep_preal r)" for y
+  proof -
+    obtain a b where "0 < a" "0 < b" "a \<notin> Rep_preal r" "a + y + b \<in> Rep_preal s"
+      using y
+      by (simp add: diff_set_def) (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater) 
+    then have "a + (y + b) \<in> Rep_preal s"
+      by (simp add: add.assoc)
+    then have "y + b \<in> diff_set (Rep_preal s) (Rep_preal r)"
+      using \<open>0 < a\<close> \<open>0 < b\<close> \<open>a \<notin> Rep_preal r\<close> y
+      by (auto simp: diff_set_def)
+    then show ?thesis
+      using \<open>0 < b\<close> less_add_same_cancel1 by blast
+  qed
   ultimately show ?thesis
-    using assms by (simp add: preal_downwards_closed [OF Rep_preal cb])
+    by (simp add: Dedekind_Real.cut_def)
 qed
 
-lemma less_add_left_le1:
-       "R < (S::preal) ==> R + (S-R) \<le> S"
-apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff 
-                      mem_Rep_preal_diff_iff)
-apply (blast intro: less_add_left_lemma) 
-done
-
-subsection\<open>proving that \<^term>\<open>S \<le> R + D\<close> --- trickier\<close>
-
-lemma lemma_sum_mem_Rep_preal_ex:
-     "x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S"
-apply (drule Rep_preal [THEN preal_exists_greater], clarify) 
-apply (cut_tac a=x and b=u in add_eq_exists, auto) 
-done
+lemma mem_Rep_preal_diff_iff:
+  "r < s \<Longrightarrow>
+       (z \<in> Rep_preal (s - r)) \<longleftrightarrow> 
+       (\<exists>x. 0 < x \<and> 0 < z \<and> x \<notin> Rep_preal r \<and> x + z \<in> Rep_preal s)"
+  apply (simp add: preal_diff_def mem_diff_set Rep_preal)
+  apply (force simp: diff_set_def) 
+  done
 
-lemma less_add_left_lemma2:
-  assumes Rless: "R < S"
-    and x:     "x \<in> Rep_preal S"
-    and xnot: "x \<notin>  Rep_preal R"
-  shows "\<exists>u v z. 0 < v & 0 < z & u \<in> Rep_preal R & z \<notin> Rep_preal R & 
-                     z + v \<in> Rep_preal S & x = u + v"
+proposition less_add_left:
+  fixes r::preal 
+  assumes "r < s"
+  shows "r + (s-r) = s"
 proof -
-  have xpos: "0<x" by (rule preal_imp_pos [OF Rep_preal x])
-  from lemma_sum_mem_Rep_preal_ex [OF x]
-  obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal S" by blast
-  from  Gleason9_34 [OF Rep_preal epos]
-  obtain r where r: "r \<in> Rep_preal R" and notin: "r + e \<notin> Rep_preal R" ..
-  with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub)
-  from add_eq_exists [of r x]
-  obtain y where eq: "x = r+y" by auto
-  show ?thesis 
-  proof (intro exI conjI)
-    show "r \<in> Rep_preal R" by (rule r)
-    show "r + e \<notin> Rep_preal R" by (rule notin)
-    show "r + e + y \<in> Rep_preal S" using xe eq by (simp add: ac_simps)
-    show "x = r + y" by (simp add: eq)
-    show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r]
-      by simp
-    show "0 < y" using rless eq by arith
+  have "a + b \<in> Rep_preal s"
+    if "a \<in> Rep_preal r" "c + b \<in> Rep_preal s" "c \<notin> Rep_preal r"
+    and "0 < b" "0 < c" for a b c
+    by (meson cut_Rep_preal add_less_imp_less_right add_pos_pos not_in_Rep_preal_ub preal_downwards_closed preal_imp_pos that)
+  then have "r + (s-r) \<le> s"
+    using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff preal_le_def by auto
+  have "x \<in> Rep_preal (r + (s - r))" if "x \<in> Rep_preal s" for x
+  proof (cases "x \<in> Rep_preal r")
+    case True
+    then show ?thesis
+      using Rep_preal_self_subset by blast
+  next
+    case False
+    have "\<exists>u v z. 0 < v \<and> 0 < z \<and> u \<in> Rep_preal r \<and> z \<notin> Rep_preal r \<and> z + v \<in> Rep_preal s \<and> x = u + v"
+      if x: "x \<in> Rep_preal s"
+    proof -
+      have xpos: "x > 0"
+        using Rep_preal preal_imp_pos that by blast 
+      obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal s"
+        by (metis cut_Rep_preal x add_eq_exists less_add_same_cancel1 preal_exists_greater)
+      from  Gleason9_34 [OF cut_Rep_preal epos]
+      obtain u where r: "u \<in> Rep_preal r" and notin: "u + e \<notin> Rep_preal r" ..
+      with x False xpos have rless: "u < x" by (blast intro: not_in_Rep_preal_ub)
+      from add_eq_exists [of u x]
+      obtain y where eq: "x = u+y" by auto
+      show ?thesis 
+      proof (intro exI conjI)
+        show "u + e \<notin> Rep_preal r" by (rule notin)
+        show "u + e + y \<in> Rep_preal s" using xe eq by (simp add: ac_simps)
+        show "0 < u + e" 
+          using epos preal_imp_pos [OF cut_Rep_preal r] by simp
+      qed (use r rless eq in auto)
+    qed
+    then show ?thesis
+      using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff that by blast
   qed
+  then have "s \<le> r + (s-r)"
+    by (auto simp: preal_le_def)
+  then show ?thesis
+    by (simp add: \<open>r + (s - r) \<le> s\<close> antisym)
 qed
 
-lemma less_add_left_le2: "R < (S::preal) ==> S \<le> R + (S-R)"
-apply (auto simp add: preal_le_def)
-apply (case_tac "x \<in> Rep_preal R")
-apply (cut_tac Rep_preal_self_subset [of R], force)
-apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff)
-apply (blast dest: less_add_left_lemma2)
-done
+lemma preal_add_less2_mono1: "r < (s::preal) \<Longrightarrow> r + t < s + t"
+  by (metis add.assoc add.commute less_add_left preal_self_less_add_left)
 
-lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S"
-by (blast intro: antisym [OF less_add_left_le1 less_add_left_le2])
+lemma preal_add_less2_mono2: "r < (s::preal) \<Longrightarrow> t + r < t + s"
+  by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of t])
 
-lemma less_add_left_Ex: "R < (S::preal) ==> \<exists>D. R + D = S"
-by (fast dest: less_add_left)
+lemma preal_add_right_less_cancel: "r + t < s + t \<Longrightarrow> r < (s::preal)"
+  by (metis linorder_cases order.asym preal_add_less2_mono1)
 
-lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T"
-apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc)
-apply (rule_tac y1 = D in preal_add_commute [THEN subst])
-apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric])
-done
+lemma preal_add_left_less_cancel: "t + r < t + s \<Longrightarrow> r < (s::preal)"
+  by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of t])
 
-lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S"
-by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T])
+lemma preal_add_less_cancel_left [simp]: "(t + (r::preal) < t + s) \<longleftrightarrow> (r < s)"
+  by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
 
-lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)"
-apply (insert linorder_less_linear [of R S], auto)
-apply (drule_tac R = S and T = T in preal_add_less2_mono1)
-apply (blast dest: order_less_trans) 
-done
+lemma preal_add_less_cancel_right [simp]: "((r::preal) + t < s + t) = (r < s)"
+  using preal_add_less_cancel_left [symmetric, of r s t] by (simp add: ac_simps)
 
-lemma preal_add_left_less_cancel: "T + R < T + S ==> R <  (S::preal)"
-by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T])
-
-lemma preal_add_less_cancel_left [simp]: "(T + (R::preal) < T + S) = (R < S)"
-by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
+lemma preal_add_le_cancel_left [simp]: "(t + (r::preal) \<le> t + s) = (r \<le> s)"
+  by (simp add: linorder_not_less [symmetric]) 
 
-lemma preal_add_less_cancel_right [simp]: "((R::preal) + T < S + T) = (R < S)"
-  using preal_add_less_cancel_left [symmetric, of R S T] by (simp add: ac_simps)
-
-lemma preal_add_le_cancel_left [simp]: "(T + (R::preal) \<le> T + S) = (R \<le> S)"
-by (simp add: linorder_not_less [symmetric]) 
+lemma preal_add_le_cancel_right [simp]: "((r::preal) + t \<le> s + t) = (r \<le> s)"
+  using preal_add_le_cancel_left [symmetric, of r s t] by (simp add: ac_simps)
 
-lemma preal_add_le_cancel_right [simp]: "((R::preal) + T \<le> S + T) = (R \<le> S)"
-  using preal_add_le_cancel_left [symmetric, of R S T] by (simp add: ac_simps)
+lemma preal_add_right_cancel: "(r::preal) + t = s + t \<Longrightarrow> r = s"
+  by (metis less_irrefl linorder_cases preal_add_less_cancel_right)
 
-lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S"
-apply (insert linorder_less_linear [of R S], safe)
-apply (drule_tac [!] T = T in preal_add_less2_mono1, auto)
-done
-
-lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)"
-by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
+lemma preal_add_left_cancel: "c + a = c + b \<Longrightarrow> a = (b::preal)"
+  by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
 
 instance preal :: linordered_ab_semigroup_add
 proof
@@ -1111,84 +903,65 @@
 
 text\<open>Part 1 of Dedekind sections definition\<close>
 
-lemma preal_sup_set_not_empty:
-     "P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
-apply auto
-apply (cut_tac X = x in mem_Rep_preal_Ex, auto)
-done
-
-
-text\<open>Part 2 of Dedekind sections definition\<close>
-
-lemma preal_sup_not_exists:
-     "\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
-apply (cut_tac X = Y in Rep_preal_exists_bound)
-apply (auto simp add: preal_le_def)
-done
-
-lemma preal_sup_set_not_rat_set:
-     "\<forall>X \<in> P. X \<le> Y ==> (\<Union>X \<in> P. Rep_preal(X)) < {r. 0 < r}"
-apply (drule preal_sup_not_exists)
-apply (blast intro: preal_imp_pos [OF Rep_preal])  
-done
-
-text\<open>Part 3 of Dedekind sections definition\<close>
-lemma preal_sup_set_lemma3:
-     "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|]
-      ==> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
-by (auto elim: Rep_preal [THEN preal_downwards_closed])
-
-text\<open>Part 4 of Dedekind sections definition\<close>
-lemma preal_sup_set_lemma4:
-     "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |]
-          ==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
-by (blast dest: Rep_preal [THEN preal_exists_greater])
-
 lemma preal_sup:
-     "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> cut (\<Union>X \<in> P. Rep_preal(X))"
-apply (unfold cut_def)
-apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set
-                     preal_sup_set_lemma3 preal_sup_set_lemma4)
-done
+  assumes le: "\<And>X. X \<in> P \<Longrightarrow> X \<le> Y" and "P \<noteq> {}" 
+  shows "cut (\<Union>X \<in> P. Rep_preal(X))"
+proof -
+  have "{} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
+    using \<open>P \<noteq> {}\<close> mem_Rep_preal_Ex by fastforce
+  moreover
+  obtain q where "q > 0" and "q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
+    using Rep_preal_exists_bound [of Y] le by (auto simp: preal_le_def)
+  then have "(\<Union>X \<in> P. Rep_preal(X)) \<subset> {0<..}"
+    using cut_Rep_preal preal_imp_pos by force
+  moreover
+  have "\<And>u z. \<lbrakk>u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u\<rbrakk> \<Longrightarrow> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
+    by (auto elim: cut_Rep_preal [THEN preal_downwards_closed])
+  moreover
+  have "\<And>y. y \<in> (\<Union>X \<in> P. Rep_preal(X)) \<Longrightarrow> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
+    by (blast dest: cut_Rep_preal [THEN preal_exists_greater])
+  ultimately show ?thesis
+    by (simp add: Dedekind_Real.cut_def)
+qed
 
 lemma preal_psup_le:
-     "[| \<forall>X \<in> P. X \<le> Y;  x \<in> P |] ==> x \<le> psup P"
-apply (simp (no_asm_simp) add: preal_le_def) 
-apply (subgoal_tac "P \<noteq> {}") 
-apply (auto simp add: psup_def preal_sup) 
-done
+     "\<lbrakk>\<And>X. X \<in> P \<Longrightarrow> X \<le> Y;  x \<in> P\<rbrakk> \<Longrightarrow> x \<le> psup P"
+  using preal_sup [of P Y] unfolding preal_le_def psup_def by fastforce 
 
-lemma psup_le_ub: "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y"
-apply (simp (no_asm_simp) add: preal_le_def)
-apply (simp add: psup_def preal_sup) 
-apply (auto simp add: preal_le_def)
-done
+lemma psup_le_ub: "\<lbrakk>\<And>X. X \<in> P \<Longrightarrow> X \<le> Y; P \<noteq> {}\<rbrakk> \<Longrightarrow> psup P \<le> Y"
+  using preal_sup [of P Y] by (simp add: SUP_least preal_le_def psup_def) 
 
 text\<open>Supremum property\<close>
-lemma preal_complete:
-     "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)"
-apply (simp add: preal_less_def psup_def preal_sup)
-apply (auto simp add: preal_le_def)
-apply (rename_tac U) 
-apply (cut_tac x = U and y = Z in linorder_less_linear)
-apply (auto simp add: preal_less_def)
-done
+proposition preal_complete:
+  assumes le: "\<And>X. X \<in> P \<Longrightarrow> X \<le> Y" and "P \<noteq> {}" 
+  shows "(\<exists>X \<in> P. Z < X) \<longleftrightarrow> (Z < psup P)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    using preal_sup [OF assms] preal_less_def psup_def by auto
+next
+  assume ?rhs
+  then show ?lhs
+    by (meson \<open>P \<noteq> {}\<close> not_less psup_le_ub) 
+qed
 
-section \<open>Defining the Reals from the Positive Reals\<close>
+subsection \<open>Defining the Reals from the Positive Reals\<close>
+
+text \<open>Here we do quotients the old-fashioned way\<close>
 
 definition
   realrel   ::  "((preal * preal) * (preal * preal)) set" where
-  "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
+  "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) \<and> x1+y2 = x2+y1}"
 
 definition "Real = UNIV//realrel"
 
 typedef real = Real
   morphisms Rep_Real Abs_Real
-  unfolding Real_def by (auto simp add: quotient_def)
+  unfolding Real_def by (auto simp: quotient_def)
 
+text \<open>This doesn't involve the overloaded "real" function: users don't see it\<close>
 definition
-  (** these don't use the overloaded "real" function: users don't see them **)
-  real_of_preal :: "preal => real" where
+  real_of_preal :: "preal \<Rightarrow> real" where
   "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})"
 
 instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
@@ -1218,14 +991,14 @@
                  { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
 
 definition
-  real_inverse_def: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)"
+  real_inverse_def: "inverse (r::real) = (THE s. (r = 0 \<and> s = 0) \<or> s * r = 1)"
 
 definition
-  real_divide_def: "R div (S::real) = R * inverse S"
+  real_divide_def: "r div (s::real) = r * inverse s"
 
 definition
   real_le_def: "z \<le> (w::real) \<longleftrightarrow>
-    (\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w)"
+    (\<exists>x y u v. x+v \<le> u+y \<and> (x,y) \<in> Rep_Real z \<and> (u,v) \<in> Rep_Real w)"
 
 definition
   real_less_def: "x < (y::real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
@@ -1242,81 +1015,51 @@
 
 subsection \<open>Equivalence relation over positive reals\<close>
 
+lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"
+  by (simp add: realrel_def)
+
 lemma preal_trans_lemma:
-  assumes "x + y1 = x1 + y"
-    and "x + y2 = x2 + y"
+  assumes "x + y1 = x1 + y" and "x + y2 = x2 + y"
   shows "x1 + y2 = x2 + (y1::preal)"
-proof -
-  have "(x1 + y2) + x = (x + y2) + x1" by (simp add: ac_simps)
-  also have "... = (x2 + y) + x1"  by (simp add: assms)
-  also have "... = x2 + (x1 + y)"  by (simp add: ac_simps)
-  also have "... = x2 + (x + y1)"  by (simp add: assms)
-  also have "... = (x2 + y1) + x"  by (simp add: ac_simps)
-  finally have "(x1 + y2) + x = (x2 + y1) + x" .
-  thus ?thesis by (rule preal_add_right_cancel)
-qed
-
-
-lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"
-by (simp add: realrel_def)
+  by (metis add.left_commute assms preal_add_left_cancel)
 
 lemma equiv_realrel: "equiv UNIV realrel"
-apply (auto simp add: equiv_def refl_on_def sym_def trans_def realrel_def)
-apply (blast dest: preal_trans_lemma) 
-done
+  by (auto simp: equiv_def refl_on_def sym_def trans_def realrel_def intro: dest: preal_trans_lemma)
 
 text\<open>Reduces equality of equivalence classes to the \<^term>\<open>realrel\<close> relation:
   \<^term>\<open>(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)\<close>\<close>
-lemmas equiv_realrel_iff = 
+lemmas equiv_realrel_iff [simp] = 
        eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
 
-declare equiv_realrel_iff [simp]
-
+lemma realrel_in_real [simp]: "realrel``{(x,y)} \<in> Real"
+  by (simp add: Real_def realrel_def quotient_def, blast)
 
-lemma realrel_in_real [simp]: "realrel``{(x,y)} \<in> Real"
-by (simp add: Real_def realrel_def quotient_def, blast)
-
-declare Abs_Real_inject [simp]
-declare Abs_Real_inverse [simp]
+declare Abs_Real_inject [simp] Abs_Real_inverse [simp]
 
 
 text\<open>Case analysis on the representation of a real number as an equivalence
       class of pairs of positive reals.\<close>
 lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: 
      "(\<And>x y. z = Abs_Real(realrel``{(x,y)}) \<Longrightarrow> P) \<Longrightarrow> P"
-apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE])
-apply (drule arg_cong [where f=Abs_Real])
-apply (auto simp add: Rep_Real_inverse)
-done
-
+  by (metis Rep_Real_inverse prod.exhaust  Rep_Real [of z, unfolded Real_def, THEN quotientE])
 
 subsection \<open>Addition and Subtraction\<close>
 
-lemma real_add_congruent2_lemma:
-     "[|a + ba = aa + b; ab + bc = ac + bb|]
-      ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
-apply (simp add: add.assoc)
-apply (rule add.left_commute [of ab, THEN ssubst])
-apply (simp add: add.assoc [symmetric])
-apply (simp add: ac_simps)
-done
-
 lemma real_add:
      "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
       Abs_Real (realrel``{(x+u, y+v)})"
 proof -
   have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
         respects2 realrel"
-    by (auto simp add: congruent2_def, blast intro: real_add_congruent2_lemma) 
+  by (clarsimp simp: congruent2_def) (metis add.left_commute preal_add_assoc)
   thus ?thesis
-    by (simp add: real_add_def UN_UN_split_split_eq
-                  UN_equiv_class2 [OF equiv_realrel equiv_realrel])
+    by (simp add: real_add_def UN_UN_split_split_eq UN_equiv_class2 [OF equiv_realrel equiv_realrel])
 qed
 
 lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
 proof -
   have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
-    by (auto simp add: congruent_def add.commute) 
+    by (auto simp: congruent_def add.commute) 
   thus ?thesis
     by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])
 qed
@@ -1340,47 +1083,36 @@
 subsection \<open>Multiplication\<close>
 
 lemma real_mult_congruent2_lemma:
-     "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
+     "!!(x1::preal). \<lbrakk>x1 + y2 = x2 + y1\<rbrakk> \<Longrightarrow>
           x * x1 + y * y1 + (x * y2 + y * x2) =
           x * x2 + y * y2 + (x * y1 + y * x1)"
-apply (simp add: add.left_commute add.assoc [symmetric])
-apply (simp add: add.assoc distrib_left [symmetric])
-apply (simp add: add.commute)
-done
+  by (metis (no_types, hide_lams) add.left_commute preal_add_commute preal_add_mult_distrib2)
 
 lemma real_mult_congruent2:
-    "(%p1 p2.
-        (%(x1,y1). (%(x2,y2). 
+  "(\<lambda>p1 p2.
+        (\<lambda>(x1,y1). (\<lambda>(x2,y2). 
           { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
      respects2 realrel"
-apply (rule congruent2_commuteI [OF equiv_realrel], clarify)
-apply (simp add: mult.commute add.commute)
-apply (auto simp add: real_mult_congruent2_lemma)
-done
+  apply (rule congruent2_commuteI [OF equiv_realrel])
+  by (auto simp: mult.commute add.commute combine_common_factor preal_add_assoc preal_add_commute)
 
 lemma real_mult:
-      "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
-       Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
-by (simp add: real_mult_def UN_UN_split_split_eq
-         UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
+  "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
+   Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
+  by (simp add: real_mult_def UN_UN_split_split_eq
+      UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
 
 lemma real_mult_commute: "(z::real) * w = w * z"
 by (cases z, cases w, simp add: real_mult ac_simps)
 
 lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
-apply (cases z1, cases z2, cases z3)
-apply (simp add: real_mult algebra_simps)
-done
+  by (cases z1, cases z2, cases z3) (simp add: real_mult algebra_simps)
 
 lemma real_mult_1: "(1::real) * z = z"
-apply (cases z)
-apply (simp add: real_mult real_one_def algebra_simps)
-done
+  by (cases z) (simp add: real_mult real_one_def algebra_simps)
 
 lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
-apply (cases z1, cases z2, cases w)
-apply (simp add: real_add real_mult algebra_simps)
-done
+  by (cases z1, cases z2, cases w) (simp add: real_add real_mult algebra_simps)
 
 text\<open>one and zero are distinct\<close>
 lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
@@ -1404,32 +1136,50 @@
 subsection \<open>Inverse and Division\<close>
 
 lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
-by (simp add: real_zero_def add.commute)
-
-text\<open>Instead of using an existential quantifier and constructing the inverse
-within the proof, we could define the inverse explicitly.\<close>
+  by (simp add: real_zero_def add.commute)
 
-lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
-apply (simp add: real_zero_def real_one_def, cases x)
-apply (cut_tac x = xa and y = y in linorder_less_linear)
-apply (auto dest!: less_add_left_Ex simp add: real_zero_iff)
-apply (rule_tac
-        x = "Abs_Real (realrel``{(1, inverse (D) + 1)})"
-       in exI)
-apply (rule_tac [2]
-        x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})" 
-       in exI)
-apply (auto simp add: real_mult preal_mult_inverse_right algebra_simps)
-done
+lemma real_mult_inverse_left_ex:
+  assumes "x \<noteq> 0" obtains y::real where "y*x = 1"
+proof (cases x)
+  case (Abs_Real u v)
+  show ?thesis
+  proof (cases u v rule: linorder_cases)
+    case less
+    then have "v * inverse (v - u) = 1 + u * inverse (v - u)"
+      using less_add_left [of u v]
+      by (metis preal_add_commute preal_add_mult_distrib preal_mult_inverse_right)
+    then have "Abs_Real (realrel``{(1, inverse (v-u) + 1)}) * x - 1 = 0"
+      by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)
+    with that show thesis by auto
+  next
+    case equal
+    then show ?thesis
+      using Abs_Real assms real_zero_iff by blast
+  next
+    case greater
+    then have "u * inverse (u - v) = 1 + v * inverse (u - v)"
+      using less_add_left [of v u] by (metis add.commute distrib_right preal_mult_inverse_right)
+    then have "Abs_Real (realrel``{(inverse (u-v) + 1, 1)}) * x - 1 = 0"
+      by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)
+    with that show thesis by auto
+  qed
+qed
 
-lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
-apply (simp add: real_inverse_def)
-apply (drule real_mult_inverse_left_ex, safe)
-apply (rule theI, assumption, rename_tac z)
-apply (subgoal_tac "(z * x) * y = z * (x * y)")
-apply (simp add: mult.commute)
-apply (rule mult.assoc)
-done
+
+lemma real_mult_inverse_left:
+  fixes x :: real
+  assumes "x \<noteq> 0" shows "inverse x * x = 1"
+proof -
+  obtain y where "y*x = 1"
+    using assms real_mult_inverse_left_ex by blast
+  then have "(THE s. s * x = 1) * x = 1"
+  proof (rule theI)
+    show "y' = y" if "y' * x = 1" for y'
+      by (metis \<open>y * x = 1\<close> mult.left_commute mult.right_neutral that) 
+  qed
+  then show ?thesis
+    using assms real_inverse_def by auto
+qed
 
 
 subsection\<open>The Real Numbers form a Field\<close>
@@ -1437,7 +1187,7 @@
 instance real :: field
 proof
   fix x y z :: real
-  show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left)
+  show "x \<noteq> 0 \<Longrightarrow> inverse x * x = 1" by (rule real_mult_inverse_left)
   show "x / y = x * inverse y" by (simp add: real_divide_def)
   show "inverse 0 = (0::real)" by (simp add: real_inverse_def)
 qed
@@ -1446,7 +1196,7 @@
 subsection\<open>The \<open>\<le>\<close> Ordering\<close>
 
 lemma real_le_refl: "w \<le> (w::real)"
-by (cases w, force simp add: real_le_def)
+  by (cases w, force simp: real_le_def)
 
 text\<open>The arithmetic decision procedure is not set up for type preal.
   This lemma is currently unused, but it could simplify the proofs of the
@@ -1468,19 +1218,16 @@
 proof -
   have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: assms)
   hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: ac_simps)
-  also have "... \<le> (x2+y1) + (u2+v1)" by (simp add: assms)
+  also have "\<dots> \<le> (x2+y1) + (u2+v1)" by (simp add: assms)
   finally show ?thesis by simp
 qed
 
 lemma real_le: 
-     "(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) =  
-      (x1 + y2 \<le> x2 + y1)"
-apply (simp add: real_le_def)
-apply (auto intro: real_le_lemma)
-done
+  "Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})  \<longleftrightarrow>  x1 + y2 \<le> x2 + y1"
+  unfolding real_le_def by (auto intro: real_le_lemma)
 
-lemma real_le_antisym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
-by (cases z, cases w, simp add: real_le)
+lemma real_le_antisym: "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = (w::real)"
+  by (cases z, cases w, simp add: real_le)
 
 lemma real_trans_lemma:
   assumes "x + v \<le> u + y"
@@ -1489,73 +1236,27 @@
   shows "x + v' \<le> u' + (y::preal)"
 proof -
   have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: ac_simps)
-  also have "... \<le> (u+y) + (u+v')" by (simp add: assms)
-  also have "... \<le> (u+y) + (u'+v)" by (simp add: assms)
-  also have "... = (u'+y) + (u+v)"  by (simp add: ac_simps)
+  also have "\<dots> \<le> (u+y) + (u+v')" by (simp add: assms)
+  also have "\<dots> \<le> (u+y) + (u'+v)" by (simp add: assms)
+  also have "\<dots> = (u'+y) + (u+v)"  by (simp add: ac_simps)
   finally show ?thesis by simp
 qed
 
-lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
-apply (cases i, cases j, cases k)
-apply (simp add: real_le)
-apply (blast intro: real_trans_lemma)
-done
+lemma real_le_trans: "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> (k::real)"
+  by (cases i, cases j, cases k) (auto simp: real_le intro: real_trans_lemma)
 
 instance real :: order
 proof
-  fix u v :: real
-  show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u" 
-    by (auto simp add: real_less_def intro: real_le_antisym)
-qed (assumption | rule real_le_refl real_le_trans real_le_antisym)+
-
-(* Axiom 'linorder_linear' of class 'linorder': *)
-lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
-apply (cases z, cases w)
-apply (auto simp add: real_le real_zero_def ac_simps)
-done
+  show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u" for u v::real
+    by (auto simp: real_less_def intro: real_le_antisym)
+qed (auto intro: real_le_refl real_le_trans real_le_antisym)
 
 instance real :: linorder
-  by (intro_classes, rule real_le_linear)
-
-lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
-apply (cases x, cases y) 
-apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
-                      ac_simps)
-apply (simp_all add: add.assoc [symmetric])
-done
-
-lemma real_add_left_mono: 
-  assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
-proof -
-  have "z + x - (z + y) = (z + -z) + (x - y)" 
-    by (simp add: algebra_simps) 
-  with le show ?thesis 
-    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])
+proof
+  show "x \<le> y \<or> y \<le> x" for x y :: real
+    by (meson eq_refl le_cases real_le_def)
 qed
 
-lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
-
-lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
-
-lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
-apply (cases x, cases y)
-apply (simp add: linorder_not_le [where 'a = real, symmetric] 
-                 linorder_not_le [where 'a = preal] 
-                  real_zero_def real_le real_mult)
-  \<comment> \<open>Reduce to the (simpler) \<open>\<le>\<close> relation\<close>
-apply (auto dest!: less_add_left_Ex
-     simp add: algebra_simps preal_self_less_add_left)
-done
-
-lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
-apply (rule real_sum_gt_zero_less)
-apply (drule real_less_sum_gt_zero [of x y])
-apply (drule real_mult_order, assumption)
-apply (simp add: algebra_simps)
-done
-
 instantiation real :: distrib_lattice
 begin
 
@@ -1566,104 +1267,125 @@
   "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
 
 instance
-  by standard (auto simp add: inf_real_def sup_real_def max_min_distrib2)
+  by standard (auto simp: inf_real_def sup_real_def max_min_distrib2)
 
 end
 
+subsection\<open>The Reals Form an Ordered Field\<close>
 
-subsection\<open>The Reals Form an Ordered Field\<close>
+lemma real_le_eq_diff: "(x \<le> y) \<longleftrightarrow> (x-y \<le> (0::real))"
+  by (cases x, cases y) (simp add: real_le real_zero_def real_diff_def real_add real_minus preal_add_commute)
+
+lemma real_add_left_mono: 
+  assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
+proof -
+  have "z + x - (z + y) = (z + -z) + (x - y)" 
+    by (simp add: algebra_simps) 
+  with le show ?thesis 
+    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])
+qed
+
+lemma real_sum_gt_zero_less: "(0 < s + (-w::real)) \<Longrightarrow> (w < s)"
+  by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of s])
+
+lemma real_less_sum_gt_zero: "(w < s) \<Longrightarrow> (0 < s + (-w::real))"
+  by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of s])
+
+lemma real_mult_order: 
+  fixes x y::real
+  assumes "0 < x" "0 < y"
+  shows "0 < x * y"
+  proof (cases x, cases y)
+  show "0 < x * y"
+    if x: "x = Abs_Real (Dedekind_Real.realrel `` {(x1, x2)})"
+      and y: "y = Abs_Real (Dedekind_Real.realrel `` {(y1, y2)})"
+    for x1 x2 y1 y2
+  proof -
+    have "x2 < x1" "y2 < y1"
+      using assms not_le real_zero_def real_le x y
+      by (metis preal_add_le_cancel_left real_zero_iff)+
+    then obtain xd yd where "x1 = x2 + xd" "y1 = y2 + yd"
+      using less_add_left by metis
+    then have "\<not> (x * y \<le> 0)"
+      apply (simp add: x y real_mult real_zero_def real_le)
+      apply (simp add: not_le algebra_simps preal_self_less_add_left)
+      done
+    then show ?thesis
+      by auto
+  qed
+qed
+
+lemma real_mult_less_mono2: "\<lbrakk>(0::real) < z; x < y\<rbrakk> \<Longrightarrow> z * x < z * y"
+  by (metis add_uminus_conv_diff real_less_sum_gt_zero real_mult_order real_sum_gt_zero_less right_diff_distrib')
+
 
 instance real :: linordered_field
 proof
   fix x y z :: real
-  show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono)
-  show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2)
+  show "x \<le> y \<Longrightarrow> z + x \<le> z + y" by (rule real_add_left_mono)
   show "\<bar>x\<bar> = (if x < 0 then -x else x)" by (simp only: real_abs_def)
   show "sgn x = (if x=0 then 0 else if 0<x then 1 else - 1)"
     by (simp only: real_sgn_def)
+  show "z * x < z * y" if "x < y" "0 < z"
+    by (simp add: real_mult_less_mono2 that)
 qed
 
+
+subsection \<open>Completeness of the reals\<close>
+
 text\<open>The function \<^term>\<open>real_of_preal\<close> requires many proofs, but it seems
 to be essential for proving completeness of the reals from that of the
 positive reals.\<close>
 
 lemma real_of_preal_add:
-     "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
-by (simp add: real_of_preal_def real_add algebra_simps)
+  "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
+  by (simp add: real_of_preal_def real_add algebra_simps)
 
 lemma real_of_preal_mult:
-     "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
-by (simp add: real_of_preal_def real_mult algebra_simps)
-
+  "real_of_preal ((x::preal) * y) = real_of_preal x * real_of_preal y"
+  by (simp add: real_of_preal_def real_mult algebra_simps)
 
 text\<open>Gleason prop 9-4.4 p 127\<close>
 lemma real_of_preal_trichotomy:
-      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
-apply (simp add: real_of_preal_def real_zero_def, cases x)
-apply (auto simp add: real_minus ac_simps)
-apply (cut_tac x = xa and y = y in linorder_less_linear)
-apply (auto dest!: less_add_left_Ex simp add: add.assoc [symmetric])
-done
-
-lemma real_of_preal_leD:
-      "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2"
-by (simp add: real_of_preal_def real_le)
-
-lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
-by (auto simp add: real_of_preal_leD linorder_not_le [symmetric])
-
-lemma real_of_preal_lessD:
-      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
-by (simp add: real_of_preal_def real_le linorder_not_le [symmetric])
+  "\<exists>m. (x::real) = real_of_preal m \<or> x = 0 \<or> x = -(real_of_preal m)"
+proof (cases x)
+  case (Abs_Real u v)
+  show ?thesis
+  proof (cases u v rule: linorder_cases)
+    case less
+    then show ?thesis
+      using less_add_left
+      apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)
+      by (metis preal_add_assoc preal_add_commute)      
+  next
+    case equal
+    then show ?thesis
+      using Abs_Real real_zero_iff by blast
+  next
+    case greater
+    then show ?thesis
+      using less_add_left
+      apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)
+      by (metis preal_add_assoc preal_add_commute)      
+  qed
+qed
 
 lemma real_of_preal_less_iff [simp]:
-     "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
-by (blast intro: real_of_preal_lessI real_of_preal_lessD)
-
-lemma real_of_preal_le_iff:
-     "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
-by (simp add: linorder_not_less [symmetric])
+  "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
+  by (metis not_less preal_add_less_cancel_right real_le real_of_preal_def)
 
-lemma real_of_preal_zero_less: "0 < real_of_preal m"
-using preal_self_less_add_left [of 1 m]
-apply (auto simp add: real_zero_def real_of_preal_def real_less_def real_le_def ac_simps neq_iff)
-apply (metis Rep_preal_self_subset add.commute preal_le_def)
-done
+lemma real_of_preal_le_iff [simp]:
+  "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
+  by (simp add: linorder_not_less [symmetric])
 
-lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
-by (simp add: real_of_preal_zero_less)
-
-lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
-proof -
-  from real_of_preal_minus_less_zero
-  show ?thesis by (blast dest: order_less_trans)
-qed
+lemma real_of_preal_zero_less [simp]: "0 < real_of_preal m"
+  by (metis less_add_same_cancel2 preal_self_less_add_left real_of_preal_add real_of_preal_less_iff)
 
 
 subsection\<open>Theorems About the Ordering\<close>
 
-lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
-apply (auto simp add: real_of_preal_zero_less)
-apply (cut_tac x = x in real_of_preal_trichotomy)
-apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE])
-done
-
-lemma real_gt_preal_preal_Ex:
-     "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
-by (blast dest!: real_of_preal_zero_less [THEN order_less_trans]
-             intro: real_gt_zero_preal_Ex [THEN iffD1])
-
-lemma real_ge_preal_preal_Ex:
-     "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
-by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
-
-lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
-by (auto elim: order_le_imp_less_or_eq [THEN disjE] 
-            intro: real_of_preal_zero_less [THEN [2] order_less_trans] 
-            simp add: real_of_preal_zero_less)
-
-lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
-by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
+lemma real_gt_zero_preal_Ex: "(0 < x) \<longleftrightarrow> (\<exists>y. x = real_of_preal y)"
+  using order.asym real_of_preal_trichotomy by fastforce
 
 subsection \<open>Completeness of Positive Reals\<close>
 
@@ -1681,7 +1403,7 @@
   assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
     and not_empty_P: "\<exists>x. x \<in> P"
     and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
-  shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
+  shows "\<exists>s. \<forall>y. (\<exists>x \<in> P. y < x) = (y < s)"
 proof (rule exI, rule allI)
   fix y
   let ?pP = "{w. real_of_preal w \<in> P}"
@@ -1692,26 +1414,23 @@
     show ?thesis
     proof
       assume "\<exists>x\<in>P. y < x"
-      have "\<forall>x. y < real_of_preal x"
-        using neg_y by (rule real_less_all_real2)
-      thus "y < real_of_preal (psup ?pP)" ..
+      thus "y < real_of_preal (psup ?pP)"
+        by (metis dual_order.strict_trans neg_y not_less_iff_gr_or_eq real_of_preal_zero_less) 
     next
       assume "y < real_of_preal (psup ?pP)"
       obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
-      hence "0 < x" using positive_P by simp
-      hence "y < x" using neg_y by simp
-      thus "\<exists>x \<in> P. y < x" using x_in_P ..
+      thus "\<exists>x \<in> P. y < x" using x_in_P
+        using neg_y not_less_iff_gr_or_eq positive_P by fastforce 
     qed
   next
     assume pos_y: "0 < y"
-
     then obtain py where y_is_py: "y = real_of_preal py"
-      by (auto simp add: real_gt_zero_preal_Ex)
+      by (auto simp: real_gt_zero_preal_Ex)
 
     obtain a where "a \<in> P" using not_empty_P ..
     with positive_P have a_pos: "0 < a" ..
     then obtain pa where "a = real_of_preal pa"
-      by (auto simp add: real_gt_zero_preal_Ex)
+      by (auto simp: real_gt_zero_preal_Ex)
     hence "pa \<in> ?pP" using \<open>a \<in> P\<close> by auto
     hence pP_not_empty: "?pP \<noteq> {}" by auto
 
@@ -1720,54 +1439,51 @@
     from this and \<open>a \<in> P\<close> have "a < sup" ..
     hence "0 < sup" using a_pos by arith
     then obtain possup where "sup = real_of_preal possup"
-      by (auto simp add: real_gt_zero_preal_Ex)
+      by (auto simp: real_gt_zero_preal_Ex)
     hence "\<forall>X \<in> ?pP. X \<le> possup"
-      using sup by (auto simp add: real_of_preal_lessI)
+      using sup by auto
     with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"
-      by (rule preal_complete)
-
+      by (meson preal_complete)
     show ?thesis
     proof
       assume "\<exists>x \<in> P. y < x"
       then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..
       hence "0 < x" using pos_y by arith
       then obtain px where x_is_px: "x = real_of_preal px"
-        by (auto simp add: real_gt_zero_preal_Ex)
+        by (auto simp: real_gt_zero_preal_Ex)
 
       have py_less_X: "\<exists>X \<in> ?pP. py < X"
       proof
         show "py < px" using y_is_py and x_is_px and y_less_x
-          by (simp add: real_of_preal_lessI)
+          by simp
         show "px \<in> ?pP" using x_in_P and x_is_px by simp
       qed
 
-      have "(\<exists>X \<in> ?pP. py < X) ==> (py < psup ?pP)"
+      have "(\<exists>X \<in> ?pP. py < X) \<Longrightarrow> (py < psup ?pP)"
         using psup by simp
       hence "py < psup ?pP" using py_less_X by simp
       thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
-        using y_is_py and pos_y by (simp add: real_of_preal_lessI)
+        using y_is_py and pos_y by simp
     next
       assume y_less_psup: "y < real_of_preal (psup ?pP)"
 
       hence "py < psup ?pP" using y_is_py
-        by (simp add: real_of_preal_lessI)
+        by simp
       then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"
         using psup by auto
       then obtain x where x_is_X: "x = real_of_preal X"
         by (simp add: real_gt_zero_preal_Ex)
       hence "y < x" using py_less_X and y_is_py
-        by (simp add: real_of_preal_lessI)
-
-      moreover have "x \<in> P" using x_is_X and X_in_pP by simp
-
+        by simp
+      moreover have "x \<in> P" 
+        using x_is_X and X_in_pP by simp
       ultimately show "\<exists> x \<in> P. y < x" ..
     qed
   qed
 qed
 
-text \<open>
-  \medskip Completeness
-\<close>
+
+subsection \<open>Completeness\<close>
 
 lemma reals_complete:
   fixes S :: "real set"
@@ -1867,7 +1583,7 @@
   shows "\<exists>n. inverse (of_nat (Suc n)) < x"
 proof (rule ccontr)
   assume contr: "\<not> ?thesis"
-  have "\<forall>n. x * of_nat (Suc n) <= 1"
+  have "\<forall>n. x * of_nat (Suc n) \<le> 1"
   proof
     fix n
     from contr have "x \<le> inverse (of_nat (Suc n))"
@@ -1888,7 +1604,6 @@
     least: "\<And>y. (\<And>a. a \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> a \<le> y) \<Longrightarrow> t \<le> y"
     using reals_complete[OF 1 2] by auto
 
-
   have "t \<le> t + - x"
   proof (rule least)
     fix a assume a: "a \<in> {z. \<exists>n. z = x * (of_nat (Suc n))}"
--- a/src/HOL/ex/Tarski.thy	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/HOL/ex/Tarski.thy	Mon Apr 29 16:50:34 2019 +0100
@@ -444,34 +444,34 @@
   qed
 qed
 
-lemma flubH_le_lubH: 
+lemma lubH_is_fixp: 
   assumes "H = {x \<in> A. (x, f x) \<in> r}"
-  shows "(f (lub H cl), lub H cl) \<in> r"
+  shows "lub H cl \<in> fix f A"
 proof -
-  have "(lub H cl, f (lub H cl)) \<in> r"
-    using assms lubH_le_flubH by blast
-  then have "(f (lub H cl), f (f (lub H cl))) \<in> r"
-    by (meson PO_imp_refl_on monotoneE monotone_f refl_on_domain)
-  then have "f (lub H cl) \<in> H"
-    by (metis (no_types, lifting) PO_imp_refl_on assms mem_Collect_eq refl_on_domain)
-  then show ?thesis
-    by (simp add: assms lub_upper)
+  have "(f (lub H cl), lub H cl) \<in> r"
+  proof -
+    have "(lub H cl, f (lub H cl)) \<in> r"
+      using assms lubH_le_flubH by blast
+    then have "(f (lub H cl), f (f (lub H cl))) \<in> r"
+      by (meson PO_imp_refl_on monotoneE monotone_f refl_on_domain)
+    then have "f (lub H cl) \<in> H"
+      by (metis (no_types, lifting) PO_imp_refl_on assms mem_Collect_eq refl_on_domain)
+    then show ?thesis
+      by (simp add: assms lub_upper)
+  qed
+  with assms show ?thesis
+    by (simp add: fix_def antisymE lubH_le_flubH lub_in_lattice)
 qed
 
-
-
-lemma lubH_is_fixp: "H = {x \<in> A. (x, f x) \<in> r} \<Longrightarrow> lub H cl \<in> fix f A"
-  by (simp add: fix_def antisymE flubH_le_lubH lubH_le_flubH lub_in_lattice)
-
-lemma fix_in_H: "\<lbrakk>H = {x \<in> A. (x, f x) \<in> r}; x \<in> P\<rbrakk> \<Longrightarrow> x \<in> H"
-  by (simp add: P_def fix_imp_eq [of _ f A] reflE fix_subset [of f A, THEN subsetD])
-
-lemma fixf_le_lubH: "H = {x \<in> A. (x, f x) \<in> r} \<Longrightarrow> \<forall>x \<in> fix f A. (x, lub H cl) \<in> r"
-  by (metis (no_types, lifting) P_def fix_in_H lub_upper mem_Collect_eq subset_eq)
-
-lemma lubH_least_fixf:
-  "H = {x \<in> A. (x, f x) \<in> r} \<Longrightarrow> \<forall>L. (\<forall>y \<in> fix f A. (y,L) \<in> r) \<longrightarrow> (lub H cl, L) \<in> r"
-  using lubH_is_fixp by blast
+lemma fixf_le_lubH: 
+  assumes "H = {x \<in> A. (x, f x) \<in> r}" "x \<in> fix f A"
+  shows "(x, lub H cl) \<in> r"
+proof -
+  have "x \<in> P \<Longrightarrow> x \<in> H"
+    by (simp add: assms P_def fix_imp_eq [of _ f A] reflE fix_subset [of f A, THEN subsetD])
+  with assms show ?thesis
+    by (metis (no_types, lifting) P_def lub_upper mem_Collect_eq subset_eq)
+qed
 
 
 subsection \<open>Tarski fixpoint theorem 1, first part\<close>
@@ -623,9 +623,8 @@
   using Bot_dual_Top Bot_prop intervalI reflE by fastforce
 
 
-subsection \<open>fixed points form a partial order\<close>
-
-lemma fixf_po: "\<lparr>pset = P, order = induced P r\<rparr> \<in> PartialOrder"
+text \<open>the set of fixed points form a partial order\<close>
+proposition fixf_po: "\<lparr>pset = P, order = induced P r\<rparr> \<in> PartialOrder"
   by (simp add: P_def fix_subset po_subset_po)
 
 end
@@ -723,11 +722,9 @@
 end
 
 lemma CompleteLatticeI_simp:
-  "\<lbrakk>po \<in> PartialOrder; \<And>S. S \<subseteq> A \<Longrightarrow> \<exists>L. isLub S po  L; A = pset po\<rbrakk>
-    \<Longrightarrow> po \<in> CompleteLattice"
+  "\<lbrakk>po \<in> PartialOrder; \<And>S. S \<subseteq> A \<Longrightarrow> \<exists>L. isLub S po  L; A = pset po\<rbrakk> \<Longrightarrow> po \<in> CompleteLattice"
   by (metis CompleteLatticeI Rdual)
 
-
 theorem (in CLF) Tarski_full: "\<lparr>pset = P, order = induced P r\<rparr> \<in> CompleteLattice"
 proof (intro CompleteLatticeI_simp allI impI)
   show "\<lparr>pset = P, order = induced P r\<rparr> \<in> PartialOrder"
--- a/src/Pure/Admin/build_release.scala	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/Pure/Admin/build_release.scala	Mon Apr 29 16:50:34 2019 +0100
@@ -38,8 +38,8 @@
 
     def bundle_info(platform: Platform.Family.Value): Bundle_Info =
       platform match {
-        case Platform.Family.linux => Bundle_Info(platform, "Linux", dist_name + "_linux.tar.gz")
-        case Platform.Family.macos => Bundle_Info(platform, "Mac OS X", dist_name + "_macos.tar.gz")
+        case Platform.Family.linux => Bundle_Info(platform, "Linux", dist_name + "_linux.tar.xz")
+        case Platform.Family.macos => Bundle_Info(platform, "macOS", dist_name + "_macos.tar.xz")
         case Platform.Family.windows => Bundle_Info(platform, "Windows", dist_name + ".exe")
       }
   }
@@ -522,11 +522,11 @@
               isabelle_target + Path.explode(isabelle_name))
             Isabelle_System.rm_tree(linux_app)
 
-            val archive_name = isabelle_name + "_linux.tar.gz"
+            val archive_name = isabelle_name + "_linux.tar.xz"
             progress.echo("Packaging " + archive_name + " ...")
             execute_tar(tmp_dir,
-              "-czf " + File.bash_path(release.dist_dir + Path.explode(archive_name)) + " " +
-              Bash.string(isabelle_name))
+              "-cf- " + Bash.string(isabelle_name) +
+              " | xz > " + File.bash_path(release.dist_dir + Path.explode(archive_name)))
 
 
           case Platform.Family.macos =>
@@ -582,11 +582,11 @@
 
             // application archive
 
-            val archive_name = isabelle_name + "_macos.tar.gz"
+            val archive_name = isabelle_name + "_macos.tar.xz"
             progress.echo("Packaging " + archive_name + " ...")
             execute_tar(tmp_dir,
-              "-czf " + File.bash_path(release.dist_dir + Path.explode(archive_name)) + " " +
-              File.bash_path(isabelle_app))
+              "-cf- " + File.bash_path(isabelle_app) +
+              " | xz > " + File.bash_path(release.dist_dir + Path.explode(archive_name)))
 
 
           case Platform.Family.windows =>
--- a/src/Pure/Isar/parse.ML	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/Pure/Isar/parse.ML	Mon Apr 29 16:50:34 2019 +0100
@@ -67,6 +67,7 @@
   val name_position: (string * Position.T) parser
   val binding: binding parser
   val embedded: string parser
+  val embedded_input: Input.source parser
   val embedded_position: (string * Position.T) parser
   val text: string parser
   val path: string parser
@@ -278,7 +279,8 @@
     (cartouche || string || short_ident || long_ident || sym_ident ||
       term_var || type_ident || type_var || number);
 
-val embedded_position = input embedded >> Input.source_content;
+val embedded_input = input embedded;
+val embedded_position = embedded_input >> Input.source_content;
 
 val text = group (fn () => "text") (embedded || verbatim);
 
--- a/src/Pure/Pure.thy	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/Pure/Pure.thy	Mon Apr 29 16:50:34 2019 +0100
@@ -125,7 +125,7 @@
   val _ =
     Outer_Syntax.local_theory \<^command_keyword>\<open>generate_file\<close>
       "generate source file, with antiquotations"
-      (Parse.path_binding -- (\<^keyword>\<open>=\<close> |-- Parse.input Parse.embedded)
+      (Parse.path_binding -- (\<^keyword>\<open>=\<close> |-- Parse.embedded_input)
         >> Generated_Files.generate_file_cmd);
 
 
--- a/src/Pure/Thy/document_marker.ML	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/Pure/Thy/document_marker.ML	Mon Apr 29 16:50:34 2019 +0100
@@ -103,7 +103,7 @@
      setup0 (Binding.make ("contributor", \<^here>)) Parse.embedded (meta_data Markup.meta_contributor) #>
      setup0 (Binding.make ("date", \<^here>)) Parse.embedded (meta_data Markup.meta_date) #>
      setup0 (Binding.make ("license", \<^here>)) Parse.embedded (meta_data Markup.meta_license) #>
-     setup0 (Binding.make ("description", \<^here>)) (Parse.input Parse.embedded)
+     setup0 (Binding.make ("description", \<^here>)) Parse.embedded_input
       (fn source => fn ctxt =>
         let
           val (arg, pos) = Input.source_content source;
--- a/src/Tools/jEdit/src/isabelle_vfs.scala	Mon Apr 29 16:50:20 2019 +0100
+++ b/src/Tools/jEdit/src/isabelle_vfs.scala	Mon Apr 29 16:50:34 2019 +0100
@@ -73,7 +73,7 @@
   override def _getFile(vfs_session: AnyRef, url: String, component: Component): VFSFile =
   {
     val parent = getParentOfPath(url)
-    if (parent == prefix) null
+    if (parent == prefix) new VFSFile(prefix, prefix, prefix, VFSFile.DIRECTORY, 0L, false)
     else {
       val files = _listFiles(vfs_session, parent, component)
       if (files == null) null