src/HOL/Real/RComplete.thy
changeset 16893 0cc94e6f6ae5
parent 16827 c90a1f450327
child 19765 dfe940911617
equal deleted inserted replaced
16892:23887fee6071 16893:0cc94e6f6ae5
     1 (*  Title       : RComplete.thy
     1 (*  Title       : HOL/Real/RComplete.thy
     2     ID          : $Id$
     2     ID          : $Id$
     3     Author      : Jacques D. Fleuriot
     3     Author      : Jacques D. Fleuriot, University of Edinburgh
     4                   Converted to Isar and polished by lcp
     4     Author      : Larry Paulson, University of Cambridge
     5                   Most floor and ceiling lemmas by Jeremy Avigad
     5     Author      : Jeremy Avigad, Carnegie Mellon University
     6     Copyright   : 1998  University of Cambridge
     6     Author      : Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
     7     Copyright   : 2001,2002  University of Edinburgh
     7 *)
     8 *) 
     8 
     9 
     9 header {* Completeness of the Reals; Floor and Ceiling Functions *}
    10 header{*Completeness of the Reals; Floor and Ceiling Functions*}
       
    11 
    10 
    12 theory RComplete
    11 theory RComplete
    13 imports Lubs RealDef
    12 imports Lubs RealDef
    14 begin
    13 begin
    15 
    14 
    16 lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
    15 lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
    17 by simp
    16   by simp
    18 
    17 
    19 
    18 
    20 subsection{*Completeness of Reals by Supremum Property of type @{typ preal}*} 
    19 subsection {* Completeness of Positive Reals *}
    21 
    20 
    22  (*a few lemmas*)
    21 text {*
    23 lemma real_sup_lemma1:
    22   Supremum property for the set of positive reals
    24      "\<forall>x \<in> P. 0 < x ==>   
    23 
    25       ((\<exists>x \<in> P. y < x) = (\<exists>X. real_of_preal X \<in> P & y < real_of_preal X))"
    24   Let @{text "P"} be a non-empty set of positive reals, with an upper
    26 by (blast dest!: bspec real_gt_zero_preal_Ex [THEN iffD1])
    25   bound @{text "y"}.  Then @{text "P"} has a least upper bound
    27 
    26   (written @{text "S"}).
    28 lemma real_sup_lemma2:
    27 
    29      "[| \<forall>x \<in> P. 0 < x;  a \<in> P;   \<forall>x \<in> P. x < y |]  
    28   FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
    30       ==> (\<exists>X. X\<in> {w. real_of_preal w \<in> P}) &  
    29 *}
    31           (\<exists>Y. \<forall>X\<in> {w. real_of_preal w \<in> P}. X < Y)"
    30 
    32 apply (rule conjI)
    31 lemma posreal_complete:
    33 apply (blast dest: bspec real_gt_zero_preal_Ex [THEN iffD1], auto)
    32   assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
    34 apply (drule bspec, assumption)
    33     and not_empty_P: "\<exists>x. x \<in> P"
    35 apply (frule bspec, assumption)
    34     and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
    36 apply (drule order_less_trans, assumption)
    35   shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
    37 apply (drule real_gt_zero_preal_Ex [THEN iffD1], force) 
    36 proof (rule exI, rule allI)
    38 done
    37   fix y
    39 
    38   let ?pP = "{w. real_of_preal w \<in> P}"
    40 (*-------------------------------------------------------------
    39 
    41             Completeness of Positive Reals
    40   show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"
    42  -------------------------------------------------------------*)
    41   proof (cases "0 < y")
    43 
    42     assume neg_y: "\<not> 0 < y"
    44 (**
    43     show ?thesis
    45  Supremum property for the set of positive reals
    44     proof
    46  FIXME: long proof - should be improved
    45       assume "\<exists>x\<in>P. y < x"
    47 **)
    46       have "\<forall>x. y < real_of_preal x"
    48 
    47         using neg_y by (rule real_less_all_real2)
    49 (*Let P be a non-empty set of positive reals, with an upper bound y.
    48       thus "y < real_of_preal (psup ?pP)" ..
    50   Then P has a least upper bound (written S).  
    49     next
    51 FIXME: Can the premise be weakened to \<forall>x \<in> P. x\<le> y ??*)
    50       assume "y < real_of_preal (psup ?pP)"
    52 lemma posreal_complete: "[| \<forall>x \<in> P. (0::real) < x;  \<exists>x. x \<in> P;  \<exists>y. \<forall>x \<in> P. x<y |]  
    51       obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
    53       ==> (\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S))"
    52       hence "0 < x" using positive_P by simp
    54 apply (rule_tac x = "real_of_preal (psup ({w. real_of_preal w \<in> P}))" in exI)
    53       hence "y < x" using neg_y by simp
    55 apply clarify
    54       thus "\<exists>x \<in> P. y < x" using x_in_P ..
    56 apply (case_tac "0 < ya", auto)
    55     qed
    57 apply (frule real_sup_lemma2, assumption+)
    56   next
    58 apply (drule real_gt_zero_preal_Ex [THEN iffD1])
    57     assume pos_y: "0 < y"
    59 apply (drule_tac [3] real_less_all_real2, auto)
    58 
    60 apply (rule preal_complete [THEN iffD1])
    59     then obtain py where y_is_py: "y = real_of_preal py"
    61 apply (auto intro: order_less_imp_le)
    60       by (auto simp add: real_gt_zero_preal_Ex)
    62 apply (frule real_gt_preal_preal_Ex, force)
    61 
    63 (* second part *)
    62     obtain a where a_in_P: "a \<in> P" using not_empty_P ..
    64 apply (rule real_sup_lemma1 [THEN iffD2], assumption)
    63     have a_pos: "0 < a" using positive_P ..
    65 apply (auto dest!: real_less_all_real2 real_gt_zero_preal_Ex [THEN iffD1])
    64     then obtain pa where "a = real_of_preal pa"
    66 apply (frule_tac [2] real_sup_lemma2)
    65       by (auto simp add: real_gt_zero_preal_Ex)
    67 apply (frule real_sup_lemma2, assumption+, clarify) 
    66     hence "pa \<in> ?pP" using a_in_P by auto
    68 apply (rule preal_complete [THEN iffD2, THEN bexE])
    67     hence pP_not_empty: "?pP \<noteq> {}" by auto
    69 prefer 3 apply blast
    68 
    70 apply (blast intro!: order_less_imp_le)+
    69     obtain sup where sup: "\<forall>x \<in> P. x < sup"
    71 done
    70       using upper_bound_Ex ..
    72 
    71     hence  "a < sup" ..
    73 (*--------------------------------------------------------
    72     hence "0 < sup" using a_pos by arith
    74    Completeness properties using isUb, isLub etc.
    73     then obtain possup where "sup = real_of_preal possup"
    75  -------------------------------------------------------*)
    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 *}
    76 
   122 
    77 lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
   123 lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
    78 apply (frule isLub_isUb)
   124   apply (frule isLub_isUb)
    79 apply (frule_tac x = y in isLub_isUb)
   125   apply (frule_tac x = y in isLub_isUb)
    80 apply (blast intro!: order_antisym dest!: isLub_le_isUb)
   126   apply (blast intro!: order_antisym dest!: isLub_le_isUb)
    81 done
   127   done
    82 
   128 
    83 lemma real_order_restrict: "[| (x::real) <=* S'; S <= S' |] ==> x <=* S"
   129 
    84 by (unfold setle_def setge_def, blast)
   130 text {*
    85 
   131   \medskip Completeness theorem for the positive reals (again).
    86 (*----------------------------------------------------------------
   132 *}
    87            Completeness theorem for the positive reals(again)
       
    88  ----------------------------------------------------------------*)
       
    89 
   133 
    90 lemma posreals_complete:
   134 lemma posreals_complete:
    91      "[| \<forall>x \<in>S. 0 < x;  
   135   assumes positive_S: "\<forall>x \<in> S. 0 < x"
    92          \<exists>x. x \<in>S;  
   136     and not_empty_S: "\<exists>x. x \<in> S"
    93          \<exists>u. isUb (UNIV::real set) S u  
   137     and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
    94       |] ==> \<exists>t. isLub (UNIV::real set) S t"
   138   shows "\<exists>t. isLub (UNIV::real set) S t"
    95 apply (rule_tac x = "real_of_preal (psup ({w. real_of_preal w \<in> S}))" in exI)
   139 proof
    96 apply (auto simp add: isLub_def leastP_def isUb_def)
   140   let ?pS = "{w. real_of_preal w \<in> S}"
    97 apply (auto intro!: setleI setgeI dest!: real_gt_zero_preal_Ex [THEN iffD1])
   141 
    98 apply (frule_tac x = y in bspec, assumption)
   142   obtain u where "isUb UNIV S u" using upper_bound_Ex ..
    99 apply (drule real_gt_zero_preal_Ex [THEN iffD1])
   143   hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
   100 apply (auto simp add: real_of_preal_le_iff)
   144 
   101 apply (frule_tac y = "real_of_preal ya" in setleD, assumption)
   145   obtain x where x_in_S: "x \<in> S" using not_empty_S ..
   102 apply (frule real_ge_preal_preal_Ex, safe)
   146   hence x_gt_zero: "0 < x" using positive_S by simp
   103 apply (blast intro!: preal_psup_le dest!: setleD intro: real_of_preal_le_iff [THEN iffD1])
   147   have  "x \<le> u" using sup and x_in_S ..
   104 apply (frule_tac x = x in bspec, assumption)
   148   hence "0 < u" using x_gt_zero by arith
   105 apply (frule isUbD2)
   149 
   106 apply (drule real_gt_zero_preal_Ex [THEN iffD1])
   150   then obtain pu where u_is_pu: "u = real_of_preal pu"
   107 apply (auto dest!: isUbD real_ge_preal_preal_Ex simp add: real_of_preal_le_iff)
   151     by (auto simp add: real_gt_zero_preal_Ex)
   108 apply (blast dest!: setleD intro!: psup_le_ub intro: real_of_preal_le_iff [THEN iffD1])
   152 
   109 done
   153   have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
   110 
   154   proof
   111 
   155     fix pa
   112 (*-------------------------------
   156     assume "pa \<in> ?pS"
   113     Lemmas
   157     then obtain a where "a \<in> S" and "a = real_of_preal pa"
   114  -------------------------------*)
   158       by simp
   115 lemma real_sup_lemma3: "\<forall>y \<in> {z. \<exists>x \<in> P. z = x + (-xa) + 1} Int {x. 0 < x}. 0 < y"
   159     moreover hence "a \<le> u" using sup by simp
   116 by auto
   160     ultimately show "pa \<le> pu"
   117  
   161       using sup and u_is_pu by (simp add: real_of_preal_le_iff)
   118 lemma lemma_le_swap2: "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))"
   162   qed
   119 by auto
   163 
   120 
   164   have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
   121 lemma lemma_real_complete2b: "[| (x::real) + (-X) + 1 <= S; xa <= x |] ==> xa <= S + X + (- 1)"
   165   proof
   122 by arith
   166     fix y
   123 
   167     assume y_in_S: "y \<in> S"
   124 (*----------------------------------------------------------
   168     hence "0 < y" using positive_S by simp
   125       reals Completeness (again!)
   169     then obtain py where y_is_py: "y = real_of_preal py"
   126  ----------------------------------------------------------*)
   170       by (auto simp add: real_gt_zero_preal_Ex)
   127 lemma reals_complete: "[| \<exists>X. X \<in>S;  \<exists>Y. isUb (UNIV::real set) S Y |]   
   171     hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
   128       ==> \<exists>t. isLub (UNIV :: real set) S t"
   172     with pS_less_pu have "py \<le> psup ?pS"
   129 apply safe
   173       by (rule preal_psup_le)
   130 apply (subgoal_tac "\<exists>u. u\<in> {z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}")
   174     thus "y \<le> real_of_preal (psup ?pS)"
   131 apply (subgoal_tac "isUb (UNIV::real set) ({z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}) (Y + (-X) + 1) ")
   175       using y_is_py by (simp add: real_of_preal_le_iff)
   132 apply (cut_tac P = S and xa = X in real_sup_lemma3)
   176   qed
   133 apply (frule posreals_complete [OF _ _ exI], blast, blast, safe)
   177 
   134 apply (rule_tac x = "t + X + (- 1) " in exI)
   178   moreover {
   135 apply (rule isLubI2)
   179     fix x
   136 apply (rule_tac [2] setgeI, safe)
   180     assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
   137 apply (subgoal_tac [2] "isUb (UNIV:: real set) ({z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}) (y + (-X) + 1) ")
   181     have "real_of_preal (psup ?pS) \<le> x"
   138 apply (drule_tac [2] y = " (y + (- X) + 1) " in isLub_le_isUb)
   182     proof -
   139  prefer 2 apply assumption
   183       obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
   140  prefer 2
   184       hence s_pos: "0 < s" using positive_S by simp
   141 apply arith
   185 
   142 apply (rule setleI [THEN isUbI], safe)
   186       hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
   143 apply (rule_tac x = x and y = y in linorder_cases)
   187       then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
   144 apply (subst lemma_le_swap2)
   188       hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
   145 apply (frule isLubD2)
   189 
   146  prefer 2 apply assumption
   190       from x_ub_S have "s \<le> x" using s_in_S ..
   147 apply safe
   191       hence "0 < x" using s_pos by simp
   148 apply blast
   192       hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
   149 apply arith
   193       then obtain "px" where x_is_px: "x = real_of_preal px" ..
   150 apply (subst lemma_le_swap2)
   194 
   151 apply (frule isLubD2)
   195       have "\<forall>pe \<in> ?pS. pe \<le> px"
   152  prefer 2 apply assumption
   196       proof
   153 apply blast
   197 	fix pe
   154 apply (rule lemma_real_complete2b)
   198 	assume "pe \<in> ?pS"
   155 apply (erule_tac [2] order_less_imp_le)
   199 	hence "real_of_preal pe \<in> S" by simp
   156 apply (blast intro!: isLubD2, blast) 
   200 	hence "real_of_preal pe \<le> x" using x_ub_S by simp
   157 apply (simp (no_asm_use) add: add_assoc)
   201 	thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
   158 apply (blast dest: isUbD intro!: setleI [THEN isUbI] intro: add_right_mono)
   202       qed
   159 apply (blast dest: isUbD intro!: setleI [THEN isUbI] intro: add_right_mono, auto)
   203 
   160 done
   204       moreover have "?pS \<noteq> {}" using ps_in_pS by auto
   161 
   205       ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
   162 
   206       thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
   163 subsection{*Corollary: the Archimedean Property of the Reals*}
   207     qed
   164 
   208   }
   165 lemma reals_Archimedean: "0 < x ==> \<exists>n. inverse (real(Suc n)) < x"
   209   ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
   166 apply (rule ccontr)
   210     by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
   167 apply (subgoal_tac "\<forall>n. x * real (Suc n) <= 1")
   211 qed
   168  prefer 2
   212 
   169 apply (simp add: linorder_not_less inverse_eq_divide, clarify) 
   213 text {*
   170 apply (drule_tac x = n in spec)
   214   \medskip reals Completeness (again!)
   171 apply (drule_tac c = "real (Suc n)"  in mult_right_mono)
   215 *}
   172 apply (rule real_of_nat_ge_zero)
   216 
   173 apply (simp add: times_divide_eq_right real_of_nat_Suc_gt_zero [THEN real_not_refl2, THEN not_sym] mult_commute)
   217 lemma reals_complete:
   174 apply (subgoal_tac "isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} 1")
   218   assumes notempty_S: "\<exists>X. X \<in> S"
   175 apply (subgoal_tac "\<exists>X. X \<in> {z. \<exists>n. z = x* (real (Suc n))}")
   219     and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
   176 apply (drule reals_complete)
   220   shows "\<exists>t. isLub (UNIV :: real set) S t"
   177 apply (auto intro: isUbI setleI)
   221 proof -
   178 apply (subgoal_tac "\<forall>m. x* (real (Suc m)) <= t")
   222   obtain X where X_in_S: "X \<in> S" using notempty_S ..
   179 apply (simp add: real_of_nat_Suc right_distrib)
   223   obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
   180 prefer 2 apply (blast intro: isLubD2)
   224     using exists_Ub ..
   181 apply (simp add: le_diff_eq [symmetric] real_diff_def)
   225   let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
   182 apply (subgoal_tac "isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} (t + (-x))")
   226 
   183 prefer 2 apply (blast intro!: isUbI setleI)
   227   {
   184 apply (drule_tac y = "t+ (-x) " in isLub_le_isUb)
   228     fix x
   185 apply (auto simp add: real_of_nat_Suc right_distrib)
   229     assume "isUb (UNIV::real set) S x"
   186 done
   230     hence S_le_x: "\<forall> y \<in> S. y <= x"
   187 
   231       by (simp add: isUb_def setle_def)
   188 (*There must be other proofs, e.g. Suc of the largest integer in the
   232     {
   189   cut representing x*)
   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 
   190 lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
   362 lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
   191 apply (rule_tac x = x and y = 0 in linorder_cases)
   363 proof cases
   192 apply (rule_tac x = 0 in exI)
   364   assume "x \<le> 0"
   193 apply (rule_tac [2] x = 1 in exI)
   365   hence "x < real (1::nat)" by simp
   194 apply (auto elim: order_less_trans simp add: real_of_nat_one)
   366   thus ?thesis ..
   195 apply (frule positive_imp_inverse_positive [THEN reals_Archimedean], safe)
   367 next
   196 apply (rule_tac x = "Suc n" in exI)
   368   assume "\<not> x \<le> 0"
   197 apply (frule_tac b = "inverse x" in mult_strict_right_mono, auto)
   369   hence x_greater_zero: "0 < x" by simp
   198 done
   370   hence "0 < inverse x" by simp
   199 
   371   then obtain n where "inverse (real (Suc n)) < inverse x"
   200 lemma reals_Archimedean3: "0 < x ==> \<forall>y. \<exists>(n::nat). y < real n * x"
   372     using reals_Archimedean by blast
   201 apply safe
   373   hence "inverse (real (Suc n)) * x < inverse x * x"
   202 apply (cut_tac x = "y*inverse (x) " in reals_Archimedean2)
   374     using x_greater_zero by (rule mult_strict_right_mono)
   203 apply safe
   375   hence "inverse (real (Suc n)) * x < 1"
   204 apply (frule_tac a = "y * inverse x" in mult_strict_right_mono)
   376     using x_greater_zero by (simp add: real_mult_inverse_left mult_commute)
   205 apply (auto simp add: mult_assoc real_of_nat_def)
   377   hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1"
   206 done
   378     by (rule mult_strict_left_mono) simp
       
   379   hence "x < real (Suc n)"
       
   380     by (simp add: mult_commute ring_eq_simps real_mult_inverse_left)
       
   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: mult_commute ring_eq_simps)
       
   396   hence "y < real (n::nat) * x"
       
   397     using x_not_zero by (simp add: real_mult_inverse_left ring_eq_simps)
       
   398   thus "\<exists>(n::nat). y < real n * x" ..
       
   399 qed
   207 
   400 
   208 lemma reals_Archimedean6:
   401 lemma reals_Archimedean6:
   209      "0 \<le> r ==> \<exists>(n::nat). real (n - 1) \<le> r & r < real (n)"
   402      "0 \<le> r ==> \<exists>(n::nat). real (n - 1) \<le> r & r < real (n)"
   210 apply (insert reals_Archimedean2 [of r], safe)
   403 apply (insert reals_Archimedean2 [of r], safe)
   211 apply (frule_tac P = "%k. r < real k" and k = n and m = "%x. x"
   404 apply (frule_tac P = "%k. r < real k" and k = n and m = "%x. x"
   215 apply (rename_tac x')
   408 apply (rename_tac x')
   216 apply (drule_tac x = x' in spec, simp)
   409 apply (drule_tac x = x' in spec, simp)
   217 done
   410 done
   218 
   411 
   219 lemma reals_Archimedean6a: "0 \<le> r ==> \<exists>n. real (n) \<le> r & r < real (Suc n)"
   412 lemma reals_Archimedean6a: "0 \<le> r ==> \<exists>n. real (n) \<le> r & r < real (Suc n)"
   220 by (drule reals_Archimedean6, auto)
   413   by (drule reals_Archimedean6) auto
   221 
   414 
   222 lemma reals_Archimedean_6b_int:
   415 lemma reals_Archimedean_6b_int:
   223      "0 \<le> r ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
   416      "0 \<le> r ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
   224 apply (drule reals_Archimedean6a, auto)
   417 apply (drule reals_Archimedean6a, auto)
   225 apply (rule_tac x = "int n" in exI)
   418 apply (rule_tac x = "int n" in exI)
   239 ML
   432 ML
   240 {*
   433 {*
   241 val real_sum_of_halves = thm "real_sum_of_halves";
   434 val real_sum_of_halves = thm "real_sum_of_halves";
   242 val posreal_complete = thm "posreal_complete";
   435 val posreal_complete = thm "posreal_complete";
   243 val real_isLub_unique = thm "real_isLub_unique";
   436 val real_isLub_unique = thm "real_isLub_unique";
   244 val real_order_restrict = thm "real_order_restrict";
       
   245 val posreals_complete = thm "posreals_complete";
   437 val posreals_complete = thm "posreals_complete";
   246 val reals_complete = thm "reals_complete";
   438 val reals_complete = thm "reals_complete";
   247 val reals_Archimedean = thm "reals_Archimedean";
   439 val reals_Archimedean = thm "reals_Archimedean";
   248 val reals_Archimedean2 = thm "reals_Archimedean2";
   440 val reals_Archimedean2 = thm "reals_Archimedean2";
   249 val reals_Archimedean3 = thm "reals_Archimedean3";
   441 val reals_Archimedean3 = thm "reals_Archimedean3";
   376 lemma real_of_int_floor_cancel [simp]:
   568 lemma real_of_int_floor_cancel [simp]:
   377     "(real (floor x) = x) = (\<exists>n::int. x = real n)"
   569     "(real (floor x) = x) = (\<exists>n::int. x = real n)"
   378 apply (simp add: floor_def Least_def)
   570 apply (simp add: floor_def Least_def)
   379 apply (insert real_lb_ub_int [of x], erule exE)
   571 apply (insert real_lb_ub_int [of x], erule exE)
   380 apply (rule theI2)
   572 apply (rule theI2)
   381 apply (auto intro: lemma_floor) 
   573 apply (auto intro: lemma_floor)
   382 done
   574 done
   383 
   575 
   384 lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
   576 lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
   385 apply (simp add: floor_def)
   577 apply (simp add: floor_def)
   386 apply (rule Least_equality)
   578 apply (rule Least_equality)
   444 lemma le_floor: "real a <= x ==> a <= floor x"
   636 lemma le_floor: "real a <= x ==> a <= floor x"
   445   apply (subgoal_tac "a < floor x + 1")
   637   apply (subgoal_tac "a < floor x + 1")
   446   apply arith
   638   apply arith
   447   apply (subst real_of_int_less_iff [THEN sym])
   639   apply (subst real_of_int_less_iff [THEN sym])
   448   apply simp
   640   apply simp
   449   apply (insert real_of_int_floor_add_one_gt [of x]) 
   641   apply (insert real_of_int_floor_add_one_gt [of x])
   450   apply arith
   642   apply arith
   451 done
   643 done
   452 
   644 
   453 lemma real_le_floor: "a <= floor x ==> real a <= x"
   645 lemma real_le_floor: "a <= floor x ==> real a <= x"
   454   apply (rule order_trans)
   646   apply (rule order_trans)
   462   apply (rule iffI)
   654   apply (rule iffI)
   463   apply (erule real_le_floor)
   655   apply (erule real_le_floor)
   464   apply (erule le_floor)
   656   apply (erule le_floor)
   465 done
   657 done
   466 
   658 
   467 lemma le_floor_eq_number_of [simp]: 
   659 lemma le_floor_eq_number_of [simp]:
   468     "(number_of n <= floor x) = (number_of n <= x)"
   660     "(number_of n <= floor x) = (number_of n <= x)"
   469 by (simp add: le_floor_eq)
   661 by (simp add: le_floor_eq)
   470 
   662 
   471 lemma le_floor_eq_zero [simp]: "(0 <= floor x) = (0 <= x)"
   663 lemma le_floor_eq_zero [simp]: "(0 <= floor x) = (0 <= x)"
   472 by (simp add: le_floor_eq)
   664 by (simp add: le_floor_eq)
   478   apply (subst linorder_not_le [THEN sym])+
   670   apply (subst linorder_not_le [THEN sym])+
   479   apply simp
   671   apply simp
   480   apply (rule le_floor_eq)
   672   apply (rule le_floor_eq)
   481 done
   673 done
   482 
   674 
   483 lemma floor_less_eq_number_of [simp]: 
   675 lemma floor_less_eq_number_of [simp]:
   484     "(floor x < number_of n) = (x < number_of n)"
   676     "(floor x < number_of n) = (x < number_of n)"
   485 by (simp add: floor_less_eq)
   677 by (simp add: floor_less_eq)
   486 
   678 
   487 lemma floor_less_eq_zero [simp]: "(floor x < 0) = (x < 0)"
   679 lemma floor_less_eq_zero [simp]: "(floor x < 0) = (x < 0)"
   488 by (simp add: floor_less_eq)
   680 by (simp add: floor_less_eq)
   493 lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
   685 lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
   494   apply (insert le_floor_eq [of "a + 1" x])
   686   apply (insert le_floor_eq [of "a + 1" x])
   495   apply auto
   687   apply auto
   496 done
   688 done
   497 
   689 
   498 lemma less_floor_eq_number_of [simp]: 
   690 lemma less_floor_eq_number_of [simp]:
   499     "(number_of n < floor x) = (number_of n + 1 <= x)"
   691     "(number_of n < floor x) = (number_of n + 1 <= x)"
   500 by (simp add: less_floor_eq)
   692 by (simp add: less_floor_eq)
   501 
   693 
   502 lemma less_floor_eq_zero [simp]: "(0 < floor x) = (1 <= x)"
   694 lemma less_floor_eq_zero [simp]: "(0 < floor x) = (1 <= x)"
   503 by (simp add: less_floor_eq)
   695 by (simp add: less_floor_eq)
   508 lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
   700 lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
   509   apply (insert floor_less_eq [of x "a + 1"])
   701   apply (insert floor_less_eq [of x "a + 1"])
   510   apply auto
   702   apply auto
   511 done
   703 done
   512 
   704 
   513 lemma floor_le_eq_number_of [simp]: 
   705 lemma floor_le_eq_number_of [simp]:
   514     "(floor x <= number_of n) = (x < number_of n + 1)"
   706     "(floor x <= number_of n) = (x < number_of n + 1)"
   515 by (simp add: floor_le_eq)
   707 by (simp add: floor_le_eq)
   516 
   708 
   517 lemma floor_le_eq_zero [simp]: "(floor x <= 0) = (x < 1)"
   709 lemma floor_le_eq_zero [simp]: "(floor x <= 0) = (x < 1)"
   518 by (simp add: floor_le_eq)
   710 by (simp add: floor_le_eq)
   533   apply arith
   725   apply arith
   534   apply (rule real_of_int_floor_le)
   726   apply (rule real_of_int_floor_le)
   535   apply (rule real_of_int_floor_add_one_gt)
   727   apply (rule real_of_int_floor_add_one_gt)
   536   apply (subgoal_tac "floor (x + real a) < floor x + a + 1")
   728   apply (subgoal_tac "floor (x + real a) < floor x + a + 1")
   537   apply arith
   729   apply arith
   538   apply (subst real_of_int_less_iff [THEN sym])  
   730   apply (subst real_of_int_less_iff [THEN sym])
   539   apply simp
   731   apply simp
   540   apply (subgoal_tac "real(floor(x + real a)) <= x + real a") 
   732   apply (subgoal_tac "real(floor(x + real a)) <= x + real a")
   541   apply (subgoal_tac "x < real(floor x) + 1")
   733   apply (subgoal_tac "x < real(floor x) + 1")
   542   apply arith
   734   apply arith
   543   apply (rule real_of_int_floor_add_one_gt)
   735   apply (rule real_of_int_floor_add_one_gt)
   544   apply (rule real_of_int_floor_le)
   736   apply (rule real_of_int_floor_le)
   545 done
   737 done
   546 
   738 
   547 lemma floor_add_number_of [simp]: 
   739 lemma floor_add_number_of [simp]:
   548     "floor (x + number_of n) = floor x + number_of n"
   740     "floor (x + number_of n) = floor x + number_of n"
   549   apply (subst floor_add [THEN sym])
   741   apply (subst floor_add [THEN sym])
   550   apply simp
   742   apply simp
   551 done
   743 done
   552 
   744 
   559   apply (subst diff_minus)+
   751   apply (subst diff_minus)+
   560   apply (subst real_of_int_minus [THEN sym])
   752   apply (subst real_of_int_minus [THEN sym])
   561   apply (rule floor_add)
   753   apply (rule floor_add)
   562 done
   754 done
   563 
   755 
   564 lemma floor_subtract_number_of [simp]: "floor (x - number_of n) = 
   756 lemma floor_subtract_number_of [simp]: "floor (x - number_of n) =
   565     floor x - number_of n"
   757     floor x - number_of n"
   566   apply (subst floor_subtract [THEN sym])
   758   apply (subst floor_subtract [THEN sym])
   567   apply simp
   759   apply simp
   568 done
   760 done
   569 
   761 
   659   apply (rule iffI)
   851   apply (rule iffI)
   660   apply (erule ceiling_le_real)
   852   apply (erule ceiling_le_real)
   661   apply (erule ceiling_le)
   853   apply (erule ceiling_le)
   662 done
   854 done
   663 
   855 
   664 lemma ceiling_le_eq_number_of [simp]: 
   856 lemma ceiling_le_eq_number_of [simp]:
   665     "(ceiling x <= number_of n) = (x <= number_of n)"
   857     "(ceiling x <= number_of n) = (x <= number_of n)"
   666 by (simp add: ceiling_le_eq)
   858 by (simp add: ceiling_le_eq)
   667 
   859 
   668 lemma ceiling_le_zero_eq [simp]: "(ceiling x <= 0) = (x <= 0)" 
   860 lemma ceiling_le_zero_eq [simp]: "(ceiling x <= 0) = (x <= 0)"
   669 by (simp add: ceiling_le_eq)
   861 by (simp add: ceiling_le_eq)
   670 
   862 
   671 lemma ceiling_le_eq_one [simp]: "(ceiling x <= 1) = (x <= 1)" 
   863 lemma ceiling_le_eq_one [simp]: "(ceiling x <= 1) = (x <= 1)"
   672 by (simp add: ceiling_le_eq)
   864 by (simp add: ceiling_le_eq)
   673 
   865 
   674 lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
   866 lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
   675   apply (subst linorder_not_le [THEN sym])+
   867   apply (subst linorder_not_le [THEN sym])+
   676   apply simp
   868   apply simp
   677   apply (rule ceiling_le_eq)
   869   apply (rule ceiling_le_eq)
   678 done
   870 done
   679 
   871 
   680 lemma less_ceiling_eq_number_of [simp]: 
   872 lemma less_ceiling_eq_number_of [simp]:
   681     "(number_of n < ceiling x) = (number_of n < x)"
   873     "(number_of n < ceiling x) = (number_of n < x)"
   682 by (simp add: less_ceiling_eq)
   874 by (simp add: less_ceiling_eq)
   683 
   875 
   684 lemma less_ceiling_eq_zero [simp]: "(0 < ceiling x) = (0 < x)"
   876 lemma less_ceiling_eq_zero [simp]: "(0 < ceiling x) = (0 < x)"
   685 by (simp add: less_ceiling_eq)
   877 by (simp add: less_ceiling_eq)
   690 lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
   882 lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
   691   apply (insert ceiling_le_eq [of x "a - 1"])
   883   apply (insert ceiling_le_eq [of x "a - 1"])
   692   apply auto
   884   apply auto
   693 done
   885 done
   694 
   886 
   695 lemma ceiling_less_eq_number_of [simp]: 
   887 lemma ceiling_less_eq_number_of [simp]:
   696     "(ceiling x < number_of n) = (x <= number_of n - 1)"
   888     "(ceiling x < number_of n) = (x <= number_of n - 1)"
   697 by (simp add: ceiling_less_eq)
   889 by (simp add: ceiling_less_eq)
   698 
   890 
   699 lemma ceiling_less_eq_zero [simp]: "(ceiling x < 0) = (x <= -1)"
   891 lemma ceiling_less_eq_zero [simp]: "(ceiling x < 0) = (x <= -1)"
   700 by (simp add: ceiling_less_eq)
   892 by (simp add: ceiling_less_eq)
   705 lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
   897 lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
   706   apply (insert less_ceiling_eq [of "a - 1" x])
   898   apply (insert less_ceiling_eq [of "a - 1" x])
   707   apply auto
   899   apply auto
   708 done
   900 done
   709 
   901 
   710 lemma le_ceiling_eq_number_of [simp]: 
   902 lemma le_ceiling_eq_number_of [simp]:
   711     "(number_of n <= ceiling x) = (number_of n - 1 < x)"
   903     "(number_of n <= ceiling x) = (number_of n - 1 < x)"
   712 by (simp add: le_ceiling_eq)
   904 by (simp add: le_ceiling_eq)
   713 
   905 
   714 lemma le_ceiling_eq_zero [simp]: "(0 <= ceiling x) = (-1 < x)"
   906 lemma le_ceiling_eq_zero [simp]: "(0 <= ceiling x) = (-1 < x)"
   715 by (simp add: le_ceiling_eq)
   907 by (simp add: le_ceiling_eq)
   722   apply (subst real_of_int_minus [THEN sym])
   914   apply (subst real_of_int_minus [THEN sym])
   723   apply (subst floor_add)
   915   apply (subst floor_add)
   724   apply simp
   916   apply simp
   725 done
   917 done
   726 
   918 
   727 lemma ceiling_add_number_of [simp]: "ceiling (x + number_of n) = 
   919 lemma ceiling_add_number_of [simp]: "ceiling (x + number_of n) =
   728     ceiling x + number_of n"
   920     ceiling x + number_of n"
   729   apply (subst ceiling_add [THEN sym])
   921   apply (subst ceiling_add [THEN sym])
   730   apply simp
   922   apply simp
   731 done
   923 done
   732 
   924 
   739   apply (subst diff_minus)+
   931   apply (subst diff_minus)+
   740   apply (subst real_of_int_minus [THEN sym])
   932   apply (subst real_of_int_minus [THEN sym])
   741   apply (rule ceiling_add)
   933   apply (rule ceiling_add)
   742 done
   934 done
   743 
   935 
   744 lemma ceiling_subtract_number_of [simp]: "ceiling (x - number_of n) = 
   936 lemma ceiling_subtract_number_of [simp]: "ceiling (x - number_of n) =
   745     ceiling x - number_of n"
   937     ceiling x - number_of n"
   746   apply (subst ceiling_subtract [THEN sym])
   938   apply (subst ceiling_subtract [THEN sym])
   747   apply simp
   939   apply simp
   748 done
   940 done
   749 
   941 
   788 lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
   980 lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
   789   apply (case_tac "0 <= x")
   981   apply (case_tac "0 <= x")
   790   apply (subst natfloor_def)+
   982   apply (subst natfloor_def)+
   791   apply (subst nat_le_eq_zle)
   983   apply (subst nat_le_eq_zle)
   792   apply force
   984   apply force
   793   apply (erule floor_mono2) 
   985   apply (erule floor_mono2)
   794   apply (subst natfloor_neg)
   986   apply (subst natfloor_neg)
   795   apply simp
   987   apply simp
   796   apply simp
   988   apply simp
   797 done
   989 done
   798 
   990 
   813   apply (subst real_of_nat_le_iff)
  1005   apply (subst real_of_nat_le_iff)
   814   apply assumption
  1006   apply assumption
   815   apply (erule le_natfloor)
  1007   apply (erule le_natfloor)
   816 done
  1008 done
   817 
  1009 
   818 lemma le_natfloor_eq_number_of [simp]: 
  1010 lemma le_natfloor_eq_number_of [simp]:
   819     "~ neg((number_of n)::int) ==> 0 <= x ==>
  1011     "~ neg((number_of n)::int) ==> 0 <= x ==>
   820       (number_of n <= natfloor x) = (number_of n <= x)"
  1012       (number_of n <= natfloor x) = (number_of n <= x)"
   821   apply (subst le_natfloor_eq, assumption)
  1013   apply (subst le_natfloor_eq, assumption)
   822   apply simp
  1014   apply simp
   823 done
  1015 done
   824 
  1016 
   825 lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
  1017 lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
   826   apply (case_tac "0 <= x")
  1018   apply (case_tac "0 <= x")
   827   apply (subst le_natfloor_eq, assumption, simp)
  1019   apply (subst le_natfloor_eq, assumption, simp)
   828   apply (rule iffI)
  1020   apply (rule iffI)
   829   apply (subgoal_tac "natfloor x <= natfloor 0") 
  1021   apply (subgoal_tac "natfloor x <= natfloor 0")
   830   apply simp
  1022   apply simp
   831   apply (rule natfloor_mono)
  1023   apply (rule natfloor_mono)
   832   apply simp
  1024   apply simp
   833   apply simp
  1025   apply simp
   834 done
  1026 done
   867   apply (erule ssubst)
  1059   apply (erule ssubst)
   868   apply (simp add: nat_add_distrib)
  1060   apply (simp add: nat_add_distrib)
   869   apply simp
  1061   apply simp
   870 done
  1062 done
   871 
  1063 
   872 lemma natfloor_add_number_of [simp]: 
  1064 lemma natfloor_add_number_of [simp]:
   873     "~neg ((number_of n)::int) ==> 0 <= x ==> 
  1065     "~neg ((number_of n)::int) ==> 0 <= x ==>
   874       natfloor (x + number_of n) = natfloor x + number_of n"
  1066       natfloor (x + number_of n) = natfloor x + number_of n"
   875   apply (subst natfloor_add [THEN sym])
  1067   apply (subst natfloor_add [THEN sym])
   876   apply simp_all
  1068   apply simp_all
   877 done
  1069 done
   878 
  1070 
   880   apply (subst natfloor_add [THEN sym])
  1072   apply (subst natfloor_add [THEN sym])
   881   apply assumption
  1073   apply assumption
   882   apply simp
  1074   apply simp
   883 done
  1075 done
   884 
  1076 
   885 lemma natfloor_subtract [simp]: "real a <= x ==> 
  1077 lemma natfloor_subtract [simp]: "real a <= x ==>
   886     natfloor(x - real a) = natfloor x - a"
  1078     natfloor(x - real a) = natfloor x - a"
   887   apply (unfold natfloor_def)
  1079   apply (unfold natfloor_def)
   888   apply (subgoal_tac "real a = real (int a)")
  1080   apply (subgoal_tac "real a = real (int a)")
   889   apply (erule ssubst)
  1081   apply (erule ssubst)
   890   apply simp
  1082   apply simp
   960   apply (subst real_of_nat_le_iff)
  1152   apply (subst real_of_nat_le_iff)
   961   apply assumption
  1153   apply assumption
   962   apply (erule natceiling_le)
  1154   apply (erule natceiling_le)
   963 done
  1155 done
   964 
  1156 
   965 lemma natceiling_le_eq_number_of [simp]: 
  1157 lemma natceiling_le_eq_number_of [simp]:
   966     "~ neg((number_of n)::int) ==> 0 <= x ==>
  1158     "~ neg((number_of n)::int) ==> 0 <= x ==>
   967       (natceiling x <= number_of n) = (x <= number_of n)"
  1159       (natceiling x <= number_of n) = (x <= number_of n)"
   968   apply (subst natceiling_le_eq, assumption)
  1160   apply (subst natceiling_le_eq, assumption)
   969   apply simp
  1161   apply simp
   970 done
  1162 done
   994   apply (simp, simp)
  1186   apply (simp, simp)
   995   apply (subst nat_add_distrib)
  1187   apply (subst nat_add_distrib)
   996   apply auto
  1188   apply auto
   997 done
  1189 done
   998 
  1190 
   999 lemma natceiling_add [simp]: "0 <= x ==> 
  1191 lemma natceiling_add [simp]: "0 <= x ==>
  1000     natceiling (x + real a) = natceiling x + a"
  1192     natceiling (x + real a) = natceiling x + a"
  1001   apply (unfold natceiling_def)
  1193   apply (unfold natceiling_def)
  1002   apply (subgoal_tac "real a = real (int a)")
  1194   apply (subgoal_tac "real a = real (int a)")
  1003   apply (erule ssubst)
  1195   apply (erule ssubst)
  1004   apply simp
  1196   apply simp
  1007   apply (erule ssubst)
  1199   apply (erule ssubst)
  1008   apply (erule ceiling_mono2)
  1200   apply (erule ceiling_mono2)
  1009   apply simp_all
  1201   apply simp_all
  1010 done
  1202 done
  1011 
  1203 
  1012 lemma natceiling_add_number_of [simp]: 
  1204 lemma natceiling_add_number_of [simp]:
  1013     "~ neg ((number_of n)::int) ==> 0 <= x ==> 
  1205     "~ neg ((number_of n)::int) ==> 0 <= x ==>
  1014       natceiling (x + number_of n) = natceiling x + number_of n"
  1206       natceiling (x + number_of n) = natceiling x + number_of n"
  1015   apply (subst natceiling_add [THEN sym])
  1207   apply (subst natceiling_add [THEN sym])
  1016   apply simp_all
  1208   apply simp_all
  1017 done
  1209 done
  1018 
  1210 
  1020   apply (subst natceiling_add [THEN sym])
  1212   apply (subst natceiling_add [THEN sym])
  1021   apply assumption
  1213   apply assumption
  1022   apply simp
  1214   apply simp
  1023 done
  1215 done
  1024 
  1216 
  1025 lemma natceiling_subtract [simp]: "real a <= x ==> 
  1217 lemma natceiling_subtract [simp]: "real a <= x ==>
  1026     natceiling(x - real a) = natceiling x - a"
  1218     natceiling(x - real a) = natceiling x - a"
  1027   apply (unfold natceiling_def)
  1219   apply (unfold natceiling_def)
  1028   apply (subgoal_tac "real a = real (int a)")
  1220   apply (subgoal_tac "real a = real (int a)")
  1029   apply (erule ssubst)
  1221   apply (erule ssubst)
  1030   apply simp
  1222   apply simp
  1035   apply (rule ceiling_mono2)
  1227   apply (rule ceiling_mono2)
  1036   apply assumption
  1228   apply assumption
  1037   apply simp_all
  1229   apply simp_all
  1038 done
  1230 done
  1039 
  1231 
  1040 lemma natfloor_div_nat: "1 <= x ==> 0 < y ==> 
  1232 lemma natfloor_div_nat: "1 <= x ==> 0 < y ==>
  1041   natfloor (x / real y) = natfloor x div y"
  1233   natfloor (x / real y) = natfloor x div y"
  1042 proof -
  1234 proof -
  1043   assume "1 <= (x::real)" and "0 < (y::nat)"
  1235   assume "1 <= (x::real)" and "0 < (y::nat)"
  1044   have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y"
  1236   have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y"
  1045     by simp
  1237     by simp
  1046   then have a: "real(natfloor x) = real ((natfloor x) div y) * real y + 
  1238   then have a: "real(natfloor x) = real ((natfloor x) div y) * real y +
  1047     real((natfloor x) mod y)"
  1239     real((natfloor x) mod y)"
  1048     by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym])
  1240     by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym])
  1049   have "x = real(natfloor x) + (x - real(natfloor x))"
  1241   have "x = real(natfloor x) + (x - real(natfloor x))"
  1050     by simp
  1242     by simp
  1051   then have "x = real ((natfloor x) div y) * real y + 
  1243   then have "x = real ((natfloor x) div y) * real y +
  1052       real((natfloor x) mod y) + (x - real(natfloor x))"
  1244       real((natfloor x) mod y) + (x - real(natfloor x))"
  1053     by (simp add: a)
  1245     by (simp add: a)
  1054   then have "x / real y = ... / real y"
  1246   then have "x / real y = ... / real y"
  1055     by simp
  1247     by simp
  1056   also have "... = real((natfloor x) div y) + real((natfloor x) mod y) / 
  1248   also have "... = real((natfloor x) div y) + real((natfloor x) mod y) /
  1057     real y + (x - real(natfloor x)) / real y"
  1249     real y + (x - real(natfloor x)) / real y"
  1058     by (auto simp add: ring_distrib ring_eq_simps add_divide_distrib
  1250     by (auto simp add: ring_distrib ring_eq_simps add_divide_distrib
  1059       diff_divide_distrib prems)
  1251       diff_divide_distrib prems)
  1060   finally have "natfloor (x / real y) = natfloor(...)" by simp
  1252   finally have "natfloor (x / real y) = natfloor(...)" by simp
  1061   also have "... = natfloor(real((natfloor x) mod y) / 
  1253   also have "... = natfloor(real((natfloor x) mod y) /
  1062     real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))"
  1254     real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))"
  1063     by (simp add: add_ac)
  1255     by (simp add: add_ac)
  1064   also have "... = natfloor(real((natfloor x) mod y) / 
  1256   also have "... = natfloor(real((natfloor x) mod y) /
  1065     real y + (x - real(natfloor x)) / real y) + (natfloor x) div y"
  1257     real y + (x - real(natfloor x)) / real y) + (natfloor x) div y"
  1066     apply (rule natfloor_add)
  1258     apply (rule natfloor_add)
  1067     apply (rule add_nonneg_nonneg)
  1259     apply (rule add_nonneg_nonneg)
  1068     apply (rule divide_nonneg_pos)
  1260     apply (rule divide_nonneg_pos)
  1069     apply simp
  1261     apply simp
  1071     apply (rule divide_nonneg_pos)
  1263     apply (rule divide_nonneg_pos)
  1072     apply (simp add: compare_rls)
  1264     apply (simp add: compare_rls)
  1073     apply (rule real_natfloor_le)
  1265     apply (rule real_natfloor_le)
  1074     apply (insert prems, auto)
  1266     apply (insert prems, auto)
  1075     done
  1267     done
  1076   also have "natfloor(real((natfloor x) mod y) / 
  1268   also have "natfloor(real((natfloor x) mod y) /
  1077     real y + (x - real(natfloor x)) / real y) = 0"
  1269     real y + (x - real(natfloor x)) / real y) = 0"
  1078     apply (rule natfloor_eq)
  1270     apply (rule natfloor_eq)
  1079     apply simp
  1271     apply simp
  1080     apply (rule add_nonneg_nonneg)
  1272     apply (rule add_nonneg_nonneg)
  1081     apply (rule divide_nonneg_pos)
  1273     apply (rule divide_nonneg_pos)
  1089     apply (simp add: add_divide_distrib [THEN sym])
  1281     apply (simp add: add_divide_distrib [THEN sym])
  1090     apply (subgoal_tac "real y = real y - 1 + 1")
  1282     apply (subgoal_tac "real y = real y - 1 + 1")
  1091     apply (erule ssubst)
  1283     apply (erule ssubst)
  1092     apply (rule add_le_less_mono)
  1284     apply (rule add_le_less_mono)
  1093     apply (simp add: compare_rls)
  1285     apply (simp add: compare_rls)
  1094     apply (subgoal_tac "real(natfloor x mod y) + 1 = 
  1286     apply (subgoal_tac "real(natfloor x mod y) + 1 =
  1095       real(natfloor x mod y + 1)")
  1287       real(natfloor x mod y + 1)")
  1096     apply (erule ssubst)
  1288     apply (erule ssubst)
  1097     apply (subst real_of_nat_le_iff)
  1289     apply (subst real_of_nat_le_iff)
  1098     apply (subgoal_tac "natfloor x mod y < y")
  1290     apply (subgoal_tac "natfloor x mod y < y")
  1099     apply arith
  1291     apply arith
  1157 val ceiling_number_of_eq = thm "ceiling_number_of_eq";
  1349 val ceiling_number_of_eq = thm "ceiling_number_of_eq";
  1158 val real_of_int_ceiling_diff_one_le = thm "real_of_int_ceiling_diff_one_le";
  1350 val real_of_int_ceiling_diff_one_le = thm "real_of_int_ceiling_diff_one_le";
  1159 val real_of_int_ceiling_le_add_one = thm "real_of_int_ceiling_le_add_one";
  1351 val real_of_int_ceiling_le_add_one = thm "real_of_int_ceiling_le_add_one";
  1160 *}
  1352 *}
  1161 
  1353 
  1162 
       
  1163 end
  1354 end