src/HOL/RComplete.thy
author haftmann
Thu, 16 Sep 2010 16:51:34 +0200
changeset 39475 9cc1ba3c5706
parent 37887 2ae085b07f2f
child 41550 efa734d9b221
permissions -rw-r--r--
separation of static and dynamic thy context
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30122
1c912a9d8200 standard headers;
wenzelm
parents: 30102
diff changeset
     1
(*  Title:      HOL/RComplete.thy
1c912a9d8200 standard headers;
wenzelm
parents: 30102
diff changeset
     2
    Author:     Jacques D. Fleuriot, University of Edinburgh
1c912a9d8200 standard headers;
wenzelm
parents: 30102
diff changeset
     3
    Author:     Larry Paulson, University of Cambridge
1c912a9d8200 standard headers;
wenzelm
parents: 30102
diff changeset
     4
    Author:     Jeremy Avigad, Carnegie Mellon University
1c912a9d8200 standard headers;
wenzelm
parents: 30102
diff changeset
     5
    Author:     Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
     6
*)
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
     8
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
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14641
diff changeset
    10
theory RComplete
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    11
imports Lubs RealDef
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14641
diff changeset
    12
begin
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    13
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    14
lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    15
  by simp
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    16
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    17
lemma abs_diff_less_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 32960
diff changeset
    18
  "(\<bar>x - a\<bar> < (r::'a::linordered_idom)) = (a - r < x \<and> x < a + r)"
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    19
  by auto
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    20
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    21
subsection {* Completeness of Positive Reals *}
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    22
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    23
text {*
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    24
  Supremum property for the set of positive reals
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    25
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    26
  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
    27
  bound @{text "y"}.  Then @{text "P"} has a least upper bound
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    28
  (written @{text "S"}).
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    29
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    30
  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
    31
*}
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    32
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    33
text {* Only used in HOL/Import/HOL4Compat.thy; delete? *}
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    34
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    35
lemma posreal_complete:
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    36
  assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    37
    and not_empty_P: "\<exists>x. x \<in> P"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    38
    and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    39
  shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    40
