src/HOL/Real.thy
author hoelzl
Tue, 05 Nov 2013 09:45:02 +0100
changeset 55715 c4159fe6fa46
parent 55714 326fd7103cb4
child 55733 b01057e72233
permissions -rw-r--r--
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
     1 (*  Title:      HOL/Real.thy
     2     Author:     Jacques D. Fleuriot, University of Edinburgh, 1998
     3     Author:     Larry Paulson, University of Cambridge
     4     Author:     Jeremy Avigad, Carnegie Mellon University
     5     Author:     Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
     6     Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
     7     Construction of Cauchy Reals by Brian Huffman, 2010
     8 *)
     9 
    10 header {* Development of the Reals using Cauchy Sequences *}
    11 
    12 theory Real
    13 imports Rat Conditionally_Complete_Lattices
    14 begin
    15 
    16 text {*
    17   This theory contains a formalization of the real numbers as
    18   equivalence classes of Cauchy sequences of rationals.  See
    19   @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
    20   construction using Dedekind cuts.
    21 *}
    22 
    23 subsection {* Preliminary lemmas *}
    24 
    25 lemma add_diff_add:
    26   fixes a b c d :: "'a::ab_group_add"
    27   shows "(a + c) - (b + d) = (a - b) + (c - d)"
    28   by simp
    29 
    30 lemma minus_diff_minus:
    31   fixes a b :: "'a::ab_group_add"
    32   shows "- a - - b = - (a - b)"
    33   by simp
    34 
    35 lemma mult_diff_mult:
    36   fixes x y a b :: "'a::ring"
    37   shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
    38   by (simp add: algebra_simps)
    39 
    40 lemma inverse_diff_inverse:
    41   fixes a b :: "'a::division_ring"
    42   assumes "a \<noteq> 0" and "b \<noteq> 0"
    43   shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
    44   using assms by (simp add: algebra_simps)
    45 
    46 lemma obtain_pos_sum:
    47   fixes r :: rat assumes r: "0 < r"
    48   obtains s t where "0 < s" and "0 < t" and "r = s + t"
    49 proof
    50     from r show "0 < r/2" by simp
    51     from r show "0 < r/2" by simp
    52     show "r = r/2 + r/2" by simp
    53 qed
    54 
    55 subsection {* Sequences that converge to zero *}
    56 
    57 definition
    58   vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
    59 where
    60   "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
    61 
    62 lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
    63   unfolding vanishes_def by simp
    64 
    65 lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
    66   unfolding vanishes_def by simp
    67 
    68 lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
    69   unfolding vanishes_def
    70   apply (cases "c = 0", auto)
    71   apply (rule exI [where x="\<bar>c\<bar>"], auto)
    72   done
    73 
    74 lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
    75   unfolding vanishes_def by simp
    76 
    77 lemma vanishes_add:
    78   assumes X: "vanishes X" and Y: "vanishes Y"
    79   shows "vanishes (\<lambda>n. X n + Y n)"
    80 proof (rule vanishesI)
    81   fix r :: rat assume "0 < r"
    82   then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
    83     by (rule obtain_pos_sum)
    84   obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
    85     using vanishesD [OF X s] ..
    86   obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
    87     using vanishesD [OF Y t] ..
    88   have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
    89   proof (clarsimp)
    90     fix n assume n: "i \<le> n" "j \<le> n"
    91     have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
    92     also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
    93     finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
    94   qed
    95   thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
    96 qed
    97 
    98 lemma vanishes_diff:
    99   assumes X: "vanishes X" and Y: "vanishes Y"
   100   shows "vanishes (\<lambda>n. X n - Y n)"
   101   unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus X Y)
   102 
   103 lemma vanishes_mult_bounded:
   104   assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
   105   assumes Y: "vanishes (\<lambda>n. Y n)"
   106   shows "vanishes (\<lambda>n. X n * Y n)"
   107 proof (rule vanishesI)
   108   fix r :: rat assume r: "0 < r"
   109   obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
   110     using X by fast
   111   obtain b where b: "0 < b" "r = a * b"
   112   proof
   113     show "0 < r / a" using r a by (simp add: divide_pos_pos)
   114     show "r = a * (r / a)" using a by simp
   115   qed
   116   obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
   117     using vanishesD [OF Y b(1)] ..
   118   have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
   119     by (simp add: b(2) abs_mult mult_strict_mono' a k)
   120   thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
   121 qed
   122 
   123 subsection {* Cauchy sequences *}
   124 
   125 definition
   126   cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
   127 where
   128   "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
   129 
   130 lemma cauchyI:
   131   "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r) \<Longrightarrow> cauchy X"
   132   unfolding cauchy_def by simp
   133 
   134 lemma cauchyD:
   135   "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
   136   unfolding cauchy_def by simp
   137 
   138 lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
   139   unfolding cauchy_def by simp
   140 
   141 lemma cauchy_add [simp]:
   142   assumes X: "cauchy X" and Y: "cauchy Y"
   143   shows "cauchy (\<lambda>n. X n + Y n)"
   144 proof (rule cauchyI)
   145   fix r :: rat assume "0 < r"
   146   then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
   147     by (rule obtain_pos_sum)
   148   obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
   149     using cauchyD [OF X s] ..
   150   obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
   151     using cauchyD [OF Y t] ..
   152   have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
   153   proof (clarsimp)
   154     fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
   155     have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
   156       unfolding add_diff_add by (rule abs_triangle_ineq)
   157     also have "\<dots> < s + t"
   158       by (rule add_strict_mono, simp_all add: i j *)
   159     finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
   160   qed
   161   thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
   162 qed
   163 
   164 lemma cauchy_minus [simp]:
   165   assumes X: "cauchy X"
   166   shows "cauchy (\<lambda>n. - X n)"
   167 using assms unfolding cauchy_def
   168 unfolding minus_diff_minus abs_minus_cancel .
   169 
   170 lemma cauchy_diff [simp]:
   171   assumes X: "cauchy X" and Y: "cauchy Y"
   172   shows "cauchy (\<lambda>n. X n - Y n)"
   173   using assms unfolding diff_conv_add_uminus by (simp del: add_uminus_conv_diff)
   174 
   175 lemma cauchy_imp_bounded:
   176   assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
   177 proof -
   178   obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
   179     using cauchyD [OF assms zero_less_one] ..
   180   show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
   181   proof (intro exI conjI allI)
   182     have "0 \<le> \<bar>X 0\<bar>" by simp
   183     also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
   184     finally have "0 \<le> Max (abs ` X ` {..k})" .
   185     thus "0 < Max (abs ` X ` {..k}) + 1" by simp
   186   next
   187     fix n :: nat
   188     show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
   189     proof (rule linorder_le_cases)
   190       assume "n \<le> k"
   191       hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
   192       thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
   193     next
   194       assume "k \<le> n"
   195       have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
   196       also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
   197         by (rule abs_triangle_ineq)
   198       also have "\<dots> < Max (abs ` X ` {..k}) + 1"
   199         by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
   200       finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
   201     qed
   202   qed
   203 qed
   204 
   205 lemma cauchy_mult [simp]:
   206   assumes X: "cauchy X" and Y: "cauchy Y"
   207   shows "cauchy (\<lambda>n. X n * Y n)"
   208 proof (rule cauchyI)
   209   fix r :: rat assume "0 < r"
   210   then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
   211     by (rule obtain_pos_sum)
   212   obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
   213     using cauchy_imp_bounded [OF X] by fast
   214   obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
   215     using cauchy_imp_bounded [OF Y] by fast
   216   obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
   217   proof
   218     show "0 < v/b" using v b(1) by (rule divide_pos_pos)
   219     show "0 < u/a" using u a(1) by (rule divide_pos_pos)
   220     show "r = a * (u/a) + (v/b) * b"
   221       using a(1) b(1) `r = u + v` by simp
   222   qed
   223   obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
   224     using cauchyD [OF X s] ..
   225   obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
   226     using cauchyD [OF Y t] ..
   227   have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
   228   proof (clarsimp)
   229     fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
   230     have "\<bar>X m * Y m - X n * Y n\<bar> = \<bar>X m * (Y m - Y n) + (X m - X n) * Y n\<bar>"
   231       unfolding mult_diff_mult ..
   232     also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
   233       by (rule abs_triangle_ineq)
   234     also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
   235       unfolding abs_mult ..
   236     also have "\<dots> < a * t + s * b"
   237       by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
   238     finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
   239   qed
   240   thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
   241 qed
   242 
   243 lemma cauchy_not_vanishes_cases:
   244   assumes X: "cauchy X"
   245   assumes nz: "\<not> vanishes X"
   246   shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
   247 proof -
   248   obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
   249     using nz unfolding vanishes_def by (auto simp add: not_less)
   250   obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
   251     using `0 < r` by (rule obtain_pos_sum)
   252   obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
   253     using cauchyD [OF X s] ..
   254   obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
   255     using r by fast
   256   have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
   257     using i `i \<le> k` by auto
   258   have "X k \<le> - r \<or> r \<le> X k"
   259     using `r \<le> \<bar>X k\<bar>` by auto
   260   hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
   261     unfolding `r = s + t` using k by auto
   262   hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
   263   thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
   264     using t by auto
   265 qed
   266 
   267 lemma cauchy_not_vanishes:
   268   assumes X: "cauchy X"
   269   assumes nz: "\<not> vanishes X"
   270   shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
   271 using cauchy_not_vanishes_cases [OF assms]
   272 by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
   273 
   274 lemma cauchy_inverse [simp]:
   275   assumes X: "cauchy X"
   276   assumes nz: "\<not> vanishes X"
   277   shows "cauchy (\<lambda>n. inverse (X n))"
   278 proof (rule cauchyI)
   279   fix r :: rat assume "0 < r"
   280   obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
   281     using cauchy_not_vanishes [OF X nz] by fast
   282   from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
   283   obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
   284   proof
   285     show "0 < b * r * b"
   286       by (simp add: `0 < r` b mult_pos_pos)
   287     show "r = inverse b * (b * r * b) * inverse b"
   288       using b by simp
   289   qed
   290   obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
   291     using cauchyD [OF X s] ..
   292   have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
   293   proof (clarsimp)
   294     fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
   295     have "\<bar>inverse (X m) - inverse (X n)\<bar> =
   296           inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
   297       by (simp add: inverse_diff_inverse nz * abs_mult)
   298     also have "\<dots> < inverse b * s * inverse b"
   299       by (simp add: mult_strict_mono less_imp_inverse_less
   300                     mult_pos_pos i j b * s)
   301     finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
   302   qed
   303   thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
   304 qed
   305 
   306 lemma vanishes_diff_inverse:
   307   assumes X: "cauchy X" "\<not> vanishes X"
   308   assumes Y: "cauchy Y" "\<not> vanishes Y"
   309   assumes XY: "vanishes (\<lambda>n. X n - Y n)"
   310   shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
   311 proof (rule vanishesI)
   312   fix r :: rat assume r: "0 < r"
   313   obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
   314     using cauchy_not_vanishes [OF X] by fast
   315   obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
   316     using cauchy_not_vanishes [OF Y] by fast
   317   obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
   318   proof
   319     show "0 < a * r * b"
   320       using a r b by (simp add: mult_pos_pos)
   321     show "inverse a * (a * r * b) * inverse b = r"
   322       using a r b by simp
   323   qed
   324   obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
   325     using vanishesD [OF XY s] ..
   326   have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
   327   proof (clarsimp)
   328     fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
   329     have "X n \<noteq> 0" and "Y n \<noteq> 0"
   330       using i j a b n by auto
   331     hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
   332         inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
   333       by (simp add: inverse_diff_inverse abs_mult)
   334     also have "\<dots> < inverse a * s * inverse b"
   335       apply (intro mult_strict_mono' less_imp_inverse_less)
   336       apply (simp_all add: a b i j k n mult_nonneg_nonneg)
   337       done
   338     also note `inverse a * s * inverse b = r`
   339     finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
   340   qed
   341   thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
   342 qed
   343 
   344 subsection {* Equivalence relation on Cauchy sequences *}
   345 
   346 definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
   347   where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
   348 
   349 lemma realrelI [intro?]:
   350   assumes "cauchy X" and "cauchy Y" and "vanishes (\<lambda>n. X n - Y n)"
   351   shows "realrel X Y"
   352   using assms unfolding realrel_def by simp
   353 
   354 lemma realrel_refl: "cauchy X \<Longrightarrow> realrel X X"
   355   unfolding realrel_def by simp
   356 
   357 lemma symp_realrel: "symp realrel"
   358   unfolding realrel_def
   359   by (rule sympI, clarify, drule vanishes_minus, simp)
   360 
   361 lemma transp_realrel: "transp realrel"
   362   unfolding realrel_def
   363   apply (rule transpI, clarify)
   364   apply (drule (1) vanishes_add)
   365   apply (simp add: algebra_simps)
   366   done
   367 
   368 lemma part_equivp_realrel: "part_equivp realrel"
   369   by (fast intro: part_equivpI symp_realrel transp_realrel
   370     realrel_refl cauchy_const)
   371 
   372 subsection {* The field of real numbers *}
   373 
   374 quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
   375   morphisms rep_real Real
   376   by (rule part_equivp_realrel)
   377 
   378 lemma cr_real_eq: "pcr_real = (\<lambda>x y. cauchy x \<and> Real x = y)"
   379   unfolding real.pcr_cr_eq cr_real_def realrel_def by auto
   380 
   381 lemma Real_induct [induct type: real]: (* TODO: generate automatically *)
   382   assumes "\<And>X. cauchy X \<Longrightarrow> P (Real X)" shows "P x"
   383 proof (induct x)
   384   case (1 X)
   385   hence "cauchy X" by (simp add: realrel_def)
   386   thus "P (Real X)" by (rule assms)
   387 qed
   388 
   389 lemma eq_Real:
   390   "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
   391   using real.rel_eq_transfer
   392   unfolding real.pcr_cr_eq cr_real_def fun_rel_def realrel_def by simp
   393 
   394 lemma Domainp_pcr_real [transfer_domain_rule]: "Domainp pcr_real = cauchy"
   395 by (simp add: real.domain_eq realrel_def)
   396 
   397 instantiation real :: field_inverse_zero
   398 begin
   399 
   400 lift_definition zero_real :: "real" is "\<lambda>n. 0"
   401   by (simp add: realrel_refl)
   402 
   403 lift_definition one_real :: "real" is "\<lambda>n. 1"
   404   by (simp add: realrel_refl)
   405 
   406 lift_definition plus_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n + Y n"
   407   unfolding realrel_def add_diff_add
   408   by (simp only: cauchy_add vanishes_add simp_thms)
   409 
   410 lift_definition uminus_real :: "real \<Rightarrow> real" is "\<lambda>X n. - X n"
   411   unfolding realrel_def minus_diff_minus
   412   by (simp only: cauchy_minus vanishes_minus simp_thms)
   413 
   414 lift_definition times_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n * Y n"
   415   unfolding realrel_def mult_diff_mult
   416   by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add
   417     vanishes_mult_bounded cauchy_imp_bounded simp_thms)
   418 
   419 lift_definition inverse_real :: "real \<Rightarrow> real"
   420   is "\<lambda>X. if vanishes X then (\<lambda>n. 0) else (\<lambda>n. inverse (X n))"
   421 proof -
   422   fix X Y assume "realrel X Y"
   423   hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
   424     unfolding realrel_def by simp_all
   425   have "vanishes X \<longleftrightarrow> vanishes Y"
   426   proof
   427     assume "vanishes X"
   428     from vanishes_diff [OF this XY] show "vanishes Y" by simp
   429   next
   430     assume "vanishes Y"
   431     from vanishes_add [OF this XY] show "vanishes X" by simp
   432   qed
   433   thus "?thesis X Y"
   434     unfolding realrel_def
   435     by (simp add: vanishes_diff_inverse X Y XY)
   436 qed
   437 
   438 definition
   439   "x - y = (x::real) + - y"
   440 
   441 definition
   442   "x / y = (x::real) * inverse y"
   443 
   444 lemma add_Real:
   445   assumes X: "cauchy X" and Y: "cauchy Y"
   446   shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
   447   using assms plus_real.transfer
   448   unfolding cr_real_eq fun_rel_def by simp
   449 
   450 lemma minus_Real:
   451   assumes X: "cauchy X"
   452   shows "- Real X = Real (\<lambda>n. - X n)"
   453   using assms uminus_real.transfer
   454   unfolding cr_real_eq fun_rel_def by simp
   455 
   456 lemma diff_Real:
   457   assumes X: "cauchy X" and Y: "cauchy Y"
   458   shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
   459   unfolding minus_real_def
   460   by (simp add: minus_Real add_Real X Y)
   461 
   462 lemma mult_Real:
   463   assumes X: "cauchy X" and Y: "cauchy Y"
   464   shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
   465   using assms times_real.transfer
   466   unfolding cr_real_eq fun_rel_def by simp
   467 
   468 lemma inverse_Real:
   469   assumes X: "cauchy X"
   470   shows "inverse (Real X) =
   471     (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
   472   using assms inverse_real.transfer zero_real.transfer
   473   unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis)
   474 
   475 instance proof
   476   fix a b c :: real
   477   show "a + b = b + a"
   478     by transfer (simp add: add_ac realrel_def)
   479   show "(a + b) + c = a + (b + c)"
   480     by transfer (simp add: add_ac realrel_def)
   481   show "0 + a = a"
   482     by transfer (simp add: realrel_def)
   483   show "- a + a = 0"
   484     by transfer (simp add: realrel_def)
   485   show "a - b = a + - b"
   486     by (rule minus_real_def)
   487   show "(a * b) * c = a * (b * c)"
   488     by transfer (simp add: mult_ac realrel_def)
   489   show "a * b = b * a"
   490     by transfer (simp add: mult_ac realrel_def)
   491   show "1 * a = a"
   492     by transfer (simp add: mult_ac realrel_def)
   493   show "(a + b) * c = a * c + b * c"
   494     by transfer (simp add: distrib_right realrel_def)
   495   show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
   496     by transfer (simp add: realrel_def)
   497   show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
   498     apply transfer
   499     apply (simp add: realrel_def)
   500     apply (rule vanishesI)
   501     apply (frule (1) cauchy_not_vanishes, clarify)
   502     apply (rule_tac x=k in exI, clarify)
   503     apply (drule_tac x=n in spec, simp)
   504     done
   505   show "a / b = a * inverse b"
   506     by (rule divide_real_def)
   507   show "inverse (0::real) = 0"
   508     by transfer (simp add: realrel_def)
   509 qed
   510 
   511 end
   512 
   513 subsection {* Positive reals *}
   514 
   515 lift_definition positive :: "real \<Rightarrow> bool"
   516   is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
   517 proof -
   518   { fix X Y
   519     assume "realrel X Y"
   520     hence XY: "vanishes (\<lambda>n. X n - Y n)"
   521       unfolding realrel_def by simp_all
   522     assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
   523     then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
   524       by fast
   525     obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
   526       using `0 < r` by (rule obtain_pos_sum)
   527     obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
   528       using vanishesD [OF XY s] ..
   529     have "\<forall>n\<ge>max i j. t < Y n"
   530     proof (clarsimp)
   531       fix n assume n: "i \<le> n" "j \<le> n"
   532       have "\<bar>X n - Y n\<bar> < s" and "r < X n"
   533         using i j n by simp_all
   534       thus "t < Y n" unfolding r by simp
   535     qed
   536     hence "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
   537   } note 1 = this
   538   fix X Y assume "realrel X Y"
   539   hence "realrel X Y" and "realrel Y X"
   540     using symp_realrel unfolding symp_def by auto
   541   thus "?thesis X Y"
   542     by (safe elim!: 1)
   543 qed
   544 
   545 lemma positive_Real:
   546   assumes X: "cauchy X"
   547   shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
   548   using assms positive.transfer
   549   unfolding cr_real_eq fun_rel_def by simp
   550 
   551 lemma positive_zero: "\<not> positive 0"
   552   by transfer auto
   553 
   554 lemma positive_add:
   555   "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
   556 apply transfer
   557 apply (clarify, rename_tac a b i j)
   558 apply (rule_tac x="a + b" in exI, simp)
   559 apply (rule_tac x="max i j" in exI, clarsimp)
   560 apply (simp add: add_strict_mono)
   561 done
   562 
   563 lemma positive_mult:
   564   "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
   565 apply transfer
   566 apply (clarify, rename_tac a b i j)
   567 apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
   568 apply (rule_tac x="max i j" in exI, clarsimp)
   569 apply (rule mult_strict_mono, auto)
   570 done
   571 
   572 lemma positive_minus:
   573   "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
   574 apply transfer
   575 apply (simp add: realrel_def)
   576 apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
   577 done
   578 
   579 instantiation real :: linordered_field_inverse_zero
   580 begin
   581 
   582 definition
   583   "x < y \<longleftrightarrow> positive (y - x)"
   584 
   585 definition
   586   "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
   587 
   588 definition
   589   "abs (a::real) = (if a < 0 then - a else a)"
   590 
   591 definition
   592   "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
   593 
   594 instance proof
   595   fix a b c :: real
   596   show "\<bar>a\<bar> = (if a < 0 then - a else a)"
   597     by (rule abs_real_def)
   598   show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
   599     unfolding less_eq_real_def less_real_def
   600     by (auto, drule (1) positive_add, simp_all add: positive_zero)
   601   show "a \<le> a"
   602     unfolding less_eq_real_def by simp
   603   show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
   604     unfolding less_eq_real_def less_real_def
   605     by (auto, drule (1) positive_add, simp add: algebra_simps)
   606   show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
   607     unfolding less_eq_real_def less_real_def
   608     by (auto, drule (1) positive_add, simp add: positive_zero)
   609   show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
   610     unfolding less_eq_real_def less_real_def by auto
   611     (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
   612     (* Should produce c + b - (c + a) \<equiv> b - a *)
   613   show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
   614     by (rule sgn_real_def)
   615   show "a \<le> b \<or> b \<le> a"
   616     unfolding less_eq_real_def less_real_def
   617     by (auto dest!: positive_minus)
   618   show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
   619     unfolding less_real_def
   620     by (drule (1) positive_mult, simp add: algebra_simps)
   621 qed
   622 
   623 end
   624 
   625 instantiation real :: distrib_lattice
   626 begin
   627 
   628 definition
   629   "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
   630 
   631 definition
   632   "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
   633 
   634 instance proof
   635 qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
   636 
   637 end
   638 
   639 lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
   640 apply (induct x)
   641 apply (simp add: zero_real_def)
   642 apply (simp add: one_real_def add_Real)
   643 done
   644 
   645 lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
   646 apply (cases x rule: int_diff_cases)
   647 apply (simp add: of_nat_Real diff_Real)
   648 done
   649 
   650 lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
   651 apply (induct x)
   652 apply (simp add: Fract_of_int_quotient of_rat_divide)
   653 apply (simp add: of_int_Real divide_inverse)
   654 apply (simp add: inverse_Real mult_Real)
   655 done
   656 
   657 instance real :: archimedean_field
   658 proof
   659   fix x :: real
   660   show "\<exists>z. x \<le> of_int z"
   661     apply (induct x)
   662     apply (frule cauchy_imp_bounded, clarify)
   663     apply (rule_tac x="ceiling b + 1" in exI)
   664     apply (rule less_imp_le)
   665     apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
   666     apply (rule_tac x=1 in exI, simp add: algebra_simps)
   667     apply (rule_tac x=0 in exI, clarsimp)
   668     apply (rule le_less_trans [OF abs_ge_self])
   669     apply (rule less_le_trans [OF _ le_of_int_ceiling])
   670     apply simp
   671     done
   672 qed
   673 
   674 instantiation real :: floor_ceiling
   675 begin
   676 
   677 definition [code del]:
   678   "floor (x::real) = (THE z. of_int z \<le> x \<and> x < of_int (z + 1))"
   679 
   680 instance proof
   681   fix x :: real
   682   show "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
   683     unfolding floor_real_def using floor_exists1 by (rule theI')
   684 qed
   685 
   686 end
   687 
   688 subsection {* Completeness *}
   689 
   690 lemma not_positive_Real:
   691   assumes X: "cauchy X"
   692   shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
   693 unfolding positive_Real [OF X]
   694 apply (auto, unfold not_less)
   695 apply (erule obtain_pos_sum)
   696 apply (drule_tac x=s in spec, simp)
   697 apply (drule_tac r=t in cauchyD [OF X], clarify)
   698 apply (drule_tac x=k in spec, clarsimp)
   699 apply (rule_tac x=n in exI, clarify, rename_tac m)
   700 apply (drule_tac x=m in spec, simp)
   701 apply (drule_tac x=n in spec, simp)
   702 apply (drule spec, drule (1) mp, clarify, rename_tac i)
   703 apply (rule_tac x="max i k" in exI, simp)
   704 done
   705 
   706 lemma le_Real:
   707   assumes X: "cauchy X" and Y: "cauchy Y"
   708   shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
   709 unfolding not_less [symmetric, where 'a=real] less_real_def
   710 apply (simp add: diff_Real not_positive_Real X Y)
   711 apply (simp add: diff_le_eq add_ac)
   712 done
   713 
   714 lemma le_RealI:
   715   assumes Y: "cauchy Y"
   716   shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
   717 proof (induct x)
   718   fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
   719   hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
   720     by (simp add: of_rat_Real le_Real)
   721   {
   722     fix r :: rat assume "0 < r"
   723     then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
   724       by (rule obtain_pos_sum)
   725     obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
   726       using cauchyD [OF Y s] ..
   727     obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
   728       using le [OF t] ..
   729     have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
   730     proof (clarsimp)
   731       fix n assume n: "i \<le> n" "j \<le> n"
   732       have "X n \<le> Y i + t" using n j by simp
   733       moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
   734       ultimately show "X n \<le> Y n + r" unfolding r by simp
   735     qed
   736     hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
   737   }
   738   thus "Real X \<le> Real Y"
   739     by (simp add: of_rat_Real le_Real X Y)
   740 qed
   741 
   742 lemma Real_leI:
   743   assumes X: "cauchy X"
   744   assumes le: "\<forall>n. of_rat (X n) \<le> y"
   745   shows "Real X \<le> y"
   746 proof -
   747   have "- y \<le> - Real X"
   748     by (simp add: minus_Real X le_RealI of_rat_minus le)
   749   thus ?thesis by simp
   750 qed
   751 
   752 lemma less_RealD:
   753   assumes Y: "cauchy Y"
   754   shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
   755 by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
   756 
   757 lemma of_nat_less_two_power:
   758   "of_nat n < (2::'a::linordered_idom) ^ n"
   759 apply (induct n)
   760 apply simp
   761 apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
   762 apply (drule (1) add_le_less_mono, simp)
   763 apply simp
   764 done
   765 
   766 lemma complete_real:
   767   fixes S :: "real set"
   768   assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
   769   shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
   770 proof -
   771   obtain x where x: "x \<in> S" using assms(1) ..
   772   obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
   773 
   774   def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
   775   obtain a where a: "\<not> P a"
   776   proof
   777     have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
   778     also have "x - 1 < x" by simp
   779     finally have "of_int (floor (x - 1)) < x" .
   780     hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
   781     then show "\<not> P (of_int (floor (x - 1)))"
   782       unfolding P_def of_rat_of_int_eq using x by fast
   783   qed
   784   obtain b where b: "P b"
   785   proof
   786     show "P (of_int (ceiling z))"
   787     unfolding P_def of_rat_of_int_eq
   788     proof
   789       fix y assume "y \<in> S"
   790       hence "y \<le> z" using z by simp
   791       also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
   792       finally show "y \<le> of_int (ceiling z)" .
   793     qed
   794   qed
   795 
   796   def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
   797   def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
   798   def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
   799   def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
   800   def C \<equiv> "\<lambda>n. avg (A n) (B n)"
   801   have A_0 [simp]: "A 0 = a" unfolding A_def by simp
   802   have B_0 [simp]: "B 0 = b" unfolding B_def by simp
   803   have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
   804     unfolding A_def B_def C_def bisect_def split_def by simp
   805   have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
   806     unfolding A_def B_def C_def bisect_def split_def by simp
   807 
   808   have width: "\<And>n. B n - A n = (b - a) / 2^n"
   809     apply (simp add: eq_divide_eq)
   810     apply (induct_tac n, simp)
   811     apply (simp add: C_def avg_def algebra_simps)
   812     done
   813 
   814   have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
   815     apply (simp add: divide_less_eq)
   816     apply (subst mult_commute)
   817     apply (frule_tac y=y in ex_less_of_nat_mult)
   818     apply clarify
   819     apply (rule_tac x=n in exI)
   820     apply (erule less_trans)
   821     apply (rule mult_strict_right_mono)
   822     apply (rule le_less_trans [OF _ of_nat_less_two_power])
   823     apply simp
   824     apply assumption
   825     done
   826 
   827   have PA: "\<And>n. \<not> P (A n)"
   828     by (induct_tac n, simp_all add: a)
   829   have PB: "\<And>n. P (B n)"
   830     by (induct_tac n, simp_all add: b)
   831   have ab: "a < b"
   832     using a b unfolding P_def
   833     apply (clarsimp simp add: not_le)
   834     apply (drule (1) bspec)
   835     apply (drule (1) less_le_trans)
   836     apply (simp add: of_rat_less)
   837     done
   838   have AB: "\<And>n. A n < B n"
   839     by (induct_tac n, simp add: ab, simp add: C_def avg_def)
   840   have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
   841     apply (auto simp add: le_less [where 'a=nat])
   842     apply (erule less_Suc_induct)
   843     apply (clarsimp simp add: C_def avg_def)
   844     apply (simp add: add_divide_distrib [symmetric])
   845     apply (rule AB [THEN less_imp_le])
   846     apply simp
   847     done
   848   have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
   849     apply (auto simp add: le_less [where 'a=nat])
   850     apply (erule less_Suc_induct)
   851     apply (clarsimp simp add: C_def avg_def)
   852     apply (simp add: add_divide_distrib [symmetric])
   853     apply (rule AB [THEN less_imp_le])
   854     apply simp
   855     done
   856   have cauchy_lemma:
   857     "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
   858     apply (rule cauchyI)
   859     apply (drule twos [where y="b - a"])
   860     apply (erule exE)
   861     apply (rule_tac x=n in exI, clarify, rename_tac i j)
   862     apply (rule_tac y="B n - A n" in le_less_trans) defer
   863     apply (simp add: width)
   864     apply (drule_tac x=n in spec)
   865     apply (frule_tac x=i in spec, drule (1) mp)
   866     apply (frule_tac x=j in spec, drule (1) mp)
   867     apply (frule A_mono, drule B_mono)
   868     apply (frule A_mono, drule B_mono)
   869     apply arith
   870     done
   871   have "cauchy A"
   872     apply (rule cauchy_lemma [rule_format])
   873     apply (simp add: A_mono)
   874     apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
   875     done
   876   have "cauchy B"
   877     apply (rule cauchy_lemma [rule_format])
   878     apply (simp add: B_mono)
   879     apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
   880     done
   881   have 1: "\<forall>x\<in>S. x \<le> Real B"
   882   proof
   883     fix x assume "x \<in> S"
   884     then show "x \<le> Real B"
   885       using PB [unfolded P_def] `cauchy B`
   886       by (simp add: le_RealI)
   887   qed
   888   have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
   889     apply clarify
   890     apply (erule contrapos_pp)
   891     apply (simp add: not_le)
   892     apply (drule less_RealD [OF `cauchy A`], clarify)
   893     apply (subgoal_tac "\<not> P (A n)")
   894     apply (simp add: P_def not_le, clarify)
   895     apply (erule rev_bexI)
   896     apply (erule (1) less_trans)
   897     apply (simp add: PA)
   898     done
   899   have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
   900   proof (rule vanishesI)
   901     fix r :: rat assume "0 < r"
   902     then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
   903       using twos by fast
   904     have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
   905     proof (clarify)
   906       fix n assume n: "k \<le> n"
   907       have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
   908         by simp
   909       also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
   910         using n by (simp add: divide_left_mono mult_pos_pos)
   911       also note k
   912       finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
   913     qed
   914     thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
   915   qed
   916   hence 3: "Real B = Real A"
   917     by (simp add: eq_Real `cauchy A` `cauchy B` width)
   918   show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
   919     using 1 2 3 by (rule_tac x="Real B" in exI, simp)
   920 qed
   921 
   922 instantiation real :: linear_continuum
   923 begin
   924 
   925 subsection{*Supremum of a set of reals*}
   926 
   927 definition
   928   Sup_real_def: "Sup X \<equiv> LEAST z::real. \<forall>x\<in>X. x\<le>z"
   929 
   930 definition
   931   Inf_real_def: "Inf (X::real set) \<equiv> - Sup (uminus ` X)"
   932 
   933 instance
   934 proof
   935   { fix x :: real and X :: "real set"
   936     assume x: "x \<in> X" "bdd_above X"
   937     then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
   938       using complete_real[of X] unfolding bdd_above_def by blast
   939     then show "x \<le> Sup X"
   940       unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) }
   941   note Sup_upper = this
   942 
   943   { fix z :: real and X :: "real set"
   944     assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
   945     then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
   946       using complete_real[of X] by blast
   947     then have "Sup X = s"
   948       unfolding Sup_real_def by (best intro: Least_equality)  
   949     also from s z have "... \<le> z"
   950       by blast
   951     finally show "Sup X \<le> z" . }
   952   note Sup_least = this
   953 
   954   { fix x z :: real and X :: "real set"
   955     assume x: "x \<in> X" and [simp]: "bdd_below X"
   956     have "-x \<le> Sup (uminus ` X)"
   957       by (rule Sup_upper) (auto simp add: image_iff x)
   958     then show "Inf X \<le> x" 
   959       by (auto simp add: Inf_real_def) }
   960 
   961   { fix z :: real and X :: "real set"
   962     assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
   963     have "Sup (uminus ` X) \<le> -z"
   964       using x z by (force intro: Sup_least)
   965     then show "z \<le> Inf X" 
   966         by (auto simp add: Inf_real_def) }
   967 
   968   show "\<exists>a b::real. a \<noteq> b"
   969     using zero_neq_one by blast
   970 qed
   971 end
   972 
   973 
   974 subsection {* Hiding implementation details *}
   975 
   976 hide_const (open) vanishes cauchy positive Real
   977 
   978 declare Real_induct [induct del]
   979 declare Abs_real_induct [induct del]
   980 declare Abs_real_cases [cases del]
   981 
   982 lifting_update real.lifting
   983 lifting_forget real.lifting
   984   
   985 subsection{*More Lemmas*}
   986 
   987 text {* BH: These lemmas should not be necessary; they should be
   988 covered by existing simp rules and simplification procedures. *}
   989 
   990 lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
   991 by simp (* redundant with mult_cancel_left *)
   992 
   993 lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
   994 by simp (* redundant with mult_cancel_right *)
   995 
   996 lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
   997 by simp (* solved by linordered_ring_less_cancel_factor simproc *)
   998 
   999 lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
  1000 by simp (* solved by linordered_ring_le_cancel_factor simproc *)
  1001 
  1002 lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
  1003 by simp (* solved by linordered_ring_le_cancel_factor simproc *)
  1004 
  1005 
  1006 subsection {* Embedding numbers into the Reals *}
  1007 
  1008 abbreviation
  1009   real_of_nat :: "nat \<Rightarrow> real"
  1010 where
  1011   "real_of_nat \<equiv> of_nat"
  1012 
  1013 abbreviation
  1014   real_of_int :: "int \<Rightarrow> real"
  1015 where
  1016   "real_of_int \<equiv> of_int"
  1017 
  1018 abbreviation
  1019   real_of_rat :: "rat \<Rightarrow> real"
  1020 where
  1021   "real_of_rat \<equiv> of_rat"
  1022 
  1023 consts
  1024   (*overloaded constant for injecting other types into "real"*)
  1025   real :: "'a => real"
  1026 
  1027 defs (overloaded)
  1028   real_of_nat_def [code_unfold]: "real == real_of_nat"
  1029   real_of_int_def [code_unfold]: "real == real_of_int"
  1030 
  1031 declare [[coercion_enabled]]
  1032 declare [[coercion "real::nat\<Rightarrow>real"]]
  1033 declare [[coercion "real::int\<Rightarrow>real"]]
  1034 declare [[coercion "int"]]
  1035 
  1036 declare [[coercion_map map]]
  1037 declare [[coercion_map "% f g h x. g (h (f x))"]]
  1038 declare [[coercion_map "% f g (x,y) . (f x, g y)"]]
  1039 
  1040 lemma real_eq_of_nat: "real = of_nat"
  1041   unfolding real_of_nat_def ..
  1042 
  1043 lemma real_eq_of_int: "real = of_int"
  1044   unfolding real_of_int_def ..
  1045 
  1046 lemma real_of_int_zero [simp]: "real (0::int) = 0"  
  1047 by (simp add: real_of_int_def) 
  1048 
  1049 lemma real_of_one [simp]: "real (1::int) = (1::real)"
  1050 by (simp add: real_of_int_def) 
  1051 
  1052 lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y"
  1053 by (simp add: real_of_int_def) 
  1054 
  1055 lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)"
  1056 by (simp add: real_of_int_def) 
  1057 
  1058 lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y"
  1059 by (simp add: real_of_int_def) 
  1060 
  1061 lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y"
  1062 by (simp add: real_of_int_def) 
  1063 
  1064 lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n"
  1065 by (simp add: real_of_int_def of_int_power)
  1066 
  1067 lemmas power_real_of_int = real_of_int_power [symmetric]
  1068 
  1069 lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))"
  1070   apply (subst real_eq_of_int)+
  1071   apply (rule of_int_setsum)
  1072 done
  1073 
  1074 lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) = 
  1075     (PROD x:A. real(f x))"
  1076   apply (subst real_eq_of_int)+
  1077   apply (rule of_int_setprod)
  1078 done
  1079 
  1080 lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))"
  1081 by (simp add: real_of_int_def) 
  1082 
  1083 lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)"
  1084 by (simp add: real_of_int_def) 
  1085 
  1086 lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)"
  1087 by (simp add: real_of_int_def) 
  1088 
  1089 lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \<le> real y) = (x \<le> y)"
  1090 by (simp add: real_of_int_def) 
  1091 
  1092 lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)"
  1093 by (simp add: real_of_int_def) 
  1094 
  1095 lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)"
  1096 by (simp add: real_of_int_def) 
  1097 
  1098 lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)" 
  1099 by (simp add: real_of_int_def)
  1100 
  1101 lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)"
  1102 by (simp add: real_of_int_def)
  1103 
  1104 lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \<longleftrightarrow> 1 < i"
  1105   unfolding real_of_one[symmetric] real_of_int_less_iff ..
  1106 
  1107 lemma one_le_real_of_int_cancel_iff: "1 \<le> real (i :: int) \<longleftrightarrow> 1 \<le> i"
  1108   unfolding real_of_one[symmetric] real_of_int_le_iff ..
  1109 
  1110 lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \<longleftrightarrow> i < 1"
  1111   unfolding real_of_one[symmetric] real_of_int_less_iff ..
  1112 
  1113 lemma real_of_int_le_one_cancel_iff: "real (i :: int) \<le> 1 \<longleftrightarrow> i \<le> 1"
  1114   unfolding real_of_one[symmetric] real_of_int_le_iff ..
  1115 
  1116 lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))"
  1117 by (auto simp add: abs_if)
  1118 
  1119 lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)"
  1120   apply (subgoal_tac "real n + 1 = real (n + 1)")
  1121   apply (simp del: real_of_int_add)
  1122   apply auto
  1123 done
  1124 
  1125 lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)"
  1126   apply (subgoal_tac "real m + 1 = real (m + 1)")
  1127   apply (simp del: real_of_int_add)
  1128   apply simp
  1129 done
  1130 
  1131 lemma real_of_int_div_aux: "(real (x::int)) / (real d) = 
  1132     real (x div d) + (real (x mod d)) / (real d)"
  1133 proof -
  1134   have "x = (x div d) * d + x mod d"
  1135     by auto
  1136   then have "real x = real (x div d) * real d + real(x mod d)"
  1137     by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym])
  1138   then have "real x / real d = ... / real d"
  1139     by simp
  1140   then show ?thesis
  1141     by (auto simp add: add_divide_distrib algebra_simps)
  1142 qed
  1143 
  1144 lemma real_of_int_div: "(d :: int) dvd n ==>
  1145     real(n div d) = real n / real d"
  1146   apply (subst real_of_int_div_aux)
  1147   apply simp
  1148   apply (simp add: dvd_eq_mod_eq_0)
  1149 done
  1150 
  1151 lemma real_of_int_div2:
  1152   "0 <= real (n::int) / real (x) - real (n div x)"
  1153   apply (case_tac "x = 0")
  1154   apply simp
  1155   apply (case_tac "0 < x")
  1156   apply (simp add: algebra_simps)
  1157   apply (subst real_of_int_div_aux)
  1158   apply simp
  1159   apply (subst zero_le_divide_iff)
  1160   apply auto
  1161   apply (simp add: algebra_simps)
  1162   apply (subst real_of_int_div_aux)
  1163   apply simp
  1164   apply (subst zero_le_divide_iff)
  1165   apply auto
  1166 done
  1167 
  1168 lemma real_of_int_div3:
  1169   "real (n::int) / real (x) - real (n div x) <= 1"
  1170   apply (simp add: algebra_simps)
  1171   apply (subst real_of_int_div_aux)
  1172   apply (auto simp add: divide_le_eq intro: order_less_imp_le)
  1173 done
  1174 
  1175 lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" 
  1176 by (insert real_of_int_div2 [of n x], simp)
  1177 
  1178 lemma Ints_real_of_int [simp]: "real (x::int) \<in> Ints"
  1179 unfolding real_of_int_def by (rule Ints_of_int)
  1180 
  1181 
  1182 subsection{*Embedding the Naturals into the Reals*}
  1183 
  1184 lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
  1185 by (simp add: real_of_nat_def)
  1186 
  1187 lemma real_of_nat_1 [simp]: "real (1::nat) = 1"
  1188 by (simp add: real_of_nat_def)
  1189 
  1190 lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
  1191 by (simp add: real_of_nat_def)
  1192 
  1193 lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n"
  1194 by (simp add: real_of_nat_def)
  1195 
  1196 (*Not for addsimps: often the LHS is used to represent a positive natural*)
  1197 lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
  1198 by (simp add: real_of_nat_def)
  1199 
  1200 lemma real_of_nat_less_iff [iff]: 
  1201      "(real (n::nat) < real m) = (n < m)"
  1202 by (simp add: real_of_nat_def)
  1203 
  1204 lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
  1205 by (simp add: real_of_nat_def)
  1206 
  1207 lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
  1208 by (simp add: real_of_nat_def)
  1209 
  1210 lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)"
  1211 by (simp add: real_of_nat_def del: of_nat_Suc)
  1212 
  1213 lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
  1214 by (simp add: real_of_nat_def of_nat_mult)
  1215 
  1216 lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n"
  1217 by (simp add: real_of_nat_def of_nat_power)
  1218 
  1219 lemmas power_real_of_nat = real_of_nat_power [symmetric]
  1220 
  1221 lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) = 
  1222     (SUM x:A. real(f x))"
  1223   apply (subst real_eq_of_nat)+
  1224   apply (rule of_nat_setsum)
  1225 done
  1226 
  1227 lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) = 
  1228     (PROD x:A. real(f x))"
  1229   apply (subst real_eq_of_nat)+
  1230   apply (rule of_nat_setprod)
  1231 done
  1232 
  1233 lemma real_of_card: "real (card A) = setsum (%x.1) A"
  1234   apply (subst card_eq_setsum)
  1235   apply (subst real_of_nat_setsum)
  1236   apply simp
  1237 done
  1238 
  1239 lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
  1240 by (simp add: real_of_nat_def)
  1241 
  1242 lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)"
  1243 by (simp add: real_of_nat_def)
  1244 
  1245 lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
  1246 by (simp add: add: real_of_nat_def of_nat_diff)
  1247 
  1248 lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
  1249 by (auto simp: real_of_nat_def)
  1250 
  1251 lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)"
  1252 by (simp add: add: real_of_nat_def)
  1253 
  1254 lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0"
  1255 by (simp add: add: real_of_nat_def)
  1256 
  1257 lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)"
  1258   apply (subgoal_tac "real n + 1 = real (Suc n)")
  1259   apply simp
  1260   apply (auto simp add: real_of_nat_Suc)
  1261 done
  1262 
  1263 lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)"
  1264   apply (subgoal_tac "real m + 1 = real (Suc m)")
  1265   apply (simp add: less_Suc_eq_le)
  1266   apply (simp add: real_of_nat_Suc)
  1267 done
  1268 
  1269 lemma real_of_nat_div_aux: "(real (x::nat)) / (real d) = 
  1270     real (x div d) + (real (x mod d)) / (real d)"
  1271 proof -
  1272   have "x = (x div d) * d + x mod d"
  1273     by auto
  1274   then have "real x = real (x div d) * real d + real(x mod d)"
  1275     by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym])
  1276   then have "real x / real d = \<dots> / real d"
  1277     by simp
  1278   then show ?thesis
  1279     by (auto simp add: add_divide_distrib algebra_simps)
  1280 qed
  1281 
  1282 lemma real_of_nat_div: "(d :: nat) dvd n ==>
  1283     real(n div d) = real n / real d"
  1284   by (subst real_of_nat_div_aux)
  1285     (auto simp add: dvd_eq_mod_eq_0 [symmetric])
  1286 
  1287 lemma real_of_nat_div2:
  1288   "0 <= real (n::nat) / real (x) - real (n div x)"
  1289 apply (simp add: algebra_simps)
  1290 apply (subst real_of_nat_div_aux)
  1291 apply simp
  1292 apply (subst zero_le_divide_iff)
  1293 apply simp
  1294 done
  1295 
  1296 lemma real_of_nat_div3:
  1297   "real (n::nat) / real (x) - real (n div x) <= 1"
  1298 apply(case_tac "x = 0")
  1299 apply (simp)
  1300 apply (simp add: algebra_simps)
  1301 apply (subst real_of_nat_div_aux)
  1302 apply simp
  1303 done
  1304 
  1305 lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" 
  1306 by (insert real_of_nat_div2 [of n x], simp)
  1307 
  1308 lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n"
  1309 by (simp add: real_of_int_def real_of_nat_def)
  1310 
  1311 lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x"
  1312   apply (subgoal_tac "real(int(nat x)) = real(nat x)")
  1313   apply force
  1314   apply (simp only: real_of_int_of_nat_eq)
  1315 done
  1316 
  1317 lemma Nats_real_of_nat [simp]: "real (n::nat) \<in> Nats"
  1318 unfolding real_of_nat_def by (rule of_nat_in_Nats)
  1319 
  1320 lemma Ints_real_of_nat [simp]: "real (n::nat) \<in> Ints"
  1321 unfolding real_of_nat_def by (rule Ints_of_nat)
  1322 
  1323 subsection {* The Archimedean Property of the Reals *}
  1324 
  1325 theorem reals_Archimedean:
  1326   assumes x_pos: "0 < x"
  1327   shows "\<exists>n. inverse (real (Suc n)) < x"
  1328   unfolding real_of_nat_def using x_pos
  1329   by (rule ex_inverse_of_nat_Suc_less)
  1330 
  1331 lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
  1332   unfolding real_of_nat_def by (rule ex_less_of_nat)
  1333 
  1334 lemma reals_Archimedean3:
  1335   assumes x_greater_zero: "0 < x"
  1336   shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
  1337   unfolding real_of_nat_def using `0 < x`
  1338   by (auto intro: ex_less_of_nat_mult)
  1339 
  1340 
  1341 subsection{* Rationals *}
  1342 
  1343 lemma Rats_real_nat[simp]: "real(n::nat) \<in> \<rat>"
  1344 by (simp add: real_eq_of_nat)
  1345 
  1346 
  1347 lemma Rats_eq_int_div_int:
  1348   "\<rat> = { real(i::int)/real(j::int) |i j. j \<noteq> 0}" (is "_ = ?S")
  1349 proof
  1350   show "\<rat> \<subseteq> ?S"
  1351   proof
  1352     fix x::real assume "x : \<rat>"
  1353     then obtain r where "x = of_rat r" unfolding Rats_def ..
  1354     have "of_rat r : ?S"
  1355       by (cases r)(auto simp add:of_rat_rat real_eq_of_int)
  1356     thus "x : ?S" using `x = of_rat r` by simp
  1357   qed
  1358 next
  1359   show "?S \<subseteq> \<rat>"
  1360   proof(auto simp:Rats_def)
  1361     fix i j :: int assume "j \<noteq> 0"
  1362     hence "real i / real j = of_rat(Fract i j)"
  1363       by (simp add:of_rat_rat real_eq_of_int)
  1364     thus "real i / real j \<in> range of_rat" by blast
  1365   qed
  1366 qed
  1367 
  1368 lemma Rats_eq_int_div_nat:
  1369   "\<rat> = { real(i::int)/real(n::nat) |i n. n \<noteq> 0}"
  1370 proof(auto simp:Rats_eq_int_div_int)
  1371   fix i j::int assume "j \<noteq> 0"
  1372   show "EX (i'::int) (n::nat). real i/real j = real i'/real n \<and> 0<n"
  1373   proof cases
  1374     assume "j>0"
  1375     hence "real i/real j = real i/real(nat j) \<and> 0<nat j"
  1376       by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
  1377     thus ?thesis by blast
  1378   next
  1379     assume "~ j>0"
  1380     hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using `j\<noteq>0`
  1381       by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
  1382     thus ?thesis by blast
  1383   qed
  1384 next
  1385   fix i::int and n::nat assume "0 < n"
  1386   hence "real i/real n = real i/real(int n) \<and> int n \<noteq> 0" by simp
  1387   thus "\<exists>(i'::int) j::int. real i/real n = real i'/real j \<and> j \<noteq> 0" by blast
  1388 qed
  1389 
  1390 lemma Rats_abs_nat_div_natE:
  1391   assumes "x \<in> \<rat>"
  1392   obtains m n :: nat
  1393   where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
  1394 proof -
  1395   from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
  1396     by(auto simp add: Rats_eq_int_div_nat)
  1397   hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
  1398   then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
  1399   let ?gcd = "gcd m n"
  1400   from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
  1401   let ?k = "m div ?gcd"
  1402   let ?l = "n div ?gcd"
  1403   let ?gcd' = "gcd ?k ?l"
  1404   have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
  1405     by (rule dvd_mult_div_cancel)
  1406   have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
  1407     by (rule dvd_mult_div_cancel)
  1408   from `n\<noteq>0` and gcd_l have "?l \<noteq> 0" by (auto iff del: neq0_conv)
  1409   moreover
  1410   have "\<bar>x\<bar> = real ?k / real ?l"
  1411   proof -
  1412     from gcd have "real ?k / real ?l =
  1413         real (?gcd * ?k) / real (?gcd * ?l)" by simp
  1414     also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
  1415     also from x_rat have "\<dots> = \<bar>x\<bar>" ..
  1416     finally show ?thesis ..
  1417   qed
  1418   moreover
  1419   have "?gcd' = 1"
  1420   proof -
  1421     have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
  1422       by (rule gcd_mult_distrib_nat)
  1423     with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
  1424     with gcd show ?thesis by auto
  1425   qed
  1426   ultimately show ?thesis ..
  1427 qed
  1428 
  1429 subsection{*Density of the Rational Reals in the Reals*}
  1430 
  1431 text{* This density proof is due to Stefan Richter and was ported by TN.  The
  1432 original source is \emph{Real Analysis} by H.L. Royden.
  1433 It employs the Archimedean property of the reals. *}
  1434 
  1435 lemma Rats_dense_in_real:
  1436   fixes x :: real
  1437   assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"
  1438 proof -
  1439   from `x<y` have "0 < y-x" by simp
  1440   with reals_Archimedean obtain q::nat 
  1441     where q: "inverse (real q) < y-x" and "0 < q" by auto
  1442   def p \<equiv> "ceiling (y * real q) - 1"
  1443   def r \<equiv> "of_int p / real q"
  1444   from q have "x < y - inverse (real q)" by simp
  1445   also have "y - inverse (real q) \<le> r"
  1446     unfolding r_def p_def
  1447     by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)
  1448   finally have "x < r" .
  1449   moreover have "r < y"
  1450     unfolding r_def p_def
  1451     by (simp add: divide_less_eq diff_less_eq `0 < q`
  1452       less_ceiling_iff [symmetric])
  1453   moreover from r_def have "r \<in> \<rat>" by simp
  1454   ultimately show ?thesis by fast
  1455 qed
  1456 
  1457 
  1458 
  1459 subsection{*Numerals and Arithmetic*}
  1460 
  1461 lemma [code_abbrev]:
  1462   "real_of_int (numeral k) = numeral k"
  1463   "real_of_int (neg_numeral k) = neg_numeral k"
  1464   by simp_all
  1465 
  1466 text{*Collapse applications of @{term real} to @{term number_of}*}
  1467 lemma real_numeral [simp]:
  1468   "real (numeral v :: int) = numeral v"
  1469   "real (neg_numeral v :: int) = neg_numeral v"
  1470 by (simp_all add: real_of_int_def)
  1471 
  1472 lemma real_of_nat_numeral [simp]:
  1473   "real (numeral v :: nat) = numeral v"
  1474 by (simp add: real_of_nat_def)
  1475 
  1476 declaration {*
  1477   K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2]
  1478     (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *)
  1479   #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2]
  1480     (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *)
  1481   #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add},
  1482       @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one},
  1483       @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
  1484       @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
  1485       @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}]
  1486   #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
  1487   #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
  1488 *}
  1489 
  1490 
  1491 subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
  1492 
  1493 lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" 
  1494 by arith
  1495 
  1496 text {* FIXME: redundant with @{text add_eq_0_iff} below *}
  1497 lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)"
  1498 by auto
  1499 
  1500 lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)"
  1501 by auto
  1502 
  1503 lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)"
  1504 by auto
  1505 
  1506 lemma real_add_le_0_iff: "(x+y \<le> (0::real)) = (y \<le> -x)"
  1507 by auto
  1508 
  1509 lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)"
  1510 by auto
  1511 
  1512 subsection {* Lemmas about powers *}
  1513 
  1514 text {* FIXME: declare this in Rings.thy or not at all *}
  1515 declare abs_mult_self [simp]
  1516 
  1517 (* used by Import/HOL/real.imp *)
  1518 lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
  1519 by simp
  1520 
  1521 lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
  1522 apply (induct "n")
  1523 apply (auto simp add: real_of_nat_Suc)
  1524 apply (subst mult_2)
  1525 apply (erule add_less_le_mono)
  1526 apply (rule two_realpow_ge_one)
  1527 done
  1528 
  1529 text {* TODO: no longer real-specific; rename and move elsewhere *}
  1530 lemma realpow_Suc_le_self:
  1531   fixes r :: "'a::linordered_semidom"
  1532   shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
  1533 by (insert power_decreasing [of 1 "Suc n" r], simp)
  1534 
  1535 text {* TODO: no longer real-specific; rename and move elsewhere *}
  1536 lemma realpow_minus_mult:
  1537   fixes x :: "'a::monoid_mult"
  1538   shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
  1539 by (simp add: power_commutes split add: nat_diff_split)
  1540 
  1541 text {* FIXME: declare this [simp] for all types, or not at all *}
  1542 lemma real_two_squares_add_zero_iff [simp]:
  1543   "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
  1544 by (rule sum_squares_eq_zero_iff)
  1545 
  1546 text {* FIXME: declare this [simp] for all types, or not at all *}
  1547 lemma realpow_two_sum_zero_iff [simp]:
  1548      "(x\<^sup>2 + y\<^sup>2 = (0::real)) = (x = 0 & y = 0)"
  1549 by (rule sum_power2_eq_zero_iff)
  1550 
  1551 lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
  1552 by (rule_tac y = 0 in order_trans, auto)
  1553 
  1554 lemma realpow_square_minus_le [simp]: "- u\<^sup>2 \<le> (x::real)\<^sup>2"
  1555 by (auto simp add: power2_eq_square)
  1556 
  1557 
  1558 lemma numeral_power_le_real_of_nat_cancel_iff[simp]:
  1559   "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::nat) ^ n \<le> a"
  1560   unfolding real_of_nat_le_iff[symmetric] by simp
  1561 
  1562 lemma real_of_nat_le_numeral_power_cancel_iff[simp]:
  1563   "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::nat) ^ n"
  1564   unfolding real_of_nat_le_iff[symmetric] by simp
  1565 
  1566 lemma numeral_power_le_real_of_int_cancel_iff[simp]:
  1567   "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::int) ^ n \<le> a"
  1568   unfolding real_of_int_le_iff[symmetric] by simp
  1569 
  1570 lemma real_of_int_le_numeral_power_cancel_iff[simp]:
  1571   "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::int) ^ n"
  1572   unfolding real_of_int_le_iff[symmetric] by simp
  1573 
  1574 lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]:
  1575   "(neg_numeral x::real) ^ n \<le> real a \<longleftrightarrow> (neg_numeral x::int) ^ n \<le> a"
  1576   unfolding real_of_int_le_iff[symmetric] by simp
  1577 
  1578 lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]:
  1579   "real a \<le> (neg_numeral x::real) ^ n \<longleftrightarrow> a \<le> (neg_numeral x::int) ^ n"
  1580   unfolding real_of_int_le_iff[symmetric] by simp
  1581 
  1582 subsection{*Density of the Reals*}
  1583 
  1584 lemma real_lbound_gt_zero:
  1585      "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
  1586 apply (rule_tac x = " (min d1 d2) /2" in exI)
  1587 apply (simp add: min_def)
  1588 done
  1589 
  1590 
  1591 text{*Similar results are proved in @{text Fields}*}
  1592 lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
  1593   by auto
  1594 
  1595 lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
  1596   by auto
  1597 
  1598 lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
  1599   by simp
  1600 
  1601 subsection{*Absolute Value Function for the Reals*}
  1602 
  1603 lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
  1604 by (simp add: abs_if)
  1605 
  1606 (* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
  1607 lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
  1608 by (force simp add: abs_le_iff)
  1609 
  1610 lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)"
  1611 by (simp add: abs_if)
  1612 
  1613 lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
  1614 by (rule abs_of_nonneg [OF real_of_nat_ge_zero])
  1615 
  1616 lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x"
  1617 by simp
  1618  
  1619 lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
  1620 by simp
  1621 
  1622 
  1623 subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
  1624 
  1625 (* FIXME: theorems for negative numerals *)
  1626 lemma numeral_less_real_of_int_iff [simp]:
  1627      "((numeral n) < real (m::int)) = (numeral n < m)"
  1628 apply auto
  1629 apply (rule real_of_int_less_iff [THEN iffD1])
  1630 apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
  1631 done
  1632 
  1633 lemma numeral_less_real_of_int_iff2 [simp]:
  1634      "(real (m::int) < (numeral n)) = (m < numeral n)"
  1635 apply auto
  1636 apply (rule real_of_int_less_iff [THEN iffD1])
  1637 apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
  1638 done
  1639 
  1640 lemma numeral_le_real_of_int_iff [simp]:
  1641      "((numeral n) \<le> real (m::int)) = (numeral n \<le> m)"
  1642 by (simp add: linorder_not_less [symmetric])
  1643 
  1644 lemma numeral_le_real_of_int_iff2 [simp]:
  1645      "(real (m::int) \<le> (numeral n)) = (m \<le> numeral n)"
  1646 by (simp add: linorder_not_less [symmetric])
  1647 
  1648 lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
  1649 unfolding real_of_nat_def by simp
  1650 
  1651 lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
  1652 unfolding real_of_nat_def by (simp add: floor_minus)
  1653 
  1654 lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
  1655 unfolding real_of_int_def by simp
  1656 
  1657 lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
  1658 unfolding real_of_int_def by (simp add: floor_minus)
  1659 
  1660 lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
  1661 unfolding real_of_int_def by (rule floor_exists)
  1662 
  1663 lemma lemma_floor:
  1664   assumes a1: "real m \<le> r" and a2: "r < real n + 1"
  1665   shows "m \<le> (n::int)"
  1666 proof -
  1667   have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
  1668   also have "... = real (n + 1)" by simp
  1669   finally have "m < n + 1" by (simp only: real_of_int_less_iff)
  1670   thus ?thesis by arith
  1671 qed
  1672 
  1673 lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
  1674 unfolding real_of_int_def by (rule of_int_floor_le)
  1675 
  1676 lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
  1677 by (auto intro: lemma_floor)
  1678 
  1679 lemma real_of_int_floor_cancel [simp]:
  1680     "(real (floor x) = x) = (\<exists>n::int. x = real n)"
  1681   using floor_real_of_int by metis
  1682 
  1683 lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
  1684   unfolding real_of_int_def using floor_unique [of n x] by simp
  1685 
  1686 lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
  1687   unfolding real_of_int_def by (rule floor_unique)
  1688 
  1689 lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
  1690 apply (rule inj_int [THEN injD])
  1691 apply (simp add: real_of_nat_Suc)
  1692 apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
  1693 done
  1694 
  1695 lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
  1696 apply (drule order_le_imp_less_or_eq)
  1697 apply (auto intro: floor_eq3)
  1698 done
  1699 
  1700 lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
  1701   unfolding real_of_int_def using floor_correct [of r] by simp
  1702 
  1703 lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
  1704   unfolding real_of_int_def using floor_correct [of r] by simp
  1705 
  1706 lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
  1707   unfolding real_of_int_def using floor_correct [of r] by simp
  1708 
  1709 lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
  1710   unfolding real_of_int_def using floor_correct [of r] by simp
  1711 
  1712 lemma le_floor: "real a <= x ==> a <= floor x"
  1713   unfolding real_of_int_def by (simp add: le_floor_iff)
  1714 
  1715 lemma real_le_floor: "a <= floor x ==> real a <= x"
  1716   unfolding real_of_int_def by (simp add: le_floor_iff)
  1717 
  1718 lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
  1719   unfolding real_of_int_def by (rule le_floor_iff)
  1720 
  1721 lemma floor_less_eq: "(floor x < a) = (x < real a)"
  1722   unfolding real_of_int_def by (rule floor_less_iff)
  1723 
  1724 lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
  1725   unfolding real_of_int_def by (rule less_floor_iff)
  1726 
  1727 lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
  1728   unfolding real_of_int_def by (rule floor_le_iff)
  1729 
  1730 lemma floor_add [simp]: "floor (x + real a) = floor x + a"
  1731   unfolding real_of_int_def by (rule floor_add_of_int)
  1732 
  1733 lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
  1734   unfolding real_of_int_def by (rule floor_diff_of_int)
  1735 
  1736 lemma le_mult_floor:
  1737   assumes "0 \<le> (a :: real)" and "0 \<le> b"
  1738   shows "floor a * floor b \<le> floor (a * b)"
  1739 proof -
  1740   have "real (floor a) \<le> a"
  1741     and "real (floor b) \<le> b" by auto
  1742   hence "real (floor a * floor b) \<le> a * b"
  1743     using assms by (auto intro!: mult_mono)
  1744   also have "a * b < real (floor (a * b) + 1)" by auto
  1745   finally show ?thesis unfolding real_of_int_less_iff by simp
  1746 qed
  1747 
  1748 lemma floor_divide_eq_div:
  1749   "floor (real a / real b) = a div b"
  1750 proof cases
  1751   assume "b \<noteq> 0 \<or> b dvd a"
  1752   with real_of_int_div3[of a b] show ?thesis
  1753     by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans)
  1754        (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject
  1755               real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial)
  1756 qed (auto simp: real_of_int_div)
  1757 
  1758 lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
  1759   unfolding real_of_nat_def by simp
  1760 
  1761 lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
  1762   unfolding real_of_int_def by (rule le_of_int_ceiling)
  1763 
  1764 lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
  1765   unfolding real_of_int_def by simp
  1766 
  1767 lemma real_of_int_ceiling_cancel [simp]:
  1768      "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
  1769   using ceiling_real_of_int by metis
  1770 
  1771 lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
  1772   unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
  1773 
  1774 lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
  1775   unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
  1776 
  1777 lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n  |] ==> ceiling x = n"
  1778   unfolding real_of_int_def using ceiling_unique [of n x] by simp
  1779 
  1780 lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
  1781   unfolding real_of_int_def using ceiling_correct [of r] by simp
  1782 
  1783 lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
  1784   unfolding real_of_int_def using ceiling_correct [of r] by simp
  1785 
  1786 lemma ceiling_le: "x <= real a ==> ceiling x <= a"
  1787   unfolding real_of_int_def by (simp add: ceiling_le_iff)
  1788 
  1789 lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
  1790   unfolding real_of_int_def by (simp add: ceiling_le_iff)
  1791 
  1792 lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
  1793   unfolding real_of_int_def by (rule ceiling_le_iff)
  1794 
  1795 lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
  1796   unfolding real_of_int_def by (rule less_ceiling_iff)
  1797 
  1798 lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
  1799   unfolding real_of_int_def by (rule ceiling_less_iff)
  1800 
  1801 lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
  1802   unfolding real_of_int_def by (rule le_ceiling_iff)
  1803 
  1804 lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
  1805   unfolding real_of_int_def by (rule ceiling_add_of_int)
  1806 
  1807 lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
  1808   unfolding real_of_int_def by (rule ceiling_diff_of_int)
  1809 
  1810 
  1811 subsubsection {* Versions for the natural numbers *}
  1812 
  1813 definition
  1814   natfloor :: "real => nat" where
  1815   "natfloor x = nat(floor x)"
  1816 
  1817 definition
  1818   natceiling :: "real => nat" where
  1819   "natceiling x = nat(ceiling x)"
  1820 
  1821 lemma natfloor_zero [simp]: "natfloor 0 = 0"
  1822   by (unfold natfloor_def, simp)
  1823 
  1824 lemma natfloor_one [simp]: "natfloor 1 = 1"
  1825   by (unfold natfloor_def, simp)
  1826 
  1827 lemma zero_le_natfloor [simp]: "0 <= natfloor x"
  1828   by (unfold natfloor_def, simp)
  1829 
  1830 lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n"
  1831   by (unfold natfloor_def, simp)
  1832 
  1833 lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
  1834   by (unfold natfloor_def, simp)
  1835 
  1836 lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
  1837   by (unfold natfloor_def, simp)
  1838 
  1839 lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
  1840   unfolding natfloor_def by simp
  1841 
  1842 lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
  1843   unfolding natfloor_def by (intro nat_mono floor_mono)
  1844 
  1845 lemma le_natfloor: "real x <= a ==> x <= natfloor a"
  1846   apply (unfold natfloor_def)
  1847   apply (subst nat_int [THEN sym])
  1848   apply (rule nat_mono)
  1849   apply (rule le_floor)
  1850   apply simp
  1851 done
  1852 
  1853 lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"
  1854   unfolding natfloor_def real_of_nat_def
  1855   by (simp add: nat_less_iff floor_less_iff)
  1856 
  1857 lemma less_natfloor:
  1858   assumes "0 \<le> x" and "x < real (n :: nat)"
  1859   shows "natfloor x < n"
  1860   using assms by (simp add: natfloor_less_iff)
  1861 
  1862 lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
  1863   apply (rule iffI)
  1864   apply (rule order_trans)
  1865   prefer 2
  1866   apply (erule real_natfloor_le)
  1867   apply (subst real_of_nat_le_iff)
  1868   apply assumption
  1869   apply (erule le_natfloor)
  1870 done
  1871 
  1872 lemma le_natfloor_eq_numeral [simp]:
  1873     "~ neg((numeral n)::int) ==> 0 <= x ==>
  1874       (numeral n <= natfloor x) = (numeral n <= x)"
  1875   apply (subst le_natfloor_eq, assumption)
  1876   apply simp
  1877 done
  1878 
  1879 lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
  1880   apply (case_tac "0 <= x")
  1881   apply (subst le_natfloor_eq, assumption, simp)
  1882   apply (rule iffI)
  1883   apply (subgoal_tac "natfloor x <= natfloor 0")
  1884   apply simp
  1885   apply (rule natfloor_mono)
  1886   apply simp
  1887   apply simp
  1888 done
  1889 
  1890 lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
  1891   unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])
  1892 
  1893 lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
  1894   apply (case_tac "0 <= x")
  1895   apply (unfold natfloor_def)
  1896   apply simp
  1897   apply simp_all
  1898 done
  1899 
  1900 lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
  1901 using real_natfloor_add_one_gt by (simp add: algebra_simps)
  1902 
  1903 lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
  1904   apply (subgoal_tac "z < real(natfloor z) + 1")
  1905   apply arith
  1906   apply (rule real_natfloor_add_one_gt)
  1907 done
  1908 
  1909 lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
  1910   unfolding natfloor_def
  1911   unfolding real_of_int_of_nat_eq [symmetric] floor_add
  1912   by (simp add: nat_add_distrib)
  1913 
  1914 lemma natfloor_add_numeral [simp]:
  1915     "~neg ((numeral n)::int) ==> 0 <= x ==>
  1916       natfloor (x + numeral n) = natfloor x + numeral n"
  1917   by (simp add: natfloor_add [symmetric])
  1918 
  1919 lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
  1920   by (simp add: natfloor_add [symmetric] del: One_nat_def)
  1921 
  1922 lemma natfloor_subtract [simp]:
  1923     "natfloor(x - real a) = natfloor x - a"
  1924   unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract
  1925   by simp
  1926 
  1927 lemma natfloor_div_nat:
  1928   assumes "1 <= x" and "y > 0"
  1929   shows "natfloor (x / real y) = natfloor x div y"
  1930 proof (rule natfloor_eq)
  1931   have "(natfloor x) div y * y \<le> natfloor x"
  1932     by (rule add_leD1 [where k="natfloor x mod y"], simp)
  1933   thus "real (natfloor x div y) \<le> x / real y"
  1934     using assms by (simp add: le_divide_eq le_natfloor_eq)
  1935   have "natfloor x < (natfloor x) div y * y + y"
  1936     apply (subst mod_div_equality [symmetric])
  1937     apply (rule add_strict_left_mono)
  1938     apply (rule mod_less_divisor)
  1939     apply fact
  1940     done
  1941   thus "x / real y < real (natfloor x div y) + 1"
  1942     using assms
  1943     by (simp add: divide_less_eq natfloor_less_iff distrib_right)
  1944 qed
  1945 
  1946 lemma le_mult_natfloor:
  1947   shows "natfloor a * natfloor b \<le> natfloor (a * b)"
  1948   by (cases "0 <= a & 0 <= b")
  1949     (auto simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le natfloor_neg)
  1950 
  1951 lemma natceiling_zero [simp]: "natceiling 0 = 0"
  1952   by (unfold natceiling_def, simp)
  1953 
  1954 lemma natceiling_one [simp]: "natceiling 1 = 1"
  1955   by (unfold natceiling_def, simp)
  1956 
  1957 lemma zero_le_natceiling [simp]: "0 <= natceiling x"
  1958   by (unfold natceiling_def, simp)
  1959 
  1960 lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n"
  1961   by (unfold natceiling_def, simp)
  1962 
  1963 lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
  1964   by (unfold natceiling_def, simp)
  1965 
  1966 lemma real_natceiling_ge: "x <= real(natceiling x)"
  1967   unfolding natceiling_def by (cases "x < 0", simp_all)
  1968 
  1969 lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
  1970   unfolding natceiling_def by simp
  1971 
  1972 lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
  1973   unfolding natceiling_def by (intro nat_mono ceiling_mono)
  1974 
  1975 lemma natceiling_le: "x <= real a ==> natceiling x <= a"
  1976   unfolding natceiling_def real_of_nat_def
  1977   by (simp add: nat_le_iff ceiling_le_iff)
  1978 
  1979 lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"
  1980   unfolding natceiling_def real_of_nat_def
  1981   by (simp add: nat_le_iff ceiling_le_iff)
  1982 
  1983 lemma natceiling_le_eq_numeral [simp]:
  1984     "~ neg((numeral n)::int) ==>
  1985       (natceiling x <= numeral n) = (x <= numeral n)"
  1986   by (simp add: natceiling_le_eq)
  1987 
  1988 lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
  1989   unfolding natceiling_def
  1990   by (simp add: nat_le_iff ceiling_le_iff)
  1991 
  1992 lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
  1993   unfolding natceiling_def
  1994   by (simp add: ceiling_eq2 [where n="int n"])
  1995 
  1996 lemma natceiling_add [simp]: "0 <= x ==>
  1997     natceiling (x + real a) = natceiling x + a"
  1998   unfolding natceiling_def
  1999   unfolding real_of_int_of_nat_eq [symmetric] ceiling_add
  2000   by (simp add: nat_add_distrib)
  2001 
  2002 lemma natceiling_add_numeral [simp]:
  2003     "~ neg ((numeral n)::int) ==> 0 <= x ==>
  2004       natceiling (x + numeral n) = natceiling x + numeral n"
  2005   by (simp add: natceiling_add [symmetric])
  2006 
  2007 lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
  2008   by (simp add: natceiling_add [symmetric] del: One_nat_def)
  2009 
  2010 lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a"
  2011   unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract
  2012   by simp
  2013 
  2014 subsection {* Exponentiation with floor *}
  2015 
  2016 lemma floor_power:
  2017   assumes "x = real (floor x)"
  2018   shows "floor (x ^ n) = floor x ^ n"
  2019 proof -
  2020   have *: "x ^ n = real (floor x ^ n)"
  2021     using assms by (induct n arbitrary: x) simp_all
  2022   show ?thesis unfolding real_of_int_inject[symmetric]
  2023     unfolding * floor_real_of_int ..
  2024 qed
  2025 
  2026 lemma natfloor_power:
  2027   assumes "x = real (natfloor x)"
  2028   shows "natfloor (x ^ n) = natfloor x ^ n"
  2029 proof -
  2030   from assms have "0 \<le> floor x" by auto
  2031   note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
  2032   from floor_power[OF this]
  2033   show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
  2034     by simp
  2035 qed
  2036 
  2037 
  2038 subsection {* Implementation of rational real numbers *}
  2039 
  2040 text {* Formal constructor *}
  2041 
  2042 definition Ratreal :: "rat \<Rightarrow> real" where
  2043   [code_abbrev, simp]: "Ratreal = of_rat"
  2044 
  2045 code_datatype Ratreal
  2046 
  2047 
  2048 text {* Numerals *}
  2049 
  2050 lemma [code_abbrev]:
  2051   "(of_rat (of_int a) :: real) = of_int a"
  2052   by simp
  2053 
  2054 lemma [code_abbrev]:
  2055   "(of_rat 0 :: real) = 0"
  2056   by simp
  2057 
  2058 lemma [code_abbrev]:
  2059   "(of_rat 1 :: real) = 1"
  2060   by simp
  2061 
  2062 lemma [code_abbrev]:
  2063   "(of_rat (numeral k) :: real) = numeral k"
  2064   by simp
  2065 
  2066 lemma [code_abbrev]:
  2067   "(of_rat (neg_numeral k) :: real) = neg_numeral k"
  2068   by simp
  2069 
  2070 lemma [code_post]:
  2071   "(of_rat (0 / r)  :: real) = 0"
  2072   "(of_rat (r / 0)  :: real) = 0"
  2073   "(of_rat (1 / 1)  :: real) = 1"
  2074   "(of_rat (numeral k / 1) :: real) = numeral k"
  2075   "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k"
  2076   "(of_rat (1 / numeral k) :: real) = 1 / numeral k"
  2077   "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k"
  2078   "(of_rat (numeral k / numeral l)  :: real) = numeral k / numeral l"
  2079   "(of_rat (numeral k / neg_numeral l)  :: real) = numeral k / neg_numeral l"
  2080   "(of_rat (neg_numeral k / numeral l)  :: real) = neg_numeral k / numeral l"
  2081   "(of_rat (neg_numeral k / neg_numeral l)  :: real) = neg_numeral k / neg_numeral l"
  2082   by (simp_all add: of_rat_divide)
  2083 
  2084 
  2085 text {* Operations *}
  2086 
  2087 lemma zero_real_code [code]:
  2088   "0 = Ratreal 0"
  2089 by simp
  2090 
  2091 lemma one_real_code [code]:
  2092   "1 = Ratreal 1"
  2093 by simp
  2094 
  2095 instantiation real :: equal
  2096 begin
  2097 
  2098 definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
  2099 
  2100 instance proof
  2101 qed (simp add: equal_real_def)
  2102 
  2103 lemma real_equal_code [code]:
  2104   "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
  2105   by (simp add: equal_real_def equal)
  2106 
  2107 lemma [code nbe]:
  2108   "HOL.equal (x::real) x \<longleftrightarrow> True"
  2109   by (rule equal_refl)
  2110 
  2111 end
  2112 
  2113 lemma real_less_eq_code [code]: "Ratreal x \<le> Ratreal y \<longleftrightarrow> x \<le> y"
  2114   by (simp add: of_rat_less_eq)
  2115 
  2116 lemma real_less_code [code]: "Ratreal x < Ratreal y \<longleftrightarrow> x < y"
  2117   by (simp add: of_rat_less)
  2118 
  2119 lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)"
  2120   by (simp add: of_rat_add)
  2121 
  2122 lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)"
  2123   by (simp add: of_rat_mult)
  2124 
  2125 lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)"
  2126   by (simp add: of_rat_minus)
  2127 
  2128 lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)"
  2129   by (simp add: of_rat_diff)
  2130 
  2131 lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)"
  2132   by (simp add: of_rat_inverse)
  2133  
  2134 lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)"
  2135   by (simp add: of_rat_divide)
  2136 
  2137 lemma real_floor_code [code]: "floor (Ratreal x) = floor x"
  2138   by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code)
  2139 
  2140 
  2141 text {* Quickcheck *}
  2142 
  2143 definition (in term_syntax)
  2144   valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
  2145   [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
  2146 
  2147 notation fcomp (infixl "\<circ>>" 60)
  2148 notation scomp (infixl "\<circ>\<rightarrow>" 60)
  2149 
  2150 instantiation real :: random
  2151 begin
  2152 
  2153 definition
  2154   "Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
  2155 
  2156 instance ..
  2157 
  2158 end
  2159 
  2160 no_notation fcomp (infixl "\<circ>>" 60)
  2161 no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
  2162 
  2163 instantiation real :: exhaustive
  2164 begin
  2165 
  2166 definition
  2167   "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d"
  2168 
  2169 instance ..
  2170 
  2171 end
  2172 
  2173 instantiation real :: full_exhaustive
  2174 begin
  2175 
  2176 definition
  2177   "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d"
  2178 
  2179 instance ..
  2180 
  2181 end
  2182 
  2183 instantiation real :: narrowing
  2184 begin
  2185 
  2186 definition
  2187   "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing"
  2188 
  2189 instance ..
  2190 
  2191 end
  2192 
  2193 
  2194 subsection {* Setup for Nitpick *}
  2195 
  2196 declaration {*
  2197   Nitpick_HOL.register_frac_type @{type_name real}
  2198    [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}),
  2199     (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}),
  2200     (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}),
  2201     (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}),
  2202     (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}),
  2203     (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
  2204     (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}),
  2205     (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
  2206 *}
  2207 
  2208 lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real
  2209     ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
  2210     times_real_inst.times_real uminus_real_inst.uminus_real
  2211     zero_real_inst.zero_real
  2212 
  2213 ML_file "Tools/SMT/smt_real.ML"
  2214 setup SMT_Real.setup
  2215 
  2216 end