src/HOL/Real/RComplete.thy
changeset 28952 15a4b2cf8c34
parent 28948 1860f016886d
child 28953 48cd567f6940
equal deleted inserted replaced
28948:1860f016886d 28952:15a4b2cf8c34
     1 (*  Title       : HOL/Real/RComplete.thy
       
     2     ID          : $Id$
       
     3     Author      : Jacques D. Fleuriot, University of Edinburgh
       
     4     Author      : Larry Paulson, University of Cambridge
       
     5     Author      : Jeremy Avigad, Carnegie Mellon University
       
     6     Author      : Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
       
     7 *)
       
     8 
       
     9 header {* Completeness of the Reals; Floor and Ceiling Functions *}
       
    10 
       
    11 theory RComplete
       
    12 imports Lubs RealDef
       
    13 begin
       
    14 
       
    15 lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
       
    16   by simp
       
    17 
       
    18 
       
    19 subsection {* Completeness of Positive Reals *}
       
    20 
       
    21 text {*
       
    22   Supremum property for the set of positive reals
       
    23 
       
    24   Let @{text "P"} be a non-empty set of positive reals, with an upper
       
    25   bound @{text "y"}.  Then @{text "P"} has a least upper bound
       
    26   (written @{text "S"}).
       
    27 
       
    28   FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
       
    29 *}
       
    30 
       
    31 lemma posreal_complete:
       
    32   assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
       
    33     and not_empty_P: "\<exists>x. x \<in> P"
       
    34     and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
       
    35   shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
       
    36 proof (rule exI, rule allI)
       
    37   fix y
       
    38   let ?pP = "{w. real_of_preal w \<in> P}"
       
    39 
       
    40   show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"
       
    41   proof (cases "0 < y")
       
    42     assume neg_y: "\<not> 0 < y"
       
    43     show ?thesis
       
    44     proof
       
    45       assume "\<exists>x\<in>P. y < x"
       
    46       have "\<forall>x. y < real_of_preal x"
       
    47         using neg_y by (rule real_less_all_real2)
       
    48       thus "y < real_of_preal (psup ?pP)" ..
       
    49     next
       
    50       assume "y < real_of_preal (psup ?pP)"
       
    51       obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
       
    52       hence "0 < x" using positive_P by simp
       
    53       hence "y < x" using neg_y by simp
       
    54       thus "\<exists>x \<in> P. y < x" using x_in_P ..
       
    55     qed
       
    56   next
       
    57     assume pos_y: "0 < y"
       
    58 
       
    59     then obtain py where y_is_py: "y = real_of_preal py"
       
    60       by (auto simp add: real_gt_zero_preal_Ex)
       
    61 
       
    62     obtain a where "a \<in> P" using not_empty_P ..
       
    63     with positive_P have a_pos: "0 < a" ..
       
    64     then obtain pa where "a = real_of_preal pa"
       
    65       by (auto simp add: real_gt_zero_preal_Ex)
       
    66     hence "pa \<in> ?pP" using `a \<in> P` by auto
       
    67     hence pP_not_empty: "?pP \<noteq> {}" by auto
       
    68 
       
    69     obtain sup where sup: "\<forall>x \<in> P. x < sup"
       
    70       using upper_bound_Ex ..
       
    71     from this and `a \<in> P` have "a < sup" ..
       
    72     hence "0 < sup" using a_pos by arith
       
    73     then obtain possup where "sup = real_of_preal possup"
       
    74       by (auto simp add: real_gt_zero_preal_Ex)
       
    75     hence "\<forall>X \<in> ?pP. X \<le> possup"
       
    76       using sup by (auto simp add: real_of_preal_lessI)
       
    77     with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"
       
    78       by (rule preal_complete)
       
    79 
       
    80     show ?thesis
       
    81     proof
       
    82       assume "\<exists>x \<in> P. y < x"
       
    83       then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..
       
    84       hence "0 < x" using pos_y by arith
       
    85       then obtain px where x_is_px: "x = real_of_preal px"
       
    86         by (auto simp add: real_gt_zero_preal_Ex)
       
    87 
       
    88       have py_less_X: "\<exists>X \<in> ?pP. py < X"
       
    89       proof
       
    90         show "py < px" using y_is_py and x_is_px and y_less_x
       
    91           by (simp add: real_of_preal_lessI)
       
    92         show "px \<in> ?pP" using x_in_P and x_is_px by simp
       
    93       qed
       
    94 
       
    95       have "(\<exists>X \<in> ?pP. py < X) ==> (py < psup ?pP)"
       
    96         using psup by simp
       
    97       hence "py < psup ?pP" using py_less_X by simp
       
    98       thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
       
    99         using y_is_py and pos_y by (simp add: real_of_preal_lessI)
       
   100     next
       
   101       assume y_less_psup: "y < real_of_preal (psup ?pP)"
       
   102 
       
   103       hence "py < psup ?pP" using y_is_py
       
   104         by (simp add: real_of_preal_lessI)
       
   105       then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"
       
   106         using psup by auto
       
   107       then obtain x where x_is_X: "x = real_of_preal X"
       
   108         by (simp add: real_gt_zero_preal_Ex)
       
   109       hence "y < x" using py_less_X and y_is_py
       
   110         by (simp add: real_of_preal_lessI)
       
   111 
       
   112       moreover have "x \<in> P" using x_is_X and X_in_pP by simp
       
   113 
       
   114       ultimately show "\<exists> x \<in> P. y < x" ..
       
   115     qed
       
   116   qed
       
   117 qed
       
   118 
       
   119 text {*
       
   120   \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
       
   121 *}
       
   122 
       
   123 lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
       
   124   apply (frule isLub_isUb)
       
   125   apply (frule_tac x = y in isLub_isUb)
       
   126   apply (blast intro!: order_antisym dest!: isLub_le_isUb)
       
   127   done
       
   128 
       
   129 
       
   130 text {*
       
   131   \medskip Completeness theorem for the positive reals (again).
       
   132 *}
       
   133 
       
   134 lemma posreals_complete:
       
   135   assumes positive_S: "\<forall>x \<in> S. 0 < x"
       
   136     and not_empty_S: "\<exists>x. x \<in> S"
       
   137     and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
       
   138   shows "\<exists>t. isLub (UNIV::real set) S t"
       
   139 proof
       
   140   let ?pS = "{w. real_of_preal w \<in> S}"
       
   141 
       
   142   obtain u where "isUb UNIV S u" using upper_bound_Ex ..
       
   143   hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
       
   144 
       
   145   obtain x where x_in_S: "x \<in> S" using not_empty_S ..
       
   146   hence x_gt_zero: "0 < x" using positive_S by simp
       
   147   have  "x \<le> u" using sup and x_in_S ..
       
   148   hence "0 < u" using x_gt_zero by arith
       
   149 
       
   150   then obtain pu where u_is_pu: "u = real_of_preal pu"
       
   151     by (auto simp add: real_gt_zero_preal_Ex)
       
   152 
       
   153   have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
       
   154   proof
       
   155     fix pa
       
   156     assume "pa \<in> ?pS"
       
   157     then obtain a where "a \<in> S" and "a = real_of_preal pa"
       
   158       by simp
       
   159     moreover hence "a \<le> u" using sup by simp
       
   160     ultimately show "pa \<le> pu"
       
   161       using sup and u_is_pu by (simp add: real_of_preal_le_iff)
       
   162   qed
       
   163 
       
   164   have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
       
   165   proof
       
   166     fix y
       
   167     assume y_in_S: "y \<in> S"
       
   168     hence "0 < y" using positive_S by simp
       
   169     then obtain py where y_is_py: "y = real_of_preal py"
       
   170       by (auto simp add: real_gt_zero_preal_Ex)
       
   171     hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
       
   172     with pS_less_pu have "py \<le> psup ?pS"
       
   173       by (rule preal_psup_le)
       
   174     thus "y \<le> real_of_preal (psup ?pS)"
       
   175       using y_is_py by (simp add: real_of_preal_le_iff)
       
   176   qed
       
   177 
       
   178   moreover {
       
   179     fix x
       
   180     assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
       
   181     have "real_of_preal (psup ?pS) \<le> x"
       
   182     proof -
       
   183       obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
       
   184       hence s_pos: "0 < s" using positive_S by simp
       
   185 
       
   186       hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
       
   187       then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
       
   188       hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
       
   189 
       
   190       from x_ub_S have "s \<le> x" using s_in_S ..
       
   191       hence "0 < x" using s_pos by simp
       
   192       hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
       
   193       then obtain "px" where x_is_px: "x = real_of_preal px" ..
       
   194 
       
   195       have "\<forall>pe \<in> ?pS. pe \<le> px"
       
   196       proof
       
   197 	fix pe
       
   198 	assume "pe \<in> ?pS"
       
   199 	hence "real_of_preal pe \<in> S" by simp
       
   200 	hence "real_of_preal pe \<le> x" using x_ub_S by simp
       
   201 	thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
       
   202       qed
       
   203 
       
   204       moreover have "?pS \<noteq> {}" using ps_in_pS by auto
       
   205       ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
       
   206       thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
       
   207     qed
       
   208   }
       
   209   ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
       
   210     by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
       
   211 qed
       
   212 
       
   213 text {*
       
   214   \medskip reals Completeness (again!)
       
   215 *}
       
   216 
       
   217 lemma reals_complete:
       
   218   assumes notempty_S: "\<exists>X. X \<in> S"
       
   219     and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
       
   220   shows "\<exists>t. isLub (UNIV :: real set) S t"
       
   221 proof -
       
   222   obtain X where X_in_S: "X \<in> S" using notempty_S ..
       
   223   obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
       
   224     using exists_Ub ..
       
   225   let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
       
   226 
       
   227   {
       
   228     fix x
       
   229     assume "isUb (UNIV::real set) S x"
       
   230     hence S_le_x: "\<forall> y \<in> S. y <= x"
       
   231       by (simp add: isUb_def setle_def)
       
   232     {
       
   233       fix s
       
   234       assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
       
   235       hence "\<exists> x \<in> S. s = x + -X + 1" ..
       
   236       then obtain x1 where "x1 \<in> S" and "s = x1 + (-X) + 1" ..
       
   237       moreover hence "x1 \<le> x" using S_le_x by simp
       
   238       ultimately have "s \<le> x + - X + 1" by arith
       
   239     }
       
   240     then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)"
       
   241       by (auto simp add: isUb_def setle_def)
       
   242   } note S_Ub_is_SHIFT_Ub = this
       
   243 
       
   244   hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp
       
   245   hence "\<exists>Z. isUb UNIV ?SHIFT Z" ..
       
   246   moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
       
   247   moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
       
   248     using X_in_S and Y_isUb by auto
       
   249   ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t"
       
   250     using posreals_complete [of ?SHIFT] by blast
       
   251 
       
   252   show ?thesis
       
   253   proof
       
   254     show "isLub UNIV S (t + X + (-1))"
       
   255     proof (rule isLubI2)
       
   256       {
       
   257         fix x
       
   258         assume "isUb (UNIV::real set) S x"
       
   259         hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
       
   260 	  using S_Ub_is_SHIFT_Ub by simp
       
   261         hence "t \<le> (x + (-X) + 1)"
       
   262 	  using t_is_Lub by (simp add: isLub_le_isUb)
       
   263         hence "t + X + -1 \<le> x" by arith
       
   264       }
       
   265       then show "(t + X + -1) <=* Collect (isUb UNIV S)"
       
   266 	by (simp add: setgeI)
       
   267     next
       
   268       show "isUb UNIV S (t + X + -1)"
       
   269       proof -
       
   270         {
       
   271           fix y
       
   272           assume y_in_S: "y \<in> S"
       
   273           have "y \<le> t + X + -1"
       
   274           proof -
       
   275             obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
       
   276             hence "\<exists> x \<in> S. u = x + - X + 1" by simp
       
   277             then obtain "x" where x_and_u: "u = x + - X + 1" ..
       
   278             have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2)
       
   279 
       
   280             show ?thesis
       
   281             proof cases
       
   282               assume "y \<le> x"
       
   283               moreover have "x = u + X + - 1" using x_and_u by arith
       
   284               moreover have "u + X + - 1  \<le> t + X + -1" using u_le_t by arith
       
   285               ultimately show "y  \<le> t + X + -1" by arith
       
   286             next
       
   287               assume "~(y \<le> x)"
       
   288               hence x_less_y: "x < y" by arith
       
   289 
       
   290               have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
       
   291               hence "0 < x + (-X) + 1" by simp
       
   292               hence "0 < y + (-X) + 1" using x_less_y by arith
       
   293               hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
       
   294               hence "y + (-X) + 1 \<le> t" using t_is_Lub  by (simp add: isLubD2)
       
   295               thus ?thesis by simp
       
   296             qed
       
   297           qed
       
   298         }
       
   299         then show ?thesis by (simp add: isUb_def setle_def)
       
   300       qed
       
   301     qed
       
   302   qed
       
   303 qed
       
   304 
       
   305 
       
   306 subsection {* The Archimedean Property of the Reals *}
       
   307 
       
   308 theorem reals_Archimedean:
       
   309   assumes x_pos: "0 < x"
       
   310   shows "\<exists>n. inverse (real (Suc n)) < x"
       
   311 proof (rule ccontr)
       
   312   assume contr: "\<not> ?thesis"
       
   313   have "\<forall>n. x * real (Suc n) <= 1"
       
   314   proof
       
   315     fix n
       
   316     from contr have "x \<le> inverse (real (Suc n))"
       
   317       by (simp add: linorder_not_less)
       
   318     hence "x \<le> (1 / (real (Suc n)))"
       
   319       by (simp add: inverse_eq_divide)
       
   320     moreover have "0 \<le> real (Suc n)"
       
   321       by (rule real_of_nat_ge_zero)
       
   322     ultimately have "x * real (Suc n) \<le> (1 / real (Suc n)) * real (Suc n)"
       
   323       by (rule mult_right_mono)
       
   324     thus "x * real (Suc n) \<le> 1" by simp
       
   325   qed
       
   326   hence "{z. \<exists>n. z = x * (real (Suc n))} *<= 1"
       
   327     by (simp add: setle_def, safe, rule spec)
       
   328   hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (real (Suc n))} 1"
       
   329     by (simp add: isUbI)
       
   330   hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} Y" ..
       
   331   moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (real (Suc n))}" by auto
       
   332   ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * real (Suc n)} t"
       
   333     by (simp add: reals_complete)
       
   334   then obtain "t" where
       
   335     t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * real (Suc n)} t" ..
       
   336 
       
   337   have "\<forall>n::nat. x * real n \<le> t + - x"
       
   338   proof
       
   339     fix n
       
   340     from t_is_Lub have "x * real (Suc n) \<le> t"
       
   341       by (simp add: isLubD2)
       
   342     hence  "x * (real n) + x \<le> t"
       
   343       by (simp add: right_distrib real_of_nat_Suc)
       
   344     thus  "x * (real n) \<le> t + - x" by arith
       
   345   qed
       
   346 
       
   347   hence "\<forall>m. x * real (Suc m) \<le> t + - x" by simp
       
   348   hence "{z. \<exists>n. z = x * (real (Suc n))}  *<= (t + - x)"
       
   349     by (auto simp add: setle_def)
       
   350   hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (real (Suc n))} (t + (-x))"
       
   351     by (simp add: isUbI)
       
   352   hence "t \<le> t + - x"
       
   353     using t_is_Lub by (simp add: isLub_le_isUb)
       
   354   thus False using x_pos by arith
       
   355 qed
       
   356 
       
   357 text {*
       
   358   There must be other proofs, e.g. @{text "Suc"} of the largest
       
   359   integer in the cut representing @{text "x"}.
       
   360 *}
       
   361 
       
   362 lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
       
   363 proof cases
       
   364   assume "x \<le> 0"
       
   365   hence "x < real (1::nat)" by simp
       
   366   thus ?thesis ..
       
   367 next
       
   368   assume "\<not> x \<le> 0"
       
   369   hence x_greater_zero: "0 < x" by simp
       
   370   hence "0 < inverse x" by simp
       
   371   then obtain n where "inverse (real (Suc n)) < inverse x"
       
   372     using reals_Archimedean by blast
       
   373   hence "inverse (real (Suc n)) * x < inverse x * x"
       
   374     using x_greater_zero by (rule mult_strict_right_mono)
       
   375   hence "inverse (real (Suc n)) * x < 1"
       
   376     using x_greater_zero by simp
       
   377   hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1"
       
   378     by (rule mult_strict_left_mono) simp
       
   379   hence "x < real (Suc n)"
       
   380     by (simp add: ring_simps)
       
   381   thus "\<exists>(n::nat). x < real n" ..
       
   382 qed
       
   383 
       
   384 lemma reals_Archimedean3:
       
   385   assumes x_greater_zero: "0 < x"
       
   386   shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
       
   387 proof
       
   388   fix y
       
   389   have x_not_zero: "x \<noteq> 0" using x_greater_zero by simp
       
   390   obtain n where "y * inverse x < real (n::nat)"
       
   391     using reals_Archimedean2 ..
       
   392   hence "y * inverse x * x < real n * x"
       
   393     using x_greater_zero by (simp add: mult_strict_right_mono)
       
   394   hence "x * inverse x * y < x * real n"
       
   395     by (simp add: ring_simps)
       
   396   hence "y < real (n::nat) * x"
       
   397     using x_not_zero by (simp add: ring_simps)
       
   398   thus "\<exists>(n::nat). y < real n * x" ..
       
   399 qed
       
   400 
       
   401 lemma reals_Archimedean6:
       
   402      "0 \<le> r ==> \<exists>(n::nat). real (n - 1) \<le> r & r < real (n)"
       
   403 apply (insert reals_Archimedean2 [of r], safe)
       
   404 apply (subgoal_tac "\<exists>x::nat. r < real x \<and> (\<forall>y. r < real y \<longrightarrow> x \<le> y)", auto)
       
   405 apply (rule_tac x = x in exI)
       
   406 apply (case_tac x, simp)
       
   407 apply (rename_tac x')
       
   408 apply (drule_tac x = x' in spec, simp)
       
   409 apply (rule_tac x="LEAST n. r < real n" in exI, safe)
       
   410 apply (erule LeastI, erule Least_le)
       
   411 done
       
   412 
       
   413 lemma reals_Archimedean6a: "0 \<le> r ==> \<exists>n. real (n) \<le> r & r < real (Suc n)"
       
   414   by (drule reals_Archimedean6) auto
       
   415 
       
   416 lemma reals_Archimedean_6b_int:
       
   417      "0 \<le> r ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
       
   418 apply (drule reals_Archimedean6a, auto)
       
   419 apply (rule_tac x = "int n" in exI)
       
   420 apply (simp add: real_of_int_real_of_nat real_of_nat_Suc)
       
   421 done
       
   422 
       
   423 lemma reals_Archimedean_6c_int:
       
   424      "r < 0 ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
       
   425 apply (rule reals_Archimedean_6b_int [of "-r", THEN exE], simp, auto)
       
   426 apply (rename_tac n)
       
   427 apply (drule order_le_imp_less_or_eq, auto)
       
   428 apply (rule_tac x = "- n - 1" in exI)
       
   429 apply (rule_tac [2] x = "- n" in exI, auto)
       
   430 done
       
   431 
       
   432 
       
   433 subsection{*Density of the Rational Reals in the Reals*}
       
   434 
       
   435 text{* This density proof is due to Stefan Richter and was ported by TN.  The
       
   436 original source is \emph{Real Analysis} by H.L. Royden.
       
   437 It employs the Archimedean property of the reals. *}
       
   438 
       
   439 lemma Rats_dense_in_nn_real: fixes x::real
       
   440 assumes "0\<le>x" and "x<y" shows "\<exists>r \<in> \<rat>.  x<r \<and> r<y"
       
   441 proof -
       
   442   from `x<y` have "0 < y-x" by simp
       
   443   with reals_Archimedean obtain q::nat 
       
   444     where q: "inverse (real q) < y-x" and "0 < real q" by auto  
       
   445   def p \<equiv> "LEAST n.  y \<le> real (Suc n)/real q"  
       
   446   from reals_Archimedean2 obtain n::nat where "y * real q < real n" by auto
       
   447   with `0 < real q` have ex: "y \<le> real n/real q" (is "?P n")
       
   448     by (simp add: pos_less_divide_eq[THEN sym])
       
   449   also from assms have "\<not> y \<le> real (0::nat) / real q" by simp
       
   450   ultimately have main: "(LEAST n. y \<le> real n/real q) = Suc p"
       
   451     by (unfold p_def) (rule Least_Suc)
       
   452   also from ex have "?P (LEAST x. ?P x)" by (rule LeastI)
       
   453   ultimately have suc: "y \<le> real (Suc p) / real q" by simp
       
   454   def r \<equiv> "real p/real q"
       
   455   have "x = y-(y-x)" by simp
       
   456   also from suc q have "\<dots> < real (Suc p)/real q - inverse (real q)" by arith
       
   457   also have "\<dots> = real p / real q"
       
   458     by (simp only: inverse_eq_divide real_diff_def real_of_nat_Suc 
       
   459     minus_divide_left add_divide_distrib[THEN sym]) simp
       
   460   finally have "x<r" by (unfold r_def)
       
   461   have "p<Suc p" .. also note main[THEN sym]
       
   462   finally have "\<not> ?P p"  by (rule not_less_Least)
       
   463   hence "r<y" by (simp add: r_def)
       
   464   from r_def have "r \<in> \<rat>" by simp
       
   465   with `x<r` `r<y` show ?thesis by fast
       
   466 qed
       
   467 
       
   468 theorem Rats_dense_in_real: fixes x y :: real
       
   469 assumes "x<y" shows "\<exists>r \<in> \<rat>.  x<r \<and> r<y"
       
   470 proof -
       
   471   from reals_Archimedean2 obtain n::nat where "-x < real n" by auto
       
   472   hence "0 \<le> x + real n" by arith
       
   473   also from `x<y` have "x + real n < y + real n" by arith
       
   474   ultimately have "\<exists>r \<in> \<rat>. x + real n < r \<and> r < y + real n"
       
   475     by(rule Rats_dense_in_nn_real)
       
   476   then obtain r where "r \<in> \<rat>" and r2: "x + real n < r" 
       
   477     and r3: "r < y + real n"
       
   478     by blast
       
   479   have "r - real n = r + real (int n)/real (-1::int)" by simp
       
   480   also from `r\<in>\<rat>` have "r + real (int n)/real (-1::int) \<in> \<rat>" by simp
       
   481   also from r2 have "x < r - real n" by arith
       
   482   moreover from r3 have "r - real n < y" by arith
       
   483   ultimately show ?thesis by fast
       
   484 qed
       
   485 
       
   486 
       
   487 subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
       
   488 
       
   489 definition
       
   490   floor :: "real => int" where
       
   491   [code del]: "floor r = (LEAST n::int. r < real (n+1))"
       
   492 
       
   493 definition
       
   494   ceiling :: "real => int" where
       
   495   "ceiling r = - floor (- r)"
       
   496 
       
   497 notation (xsymbols)
       
   498   floor  ("\<lfloor>_\<rfloor>") and
       
   499   ceiling  ("\<lceil>_\<rceil>")
       
   500 
       
   501 notation (HTML output)
       
   502   floor  ("\<lfloor>_\<rfloor>") and
       
   503   ceiling  ("\<lceil>_\<rceil>")
       
   504 
       
   505 
       
   506 lemma number_of_less_real_of_int_iff [simp]:
       
   507      "((number_of n) < real (m::int)) = (number_of n < m)"
       
   508 apply auto
       
   509 apply (rule real_of_int_less_iff [THEN iffD1])
       
   510 apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
       
   511 done
       
   512 
       
   513 lemma number_of_less_real_of_int_iff2 [simp]:
       
   514      "(real (m::int) < (number_of n)) = (m < number_of n)"
       
   515 apply auto
       
   516 apply (rule real_of_int_less_iff [THEN iffD1])
       
   517 apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
       
   518 done
       
   519 
       
   520 lemma number_of_le_real_of_int_iff [simp]:
       
   521      "((number_of n) \<le> real (m::int)) = (number_of n \<le> m)"
       
   522 by (simp add: linorder_not_less [symmetric])
       
   523 
       
   524 lemma number_of_le_real_of_int_iff2 [simp]:
       
   525      "(real (m::int) \<le> (number_of n)) = (m \<le> number_of n)"
       
   526 by (simp add: linorder_not_less [symmetric])
       
   527 
       
   528 lemma floor_zero [simp]: "floor 0 = 0"
       
   529 apply (simp add: floor_def del: real_of_int_add)
       
   530 apply (rule Least_equality)
       
   531 apply simp_all
       
   532 done
       
   533 
       
   534 lemma floor_real_of_nat_zero [simp]: "floor (real (0::nat)) = 0"
       
   535 by auto
       
   536 
       
   537 lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
       
   538 apply (simp only: floor_def)
       
   539 apply (rule Least_equality)
       
   540 apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst])
       
   541 apply (drule_tac [2] real_of_int_less_iff [THEN iffD1])
       
   542 apply simp_all
       
   543 done
       
   544 
       
   545 lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
       
   546 apply (simp only: floor_def)
       
   547 apply (rule Least_equality)
       
   548 apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst])
       
   549 apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst])
       
   550 apply (drule_tac [2] real_of_int_less_iff [THEN iffD1])
       
   551 apply simp_all
       
   552 done
       
   553 
       
   554 lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
       
   555 apply (simp only: floor_def)
       
   556 apply (rule Least_equality)
       
   557 apply auto
       
   558 done
       
   559 
       
   560 lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
       
   561 apply (simp only: floor_def)
       
   562 apply (rule Least_equality)
       
   563 apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst])
       
   564 apply auto
       
   565 done
       
   566 
       
   567 lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
       
   568 apply (case_tac "r < 0")
       
   569 apply (blast intro: reals_Archimedean_6c_int)
       
   570 apply (simp only: linorder_not_less)
       
   571 apply (blast intro: reals_Archimedean_6b_int reals_Archimedean_6c_int)
       
   572 done
       
   573 
       
   574 lemma lemma_floor:
       
   575   assumes a1: "real m \<le> r" and a2: "r < real n + 1"
       
   576   shows "m \<le> (n::int)"
       
   577 proof -
       
   578   have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
       
   579   also have "... = real (n + 1)" by simp
       
   580   finally have "m < n + 1" by (simp only: real_of_int_less_iff)
       
   581   thus ?thesis by arith
       
   582 qed
       
   583 
       
   584 lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
       
   585 apply (simp add: floor_def Least_def)
       
   586 apply (insert real_lb_ub_int [of r], safe)
       
   587 apply (rule theI2)
       
   588 apply auto
       
   589 done
       
   590 
       
   591 lemma floor_mono: "x < y ==> floor x \<le> floor y"
       
   592 apply (simp add: floor_def Least_def)
       
   593 apply (insert real_lb_ub_int [of x])
       
   594 apply (insert real_lb_ub_int [of y], safe)
       
   595 apply (rule theI2)
       
   596 apply (rule_tac [3] theI2)
       
   597 apply simp
       
   598 apply (erule conjI)
       
   599 apply (auto simp add: order_eq_iff int_le_real_less)
       
   600 done
       
   601 
       
   602 lemma floor_mono2: "x \<le> y ==> floor x \<le> floor y"
       
   603 by (auto dest: order_le_imp_less_or_eq simp add: floor_mono)
       
   604 
       
   605 lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
       
   606 by (auto intro: lemma_floor)
       
   607 
       
   608 lemma real_of_int_floor_cancel [simp]:
       
   609     "(real (floor x) = x) = (\<exists>n::int. x = real n)"
       
   610 apply (simp add: floor_def Least_def)
       
   611 apply (insert real_lb_ub_int [of x], erule exE)
       
   612 apply (rule theI2)
       
   613 apply (auto intro: lemma_floor)
       
   614 done
       
   615 
       
   616 lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
       
   617 apply (simp add: floor_def)
       
   618 apply (rule Least_equality)
       
   619 apply (auto intro: lemma_floor)
       
   620 done
       
   621 
       
   622 lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
       
   623 apply (simp add: floor_def)
       
   624 apply (rule Least_equality)
       
   625 apply (auto intro: lemma_floor)
       
   626 done
       
   627 
       
   628 lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
       
   629 apply (rule inj_int [THEN injD])
       
   630 apply (simp add: real_of_nat_Suc)
       
   631 apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
       
   632 done
       
   633 
       
   634 lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
       
   635 apply (drule order_le_imp_less_or_eq)
       
   636 apply (auto intro: floor_eq3)
       
   637 done
       
   638 
       
   639 lemma floor_number_of_eq [simp]:
       
   640      "floor(number_of n :: real) = (number_of n :: int)"
       
   641 apply (subst real_number_of [symmetric])
       
   642 apply (rule floor_real_of_int)
       
   643 done
       
   644 
       
   645 lemma floor_one [simp]: "floor 1 = 1"
       
   646   apply (rule trans)
       
   647   prefer 2
       
   648   apply (rule floor_real_of_int)
       
   649   apply simp
       
   650 done
       
   651 
       
   652 lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
       
   653 apply (simp add: floor_def Least_def)
       
   654 apply (insert real_lb_ub_int [of r], safe)
       
   655 apply (rule theI2)
       
   656 apply (auto intro: lemma_floor)
       
   657 done
       
   658 
       
   659 lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
       
   660 apply (simp add: floor_def Least_def)
       
   661 apply (insert real_lb_ub_int [of r], safe)
       
   662 apply (rule theI2)
       
   663 apply (auto intro: lemma_floor)
       
   664 done
       
   665 
       
   666 lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
       
   667 apply (insert real_of_int_floor_ge_diff_one [of r])
       
   668 apply (auto simp del: real_of_int_floor_ge_diff_one)
       
   669 done
       
   670 
       
   671 lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
       
   672 apply (insert real_of_int_floor_gt_diff_one [of r])
       
   673 apply (auto simp del: real_of_int_floor_gt_diff_one)
       
   674 done
       
   675 
       
   676 lemma le_floor: "real a <= x ==> a <= floor x"
       
   677   apply (subgoal_tac "a < floor x + 1")
       
   678   apply arith
       
   679   apply (subst real_of_int_less_iff [THEN sym])
       
   680   apply simp
       
   681   apply (insert real_of_int_floor_add_one_gt [of x])
       
   682   apply arith
       
   683 done
       
   684 
       
   685 lemma real_le_floor: "a <= floor x ==> real a <= x"
       
   686   apply (rule order_trans)
       
   687   prefer 2
       
   688   apply (rule real_of_int_floor_le)
       
   689   apply (subst real_of_int_le_iff)
       
   690   apply assumption
       
   691 done
       
   692 
       
   693 lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
       
   694   apply (rule iffI)
       
   695   apply (erule real_le_floor)
       
   696   apply (erule le_floor)
       
   697 done
       
   698 
       
   699 lemma le_floor_eq_number_of [simp]:
       
   700     "(number_of n <= floor x) = (number_of n <= x)"
       
   701 by (simp add: le_floor_eq)
       
   702 
       
   703 lemma le_floor_eq_zero [simp]: "(0 <= floor x) = (0 <= x)"
       
   704 by (simp add: le_floor_eq)
       
   705 
       
   706 lemma le_floor_eq_one [simp]: "(1 <= floor x) = (1 <= x)"
       
   707 by (simp add: le_floor_eq)
       
   708 
       
   709 lemma floor_less_eq: "(floor x < a) = (x < real a)"
       
   710   apply (subst linorder_not_le [THEN sym])+
       
   711   apply simp
       
   712   apply (rule le_floor_eq)
       
   713 done
       
   714 
       
   715 lemma floor_less_eq_number_of [simp]:
       
   716     "(floor x < number_of n) = (x < number_of n)"
       
   717 by (simp add: floor_less_eq)
       
   718 
       
   719 lemma floor_less_eq_zero [simp]: "(floor x < 0) = (x < 0)"
       
   720 by (simp add: floor_less_eq)
       
   721 
       
   722 lemma floor_less_eq_one [simp]: "(floor x < 1) = (x < 1)"
       
   723 by (simp add: floor_less_eq)
       
   724 
       
   725 lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
       
   726   apply (insert le_floor_eq [of "a + 1" x])
       
   727   apply auto
       
   728 done
       
   729 
       
   730 lemma less_floor_eq_number_of [simp]:
       
   731     "(number_of n < floor x) = (number_of n + 1 <= x)"
       
   732 by (simp add: less_floor_eq)
       
   733 
       
   734 lemma less_floor_eq_zero [simp]: "(0 < floor x) = (1 <= x)"
       
   735 by (simp add: less_floor_eq)
       
   736 
       
   737 lemma less_floor_eq_one [simp]: "(1 < floor x) = (2 <= x)"
       
   738 by (simp add: less_floor_eq)
       
   739 
       
   740 lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
       
   741   apply (insert floor_less_eq [of x "a + 1"])
       
   742   apply auto
       
   743 done
       
   744 
       
   745 lemma floor_le_eq_number_of [simp]:
       
   746     "(floor x <= number_of n) = (x < number_of n + 1)"
       
   747 by (simp add: floor_le_eq)
       
   748 
       
   749 lemma floor_le_eq_zero [simp]: "(floor x <= 0) = (x < 1)"
       
   750 by (simp add: floor_le_eq)
       
   751 
       
   752 lemma floor_le_eq_one [simp]: "(floor x <= 1) = (x < 2)"
       
   753 by (simp add: floor_le_eq)
       
   754 
       
   755 lemma floor_add [simp]: "floor (x + real a) = floor x + a"
       
   756   apply (subst order_eq_iff)
       
   757   apply (rule conjI)
       
   758   prefer 2
       
   759   apply (subgoal_tac "floor x + a < floor (x + real a) + 1")
       
   760   apply arith
       
   761   apply (subst real_of_int_less_iff [THEN sym])
       
   762   apply simp
       
   763   apply (subgoal_tac "x + real a < real(floor(x + real a)) + 1")
       
   764   apply (subgoal_tac "real (floor x) <= x")
       
   765   apply arith
       
   766   apply (rule real_of_int_floor_le)
       
   767   apply (rule real_of_int_floor_add_one_gt)
       
   768   apply (subgoal_tac "floor (x + real a) < floor x + a + 1")
       
   769   apply arith
       
   770   apply (subst real_of_int_less_iff [THEN sym])
       
   771   apply simp
       
   772   apply (subgoal_tac "real(floor(x + real a)) <= x + real a")
       
   773   apply (subgoal_tac "x < real(floor x) + 1")
       
   774   apply arith
       
   775   apply (rule real_of_int_floor_add_one_gt)
       
   776   apply (rule real_of_int_floor_le)
       
   777 done
       
   778 
       
   779 lemma floor_add_number_of [simp]:
       
   780     "floor (x + number_of n) = floor x + number_of n"
       
   781   apply (subst floor_add [THEN sym])
       
   782   apply simp
       
   783 done
       
   784 
       
   785 lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1"
       
   786   apply (subst floor_add [THEN sym])
       
   787   apply simp
       
   788 done
       
   789 
       
   790 lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
       
   791   apply (subst diff_minus)+
       
   792   apply (subst real_of_int_minus [THEN sym])
       
   793   apply (rule floor_add)
       
   794 done
       
   795 
       
   796 lemma floor_subtract_number_of [simp]: "floor (x - number_of n) =
       
   797     floor x - number_of n"
       
   798   apply (subst floor_subtract [THEN sym])
       
   799   apply simp
       
   800 done
       
   801 
       
   802 lemma floor_subtract_one [simp]: "floor (x - 1) = floor x - 1"
       
   803   apply (subst floor_subtract [THEN sym])
       
   804   apply simp
       
   805 done
       
   806 
       
   807 lemma ceiling_zero [simp]: "ceiling 0 = 0"
       
   808 by (simp add: ceiling_def)
       
   809 
       
   810 lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
       
   811 by (simp add: ceiling_def)
       
   812 
       
   813 lemma ceiling_real_of_nat_zero [simp]: "ceiling (real (0::nat)) = 0"
       
   814 by auto
       
   815 
       
   816 lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r"
       
   817 by (simp add: ceiling_def)
       
   818 
       
   819 lemma floor_ceiling [simp]: "floor (real (ceiling r)) = ceiling r"
       
   820 by (simp add: ceiling_def)
       
   821 
       
   822 lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
       
   823 apply (simp add: ceiling_def)
       
   824 apply (subst le_minus_iff, simp)
       
   825 done
       
   826 
       
   827 lemma ceiling_mono: "x < y ==> ceiling x \<le> ceiling y"
       
   828 by (simp add: floor_mono ceiling_def)
       
   829 
       
   830 lemma ceiling_mono2: "x \<le> y ==> ceiling x \<le> ceiling y"
       
   831 by (simp add: floor_mono2 ceiling_def)
       
   832 
       
   833 lemma real_of_int_ceiling_cancel [simp]:
       
   834      "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
       
   835 apply (auto simp add: ceiling_def)
       
   836 apply (drule arg_cong [where f = uminus], auto)
       
   837 apply (rule_tac x = "-n" in exI, auto)
       
   838 done
       
   839 
       
   840 lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
       
   841 apply (simp add: ceiling_def)
       
   842 apply (rule minus_equation_iff [THEN iffD1])
       
   843 apply (simp add: floor_eq [where n = "-(n+1)"])
       
   844 done
       
   845 
       
   846 lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
       
   847 by (simp add: ceiling_def floor_eq2 [where n = "-(n+1)"])
       
   848 
       
   849 lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n  |] ==> ceiling x = n"
       
   850 by (simp add: ceiling_def floor_eq2 [where n = "-n"])
       
   851 
       
   852 lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
       
   853 by (simp add: ceiling_def)
       
   854 
       
   855 lemma ceiling_number_of_eq [simp]:
       
   856      "ceiling (number_of n :: real) = (number_of n)"
       
   857 apply (subst real_number_of [symmetric])
       
   858 apply (rule ceiling_real_of_int)
       
   859 done
       
   860 
       
   861 lemma ceiling_one [simp]: "ceiling 1 = 1"
       
   862   by (unfold ceiling_def, simp)
       
   863 
       
   864 lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
       
   865 apply (rule neg_le_iff_le [THEN iffD1])
       
   866 apply (simp add: ceiling_def diff_minus)
       
   867 done
       
   868 
       
   869 lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
       
   870 apply (insert real_of_int_ceiling_diff_one_le [of r])
       
   871 apply (simp del: real_of_int_ceiling_diff_one_le)
       
   872 done
       
   873 
       
   874 lemma ceiling_le: "x <= real a ==> ceiling x <= a"
       
   875   apply (unfold ceiling_def)
       
   876   apply (subgoal_tac "-a <= floor(- x)")
       
   877   apply simp
       
   878   apply (rule le_floor)
       
   879   apply simp
       
   880 done
       
   881 
       
   882 lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
       
   883   apply (unfold ceiling_def)
       
   884   apply (subgoal_tac "real(- a) <= - x")
       
   885   apply simp
       
   886   apply (rule real_le_floor)
       
   887   apply simp
       
   888 done
       
   889 
       
   890 lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
       
   891   apply (rule iffI)
       
   892   apply (erule ceiling_le_real)
       
   893   apply (erule ceiling_le)
       
   894 done
       
   895 
       
   896 lemma ceiling_le_eq_number_of [simp]:
       
   897     "(ceiling x <= number_of n) = (x <= number_of n)"
       
   898 by (simp add: ceiling_le_eq)
       
   899 
       
   900 lemma ceiling_le_zero_eq [simp]: "(ceiling x <= 0) = (x <= 0)"
       
   901 by (simp add: ceiling_le_eq)
       
   902 
       
   903 lemma ceiling_le_eq_one [simp]: "(ceiling x <= 1) = (x <= 1)"
       
   904 by (simp add: ceiling_le_eq)
       
   905 
       
   906 lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
       
   907   apply (subst linorder_not_le [THEN sym])+
       
   908   apply simp
       
   909   apply (rule ceiling_le_eq)
       
   910 done
       
   911 
       
   912 lemma less_ceiling_eq_number_of [simp]:
       
   913     "(number_of n < ceiling x) = (number_of n < x)"
       
   914 by (simp add: less_ceiling_eq)
       
   915 
       
   916 lemma less_ceiling_eq_zero [simp]: "(0 < ceiling x) = (0 < x)"
       
   917 by (simp add: less_ceiling_eq)
       
   918 
       
   919 lemma less_ceiling_eq_one [simp]: "(1 < ceiling x) = (1 < x)"
       
   920 by (simp add: less_ceiling_eq)
       
   921 
       
   922 lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
       
   923   apply (insert ceiling_le_eq [of x "a - 1"])
       
   924   apply auto
       
   925 done
       
   926 
       
   927 lemma ceiling_less_eq_number_of [simp]:
       
   928     "(ceiling x < number_of n) = (x <= number_of n - 1)"
       
   929 by (simp add: ceiling_less_eq)
       
   930 
       
   931 lemma ceiling_less_eq_zero [simp]: "(ceiling x < 0) = (x <= -1)"
       
   932 by (simp add: ceiling_less_eq)
       
   933 
       
   934 lemma ceiling_less_eq_one [simp]: "(ceiling x < 1) = (x <= 0)"
       
   935 by (simp add: ceiling_less_eq)
       
   936 
       
   937 lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
       
   938   apply (insert less_ceiling_eq [of "a - 1" x])
       
   939   apply auto
       
   940 done
       
   941 
       
   942 lemma le_ceiling_eq_number_of [simp]:
       
   943     "(number_of n <= ceiling x) = (number_of n - 1 < x)"
       
   944 by (simp add: le_ceiling_eq)
       
   945 
       
   946 lemma le_ceiling_eq_zero [simp]: "(0 <= ceiling x) = (-1 < x)"
       
   947 by (simp add: le_ceiling_eq)
       
   948 
       
   949 lemma le_ceiling_eq_one [simp]: "(1 <= ceiling x) = (0 < x)"
       
   950 by (simp add: le_ceiling_eq)
       
   951 
       
   952 lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
       
   953   apply (unfold ceiling_def, simp)
       
   954   apply (subst real_of_int_minus [THEN sym])
       
   955   apply (subst floor_add)
       
   956   apply simp
       
   957 done
       
   958 
       
   959 lemma ceiling_add_number_of [simp]: "ceiling (x + number_of n) =
       
   960     ceiling x + number_of n"
       
   961   apply (subst ceiling_add [THEN sym])
       
   962   apply simp
       
   963 done
       
   964 
       
   965 lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1"
       
   966   apply (subst ceiling_add [THEN sym])
       
   967   apply simp
       
   968 done
       
   969 
       
   970 lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
       
   971   apply (subst diff_minus)+
       
   972   apply (subst real_of_int_minus [THEN sym])
       
   973   apply (rule ceiling_add)
       
   974 done
       
   975 
       
   976 lemma ceiling_subtract_number_of [simp]: "ceiling (x - number_of n) =
       
   977     ceiling x - number_of n"
       
   978   apply (subst ceiling_subtract [THEN sym])
       
   979   apply simp
       
   980 done
       
   981 
       
   982 lemma ceiling_subtract_one [simp]: "ceiling (x - 1) = ceiling x - 1"
       
   983   apply (subst ceiling_subtract [THEN sym])
       
   984   apply simp
       
   985 done
       
   986 
       
   987 subsection {* Versions for the natural numbers *}
       
   988 
       
   989 definition
       
   990   natfloor :: "real => nat" where
       
   991   "natfloor x = nat(floor x)"
       
   992 
       
   993 definition
       
   994   natceiling :: "real => nat" where
       
   995   "natceiling x = nat(ceiling x)"
       
   996 
       
   997 lemma natfloor_zero [simp]: "natfloor 0 = 0"
       
   998   by (unfold natfloor_def, simp)
       
   999 
       
  1000 lemma natfloor_one [simp]: "natfloor 1 = 1"
       
  1001   by (unfold natfloor_def, simp)
       
  1002 
       
  1003 lemma zero_le_natfloor [simp]: "0 <= natfloor x"
       
  1004   by (unfold natfloor_def, simp)
       
  1005 
       
  1006 lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n"
       
  1007   by (unfold natfloor_def, simp)
       
  1008 
       
  1009 lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
       
  1010   by (unfold natfloor_def, simp)
       
  1011 
       
  1012 lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
       
  1013   by (unfold natfloor_def, simp)
       
  1014 
       
  1015 lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
       
  1016   apply (unfold natfloor_def)
       
  1017   apply (subgoal_tac "floor x <= floor 0")
       
  1018   apply simp
       
  1019   apply (erule floor_mono2)
       
  1020 done
       
  1021 
       
  1022 lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
       
  1023   apply (case_tac "0 <= x")
       
  1024   apply (subst natfloor_def)+
       
  1025   apply (subst nat_le_eq_zle)
       
  1026   apply force
       
  1027   apply (erule floor_mono2)
       
  1028   apply (subst natfloor_neg)
       
  1029   apply simp
       
  1030   apply simp
       
  1031 done
       
  1032 
       
  1033 lemma le_natfloor: "real x <= a ==> x <= natfloor a"
       
  1034   apply (unfold natfloor_def)
       
  1035   apply (subst nat_int [THEN sym])
       
  1036   apply (subst nat_le_eq_zle)
       
  1037   apply simp
       
  1038   apply (rule le_floor)
       
  1039   apply simp
       
  1040 done
       
  1041 
       
  1042 lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
       
  1043   apply (rule iffI)
       
  1044   apply (rule order_trans)
       
  1045   prefer 2
       
  1046   apply (erule real_natfloor_le)
       
  1047   apply (subst real_of_nat_le_iff)
       
  1048   apply assumption
       
  1049   apply (erule le_natfloor)
       
  1050 done
       
  1051 
       
  1052 lemma le_natfloor_eq_number_of [simp]:
       
  1053     "~ neg((number_of n)::int) ==> 0 <= x ==>
       
  1054       (number_of n <= natfloor x) = (number_of n <= x)"
       
  1055   apply (subst le_natfloor_eq, assumption)
       
  1056   apply simp
       
  1057 done
       
  1058 
       
  1059 lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
       
  1060   apply (case_tac "0 <= x")
       
  1061   apply (subst le_natfloor_eq, assumption, simp)
       
  1062   apply (rule iffI)
       
  1063   apply (subgoal_tac "natfloor x <= natfloor 0")
       
  1064   apply simp
       
  1065   apply (rule natfloor_mono)
       
  1066   apply simp
       
  1067   apply simp
       
  1068 done
       
  1069 
       
  1070 lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
       
  1071   apply (unfold natfloor_def)
       
  1072   apply (subst nat_int [THEN sym]);back;
       
  1073   apply (subst eq_nat_nat_iff)
       
  1074   apply simp
       
  1075   apply simp
       
  1076   apply (rule floor_eq2)
       
  1077   apply auto
       
  1078 done
       
  1079 
       
  1080 lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
       
  1081   apply (case_tac "0 <= x")
       
  1082   apply (unfold natfloor_def)
       
  1083   apply simp
       
  1084   apply simp_all
       
  1085 done
       
  1086 
       
  1087 lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
       
  1088   apply (simp add: compare_rls)
       
  1089   apply (rule real_natfloor_add_one_gt)
       
  1090 done
       
  1091 
       
  1092 lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
       
  1093   apply (subgoal_tac "z < real(natfloor z) + 1")
       
  1094   apply arith
       
  1095   apply (rule real_natfloor_add_one_gt)
       
  1096 done
       
  1097 
       
  1098 lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
       
  1099   apply (unfold natfloor_def)
       
  1100   apply (subgoal_tac "real a = real (int a)")
       
  1101   apply (erule ssubst)
       
  1102   apply (simp add: nat_add_distrib del: real_of_int_of_nat_eq)
       
  1103   apply simp
       
  1104 done
       
  1105 
       
  1106 lemma natfloor_add_number_of [simp]:
       
  1107     "~neg ((number_of n)::int) ==> 0 <= x ==>
       
  1108       natfloor (x + number_of n) = natfloor x + number_of n"
       
  1109   apply (subst natfloor_add [THEN sym])
       
  1110   apply simp_all
       
  1111 done
       
  1112 
       
  1113 lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
       
  1114   apply (subst natfloor_add [THEN sym])
       
  1115   apply assumption
       
  1116   apply simp
       
  1117 done
       
  1118 
       
  1119 lemma natfloor_subtract [simp]: "real a <= x ==>
       
  1120     natfloor(x - real a) = natfloor x - a"
       
  1121   apply (unfold natfloor_def)
       
  1122   apply (subgoal_tac "real a = real (int a)")
       
  1123   apply (erule ssubst)
       
  1124   apply (simp del: real_of_int_of_nat_eq)
       
  1125   apply simp
       
  1126 done
       
  1127 
       
  1128 lemma natceiling_zero [simp]: "natceiling 0 = 0"
       
  1129   by (unfold natceiling_def, simp)
       
  1130 
       
  1131 lemma natceiling_one [simp]: "natceiling 1 = 1"
       
  1132   by (unfold natceiling_def, simp)
       
  1133 
       
  1134 lemma zero_le_natceiling [simp]: "0 <= natceiling x"
       
  1135   by (unfold natceiling_def, simp)
       
  1136 
       
  1137 lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n"
       
  1138   by (unfold natceiling_def, simp)
       
  1139 
       
  1140 lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
       
  1141   by (unfold natceiling_def, simp)
       
  1142 
       
  1143 lemma real_natceiling_ge: "x <= real(natceiling x)"
       
  1144   apply (unfold natceiling_def)
       
  1145   apply (case_tac "x < 0")
       
  1146   apply simp
       
  1147   apply (subst real_nat_eq_real)
       
  1148   apply (subgoal_tac "ceiling 0 <= ceiling x")
       
  1149   apply simp
       
  1150   apply (rule ceiling_mono2)
       
  1151   apply simp
       
  1152   apply simp
       
  1153 done
       
  1154 
       
  1155 lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
       
  1156   apply (unfold natceiling_def)
       
  1157   apply simp
       
  1158 done
       
  1159 
       
  1160 lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
       
  1161   apply (case_tac "0 <= x")
       
  1162   apply (subst natceiling_def)+
       
  1163   apply (subst nat_le_eq_zle)
       
  1164   apply (rule disjI2)
       
  1165   apply (subgoal_tac "real (0::int) <= real(ceiling y)")
       
  1166   apply simp
       
  1167   apply (rule order_trans)
       
  1168   apply simp
       
  1169   apply (erule order_trans)
       
  1170   apply simp
       
  1171   apply (erule ceiling_mono2)
       
  1172   apply (subst natceiling_neg)
       
  1173   apply simp_all
       
  1174 done
       
  1175 
       
  1176 lemma natceiling_le: "x <= real a ==> natceiling x <= a"
       
  1177   apply (unfold natceiling_def)
       
  1178   apply (case_tac "x < 0")
       
  1179   apply simp
       
  1180   apply (subst nat_int [THEN sym]);back;
       
  1181   apply (subst nat_le_eq_zle)
       
  1182   apply simp
       
  1183   apply (rule ceiling_le)
       
  1184   apply simp
       
  1185 done
       
  1186 
       
  1187 lemma natceiling_le_eq: "0 <= x ==> (natceiling x <= a) = (x <= real a)"
       
  1188   apply (rule iffI)
       
  1189   apply (rule order_trans)
       
  1190   apply (rule real_natceiling_ge)
       
  1191   apply (subst real_of_nat_le_iff)
       
  1192   apply assumption
       
  1193   apply (erule natceiling_le)
       
  1194 done
       
  1195 
       
  1196 lemma natceiling_le_eq_number_of [simp]:
       
  1197     "~ neg((number_of n)::int) ==> 0 <= x ==>
       
  1198       (natceiling x <= number_of n) = (x <= number_of n)"
       
  1199   apply (subst natceiling_le_eq, assumption)
       
  1200   apply simp
       
  1201 done
       
  1202 
       
  1203 lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
       
  1204   apply (case_tac "0 <= x")
       
  1205   apply (subst natceiling_le_eq)
       
  1206   apply assumption
       
  1207   apply simp
       
  1208   apply (subst natceiling_neg)
       
  1209   apply simp
       
  1210   apply simp
       
  1211 done
       
  1212 
       
  1213 lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
       
  1214   apply (unfold natceiling_def)
       
  1215   apply (simplesubst nat_int [THEN sym]) back back
       
  1216   apply (subgoal_tac "nat(int n) + 1 = nat(int n + 1)")
       
  1217   apply (erule ssubst)
       
  1218   apply (subst eq_nat_nat_iff)
       
  1219   apply (subgoal_tac "ceiling 0 <= ceiling x")
       
  1220   apply simp
       
  1221   apply (rule ceiling_mono2)
       
  1222   apply force
       
  1223   apply force
       
  1224   apply (rule ceiling_eq2)
       
  1225   apply (simp, simp)
       
  1226   apply (subst nat_add_distrib)
       
  1227   apply auto
       
  1228 done
       
  1229 
       
  1230 lemma natceiling_add [simp]: "0 <= x ==>
       
  1231     natceiling (x + real a) = natceiling x + a"
       
  1232   apply (unfold natceiling_def)
       
  1233   apply (subgoal_tac "real a = real (int a)")
       
  1234   apply (erule ssubst)
       
  1235   apply (simp del: real_of_int_of_nat_eq)
       
  1236   apply (subst nat_add_distrib)
       
  1237   apply (subgoal_tac "0 = ceiling 0")
       
  1238   apply (erule ssubst)
       
  1239   apply (erule ceiling_mono2)
       
  1240   apply simp_all
       
  1241 done
       
  1242 
       
  1243 lemma natceiling_add_number_of [simp]:
       
  1244     "~ neg ((number_of n)::int) ==> 0 <= x ==>
       
  1245       natceiling (x + number_of n) = natceiling x + number_of n"
       
  1246   apply (subst natceiling_add [THEN sym])
       
  1247   apply simp_all
       
  1248 done
       
  1249 
       
  1250 lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
       
  1251   apply (subst natceiling_add [THEN sym])
       
  1252   apply assumption
       
  1253   apply simp
       
  1254 done
       
  1255 
       
  1256 lemma natceiling_subtract [simp]: "real a <= x ==>
       
  1257     natceiling(x - real a) = natceiling x - a"
       
  1258   apply (unfold natceiling_def)
       
  1259   apply (subgoal_tac "real a = real (int a)")
       
  1260   apply (erule ssubst)
       
  1261   apply (simp del: real_of_int_of_nat_eq)
       
  1262   apply simp
       
  1263 done
       
  1264 
       
  1265 lemma natfloor_div_nat: "1 <= x ==> y > 0 ==>
       
  1266   natfloor (x / real y) = natfloor x div y"
       
  1267 proof -
       
  1268   assume "1 <= (x::real)" and "(y::nat) > 0"
       
  1269   have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y"
       
  1270     by simp
       
  1271   then have a: "real(natfloor x) = real ((natfloor x) div y) * real y +
       
  1272     real((natfloor x) mod y)"
       
  1273     by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym])
       
  1274   have "x = real(natfloor x) + (x - real(natfloor x))"
       
  1275     by simp
       
  1276   then have "x = real ((natfloor x) div y) * real y +
       
  1277       real((natfloor x) mod y) + (x - real(natfloor x))"
       
  1278     by (simp add: a)
       
  1279   then have "x / real y = ... / real y"
       
  1280     by simp
       
  1281   also have "... = real((natfloor x) div y) + real((natfloor x) mod y) /
       
  1282     real y + (x - real(natfloor x)) / real y"
       
  1283     by (auto simp add: ring_simps add_divide_distrib
       
  1284       diff_divide_distrib prems)
       
  1285   finally have "natfloor (x / real y) = natfloor(...)" by simp
       
  1286   also have "... = natfloor(real((natfloor x) mod y) /
       
  1287     real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))"
       
  1288     by (simp add: add_ac)
       
  1289   also have "... = natfloor(real((natfloor x) mod y) /
       
  1290     real y + (x - real(natfloor x)) / real y) + (natfloor x) div y"
       
  1291     apply (rule natfloor_add)
       
  1292     apply (rule add_nonneg_nonneg)
       
  1293     apply (rule divide_nonneg_pos)
       
  1294     apply simp
       
  1295     apply (simp add: prems)
       
  1296     apply (rule divide_nonneg_pos)
       
  1297     apply (simp add: compare_rls)
       
  1298     apply (rule real_natfloor_le)
       
  1299     apply (insert prems, auto)
       
  1300     done
       
  1301   also have "natfloor(real((natfloor x) mod y) /
       
  1302     real y + (x - real(natfloor x)) / real y) = 0"
       
  1303     apply (rule natfloor_eq)
       
  1304     apply simp
       
  1305     apply (rule add_nonneg_nonneg)
       
  1306     apply (rule divide_nonneg_pos)
       
  1307     apply force
       
  1308     apply (force simp add: prems)
       
  1309     apply (rule divide_nonneg_pos)
       
  1310     apply (simp add: compare_rls)
       
  1311     apply (rule real_natfloor_le)
       
  1312     apply (auto simp add: prems)
       
  1313     apply (insert prems, arith)
       
  1314     apply (simp add: add_divide_distrib [THEN sym])
       
  1315     apply (subgoal_tac "real y = real y - 1 + 1")
       
  1316     apply (erule ssubst)
       
  1317     apply (rule add_le_less_mono)
       
  1318     apply (simp add: compare_rls)
       
  1319     apply (subgoal_tac "real(natfloor x mod y) + 1 =
       
  1320       real(natfloor x mod y + 1)")
       
  1321     apply (erule ssubst)
       
  1322     apply (subst real_of_nat_le_iff)
       
  1323     apply (subgoal_tac "natfloor x mod y < y")
       
  1324     apply arith
       
  1325     apply (rule mod_less_divisor)
       
  1326     apply auto
       
  1327     apply (simp add: compare_rls)
       
  1328     apply (subst add_commute)
       
  1329     apply (rule real_natfloor_add_one_gt)
       
  1330     done
       
  1331   finally show ?thesis by simp
       
  1332 qed
       
  1333 
       
  1334 end