src/HOL/RealDef.thy
author Christian Sternagel
Thu, 30 Aug 2012 15:44:03 +0900
changeset 49093 fdc301f592c4
parent 47902 34a9e81e5bfd
child 49962 a8cc904a6820
permissions -rw-r--r--
forgot to add lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28906
diff changeset
     1
(*  Title       : HOL/RealDef.thy
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
     2
    Author      : Jacques D. Fleuriot, 1998
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
     3
    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
     4
    Additional contributions by Jeremy Avigad
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
     5
    Construction of Cauchy Reals by Brian Huffman, 2010
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     6
*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     7
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
     8
header {* Development of the Reals using Cauchy Sequences *}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15086
diff changeset
    10
theory RealDef
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    11
imports Rat
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15086
diff changeset
    12
begin
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    13
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    14
text {*
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    15
  This theory contains a formalization of the real numbers as
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    16
  equivalence classes of Cauchy sequences of rationals.  See
40810
142f890ceef6 use new 'file' antiquotation for reference to Dedekind_Real.thy
huffman
parents: 38857
diff changeset
    17
  @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
142f890ceef6 use new 'file' antiquotation for reference to Dedekind_Real.thy
huffman
parents: 38857
diff changeset
    18
  construction using Dedekind cuts.
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    19
*}
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    20
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    21
subsection {* Preliminary lemmas *}
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    22
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    23
lemma add_diff_add:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    24
  fixes a b c d :: "'a::ab_group_add"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    25
  shows "(a + c) - (b + d) = (a - b) + (c - d)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    26
  by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    27
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    28
lemma minus_diff_minus:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    29
  fixes a b :: "'a::ab_group_add"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    30
  shows "- a - - b = - (a - b)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    31
  by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    32
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    33
lemma mult_diff_mult:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    34
  fixes x y a b :: "'a::ring"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    35
  shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    36
  by (simp add: algebra_simps)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    37
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    38
lemma inverse_diff_inverse:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    39
  fixes a b :: "'a::division_ring"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    40
  assumes "a \<noteq> 0" and "b \<noteq> 0"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    41
  shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    42
  using assms by (simp add: algebra_simps)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    43
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    44
lemma obtain_pos_sum:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    45
  fixes r :: rat assumes r: "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    46
  obtains s t where "0 < s" and "0 < t" and "r = s + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    47
proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    48
    from r show "0 < r/2" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    49
    from r show "0 < r/2" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    50
    show "r = r/2 + r/2" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    51
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    52
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    53
subsection {* Sequences that converge to zero *}
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    54
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    55
definition
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    56
  vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    57
where
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    58
  "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    59
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    60
lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    61
  unfolding vanishes_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    62
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    63
lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    64
  unfolding vanishes_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    65
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    66
lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    67
  unfolding vanishes_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    68
  apply (cases "c = 0", auto)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    69
  apply (rule exI [where x="\<bar>c\<bar>"], auto)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    70
  done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    71
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    72
lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    73
  unfolding vanishes_def by simp
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    74
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    75
lemma vanishes_add:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    76
  assumes X: "vanishes X" and Y: "vanishes Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    77
  shows "vanishes (\<lambda>n. X n + Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    78
proof (rule vanishesI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    79
  fix r :: rat assume "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    80
  then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    81
    by (rule obtain_pos_sum)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    82
  obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    83
    using vanishesD [OF X s] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    84
  obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    85
    using vanishesD [OF Y t] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    86
  have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    87
  proof (clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    88
    fix n assume n: "i \<le> n" "j \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    89
    have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    90
    also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    91
    finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    92
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    93
  thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    94
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    95
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    96
lemma vanishes_diff:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    97
  assumes X: "vanishes X" and Y: "vanishes Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    98
  shows "vanishes (\<lambda>n. X n - Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
    99
unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   100
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   101
lemma vanishes_mult_bounded:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   102
  assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   103
  assumes Y: "vanishes (\<lambda>n. Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   104
  shows "vanishes (\<lambda>n. X n * Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   105
proof (rule vanishesI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   106
  fix r :: rat assume r: "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   107
  obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   108
    using X by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   109
  obtain b where b: "0 < b" "r = a * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   110
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   111
    show "0 < r / a" using r a by (simp add: divide_pos_pos)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   112
    show "r = a * (r / a)" using a by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   113
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   114
  obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   115
    using vanishesD [OF Y b(1)] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   116
  have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   117
    by (simp add: b(2) abs_mult mult_strict_mono' a k)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   118
  thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   119
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   120
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   121
subsection {* Cauchy sequences *}
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
   122
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
   123
definition
44278
1220ecb81e8f observe distinction between sets and predicates more properly
haftmann
parents: 43887
diff changeset
   124
  cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   125
where
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   126
  "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   127
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   128
lemma cauchyI:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   129
  "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r) \<Longrightarrow> cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   130
  unfolding cauchy_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   131
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   132
lemma cauchyD:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   133
  "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   134
  unfolding cauchy_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   135
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   136
lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   137
  unfolding cauchy_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   138
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   139
lemma cauchy_add [simp]:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   140
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   141
  shows "cauchy (\<lambda>n. X n + Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   142
proof (rule cauchyI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   143
  fix r :: rat assume "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   144
  then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   145
    by (rule obtain_pos_sum)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   146
  obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   147
    using cauchyD [OF X s] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   148
  obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   149
    using cauchyD [OF Y t] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   150
  have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   151
  proof (clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   152
    fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   153
    have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   154
      unfolding add_diff_add by (rule abs_triangle_ineq)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   155
    also have "\<dots> < s + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   156
      by (rule add_strict_mono, simp_all add: i j *)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   157
    finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   158
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   159
  thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   160
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   161
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   162
lemma cauchy_minus [simp]:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   163
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   164
  shows "cauchy (\<lambda>n. - X n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   165
using assms unfolding cauchy_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   166
unfolding minus_diff_minus abs_minus_cancel .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   167
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   168
lemma cauchy_diff [simp]:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   169
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   170
  shows "cauchy (\<lambda>n. X n - Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   171
using assms unfolding diff_minus by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   172
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   173
lemma cauchy_imp_bounded:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   174
  assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   175
proof -
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   176
  obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   177
    using cauchyD [OF assms zero_less_one] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   178
  show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   179
  proof (intro exI conjI allI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   180
    have "0 \<le> \<bar>X 0\<bar>" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   181
    also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   182
    finally have "0 \<le> Max (abs ` X ` {..k})" .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   183
    thus "0 < Max (abs ` X ` {..k}) + 1" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   184
  next
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   185
    fix n :: nat
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   186
    show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   187
    proof (rule linorder_le_cases)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   188
      assume "n \<le> k"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   189
      hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   190
      thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   191
    next
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   192
      assume "k \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   193
      have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   194
      also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   195
        by (rule abs_triangle_ineq)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   196
      also have "\<dots> < Max (abs ` X ` {..k}) + 1"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   197
        by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   198
      finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   199
    qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   200
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   201
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   202
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   203
lemma cauchy_mult [simp]:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   204
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   205
  shows "cauchy (\<lambda>n. X n * Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   206
proof (rule cauchyI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   207
  fix r :: rat assume "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   208
  then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   209
    by (rule obtain_pos_sum)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   210
  obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   211
    using cauchy_imp_bounded [OF X] by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   212
  obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   213
    using cauchy_imp_bounded [OF Y] by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   214
  obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   215
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   216
    show "0 < v/b" using v b(1) by (rule divide_pos_pos)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   217
    show "0 < u/a" using u a(1) by (rule divide_pos_pos)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   218
    show "r = a * (u/a) + (v/b) * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   219
      using a(1) b(1) `r = u + v` by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   220
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   221
  obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   222
    using cauchyD [OF X s] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   223
  obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   224
    using cauchyD [OF Y t] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   225
  have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   226
  proof (clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   227
    fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   228
    have "\<bar>X m * Y m - X n * Y n\<bar> = \<bar>X m * (Y m - Y n) + (X m - X n) * Y n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   229
      unfolding mult_diff_mult ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   230
    also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   231
      by (rule abs_triangle_ineq)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   232
    also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   233
      unfolding abs_mult ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   234
    also have "\<dots> < a * t + s * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   235
      by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   236
    finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   237
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   238
  thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   239
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   240
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   241
lemma cauchy_not_vanishes_cases:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   242
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   243
  assumes nz: "\<not> vanishes X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   244
  shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   245
proof -
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   246
  obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   247
    using nz unfolding vanishes_def by (auto simp add: not_less)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   248
  obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   249
    using `0 < r` by (rule obtain_pos_sum)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   250
  obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   251
    using cauchyD [OF X s] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   252
  obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   253
    using r by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   254
  have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   255
    using i `i \<le> k` by auto
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   256
  have "X k \<le> - r \<or> r \<le> X k"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   257
    using `r \<le> \<bar>X k\<bar>` by auto
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   258
  hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   259
    unfolding `r = s + t` using k by auto
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   260
  hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   261
  thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   262
    using t by auto
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   263
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   264
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   265
lemma cauchy_not_vanishes:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   266
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   267
  assumes nz: "\<not> vanishes X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   268
  shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   269
using cauchy_not_vanishes_cases [OF assms]
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   270
by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   271
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   272
lemma cauchy_inverse [simp]:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   273
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   274
  assumes nz: "\<not> vanishes X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   275
  shows "cauchy (\<lambda>n. inverse (X n))"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   276
proof (rule cauchyI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   277
  fix r :: rat assume "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   278
  obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   279
    using cauchy_not_vanishes [OF X nz] by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   280
  from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   281
  obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   282
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   283
    show "0 < b * r * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   284
      by (simp add: `0 < r` b mult_pos_pos)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   285
    show "r = inverse b * (b * r * b) * inverse b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   286
      using b by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   287
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   288
  obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   289
    using cauchyD [OF X s] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   290
  have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   291
  proof (clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   292
    fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   293
    have "\<bar>inverse (X m) - inverse (X n)\<bar> =
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   294
          inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   295
      by (simp add: inverse_diff_inverse nz * abs_mult)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   296
    also have "\<dots> < inverse b * s * inverse b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   297
      by (simp add: mult_strict_mono less_imp_inverse_less
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   298
                    mult_pos_pos i j b * s)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   299
    finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   300
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   301
  thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   302
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   303
47901
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   304
lemma vanishes_diff_inverse:
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   305
  assumes X: "cauchy X" "\<not> vanishes X"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   306
  assumes Y: "cauchy Y" "\<not> vanishes Y"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   307
  assumes XY: "vanishes (\<lambda>n. X n - Y n)"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   308
  shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   309
proof (rule vanishesI)
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   310
  fix r :: rat assume r: "0 < r"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   311
  obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   312
    using cauchy_not_vanishes [OF X] by fast
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   313
  obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   314
    using cauchy_not_vanishes [OF Y] by fast
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   315
  obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   316
  proof
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   317
    show "0 < a * r * b"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   318
      using a r b by (simp add: mult_pos_pos)
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   319
    show "inverse a * (a * r * b) * inverse b = r"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   320
      using a r b by simp
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   321
  qed
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   322
  obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   323
    using vanishesD [OF XY s] ..
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   324
  have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   325
  proof (clarsimp)
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   326
    fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   327
    have "X n \<noteq> 0" and "Y n \<noteq> 0"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   328
      using i j a b n by auto
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   329
    hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   330
        inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   331
      by (simp add: inverse_diff_inverse abs_mult)
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   332
    also have "\<dots> < inverse a * s * inverse b"
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   333
      apply (intro mult_strict_mono' less_imp_inverse_less)
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   334
      apply (simp_all add: a b i j k n mult_nonneg_nonneg)
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   335
      done
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   336
    also note `inverse a * s * inverse b = r`
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   337
    finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   338
  qed
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   339
  thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   340
qed
25a6ca50fe14 tuned ordering of lemmas
huffman
parents: 47598
diff changeset
   341
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   342
subsection {* Equivalence relation on Cauchy sequences *}
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   343
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   344
definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   345
  where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   346
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   347
lemma realrelI [intro?]:
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   348
  assumes "cauchy X" and "cauchy Y" and "vanishes (\<lambda>n. X n - Y n)"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   349
  shows "realrel X Y"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   350
  using assms unfolding realrel_def by simp
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   351
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   352
lemma realrel_refl: "cauchy X \<Longrightarrow> realrel X X"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   353
  unfolding realrel_def by simp
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   354
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   355
lemma symp_realrel: "symp realrel"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   356
  unfolding realrel_def
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   357
  by (rule sympI, clarify, drule vanishes_minus, simp)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   358
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   359
lemma transp_realrel: "transp realrel"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   360
  unfolding realrel_def
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   361
  apply (rule transpI, clarify)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   362
  apply (drule (1) vanishes_add)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   363
  apply (simp add: algebra_simps)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   364
  done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   365
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   366
lemma part_equivp_realrel: "part_equivp realrel"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   367
  by (fast intro: part_equivpI symp_realrel transp_realrel
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   368
    realrel_refl cauchy_const)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   369
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   370
subsection {* The field of real numbers *}
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   371
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   372
quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   373
  morphisms rep_real Real
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   374
  by (rule part_equivp_realrel)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   375
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   376
lemma cr_real_eq: "cr_real = (\<lambda>x y. cauchy x \<and> Real x = y)"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   377
  unfolding cr_real_def realrel_def by simp
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   378
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   379
lemma Real_induct [induct type: real]: (* TODO: generate automatically *)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   380
  assumes "\<And>X. cauchy X \<Longrightarrow> P (Real X)" shows "P x"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   381
proof (induct x)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   382
  case (1 X)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   383
  hence "cauchy X" by (simp add: realrel_def)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   384
  thus "P (Real X)" by (rule assms)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   385
qed
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   386
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   387
lemma eq_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   388
  "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   389
  using real.rel_eq_transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   390
  unfolding cr_real_def fun_rel_def realrel_def by simp
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   391
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   392
declare real.forall_transfer [transfer_rule del]
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   393
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   394
lemma forall_real_transfer [transfer_rule]: (* TODO: generate automatically *)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   395
  "(fun_rel (fun_rel cr_real op =) op =)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   396
    (transfer_bforall cauchy) transfer_forall"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   397
  using Quotient_forall_transfer [OF Quotient_real]
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   398
  by (simp add: realrel_def)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   399
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   400
instantiation real :: field_inverse_zero
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   401
begin
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   402
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   403
lift_definition zero_real :: "real" is "\<lambda>n. 0"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   404
  by (simp add: realrel_refl)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   405
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   406
lift_definition one_real :: "real" is "\<lambda>n. 1"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   407
  by (simp add: realrel_refl)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   408
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   409
lift_definition plus_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n + Y n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   410
  unfolding realrel_def add_diff_add
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   411
  by (simp only: cauchy_add vanishes_add simp_thms)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   412
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   413
lift_definition uminus_real :: "real \<Rightarrow> real" is "\<lambda>X n. - X n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   414
  unfolding realrel_def minus_diff_minus
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   415
  by (simp only: cauchy_minus vanishes_minus simp_thms)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   416
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   417
lift_definition times_real :: "real \<Rightarrow> real \<Rightarrow> real" is "\<lambda>X Y n. X n * Y n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   418
  unfolding realrel_def mult_diff_mult
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   419
  by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   420
    vanishes_mult_bounded cauchy_imp_bounded simp_thms)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   421
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   422
lift_definition inverse_real :: "real \<Rightarrow> real"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   423
  is "\<lambda>X. if vanishes X then (\<lambda>n. 0) else (\<lambda>n. inverse (X n))"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   424
proof -
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   425
  fix X Y assume "realrel X Y"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   426
  hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   427
    unfolding realrel_def by simp_all
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   428
  have "vanishes X \<longleftrightarrow> vanishes Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   429
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   430
    assume "vanishes X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   431
    from vanishes_diff [OF this XY] show "vanishes Y" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   432
  next
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   433
    assume "vanishes Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   434
    from vanishes_add [OF this XY] show "vanishes X" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   435
  qed
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   436
  thus "?thesis X Y"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   437
    unfolding realrel_def
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   438
    by (simp add: vanishes_diff_inverse X Y XY)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   439
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   440
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   441
definition
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   442
  "x - y = (x::real) + - y"
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
   443
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   444
definition
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   445
  "x / y = (x::real) * inverse y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   446
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   447
lemma add_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   448
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   449
  shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   450
  using assms plus_real.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   451
  unfolding cr_real_eq fun_rel_def by simp
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   452
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   453
lemma minus_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   454
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   455
  shows "- Real X = Real (\<lambda>n. - X n)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   456
  using assms uminus_real.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   457
  unfolding cr_real_eq fun_rel_def by simp
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
   458
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   459
lemma diff_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   460
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   461
  shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   462
  unfolding minus_real_def diff_minus
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   463
  by (simp add: minus_Real add_Real X Y)
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   464
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   465
lemma mult_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   466
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   467
  shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   468
  using assms times_real.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   469
  unfolding cr_real_eq fun_rel_def by simp
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   470
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   471
lemma inverse_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   472
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   473
  shows "inverse (Real X) =
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   474
    (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   475
  using assms inverse_real.transfer zero_real.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   476
  unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   477
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   478
instance proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   479
  fix a b c :: real
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   480
  show "a + b = b + a"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   481
    by transfer (simp add: add_ac realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   482
  show "(a + b) + c = a + (b + c)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   483
    by transfer (simp add: add_ac realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   484
  show "0 + a = a"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   485
    by transfer (simp add: realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   486
  show "- a + a = 0"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   487
    by transfer (simp add: realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   488
  show "a - b = a + - b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   489
    by (rule minus_real_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   490
  show "(a * b) * c = a * (b * c)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   491
    by transfer (simp add: mult_ac realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   492
  show "a * b = b * a"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   493
    by transfer (simp add: mult_ac realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   494
  show "1 * a = a"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   495
    by transfer (simp add: mult_ac realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   496
  show "(a + b) * c = a * c + b * c"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   497
    by transfer (simp add: left_distrib realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   498
  show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   499
    by transfer (simp add: realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   500
  show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   501
    apply transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   502
    apply (simp add: realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   503
    apply (rule vanishesI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   504
    apply (frule (1) cauchy_not_vanishes, clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   505
    apply (rule_tac x=k in exI, clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   506
    apply (drule_tac x=n in spec, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   507
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   508
  show "a / b = a * inverse b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   509
    by (rule divide_real_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   510
  show "inverse (0::real) = 0"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   511
    by transfer (simp add: realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   512
qed
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   513
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   514
end
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   515
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   516
subsection {* Positive reals *}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   517
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   518
lift_definition positive :: "real \<Rightarrow> bool"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   519
  is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   520
proof -
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   521
  { fix X Y
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   522
    assume "realrel X Y"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   523
    hence XY: "vanishes (\<lambda>n. X n - Y n)"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   524
      unfolding realrel_def by simp_all
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   525
    assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   526
    then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   527
      by fast
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   528
    obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   529
      using `0 < r` by (rule obtain_pos_sum)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   530
    obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   531
      using vanishesD [OF XY s] ..
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   532
    have "\<forall>n\<ge>max i j. t < Y n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   533
    proof (clarsimp)
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   534
      fix n assume n: "i \<le> n" "j \<le> n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   535
      have "\<bar>X n - Y n\<bar> < s" and "r < X n"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   536
        using i j n by simp_all
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   537
      thus "t < Y n" unfolding r by simp
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   538
    qed
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   539
    hence "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   540
  } note 1 = this
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   541
  fix X Y assume "realrel X Y"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   542
  hence "realrel X Y" and "realrel Y X"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   543
    using symp_realrel unfolding symp_def by auto
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   544
  thus "?thesis X Y"
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   545
    by (safe elim!: 1)
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   546
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   547
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   548
lemma positive_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   549
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   550
  shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   551
  using assms positive.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   552
  unfolding cr_real_eq fun_rel_def by simp
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   553
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   554
lemma positive_zero: "\<not> positive 0"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   555
  by transfer auto
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   556
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   557
lemma positive_add:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   558
  "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   559
apply transfer
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   560
apply (clarify, rename_tac a b i j)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   561
apply (rule_tac x="a + b" in exI, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   562
apply (rule_tac x="max i j" in exI, clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   563
apply (simp add: add_strict_mono)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   564
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   565
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   566
lemma positive_mult:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   567
  "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   568
apply transfer
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   569
apply (clarify, rename_tac a b i j)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   570
apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   571
apply (rule_tac x="max i j" in exI, clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   572
apply (rule mult_strict_mono, auto)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   573
done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   574
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   575
lemma positive_minus:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   576
  "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   577
apply transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   578
apply (simp add: realrel_def)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   579
apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   580
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   581
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   582
instantiation real :: linordered_field_inverse_zero
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   583
begin
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   584
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   585
definition
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   586
  "x < y \<longleftrightarrow> positive (y - x)"
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   587
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   588
definition
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   589
  "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   590
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   591
definition
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   592
  "abs (a::real) = (if a < 0 then - a else a)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   593
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   594
definition
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   595
  "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   596
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   597
instance proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   598
  fix a b c :: real
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   599
  show "\<bar>a\<bar> = (if a < 0 then - a else a)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   600
    by (rule abs_real_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   601
  show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   602
    unfolding less_eq_real_def less_real_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   603
    by (auto, drule (1) positive_add, simp_all add: positive_zero)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   604
  show "a \<le> a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   605
    unfolding less_eq_real_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   606
  show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   607
    unfolding less_eq_real_def less_real_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   608
    by (auto, drule (1) positive_add, simp add: algebra_simps)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   609
  show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   610
    unfolding less_eq_real_def less_real_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   611
    by (auto, drule (1) positive_add, simp add: positive_zero)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   612
  show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   613
    unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   614
    (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   615
    (* Should produce c + b - (c + a) \<equiv> b - a *)
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   616
  show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   617
    by (rule sgn_real_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   618
  show "a \<le> b \<or> b \<le> a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   619
    unfolding less_eq_real_def less_real_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   620
    by (auto dest!: positive_minus)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   621
  show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   622
    unfolding less_real_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   623
    by (drule (1) positive_mult, simp add: algebra_simps)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   624
qed
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   625
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   626
end
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   627
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   628
instantiation real :: distrib_lattice
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   629
begin
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   630
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   631
definition
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   632
  "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   633
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   634
definition
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   635
  "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   636
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   637
instance proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   638
qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   639
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   640
end
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   641
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   642
lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   643
apply (induct x)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   644
apply (simp add: zero_real_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   645
apply (simp add: one_real_def add_Real)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   646
done
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   647
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   648
lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   649
apply (cases x rule: int_diff_cases)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   650
apply (simp add: of_nat_Real diff_Real)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   651
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   652
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   653
lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   654
apply (induct x)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   655
apply (simp add: Fract_of_int_quotient of_rat_divide)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   656
apply (simp add: of_int_Real divide_inverse)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   657
apply (simp add: inverse_Real mult_Real)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   658
done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   659
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   660
instance real :: archimedean_field
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   661
proof
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   662
  fix x :: real
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   663
  show "\<exists>z. x \<le> of_int z"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   664
    apply (induct x)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   665
    apply (frule cauchy_imp_bounded, clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   666
    apply (rule_tac x="ceiling b + 1" in exI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   667
    apply (rule less_imp_le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   668
    apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   669
    apply (rule_tac x=1 in exI, simp add: algebra_simps)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   670
    apply (rule_tac x=0 in exI, clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   671
    apply (rule le_less_trans [OF abs_ge_self])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   672
    apply (rule less_le_trans [OF _ le_of_int_ceiling])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   673
    apply simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   674
    done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   675
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   676
43732
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   677
instantiation real :: floor_ceiling
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   678
begin
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   679
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   680
definition [code del]:
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   681
  "floor (x::real) = (THE z. of_int z \<le> x \<and> x < of_int (z + 1))"
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   682
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   683
instance proof
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   684
  fix x :: real
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   685
  show "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   686
    unfolding floor_real_def using floor_exists1 by (rule theI')
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   687
qed
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   688
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   689
end
6b2bdc57155b adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents: 42311
diff changeset
   690
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   691
subsection {* Completeness *}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   692
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   693
lemma not_positive_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   694
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   695
  shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   696
unfolding positive_Real [OF X]
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   697
apply (auto, unfold not_less)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   698
apply (erule obtain_pos_sum)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   699
apply (drule_tac x=s in spec, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   700
apply (drule_tac r=t in cauchyD [OF X], clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   701
apply (drule_tac x=k in spec, clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   702
apply (rule_tac x=n in exI, clarify, rename_tac m)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   703
apply (drule_tac x=m in spec, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   704
apply (drule_tac x=n in spec, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   705
apply (drule spec, drule (1) mp, clarify, rename_tac i)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   706
apply (rule_tac x="max i k" in exI, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   707
done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   708
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   709
lemma le_Real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   710
  assumes X: "cauchy X" and Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   711
  shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   712
unfolding not_less [symmetric, where 'a=real] less_real_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   713
apply (simp add: diff_Real not_positive_Real X Y)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   714
apply (simp add: diff_le_eq add_ac)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   715
done
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   716
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   717
lemma le_RealI:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   718
  assumes Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   719
  shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   720
proof (induct x)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   721
  fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   722
  hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   723
    by (simp add: of_rat_Real le_Real)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   724
  {
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   725
    fix r :: rat assume "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   726
    then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   727
      by (rule obtain_pos_sum)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   728
    obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   729
      using cauchyD [OF Y s] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   730
    obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   731
      using le [OF t] ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   732
    have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   733
    proof (clarsimp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   734
      fix n assume n: "i \<le> n" "j \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   735
      have "X n \<le> Y i + t" using n j by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   736
      moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   737
      ultimately show "X n \<le> Y n + r" unfolding r by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   738
    qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   739
    hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   740
  }
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   741
  thus "Real X \<le> Real Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   742
    by (simp add: of_rat_Real le_Real X Y)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   743
qed
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   744
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   745
lemma Real_leI:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   746
  assumes X: "cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   747
  assumes le: "\<forall>n. of_rat (X n) \<le> y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   748
  shows "Real X \<le> y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   749
proof -
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   750
  have "- y \<le> - Real X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   751
    by (simp add: minus_Real X le_RealI of_rat_minus le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   752
  thus ?thesis by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   753
qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   754
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   755
lemma less_RealD:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   756
  assumes Y: "cauchy Y"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   757
  shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   758
by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   759
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   760
lemma of_nat_less_two_power:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   761
  "of_nat n < (2::'a::linordered_idom) ^ n"
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   762
apply (induct n)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   763
apply simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   764
apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   765
apply (drule (1) add_le_less_mono, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   766
apply simp
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   767
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   768
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   769
lemma complete_real:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   770
  fixes S :: "real set"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   771
  assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   772
  shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   773
proof -
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   774
  obtain x where x: "x \<in> S" using assms(1) ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   775
  obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   776
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   777
  def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   778
  obtain a where a: "\<not> P a"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   779
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   780
    have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   781
    also have "x - 1 < x" by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   782
    finally have "of_int (floor (x - 1)) < x" .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   783
    hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   784
    then show "\<not> P (of_int (floor (x - 1)))"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   785
      unfolding P_def of_rat_of_int_eq using x by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   786
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   787
  obtain b where b: "P b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   788
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   789
    show "P (of_int (ceiling z))"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   790
    unfolding P_def of_rat_of_int_eq
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   791
    proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   792
      fix y assume "y \<in> S"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   793
      hence "y \<le> z" using z by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   794
      also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   795
      finally show "y \<le> of_int (ceiling z)" .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   796
    qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   797
  qed
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   798
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   799
  def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   800
  def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   801
  def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   802
  def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   803
  def C \<equiv> "\<lambda>n. avg (A n) (B n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   804
  have A_0 [simp]: "A 0 = a" unfolding A_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   805
  have B_0 [simp]: "B 0 = b" unfolding B_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   806
  have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   807
    unfolding A_def B_def C_def bisect_def split_def by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   808
  have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   809
    unfolding A_def B_def C_def bisect_def split_def by simp
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   810
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   811
  have width: "\<And>n. B n - A n = (b - a) / 2^n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   812
    apply (simp add: eq_divide_eq)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   813
    apply (induct_tac n, simp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   814
    apply (simp add: C_def avg_def algebra_simps)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   815
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   816
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   817
  have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   818
    apply (simp add: divide_less_eq)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   819
    apply (subst mult_commute)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   820
    apply (frule_tac y=y in ex_less_of_nat_mult)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   821
    apply clarify
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   822
    apply (rule_tac x=n in exI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   823
    apply (erule less_trans)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   824
    apply (rule mult_strict_right_mono)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   825
    apply (rule le_less_trans [OF _ of_nat_less_two_power])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   826
    apply simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   827
    apply assumption
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   828
    done
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   829
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   830
  have PA: "\<And>n. \<not> P (A n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   831
    by (induct_tac n, simp_all add: a)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   832
  have PB: "\<And>n. P (B n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   833
    by (induct_tac n, simp_all add: b)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   834
  have ab: "a < b"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   835
    using a b unfolding P_def
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   836
    apply (clarsimp simp add: not_le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   837
    apply (drule (1) bspec)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   838
    apply (drule (1) less_le_trans)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   839
    apply (simp add: of_rat_less)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   840
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   841
  have AB: "\<And>n. A n < B n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   842
    by (induct_tac n, simp add: ab, simp add: C_def avg_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   843
  have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   844
    apply (auto simp add: le_less [where 'a=nat])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   845
    apply (erule less_Suc_induct)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   846
    apply (clarsimp simp add: C_def avg_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   847
    apply (simp add: add_divide_distrib [symmetric])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   848
    apply (rule AB [THEN less_imp_le])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   849
    apply simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   850
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   851
  have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   852
    apply (auto simp add: le_less [where 'a=nat])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   853
    apply (erule less_Suc_induct)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   854
    apply (clarsimp simp add: C_def avg_def)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   855
    apply (simp add: add_divide_distrib [symmetric])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   856
    apply (rule AB [THEN less_imp_le])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   857
    apply simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   858
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   859
  have cauchy_lemma:
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   860
    "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   861
    apply (rule cauchyI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   862
    apply (drule twos [where y="b - a"])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   863
    apply (erule exE)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   864
    apply (rule_tac x=n in exI, clarify, rename_tac i j)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   865
    apply (rule_tac y="B n - A n" in le_less_trans) defer
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   866
    apply (simp add: width)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   867
    apply (drule_tac x=n in spec)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   868
    apply (frule_tac x=i in spec, drule (1) mp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   869
    apply (frule_tac x=j in spec, drule (1) mp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   870
    apply (frule A_mono, drule B_mono)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   871
    apply (frule A_mono, drule B_mono)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   872
    apply arith
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   873
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   874
  have "cauchy A"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   875
    apply (rule cauchy_lemma [rule_format])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   876
    apply (simp add: A_mono)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   877
    apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   878
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   879
  have "cauchy B"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   880
    apply (rule cauchy_lemma [rule_format])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   881
    apply (simp add: B_mono)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   882
    apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   883
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   884
  have 1: "\<forall>x\<in>S. x \<le> Real B"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   885
  proof
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   886
    fix x assume "x \<in> S"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   887
    then show "x \<le> Real B"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   888
      using PB [unfolded P_def] `cauchy B`
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   889
      by (simp add: le_RealI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   890
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   891
  have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   892
    apply clarify
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   893
    apply (erule contrapos_pp)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   894
    apply (simp add: not_le)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   895
    apply (drule less_RealD [OF `cauchy A`], clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   896
    apply (subgoal_tac "\<not> P (A n)")
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   897
    apply (simp add: P_def not_le, clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   898
    apply (erule rev_bexI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   899
    apply (erule (1) less_trans)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   900
    apply (simp add: PA)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   901
    done
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   902
  have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   903
  proof (rule vanishesI)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   904
    fix r :: rat assume "0 < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   905
    then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   906
      using twos by fast
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   907
    have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   908
    proof (clarify)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   909
      fix n assume n: "k \<le> n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   910
      have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   911
        by simp
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   912
      also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   913
        using n by (simp add: divide_left_mono mult_pos_pos)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   914
      also note k
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   915
      finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   916
    qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   917
    thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   918
  qed
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   919
  hence 3: "Real B = Real A"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   920
    by (simp add: eq_Real `cauchy A` `cauchy B` width)
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   921
  show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   922
    using 1 2 3 by (rule_tac x="Real B" in exI, simp)
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   923
qed
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   924
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   925
subsection {* Hiding implementation details *}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   926
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   927
hide_const (open) vanishes cauchy positive Real
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   928
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   929
declare Real_induct [induct del]
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   930
declare Abs_real_induct [induct del]
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   931
declare Abs_real_cases [cases del]
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
   932
47902
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   933
lemmas [transfer_rule del] =
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   934
  real.All_transfer real.Ex_transfer real.rel_eq_transfer forall_real_transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   935
  zero_real.transfer one_real.transfer plus_real.transfer uminus_real.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   936
  times_real.transfer inverse_real.transfer positive.transfer
34a9e81e5bfd convert real number theory to use lifting/transfer
huffman
parents: 47901
diff changeset
   937
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   938
subsection{*More Lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   939
36776
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   940
text {* BH: These lemmas should not be necessary; they should be
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   941
covered by existing simp rules and simplification procedures. *}
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   942
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   943
lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
36776
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   944
by simp (* redundant with mult_cancel_left *)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   945
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   946
lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
36776
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   947
by simp (* redundant with mult_cancel_right *)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   948
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   949
lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
36776
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   950
by simp (* solved by linordered_ring_less_cancel_factor simproc *)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   951
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   952
lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
36776
c137ae7673d3 remove a couple of redundant lemmas; simplify some proofs
huffman
parents: 36414
diff changeset
   953
by simp (* solved by linordered_ring_le_cancel_factor simproc *)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   954
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   955
lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
47428
45b58fec9024 remove outdated comment
huffman
parents: 47108
diff changeset
   956
by simp (* solved by linordered_ring_le_cancel_factor simproc *)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   957
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   958
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   959
subsection {* Embedding numbers into the Reals *}
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   960
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   961
abbreviation
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   962
  real_of_nat :: "nat \<Rightarrow> real"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   963
where
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   964
  "real_of_nat \<equiv> of_nat"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   965
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   966
abbreviation
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   967
  real_of_int :: "int \<Rightarrow> real"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   968
where
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   969
  "real_of_int \<equiv> of_int"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   970
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   971
abbreviation
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   972
  real_of_rat :: "rat \<Rightarrow> real"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   973
where
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   974
  "real_of_rat \<equiv> of_rat"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   975
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   976
consts
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   977
  (*overloaded constant for injecting other types into "real"*)
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   978
  real :: "'a => real"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   979
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   980
defs (overloaded)
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31952
diff changeset
   981
  real_of_nat_def [code_unfold]: "real == real_of_nat"
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31952
diff changeset
   982
  real_of_int_def [code_unfold]: "real == real_of_int"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   983
40939
2c150063cd4d setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
wenzelm
parents: 40864
diff changeset
   984
declare [[coercion_enabled]]
40864
4abaaadfdaf2 moved activation of coercion inference into RealDef and declared function real a coercion.
nipkow
parents: 40826
diff changeset
   985
declare [[coercion "real::nat\<Rightarrow>real"]]
4abaaadfdaf2 moved activation of coercion inference into RealDef and declared function real a coercion.
nipkow
parents: 40826
diff changeset
   986
declare [[coercion "real::int\<Rightarrow>real"]]
41022
81d337539d57 moved coercion decl. for int
nipkow
parents: 40939
diff changeset
   987
declare [[coercion "int"]]
40864
4abaaadfdaf2 moved activation of coercion inference into RealDef and declared function real a coercion.
nipkow
parents: 40826
diff changeset
   988
41024
ba961a606c67 move coercions to appropriate places
hoelzl
parents: 41022
diff changeset
   989
declare [[coercion_map map]]
42112
9cb122742f5c Change coercion for RealDef to use function application (not composition)
noschinl
parents: 41920
diff changeset
   990
declare [[coercion_map "% f g h x. g (h (f x))"]]
41024
ba961a606c67 move coercions to appropriate places
hoelzl
parents: 41022
diff changeset
   991
declare [[coercion_map "% f g (x,y) . (f x, g y)"]]
ba961a606c67 move coercions to appropriate places
hoelzl
parents: 41022
diff changeset
   992
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   993
lemma real_eq_of_nat: "real = of_nat"
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   994
  unfolding real_of_nat_def ..
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   995
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   996
lemma real_eq_of_int: "real = of_int"
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   997
  unfolding real_of_int_def ..
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   998
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   999
lemma real_of_int_zero [simp]: "real (0::int) = 0"  
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1000
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1001
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1002
lemma real_of_one [simp]: "real (1::int) = (1::real)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1003
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1004
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1005
lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1006
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1007
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1008
lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1009
by (simp add: real_of_int_def) 
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1010
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1011
lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1012
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1013
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1014
lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1015
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1016
35344
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1017
lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n"
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1018
by (simp add: real_of_int_def of_int_power)
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1019
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1020
lemmas power_real_of_int = real_of_int_power [symmetric]
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1021
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1022
lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1023
  apply (subst real_eq_of_int)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1024
  apply (rule of_int_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1025
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1026
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1027
lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) = 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1028
    (PROD x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1029
  apply (subst real_eq_of_int)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1030
  apply (rule of_int_setprod)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1031
done
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1032
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1033
lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1034
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1035
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1036
lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1037
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1038
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1039
lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1040
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1041
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1042
lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \<le> real y) = (x \<le> y)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1043
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1044
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1045
lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1046
by (simp add: real_of_int_def) 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1047
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1048
lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1049
by (simp add: real_of_int_def) 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1050
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1051
lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)" 
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1052
by (simp add: real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1053
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
  1054
lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)"
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1055
by (simp add: real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1056
47597
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1057
lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \<longleftrightarrow> 1 < i"
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1058
  unfolding real_of_one[symmetric] real_of_int_less_iff ..
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1059
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1060
lemma one_le_real_of_int_cancel_iff: "1 \<le> real (i :: int) \<longleftrightarrow> 1 \<le> i"
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1061
  unfolding real_of_one[symmetric] real_of_int_le_iff ..
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1062
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1063
lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \<longleftrightarrow> i < 1"
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1064
  unfolding real_of_one[symmetric] real_of_int_less_iff ..
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1065
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1066
lemma real_of_int_le_one_cancel_iff: "real (i :: int) \<le> 1 \<longleftrightarrow> i \<le> 1"
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1067
  unfolding real_of_one[symmetric] real_of_int_le_iff ..
5e7e394f78d4 add simp rules to rewrite comparisons of 1 and real
hoelzl
parents: 47489
diff changeset
  1068
16888
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
  1069
lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))"
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
  1070
by (auto simp add: abs_if)
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
  1071
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1072
lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1073
  apply (subgoal_tac "real n + 1 = real (n + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1074
  apply (simp del: real_of_int_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1075
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1076
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1077
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1078
lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1079
  apply (subgoal_tac "real m + 1 = real (m + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1080
  apply (simp del: real_of_int_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1081
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1082
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1083
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1084
lemma real_of_int_div_aux: "(real (x::int)) / (real d) = 
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1085
    real (x div d) + (real (x mod d)) / (real d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1086
proof -
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1087
  have "x = (x div d) * d + x mod d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1088
    by auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1089
  then have "real x = real (x div d) * real d + real(x mod d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1090
    by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1091
  then have "real x / real d = ... / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1092
    by simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1093
  then show ?thesis
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1094
    by (auto simp add: add_divide_distrib algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1095
qed
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1096
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1097
lemma real_of_int_div: "(d :: int) dvd n ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1098
    real(n div d) = real n / real d"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1099
  apply (subst real_of_int_div_aux)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1100
  apply simp
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29667
diff changeset
  1101
  apply (simp add: dvd_eq_mod_eq_0)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1102
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1103
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1104
lemma real_of_int_div2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1105
  "0 <= real (n::int) / real (x) - real (n div x)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1106
  apply (case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1107
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1108
  apply (case_tac "0 < x")
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
  1109
  apply (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1110
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1111
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1112
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1113
  apply auto
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
  1114
  apply (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1115
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1116
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1117
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1118
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1119
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1120
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1121
lemma real_of_int_div3:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1122
  "real (n::int) / real (x) - real (n div x) <= 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
  1123
  apply (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1124
  apply (subst real_of_int_div_aux)
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1125
  apply (auto simp add: divide_le_eq intro: order_less_imp_le)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1126
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1127
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1128
lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" 
27964
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
  1129
by (insert real_of_int_div2 [of n x], simp)
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
  1130
35635
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1131
lemma Ints_real_of_int [simp]: "real (x::int) \<in> Ints"
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1132
unfolding real_of_int_def by (rule Ints_of_int)
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1133
27964
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
  1134
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1135
subsection{*Embedding the Naturals into the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1136
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1137
lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1138
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1139
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
  1140
lemma real_of_nat_1 [simp]: "real (1::nat) = 1"
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
  1141
by (simp add: real_of_nat_def)
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
  1142
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1143
lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1144
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1145
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1146
lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1147
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1148
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1149
(*Not for addsimps: often the LHS is used to represent a positive natural*)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1150
lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1151
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1152
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1153
lemma real_of_nat_less_iff [iff]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1154
     "(real (n::nat) < real m) = (n < m)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1155
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1156
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1157
lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1158
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1159
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1160
lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
47489
04e7d09ade7a tuned some proofs;
huffman
parents: 47428
diff changeset
  1161
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1162
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1163
lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1164
by (simp add: real_of_nat_def del: of_nat_Suc)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1165
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1166
lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23289
diff changeset
  1167
by (simp add: real_of_nat_def of_nat_mult)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1168
35344
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1169
lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n"
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1170
by (simp add: real_of_nat_def of_nat_power)
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1171
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1172
lemmas power_real_of_nat = real_of_nat_power [symmetric]
e0b46cd72414 moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
huffman
parents: 35216
diff changeset
  1173
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1174
lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) = 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1175
    (SUM x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1176
  apply (subst real_eq_of_nat)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1177
  apply (rule of_nat_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1178
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1179
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1180
lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) = 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1181
    (PROD x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1182
  apply (subst real_eq_of_nat)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1183
  apply (rule of_nat_setprod)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1184
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1185
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1186
lemma real_of_card: "real (card A) = setsum (%x.1) A"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1187
  apply (subst card_eq_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1188
  apply (subst real_of_nat_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1189
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1190
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1191
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1192
lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1193
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1194
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1195
lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1196
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1197
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1198
lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
23438
dd824e86fa8a remove simp attribute from of_nat_diff, for backward compatibility with zdiff_int
huffman
parents: 23431
diff changeset
  1199
by (simp add: add: real_of_nat_def of_nat_diff)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1200
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25140
diff changeset
  1201
lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
25140
273772abbea2 More changes from >0 to ~=0::nat
nipkow
parents: 25134
diff changeset
  1202
by (auto simp: real_of_nat_def)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1203
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1204
lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1205
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1206
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
  1207
lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
  1208
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1209
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1210
lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1211
  apply (subgoal_tac "real n + 1 = real (Suc n)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1212
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1213
  apply (auto simp add: real_of_nat_Suc)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1214
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1215
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1216
lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1217
  apply (subgoal_tac "real m + 1 = real (Suc m)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1218
  apply (simp add: less_Suc_eq_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1219
  apply (simp add: real_of_nat_Suc)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1220
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1221
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1222
lemma real_of_nat_div_aux: "(real (x::nat)) / (real d) = 
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1223
    real (x div d) + (real (x mod d)) / (real d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1224
proof -
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1225
  have "x = (x div d) * d + x mod d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1226
    by auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1227
  then have "real x = real (x div d) * real d + real(x mod d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1228
    by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1229
  then have "real x / real d = \<dots> / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1230
    by simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1231
  then show ?thesis
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1232
    by (auto simp add: add_divide_distrib algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1233
qed
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1234
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1235
lemma real_of_nat_div: "(d :: nat) dvd n ==>
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1236
    real(n div d) = real n / real d"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1237
  by (subst real_of_nat_div_aux)
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46028
diff changeset
  1238
    (auto simp add: dvd_eq_mod_eq_0 [symmetric])
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1239
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1240
lemma real_of_nat_div2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1241
  "0 <= real (n::nat) / real (x) - real (n div x)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
  1242
apply (simp add: algebra_simps)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1243
apply (subst real_of_nat_div_aux)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1244
apply simp
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1245
apply (subst zero_le_divide_iff)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1246
apply simp
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1247
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1248
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1249
lemma real_of_nat_div3:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1250
  "real (n::nat) / real (x) - real (n div x) <= 1"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1251
apply(case_tac "x = 0")
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1252
apply (simp)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
  1253
apply (simp add: algebra_simps)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1254
apply (subst real_of_nat_div_aux)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
  1255
apply simp
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1256
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1257
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1258
lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
  1259
by (insert real_of_nat_div2 [of n x], simp)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1260
14426
de890c247b38 fixed bugs in the setup of arithmetic procedures
paulson
parents: 14421
diff changeset
  1261
lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n"
de890c247b38 fixed bugs in the setup of arithmetic procedures
paulson
parents: 14421
diff changeset
  1262
by (simp add: real_of_int_def real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1263
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1264
lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1265
  apply (subgoal_tac "real(int(nat x)) = real(nat x)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1266
  apply force
44822
2690b6de5021 remove duplicate lemma real_of_int_real_of_nat in favor of real_of_int_of_nat_eq
huffman
parents: 44766
diff changeset
  1267
  apply (simp only: real_of_int_of_nat_eq)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
  1268
done
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1269
35635
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1270
lemma Nats_real_of_nat [simp]: "real (n::nat) \<in> Nats"
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1271
unfolding real_of_nat_def by (rule of_nat_in_Nats)
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1272
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1273
lemma Ints_real_of_nat [simp]: "real (n::nat) \<in> Ints"
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1274
unfolding real_of_nat_def by (rule Ints_of_nat)
90fffd5ff996 add simp rules about Ints, Nats
huffman
parents: 35344
diff changeset
  1275
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1276
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1277
subsection{* Rationals *}
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1278
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1279
lemma Rats_real_nat[simp]: "real(n::nat) \<in> \<rat>"
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1280
by (simp add: real_eq_of_nat)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1281
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1282
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1283
lemma Rats_eq_int_div_int:
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1284
  "\<rat> = { real(i::int)/real(j::int) |i j. j \<noteq> 0}" (is "_ = ?S")
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1285
proof
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1286
  show "\<rat> \<subseteq> ?S"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1287
  proof
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1288
    fix x::real assume "x : \<rat>"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1289
    then obtain r where "x = of_rat r" unfolding Rats_def ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1290
    have "of_rat r : ?S"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1291
      by (cases r)(auto simp add:of_rat_rat real_eq_of_int)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1292
    thus "x : ?S" using `x = of_rat r` by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1293
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1294
next
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1295
  show "?S \<subseteq> \<rat>"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1296
  proof(auto simp:Rats_def)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1297
    fix i j :: int assume "j \<noteq> 0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1298
    hence "real i / real j = of_rat(Fract i j)"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1299
      by (simp add:of_rat_rat real_eq_of_int)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1300
    thus "real i / real j \<in> range of_rat" by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1301
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1302
qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1303
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1304
lemma Rats_eq_int_div_nat:
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
  1305
  "\<rat> = { real(i::int)/real(n::nat) |i n. n \<noteq> 0}"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1306
proof(auto simp:Rats_eq_int_div_int)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1307
  fix i j::int assume "j \<noteq> 0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1308
  show "EX (i'::int) (n::nat). real i/real j = real i'/real n \<and> 0<n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1309
  proof cases
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1310
    assume "j>0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1311
    hence "real i/real j = real i/real(nat j) \<and> 0<nat j"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1312
      by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1313
    thus ?thesis by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1314
  next
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1315
    assume "~ j>0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1316
    hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using `j\<noteq>0`
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1317
      by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1318
    thus ?thesis by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1319
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1320
next
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1321
  fix i::int and n::nat assume "0 < n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1322
  hence "real i/real n = real i/real(int n) \<and> int n \<noteq> 0" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1323
  thus "\<exists>(i'::int) j::int. real i/real n = real i'/real j \<and> j \<noteq> 0" by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1324
qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1325
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1326
lemma Rats_abs_nat_div_natE:
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1327
  assumes "x \<in> \<rat>"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 31641
diff changeset
  1328
  obtains m n :: nat
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 31641
diff changeset
  1329
  where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1330
proof -
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1331
  from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1332
    by(auto simp add: Rats_eq_int_div_nat)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1333
  hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1334
  then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1335
  let ?gcd = "gcd m n"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 31641
diff changeset
  1336
  from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1337
  let ?k = "m div ?gcd"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1338
  let ?l = "n div ?gcd"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1339
  let ?gcd' = "gcd ?k ?l"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1340
  have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1341
    by (rule dvd_mult_div_cancel)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1342
  have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1343
    by (rule dvd_mult_div_cancel)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1344
  from `n\<noteq>0` and gcd_l have "?l \<noteq> 0" by (auto iff del: neq0_conv)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1345
  moreover
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1346
  have "\<bar>x\<bar> = real ?k / real ?l"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1347
  proof -
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1348
    from gcd have "real ?k / real ?l =
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1349
        real (?gcd * ?k) / real (?gcd * ?l)" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1350
    also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1351
    also from x_rat have "\<dots> = \<bar>x\<bar>" ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1352
    finally show ?thesis ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1353
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1354
  moreover
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1355
  have "?gcd' = 1"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1356
  proof -
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1357
    have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31707
diff changeset
  1358
      by (rule gcd_mult_distrib_nat)
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1359
    with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 31641
diff changeset
  1360
    with gcd show ?thesis by auto
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1361
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1362
  ultimately show ?thesis ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1363
qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1364
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
  1365
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1366
subsection{*Numerals and Arithmetic*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1367
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45859
diff changeset
  1368
lemma [code_abbrev]:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1369
  "real_of_int (numeral k) = numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1370
  "real_of_int (neg_numeral k) = neg_numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1371
  by simp_all
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1372
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1373
text{*Collapse applications of @{term real} to @{term number_of}*}
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1374
lemma real_numeral [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1375
  "real (numeral v :: int) = numeral v"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1376
  "real (neg_numeral v :: int) = neg_numeral v"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1377
by (simp_all add: real_of_int_def)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1378
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1379
lemma real_of_nat_numeral [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1380
  "real (numeral v :: nat) = numeral v"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1381
by (simp add: real_of_nat_def)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1382
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1383
declaration {*
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1384
  K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2]
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1385
    (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *)
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1386
  #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2]
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1387
    (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *)
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1388
  #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add},
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1389
      @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one},
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1390
      @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1391
      @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1392
      @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}]
36795
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
  1393
  #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
e05e1283c550 new construction of real numbers using Cauchy sequences
huffman
parents: 36776
diff changeset
  1394
  #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
  1395
*}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1396
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
  1397
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1398
subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1399
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1400
lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1401
by arith
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1402
36839
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1403
text {* FIXME: redundant with @{text add_eq_0_iff} below *}
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
  1404
lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1405
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1406
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
  1407
lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1408
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1409
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
  1410
lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1411
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1412
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
  1413
lemma real_add_le_0_iff: "(x+y \<le> (0::real)) = (y \<le> -x)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1414
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1415
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
  1416
lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1417
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1418
36839
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1419
subsection {* Lemmas about powers *}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1420
36839
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1421
text {* FIXME: declare this in Rings.thy or not at all *}
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1422
declare abs_mult_self [simp]
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1423
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1424
(* used by Import/HOL/real.imp *)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1425
lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1426
by simp
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1427
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1428
lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1429
apply (induct "n")
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1430
apply (auto simp add: real_of_nat_Suc)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1431
apply (subst mult_2)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1432
apply (erule add_less_le_mono)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1433
apply (rule two_realpow_ge_one)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1434
done
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1435
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1436
text {* TODO: no longer real-specific; rename and move elsewhere *}
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1437
lemma realpow_Suc_le_self:
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1438
  fixes r :: "'a::linordered_semidom"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1439
  shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1440
by (insert power_decreasing [of 1 "Suc n" r], simp)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1441
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1442
text {* TODO: no longer real-specific; rename and move elsewhere *}
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1443
lemma realpow_minus_mult:
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1444
  fixes x :: "'a::monoid_mult"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1445
  shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1446
by (simp add: power_commutes split add: nat_diff_split)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1447
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1448
text {* FIXME: declare this [simp] for all types, or not at all *}
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1449
lemma real_two_squares_add_zero_iff [simp]:
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1450
  "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1451
by (rule sum_squares_eq_zero_iff)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1452
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1453
text {* FIXME: declare this [simp] for all types, or not at all *}
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1454
lemma realpow_two_sum_zero_iff [simp]:
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1455
     "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1456
by (rule sum_power2_eq_zero_iff)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1457
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1458
lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1459
by (rule_tac y = 0 in order_trans, auto)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1460
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1461
lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2"
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1462
by (auto simp add: power2_eq_square)
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1463
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1464
47598
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1465
lemma numeral_power_le_real_of_nat_cancel_iff[simp]:
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1466
  "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::nat) ^ n \<le> a"
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1467
  unfolding real_of_nat_le_iff[symmetric] by simp
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1468
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1469
lemma real_of_nat_le_numeral_power_cancel_iff[simp]:
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1470
  "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::nat) ^ n"
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1471
  unfolding real_of_nat_le_iff[symmetric] by simp
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1472
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1473
lemma numeral_power_le_real_of_int_cancel_iff[simp]:
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1474
  "(numeral x::real) ^ n \<le> real a \<longleftrightarrow> (numeral x::int) ^ n \<le> a"
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1475
  unfolding real_of_int_le_iff[symmetric] by simp
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1476
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1477
lemma real_of_int_le_numeral_power_cancel_iff[simp]:
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1478
  "real a \<le> (numeral x::real) ^ n \<longleftrightarrow> a \<le> (numeral x::int) ^ n"
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1479
  unfolding real_of_int_le_iff[symmetric] by simp
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1480
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1481
lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]:
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1482
  "(neg_numeral x::real) ^ n \<le> real a \<longleftrightarrow> (neg_numeral x::int) ^ n \<le> a"
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1483
  unfolding real_of_int_le_iff[symmetric] by simp
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1484
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1485
lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]:
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1486
  "real a \<le> (neg_numeral x::real) ^ n \<longleftrightarrow> a \<le> (neg_numeral x::int) ^ n"
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1487
  unfolding real_of_int_le_iff[symmetric] by simp
d20bdee675dc add lemmas to remove real conversions when compared to power of numerals
hoelzl
parents: 47597
diff changeset
  1488
36839
34dc65df7014 no more RealPow.thy (remaining lemmas moved to RealDef.thy)
huffman
parents: 36796
diff changeset
  1489
subsection{*Density of the Reals*}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1490
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1491
lemma real_lbound_gt_zero:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1492
     "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1493
apply (rule_tac x = " (min d1 d2) /2" in exI)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1494
apply (simp add: min_def)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1495
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1496
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1497
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35032
diff changeset
  1498
text{*Similar results are proved in @{text Fields}*}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1499
lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1500
  by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1501
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1502
lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1503
  by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1504
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1505
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1506
subsection{*Absolute Value Function for the Reals*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1507
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1508
lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14754
diff changeset
  1509
by (simp add: abs_if)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1510
23289
0cf371d943b1 remove redundant lemmas
huffman
parents: 23288
diff changeset
  1511
(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1512
lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35032
diff changeset
  1513
by (force simp add: abs_le_iff)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1514
44344
49be3e7d4762 remove some over-specific rules from the simpset
huffman
parents: 44278
diff changeset
  1515
lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)"
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14754
diff changeset
  1516
by (simp add: abs_if)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1517
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1518
lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
22958
b3a5569a81e5 cleaned up
huffman
parents: 22456
diff changeset
  1519
by (rule abs_of_nonneg [OF real_of_nat_ge_zero])
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1520
44344
49be3e7d4762 remove some over-specific rules from the simpset
huffman
parents: 44278
diff changeset
  1521
lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
  1522
by simp
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1523
 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1524
lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
  1525
by simp
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1526
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1527
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1528
subsection {* Implementation of rational real numbers *}
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1529
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1530
text {* Formal constructor *}
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1531
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1532
definition Ratreal :: "rat \<Rightarrow> real" where
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1533
  [code_abbrev, simp]: "Ratreal = of_rat"
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1534
24623
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1535
code_datatype Ratreal
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1536
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1537
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1538
text {* Numerals *}
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1539
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1540
lemma [code_abbrev]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1541
  "(of_rat (of_int a) :: real) = of_int a"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1542
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1543
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1544
lemma [code_abbrev]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1545
  "(of_rat 0 :: real) = 0"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1546
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1547
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1548
lemma [code_abbrev]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1549
  "(of_rat 1 :: real) = 1"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1550
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1551
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1552
lemma [code_abbrev]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1553
  "(of_rat (numeral k) :: real) = numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1554
  by simp
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1555
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1556
lemma [code_abbrev]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1557
  "(of_rat (neg_numeral k) :: real) = neg_numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1558
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1559
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1560
lemma [code_post]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1561
  "(of_rat (0 / r)  :: real) = 0"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1562
  "(of_rat (r / 0)  :: real) = 0"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1563
  "(of_rat (1 / 1)  :: real) = 1"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1564
  "(of_rat (numeral k / 1) :: real) = numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1565
  "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1566
  "(of_rat (1 / numeral k) :: real) = 1 / numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1567
  "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1568
  "(of_rat (numeral k / numeral l)  :: real) = numeral k / numeral l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1569
  "(of_rat (numeral k / neg_numeral l)  :: real) = numeral k / neg_numeral l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1570
  "(of_rat (neg_numeral k / numeral l)  :: real) = neg_numeral k / numeral l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1571
  "(of_rat (neg_numeral k / neg_numeral l)  :: real) = neg_numeral k / neg_numeral l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1572
  by (simp_all add: of_rat_divide)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1573
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1574
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1575
text {* Operations *}
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1576
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1577
lemma zero_real_code [code]:
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1578
  "0 = Ratreal 0"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1579
by simp
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1580
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1581
lemma one_real_code [code]:
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1582
  "1 = Ratreal 1"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1583
by simp
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1584
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1585
instantiation real :: equal
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1586
begin
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1587
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1588
definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1589
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1590
instance proof
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1591
qed (simp add: equal_real_def)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1592
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1593
lemma real_equal_code [code]:
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1594
  "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1595
  by (simp add: equal_real_def equal)
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1596
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1597
lemma [code nbe]:
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1598
  "HOL.equal (x::real) x \<longleftrightarrow> True"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38287
diff changeset
  1599
  by (rule equal_refl)
28351
abfc66969d1f non left-linear equations for nbe
haftmann
parents: 28091
diff changeset
  1600
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1601
end
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1602
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1603
lemma real_less_eq_code [code]: "Ratreal x \<le> Ratreal y \<longleftrightarrow> x \<le> y"
27652
818666de6c24 refined code generator setup for rational numbers; more simplification rules for rational numbers
haftmann
parents: 27544
diff changeset
  1604
  by (simp add: of_rat_less_eq)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1605
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1606
lemma real_less_code [code]: "Ratreal x < Ratreal y \<longleftrightarrow> x < y"
27652
818666de6c24 refined code generator setup for rational numbers; more simplification rules for rational numbers
haftmann
parents: 27544
diff changeset
  1607
  by (simp add: of_rat_less)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1608
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1609
lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1610
  by (simp add: of_rat_add)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1611
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1612
lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1613
  by (simp add: of_rat_mult)
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1614
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1615
lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1616
  by (simp add: of_rat_minus)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1617
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1618
lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1619
  by (simp add: of_rat_diff)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1620
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1621
lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1622
  by (simp add: of_rat_inverse)
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1623
 
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1624
lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1625
  by (simp add: of_rat_divide)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1626
43733
a6ca7b83612f adding code equations to execute floor and ceiling on rational and real numbers
bulwahn
parents: 43732
diff changeset
  1627
lemma real_floor_code [code]: "floor (Ratreal x) = floor x"
a6ca7b83612f adding code equations to execute floor and ceiling on rational and real numbers
bulwahn
parents: 43732
diff changeset
  1628
  by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code)
a6ca7b83612f adding code equations to execute floor and ceiling on rational and real numbers
bulwahn
parents: 43732
diff changeset
  1629
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1630
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1631
text {* Quickcheck *}
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1632
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1633
definition (in term_syntax)
32657
5f13912245ff Code_Eval(uation)
haftmann
parents: 32069
diff changeset
  1634
  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
5f13912245ff Code_Eval(uation)
haftmann
parents: 32069
diff changeset
  1635
  [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1636
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37397
diff changeset
  1637
notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37397
diff changeset
  1638
notation scomp (infixl "\<circ>\<rightarrow>" 60)
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1639
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1640
instantiation real :: random
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1641
begin
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1642
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1643
definition
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37397
diff changeset
  1644
  "Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1645
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1646
instance ..
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1647
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1648
end
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1649
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37397
diff changeset
  1650
no_notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37397
diff changeset
  1651
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1652
41920
d4fb7a418152 moving exhaustive_generators.ML to Quickcheck directory
bulwahn
parents: 41792
diff changeset
  1653
instantiation real :: exhaustive
41231
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1654
begin
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1655
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1656
definition
45818
53a697f5454a hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
bulwahn
parents: 45184
diff changeset
  1657
  "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d"
42311
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1658
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1659
instance ..
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1660
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1661
end
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1662
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1663
instantiation real :: full_exhaustive
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1664
begin
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1665
eb32a8474a57 rational and real instances for new compilation scheme for exhaustive quickcheck
bulwahn
parents: 42112
diff changeset
  1666
definition
45818
53a697f5454a hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
bulwahn
parents: 45184
diff changeset
  1667
  "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d"
41231
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1668
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1669
instance ..
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1670
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1671
end
2e901158675e adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents: 41024
diff changeset
  1672
43887
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1673
instantiation real :: narrowing
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1674
begin
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1675
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1676
definition
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1677
  "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing"
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1678
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1679
instance ..
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1680
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1681
end
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1682
442aceb54969 adding narrowing instances for real and rational
bulwahn
parents: 43733
diff changeset
  1683
45184
426dbd896c9e removing old code generator setup for real numbers; tuned
bulwahn
parents: 45051
diff changeset
  1684
subsection {* Setup for Nitpick *}
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1685
38287
796302ca3611 replace "setup" with "declaration"
blanchet
parents: 38242
diff changeset
  1686
declaration {*
796302ca3611 replace "setup" with "declaration"
blanchet
parents: 38242
diff changeset
  1687
  Nitpick_HOL.register_frac_type @{type_name real}
33209
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1688
   [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}),
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1689
    (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}),
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1690
    (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}),
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1691
    (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}),
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1692
    (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}),
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1693
    (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
45859
36ff12b5663b fixed Nitpick definition of "<" on "real"s
blanchet
parents: 45818
diff changeset
  1694
    (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}),
33209
d36ca3960e33 tuned white space;
wenzelm
parents: 33197
diff changeset
  1695
    (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32657
diff changeset
  1696
*}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32657
diff changeset
  1697
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  1698
lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real
37397
18000f9d783e adjust Nitpick's handling of "<" on "rat"s and "reals"
blanchet
parents: 36977
diff changeset
  1699
    ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32657
diff changeset
  1700
    times_real_inst.times_real uminus_real_inst.uminus_real
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32657
diff changeset
  1701
    zero_real_inst.zero_real
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32657
diff changeset
  1702
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
  1703
end