proof -
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    41
  from upper_bound_Ex have "\<exists>z. \<forall>x\<in>P. x \<le> z"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    42
    by (auto intro: less_imp_le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    43
  from complete_real [OF not_empty_P this] obtain S
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    44
  where S1: "\<And>x. x \<in> P \<Longrightarrow> x \<le> S" and S2: "\<And>z. \<forall>x\<in>P. x \<le> z \<Longrightarrow> S \<le> z" by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    45
  have "\<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    46
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    47
    fix y show "(\<exists>x\<in>P. y < x) = (y < S)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    48
      apply (cases "\<exists>x\<in>P. y < x", simp_all)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    49
      apply (clarify, drule S1, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    50
      apply (simp add: not_less S2)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    51
      done
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    52
  qed
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    53
  thus ?thesis ..
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    54
qed
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    55
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    56
text {*
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    57
  \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    58
*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    59
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    60
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
    61
  apply (frule isLub_isUb)
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    62
  apply (frule_tac x = y in isLub_isUb)
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    63
  apply (blast intro!: order_antisym dest!: isLub_le_isUb)
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    64
  done
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    65
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    66
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    67
text {*
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    68
  \medskip reals Completeness (again!)
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    69
*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    70
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    71
lemma reals_complete:
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    72
  assumes notempty_S: "\<exists>X. X \<in> S"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    73
    and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    74
  shows "\<exists>t. isLub (UNIV :: real set) S t"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    75
proof -
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    76
  from assms have "\<exists>X. X \<in> S" and "\<exists>Y. \<forall>x\<in>S. x \<le> Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    77
    unfolding isUb_def setle_def by simp_all
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    78
  from complete_real [OF this] show ?thesis
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    79
    unfolding isLub_def leastP_def setle_def setge_def Ball_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    80
      Collect_def mem_def isUb_def UNIV_def by simp
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    81
qed
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    82
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    83
text{*A version of the same theorem without all those predicates!*}
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    84
lemma reals_complete2:
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    85
  fixes S :: "(real set)"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    86
  assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x"
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    87
  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) & 
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    88
               (\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    89
using assms by (rule complete_real)
32707
836ec9d0a0c8 New lemmas involving the real numbers, especially limits and series
paulson
parents: 30242
diff changeset
    90
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    91
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    92
subsection {* The Archimedean Property of the Reals *}
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    93
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    94
theorem reals_Archimedean:
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    95
  assumes x_pos: "0 < x"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
    96
  shows "\<exists>n. inverse (real (Suc n)) < x"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    97
  unfolding real_of_nat_def using x_pos
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
    98
  by (rule ex_inverse_of_nat_Suc_less)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    99
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   100
lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 35578
diff changeset
   101
  unfolding real_of_nat_def by (rule ex_less_of_nat)
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   102
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   103
lemma reals_Archimedean3:
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   104
  assumes x_greater_zero: "0 < x"
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   105
  shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   106
  unfolding real_of_nat_def using `0 < x`
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   107
  by (auto intro: ex_less_of_nat_mult)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   108
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   109
lemma reals_Archimedean6:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   110
     "0 \<le> r ==> \<exists>(n::nat). real (n - 1) \<le> r & r < real (n)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   111
unfolding real_of_nat_def
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   112
apply (rule exI [where x="nat (floor r + 1)"])
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   113
apply (insert floor_correct [of r])
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   114
apply (simp add: nat_add_distrib of_nat_nat)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   115
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   116
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   117
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
   118
  by (drule reals_Archimedean6) auto
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   119
36979
da7c06ab3169 remove several redundant lemmas about floor and ceiling
huffman
parents: 36826
diff changeset
   120
text {* TODO: delete *}
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   121
lemma reals_Archimedean_6b_int:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   122
     "0 \<le> r ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   123
  unfolding real_of_int_def by (rule floor_exists)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   124
36979
da7c06ab3169 remove several redundant lemmas about floor and ceiling
huffman
parents: 36826
diff changeset
   125
text {* TODO: delete *}
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   126
lemma reals_Archimedean_6c_int:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   127
     "r < 0 ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   128
  unfolding real_of_int_def by (rule floor_exists)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   129
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   130
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   131
subsection{*Density of the Rational Reals in the Reals*}
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   132
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   133
text{* This density proof is due to Stefan Richter and was ported by TN.  The
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   134
original source is \emph{Real Analysis} by H.L. Royden.
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   135
It employs the Archimedean property of the reals. *}
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   136
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   137
lemma Rats_dense_in_nn_real: fixes x::real
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   138
assumes "0\<le>x" and "x<y" shows "\<exists>r \<in> \<rat>.  x<r \<and> r<y"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   139
proof -
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   140
  from `x<y` have "0 < y-x" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   141
  with reals_Archimedean obtain q::nat 
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   142
    where q: "inverse (real q) < y-x" and "0 < real q" by auto  
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   143
  def p \<equiv> "LEAST n.  y \<le> real (Suc n)/real q"  
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   144
  from reals_Archimedean2 obtain n::nat where "y * real q < real n" by auto
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   145
  with `0 < real q` have ex: "y \<le> real n/real q" (is "?P n")
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   146
    by (simp add: pos_less_divide_eq[THEN sym])
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   147
  also from assms have "\<not> y \<le> real (0::nat) / real q" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   148
  ultimately have main: "(LEAST n. y \<le> real n/real q) = Suc p"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   149
    by (unfold p_def) (rule Least_Suc)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   150
  also from ex have "?P (LEAST x. ?P x)" by (rule LeastI)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   151
  ultimately have suc: "y \<le> real (Suc p) / real q" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   152
  def r \<equiv> "real p/real q"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   153
  have "x = y-(y-x)" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   154
  also from suc q have "\<dots> < real (Suc p)/real q - inverse (real q)" by arith
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   155
  also have "\<dots> = real p / real q"
37887
2ae085b07f2f diff_minus subsumes diff_def
haftmann
parents: 36979
diff changeset
   156
    by (simp only: inverse_eq_divide diff_minus real_of_nat_Suc 
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   157
    minus_divide_left add_divide_distrib[THEN sym]) simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   158
  finally have "x<r" by (unfold r_def)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   159
  have "p<Suc p" .. also note main[THEN sym]
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   160
  finally have "\<not> ?P p"  by (rule not_less_Least)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   161
  hence "r<y" by (simp add: r_def)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   162
  from r_def have "r \<in> \<rat>" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   163
  with `x<r` `r<y` show ?thesis by fast
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   164
qed
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   165
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   166
theorem Rats_dense_in_real: fixes x y :: real
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   167
assumes "x<y" shows "\<exists>r \<in> \<rat>.  x<r \<and> r<y"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   168
proof -
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   169
  from reals_Archimedean2 obtain n::nat where "-x < real n" by auto
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   170
  hence "0 \<le> x + real n" by arith
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   171
  also from `x<y` have "x + real n < y + real n" by arith
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   172
  ultimately have "\<exists>r \<in> \<rat>. x + real n < r \<and> r < y + real n"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   173
    by(rule Rats_dense_in_nn_real)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   174
  then obtain r where "r \<in> \<rat>" and r2: "x + real n < r" 
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   175
    and r3: "r < y + real n"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   176
    by blast
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   177
  have "r - real n = r + real (int n)/real (-1::int)" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   178
  also from `r\<in>\<rat>` have "r + real (int n)/real (-1::int) \<in> \<rat>" by simp
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   179
  also from r2 have "x < r - real n" by arith
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   180
  moreover from r3 have "r - real n < y" by arith
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   181
  ultimately show ?thesis by fast
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   182
qed
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   183
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 27435
diff changeset
   184
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   185
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
   186
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   187
lemma number_of_less_real_of_int_iff [simp]:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   188
     "((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
   189
apply auto
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   190
apply (rule real_of_int_less_iff [THEN iffD1])
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   191
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
   192
done
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   193
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   194
lemma number_of_less_real_of_int_iff2 [simp]:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   195
     "(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
   196
apply auto
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   197
apply (rule real_of_int_less_iff [THEN iffD1])
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   198
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
   199
done
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   200
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   201
lemma number_of_le_real_of_int_iff [simp]:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   202
     "((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
   203
by (simp add: linorder_not_less [symmetric])
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   204
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   205
lemma number_of_le_real_of_int_iff2 [simp]:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   206
     "(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
   207
by (simp add: linorder_not_less [symmetric])
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   208
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   209
lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   210
unfolding real_of_nat_def by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   211
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   212
lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
30102
799b687e4aac disable floor_minus and ceiling_minus [simp]
huffman
parents: 30097
diff changeset
   213
unfolding real_of_nat_def by (simp add: floor_minus)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   214
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   215
lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   216
unfolding real_of_int_def by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   217
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   218
lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
30102
799b687e4aac disable floor_minus and ceiling_minus [simp]
huffman
parents: 30097
diff changeset
   219
unfolding real_of_int_def by (simp add: floor_minus)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   220
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   221
lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   222
unfolding real_of_int_def by (rule floor_exists)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   223
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   224
lemma lemma_floor:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   225
  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
   226
  shows "m \<le> (n::int)"
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   227
proof -
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23309
diff changeset
   228
  have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23309
diff changeset
   229
  also have "... = real (n + 1)" by simp
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23309
diff changeset
   230
  finally have "m < n + 1" by (simp only: real_of_int_less_iff)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   231
  thus ?thesis by arith
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   232
qed
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   233
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   234
lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   235
unfolding real_of_int_def by (rule of_int_floor_le)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   236
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   237
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
   238
by (auto intro: lemma_floor)
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   239
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   240
lemma real_of_int_floor_cancel [simp]:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   241
    "(real (floor x) = x) = (\<exists>n::int. x = real n)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   242
  using floor_real_of_int by metis
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   243
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   244
lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   245
  unfolding real_of_int_def using floor_unique [of n x] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   246
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   247
lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   248
  unfolding real_of_int_def by (rule floor_unique)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   249
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   250
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
   251
apply (rule inj_int [THEN injD])
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   252
apply (simp add: real_of_nat_Suc)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15234
diff changeset
   253
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
   254
done
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   255
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   256
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
   257
apply (drule order_le_imp_less_or_eq)
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   258
apply (auto intro: floor_eq3)
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   259
done
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   260
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   261
lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   262
  unfolding real_of_int_def using floor_correct [of r] by simp
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   263
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   264
lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   265
  unfolding real_of_int_def using floor_correct [of r] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   266
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   267
lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   268
  unfolding real_of_int_def using floor_correct [of r] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   269
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   270
lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   271
  unfolding real_of_int_def using floor_correct [of r] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   272
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   273
lemma le_floor: "real a <= x ==> a <= floor x"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   274
  unfolding real_of_int_def by (simp add: le_floor_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   275
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   276
lemma real_le_floor: "a <= floor x ==> real a <= x"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   277
  unfolding real_of_int_def by (simp add: le_floor_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   278
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   279
lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   280
  unfolding real_of_int_def by (rule le_floor_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   281
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   282
lemma floor_less_eq: "(floor x < a) = (x < real a)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   283
  unfolding real_of_int_def by (rule floor_less_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   284
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   285
lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   286
  unfolding real_of_int_def by (rule less_floor_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   287
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   288
lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   289
  unfolding real_of_int_def by (rule floor_le_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   290
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   291
lemma floor_add [simp]: "floor (x + real a) = floor x + a"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   292
  unfolding real_of_int_def by (rule floor_add_of_int)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   293
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   294
lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   295
  unfolding real_of_int_def by (rule floor_diff_of_int)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   296
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   297
lemma le_mult_floor:
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   298
  assumes "0 \<le> (a :: real)" and "0 \<le> b"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   299
  shows "floor a * floor b \<le> floor (a * b)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   300
proof -
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   301
  have "real (floor a) \<le> a"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   302
    and "real (floor b) \<le> b" by auto
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   303
  hence "real (floor a * floor b) \<le> a * b"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   304
    using assms by (auto intro!: mult_mono)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   305
  also have "a * b < real (floor (a * b) + 1)" by auto
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   306
  finally show ?thesis unfolding real_of_int_less_iff by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   307
qed
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   308
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   309
lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   310
  unfolding real_of_nat_def by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   311
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   312
lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   313
  unfolding real_of_int_def by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   314
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   315
lemma floor_ceiling [simp]: "floor (real (ceiling r)) = ceiling r"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   316
  unfolding real_of_int_def by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   317
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   318
lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   319
  unfolding real_of_int_def by (rule le_of_int_ceiling)
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   320
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   321
lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   322
  unfolding real_of_int_def by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   323
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   324
lemma real_of_int_ceiling_cancel [simp]:
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   325
     "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   326
  using ceiling_real_of_int by metis
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   327
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   328
lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   329
  unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   330
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   331
lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   332
  unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   333
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   334
lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n  |] ==> ceiling x = n"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   335
  unfolding real_of_int_def using ceiling_unique [of n x] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   336
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   337
lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   338
  unfolding real_of_int_def using ceiling_correct [of r] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   339
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   340
lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   341
  unfolding real_of_int_def using ceiling_correct [of r] by simp
14641
79b7bd936264 moved Complex/NSInduct and Hyperreal/IntFloor to more appropriate
paulson
parents: 14476
diff changeset
   342
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   343
lemma ceiling_le: "x <= real a ==> ceiling x <= a"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   344
  unfolding real_of_int_def by (simp add: ceiling_le_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   345
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   346
lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   347
  unfolding real_of_int_def by (simp add: ceiling_le_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   348
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   349
lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   350
  unfolding real_of_int_def by (rule ceiling_le_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   351
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   352
lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   353
  unfolding real_of_int_def by (rule less_ceiling_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   354
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   355
lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   356
  unfolding real_of_int_def by (rule ceiling_less_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   357
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   358
lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   359
  unfolding real_of_int_def by (rule le_ceiling_iff)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   360
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   361
lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   362
  unfolding real_of_int_def by (rule ceiling_add_of_int)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   363
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   364
lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   365
  unfolding real_of_int_def by (rule ceiling_diff_of_int)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   366
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   367
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   368
subsection {* Versions for the natural numbers *}
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   369
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16893
diff changeset
   370
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   371
  natfloor :: "real => nat" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16893
diff changeset
   372
  "natfloor x = nat(floor x)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   373
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   374
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
   375
  natceiling :: "real => nat" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16893
diff changeset
   376
  "natceiling x = nat(ceiling x)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   377
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   378
lemma natfloor_zero [simp]: "natfloor 0 = 0"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   379
  by (unfold natfloor_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   380
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   381
lemma natfloor_one [simp]: "natfloor 1 = 1"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   382
  by (unfold natfloor_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   383
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   384
lemma zero_le_natfloor [simp]: "0 <= natfloor x"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   385
  by (unfold natfloor_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   386
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   387
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
   388
  by (unfold natfloor_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   389
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   390
lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   391
  by (unfold natfloor_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   392
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   393
lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   394
  by (unfold natfloor_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   395
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   396
lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   397
  apply (unfold natfloor_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   398
  apply (subgoal_tac "floor x <= floor 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   399
  apply simp
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   400
  apply (erule floor_mono)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   401
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   402
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   403
lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   404
  apply (case_tac "0 <= x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   405
  apply (subst natfloor_def)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   406
  apply (subst nat_le_eq_zle)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   407
  apply force
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   408
  apply (erule floor_mono)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   409
  apply (subst natfloor_neg)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   410
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   411
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   412
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   413
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   414
lemma le_natfloor: "real x <= a ==> x <= natfloor a"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   415
  apply (unfold natfloor_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   416
  apply (subst nat_int [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   417
  apply (subst nat_le_eq_zle)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   418
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   419
  apply (rule le_floor)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   420
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   421
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   422
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   423
lemma less_natfloor:
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   424
  assumes "0 \<le> x" and "x < real (n :: nat)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   425
  shows "natfloor x < n"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   426
proof (rule ccontr)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   427
  assume "\<not> ?thesis" hence *: "n \<le> natfloor x" by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   428
  note assms(2)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   429
  also have "real n \<le> real (natfloor x)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   430
    using * unfolding real_of_nat_le_iff .
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   431
  finally have "x < real (natfloor x)" .
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   432
  with real_natfloor_le[OF assms(1)]
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   433
  show False by auto
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   434
qed
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   435
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   436
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
   437
  apply (rule iffI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   438
  apply (rule order_trans)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   439
  prefer 2
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   440
  apply (erule real_natfloor_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   441
  apply (subst real_of_nat_le_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   442
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   443
  apply (erule le_natfloor)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   444
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   445
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   446
lemma le_natfloor_eq_number_of [simp]:
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   447
    "~ neg((number_of n)::int) ==> 0 <= x ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   448
      (number_of n <= natfloor x) = (number_of n <= x)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   449
  apply (subst le_natfloor_eq, assumption)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   450
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   451
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   452
16820
5c9d597e4cb9 fixed typos in theorem names
avigad
parents: 16819
diff changeset
   453
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
   454
  apply (case_tac "0 <= x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   455
  apply (subst le_natfloor_eq, assumption, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   456
  apply (rule iffI)
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   457
  apply (subgoal_tac "natfloor x <= natfloor 0")
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   458
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   459
  apply (rule natfloor_mono)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   460
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   461
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   462
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   463
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   464
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
   465
  apply (unfold natfloor_def)
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   466
  apply (subst (2) nat_int [THEN sym])
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   467
  apply (subst eq_nat_nat_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   468
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   469
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   470
  apply (rule floor_eq2)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   471
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   472
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   473
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   474
lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   475
  apply (case_tac "0 <= x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   476
  apply (unfold natfloor_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   477
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   478
  apply simp_all
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   479
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   480
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   481
lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   482
using real_natfloor_add_one_gt by (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   483
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   484
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
   485
  apply (subgoal_tac "z < real(natfloor z) + 1")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   486
  apply arith
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   487
  apply (rule real_natfloor_add_one_gt)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   488
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   489
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   490
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
   491
  apply (unfold natfloor_def)
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   492
  apply (subgoal_tac "real a = real (int a)")
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   493
  apply (erule ssubst)
23309
678ee30499d2 remove references to constant int::nat=>int
huffman
parents: 23012
diff changeset
   494
  apply (simp add: nat_add_distrib del: real_of_int_of_nat_eq)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   495
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   496
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   497
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   498
lemma natfloor_add_number_of [simp]:
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   499
    "~neg ((number_of n)::int) ==> 0 <= x ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   500
      natfloor (x + number_of n) = natfloor x + number_of n"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   501
  apply (subst natfloor_add [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   502
  apply simp_all
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   503
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   504
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   505
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
   506
  apply (subst natfloor_add [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   507
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   508
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   509
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   510
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   511
lemma natfloor_subtract [simp]: "real a <= x ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   512
    natfloor(x - real a) = natfloor x - a"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   513
  apply (unfold natfloor_def)
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   514
  apply (subgoal_tac "real a = real (int a)")
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   515
  apply (erule ssubst)
23309
678ee30499d2 remove references to constant int::nat=>int
huffman
parents: 23012
diff changeset
   516
  apply (simp del: real_of_int_of_nat_eq)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   517
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   518
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   519
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   520
lemma natfloor_div_nat: "1 <= x ==> y > 0 ==>
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   521
  natfloor (x / real y) = natfloor x div y"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   522
proof -
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   523
  assume "1 <= (x::real)" and "(y::nat) > 0"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   524
  have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   525
    by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   526
  then have a: "real(natfloor x) = real ((natfloor x) div y) * real y +
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   527
    real((natfloor x) mod y)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   528
    by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym])
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   529
  have "x = real(natfloor x) + (x - real(natfloor x))"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   530
    by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   531
  then have "x = real ((natfloor x) div y) * real y +
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   532
      real((natfloor x) mod y) + (x - real(natfloor x))"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   533
    by (simp add: a)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   534
  then have "x / real y = ... / real y"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   535
    by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   536
  also have "... = real((natfloor x) div y) + real((natfloor x) mod y) /
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   537
    real y + (x - real(natfloor x)) / real y"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   538
    by (auto simp add: algebra_simps add_divide_distrib
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   539
      diff_divide_distrib prems)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   540
  finally have "natfloor (x / real y) = natfloor(...)" by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   541
  also have "... = natfloor(real((natfloor x) mod y) /
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   542
    real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   543
    by (simp add: add_ac)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   544
  also have "... = natfloor(real((natfloor x) mod y) /
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   545
    real y + (x - real(natfloor x)) / real y) + (natfloor x) div y"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   546
    apply (rule natfloor_add)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   547
    apply (rule add_nonneg_nonneg)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   548
    apply (rule divide_nonneg_pos)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   549
    apply simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   550
    apply (simp add: prems)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   551
    apply (rule divide_nonneg_pos)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   552
    apply (simp add: algebra_simps)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   553
    apply (rule real_natfloor_le)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   554
    apply (insert prems, auto)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   555
    done
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   556
  also have "natfloor(real((natfloor x) mod y) /
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   557
    real y + (x - real(natfloor x)) / real y) = 0"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   558
    apply (rule natfloor_eq)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   559
    apply simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   560
    apply (rule add_nonneg_nonneg)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   561
    apply (rule divide_nonneg_pos)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   562
    apply force
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   563
    apply (force simp add: prems)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   564
    apply (rule divide_nonneg_pos)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   565
    apply (simp add: algebra_simps)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   566
    apply (rule real_natfloor_le)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   567
    apply (auto simp add: prems)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   568
    apply (insert prems, arith)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   569
    apply (simp add: add_divide_distrib [THEN sym])
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   570
    apply (subgoal_tac "real y = real y - 1 + 1")
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   571
    apply (erule ssubst)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   572
    apply (rule add_le_less_mono)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   573
    apply (simp add: algebra_simps)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   574
    apply (subgoal_tac "1 + real(natfloor x mod y) =
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   575
      real(natfloor x mod y + 1)")
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   576
    apply (erule ssubst)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   577
    apply (subst real_of_nat_le_iff)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   578
    apply (subgoal_tac "natfloor x mod y < y")
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   579
    apply arith
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   580
    apply (rule mod_less_divisor)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   581
    apply auto
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   582
    using real_natfloor_add_one_gt
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   583
    apply (simp add: algebra_simps)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   584
    done
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   585
  finally show ?thesis by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   586
qed
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   587
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   588
lemma le_mult_natfloor:
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   589
  assumes "0 \<le> (a :: real)" and "0 \<le> b"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   590
  shows "natfloor a * natfloor b \<le> natfloor (a * b)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   591
  unfolding natfloor_def
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   592
  apply (subst nat_mult_distrib[symmetric])
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   593
  using assms apply simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   594
  apply (subst nat_le_eq_zle)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   595
  using assms le_mult_floor by (auto intro!: mult_nonneg_nonneg)
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   596
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   597
lemma natceiling_zero [simp]: "natceiling 0 = 0"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   598
  by (unfold natceiling_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   599
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   600
lemma natceiling_one [simp]: "natceiling 1 = 1"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   601
  by (unfold natceiling_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   602
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   603
lemma zero_le_natceiling [simp]: "0 <= natceiling x"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   604
  by (unfold natceiling_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   605
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   606
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
   607
  by (unfold natceiling_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   608
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   609
lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   610
  by (unfold natceiling_def, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   611
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   612
lemma real_natceiling_ge: "x <= real(natceiling x)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   613
  apply (unfold natceiling_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   614
  apply (case_tac "x < 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   615
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   616
  apply (subst real_nat_eq_real)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   617
  apply (subgoal_tac "ceiling 0 <= ceiling x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   618
  apply simp
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   619
  apply (rule ceiling_mono)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   620
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   621
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   622
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   623
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   624
lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   625
  apply (unfold natceiling_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   626
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   627
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   628
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   629
lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   630
  apply (case_tac "0 <= x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   631
  apply (subst natceiling_def)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   632
  apply (subst nat_le_eq_zle)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   633
  apply (rule disjI2)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   634
  apply (subgoal_tac "real (0::int) <= real(ceiling y)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   635
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   636
  apply (rule order_trans)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   637
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   638
  apply (erule order_trans)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   639
  apply simp
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   640
  apply (erule ceiling_mono)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   641
  apply (subst natceiling_neg)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   642
  apply simp_all
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 natceiling_le: "x <= real a ==> natceiling x <= a"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   646
  apply (unfold natceiling_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   647
  apply (case_tac "x < 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   648
  apply simp
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35028
diff changeset
   649
  apply (subst (2) nat_int [THEN sym])
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   650
  apply (subst nat_le_eq_zle)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   651
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   652
  apply (rule ceiling_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   653
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   654
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   655
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   656
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
   657
  apply (rule iffI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   658
  apply (rule order_trans)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   659
  apply (rule real_natceiling_ge)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   660
  apply (subst real_of_nat_le_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   661
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   662
  apply (erule natceiling_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   663
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   664
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   665
lemma natceiling_le_eq_number_of [simp]:
16820
5c9d597e4cb9 fixed typos in theorem names
avigad
parents: 16819
diff changeset
   666
    "~ neg((number_of n)::int) ==> 0 <= x ==>
5c9d597e4cb9 fixed typos in theorem names
avigad
parents: 16819
diff changeset
   667
      (natceiling x <= number_of n) = (x <= number_of n)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   668
  apply (subst natceiling_le_eq, assumption)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   669
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   670
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   671
16820
5c9d597e4cb9 fixed typos in theorem names
avigad
parents: 16819
diff changeset
   672
lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   673
  apply (case_tac "0 <= x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   674
  apply (subst natceiling_le_eq)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   675
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   676
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   677
  apply (subst natceiling_neg)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   678
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   679
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   680
done
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 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
   683
  apply (unfold natceiling_def)
19850
29c125556d86 fixed subst step;
wenzelm
parents: 19765
diff changeset
   684
  apply (simplesubst nat_int [THEN sym]) back back
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   685
  apply (subgoal_tac "nat(int n) + 1 = nat(int n + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   686
  apply (erule ssubst)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   687
  apply (subst eq_nat_nat_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   688
  apply (subgoal_tac "ceiling 0 <= ceiling x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   689
  apply simp
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   690
  apply (rule ceiling_mono)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   691
  apply force
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   692
  apply force
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   693
  apply (rule ceiling_eq2)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   694
  apply (simp, simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   695
  apply (subst nat_add_distrib)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   696
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   697
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   698
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   699
lemma natceiling_add [simp]: "0 <= x ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   700
    natceiling (x + real a) = natceiling x + a"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   701
  apply (unfold natceiling_def)
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   702
  apply (subgoal_tac "real a = real (int a)")
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   703
  apply (erule ssubst)
23309
678ee30499d2 remove references to constant int::nat=>int
huffman
parents: 23012
diff changeset
   704
  apply (simp del: real_of_int_of_nat_eq)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   705
  apply (subst nat_add_distrib)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   706
  apply (subgoal_tac "0 = ceiling 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   707
  apply (erule ssubst)
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 29667
diff changeset
   708
  apply (erule ceiling_mono)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   709
  apply simp_all
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   710
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   711
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   712
lemma natceiling_add_number_of [simp]:
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   713
    "~ neg ((number_of n)::int) ==> 0 <= x ==>
16820
5c9d597e4cb9 fixed typos in theorem names
avigad
parents: 16819
diff changeset
   714
      natceiling (x + number_of n) = natceiling x + number_of n"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   715
  apply (subst natceiling_add [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   716
  apply simp_all
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   717
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   718
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   719
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
   720
  apply (subst natceiling_add [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   721
  apply assumption
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
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   724
16893
0cc94e6f6ae5 some structured proofs on completeness;
wenzelm
parents: 16827
diff changeset
   725
lemma natceiling_subtract [simp]: "real a <= x ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   726
    natceiling(x - real a) = natceiling x - a"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   727
  apply (unfold natceiling_def)
24355
93d78fdeb55a remove int_of_nat
huffman
parents: 23477
diff changeset
   728
  apply (subgoal_tac "real a = real (int a)")
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   729
  apply (erule ssubst)
23309
678ee30499d2 remove references to constant int::nat=>int
huffman
parents: 23012
diff changeset
   730
  apply (simp del: real_of_int_of_nat_eq)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   731
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   732
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   733
36826
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   734
subsection {* Exponentiation with floor *}
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   735
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   736
lemma floor_power:
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   737
  assumes "x = real (floor x)"
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   738
  shows "floor (x ^ n) = floor x ^ n"
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   739
proof -
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   740
  have *: "x ^ n = real (floor x ^ n)"
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   741
    using assms by (induct n arbitrary: x) simp_all
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   742
  show ?thesis unfolding real_of_int_inject[symmetric]
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   743
    unfolding * floor_real_of_int ..
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   744
qed
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   745
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   746
lemma natfloor_power:
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   747
  assumes "x = real (natfloor x)"
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   748
  shows "natfloor (x ^ n) = natfloor x ^ n"
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   749
proof -
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   750
  from assms have "0 \<le> floor x" by auto
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   751
  note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   752
  from floor_power[OF this]
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   753
  show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   754
    by simp
4d4462d644ae move floor lemmas from RealPow.thy to RComplete.thy
huffman
parents: 36795
diff changeset
   755
qed
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 15539
diff changeset
   756
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   757
end