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