src/HOL/RealDef.thy
author haftmann
Mon, 15 Jun 2009 16:13:03 +0200
changeset 31641 feea4d3d743d
parent 31203 5c8fb4fd67e0
child 31666 760c612ad800
child 31706 1db0c8f235fb
permissions -rw-r--r--
hide constant Quickcheck.random
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
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
     4
    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
     5
    Additional contributions by Jeremy Avigad
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
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
     8
header{*Defining the Reals from the Positive Reals*}
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
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    11
imports PReal
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
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    14
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20554
diff changeset
    15
  realrel   ::  "((preal * preal) * (preal * preal)) set" where
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    16
  [code del]: "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    17
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    18
typedef (Real)  real = "UNIV//realrel"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    19
  by (auto simp add: quotient_def)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    20
19765
dfe940911617 misc cleanup;
wenzelm
parents: 19023
diff changeset
    21
definition
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    22
  (** these don't use the overloaded "real" function: users don't see them **)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20554
diff changeset
    23
  real_of_preal :: "preal => real" where
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    24
  [code del]: "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    25
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 25571
diff changeset
    26
instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    27
begin
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    28
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    29
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    30
  real_zero_def [code del]: "0 = Abs_Real(realrel``{(1, 1)})"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    31
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    32
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    33
  real_one_def [code del]: "1 = Abs_Real(realrel``{(1 + 1, 1)})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    34
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    35
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    36
  real_add_def [code del]: "z + w =
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    37
       contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
27964
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
    38
                 { Abs_Real(realrel``{(x+u, y+v)}) })"
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    39
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    40
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    41
  real_minus_def [code del]: "- r =  contents (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    42
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    43
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    44
  real_diff_def [code del]: "r - (s::real) = r + - s"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    45
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    46
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    47
  real_mult_def [code del]:
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    48
    "z * w =
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    49
       contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
27964
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
    50
                 { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    51
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    52
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    53
  real_inverse_def [code del]: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    54
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    55
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    56
  real_divide_def [code del]: "R / (S::real) = R * inverse S"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    57
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    58
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    59
  real_le_def [code del]: "z \<le> (w::real) \<longleftrightarrow>
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    60
    (\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w)"
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    61
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    62
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
    63
  real_less_def [code del]: "x < (y\<Colon>real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    64
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    65
definition
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    66
  real_abs_def:  "abs (r::real) = (if r < 0 then - r else r)"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    67
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    68
definition
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    69
  real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0<x then 1 else - 1)"
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    70
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    71
instance ..
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    72
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
    73
end
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    74
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
    75
subsection {* Equivalence relation over positive reals *}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    76
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    77
lemma preal_trans_lemma:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    78
  assumes "x + y1 = x1 + y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    79
      and "x + y2 = x2 + y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    80
  shows "x1 + y2 = x2 + (y1::preal)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    81
proof -
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
    82
  have "(x1 + y2) + x = (x + y2) + x1" by (simp add: add_ac)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    83
  also have "... = (x2 + y) + x1"  by (simp add: prems)
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
    84
  also have "... = x2 + (x1 + y)"  by (simp add: add_ac)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    85
  also have "... = x2 + (x + y1)"  by (simp add: prems)
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
    86
  also have "... = (x2 + y1) + x"  by (simp add: add_ac)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    87
  finally have "(x1 + y2) + x = (x2 + y1) + x" .
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
    88
  thus ?thesis by (rule add_right_imp_eq)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    89
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    90
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    91
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    92
lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    93
by (simp add: realrel_def)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    94
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    95
lemma equiv_realrel: "equiv UNIV realrel"
30198
922f944f03b2 name changes
nipkow
parents: 30082
diff changeset
    96
apply (auto simp add: equiv_def refl_on_def sym_def trans_def realrel_def)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    97
apply (blast dest: preal_trans_lemma) 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    98
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    99
14497
paulson
parents: 14484
diff changeset
   100
text{*Reduces equality of equivalence classes to the @{term realrel} relation:
paulson
parents: 14484
diff changeset
   101
  @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   102
lemmas equiv_realrel_iff = 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   103
       eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   104
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   105
declare equiv_realrel_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   106
14497
paulson
parents: 14484
diff changeset
   107
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   108
lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   109
by (simp add: Real_def realrel_def quotient_def, blast)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   110
22958
b3a5569a81e5 cleaned up
huffman
parents: 22456
diff changeset
   111
declare Abs_Real_inject [simp]
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   112
declare Abs_Real_inverse [simp]
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   113
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   114
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   115
text{*Case analysis on the representation of a real number as an equivalence
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   116
      class of pairs of positive reals.*}
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   117
lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   118
     "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   119
apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE])
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   120
apply (drule arg_cong [where f=Abs_Real])
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   121
apply (auto simp add: Rep_Real_inverse)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   122
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   123
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   124
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   125
subsection {* Addition and Subtraction *}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   126
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   127
lemma real_add_congruent2_lemma:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   128
     "[|a + ba = aa + b; ab + bc = ac + bb|]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   129
      ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   130
apply (simp add: add_assoc)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   131
apply (rule add_left_commute [of ab, THEN ssubst])
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   132
apply (simp add: add_assoc [symmetric])
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   133
apply (simp add: add_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   134
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   135
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   136
lemma real_add:
14497
paulson
parents: 14484
diff changeset
   137
     "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
paulson
parents: 14484
diff changeset
   138
      Abs_Real (realrel``{(x+u, y+v)})"
paulson
parents: 14484
diff changeset
   139
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   140
  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   141
        respects2 realrel"
14497
paulson
parents: 14484
diff changeset
   142
    by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma) 
paulson
parents: 14484
diff changeset
   143
  thus ?thesis
paulson
parents: 14484
diff changeset
   144
    by (simp add: real_add_def UN_UN_split_split_eq
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14497
diff changeset
   145
                  UN_equiv_class2 [OF equiv_realrel equiv_realrel])
14497
paulson
parents: 14484
diff changeset
   146
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   147
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   148
lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   149
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   150
  have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   151
    by (simp add: congruent_def add_commute) 
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   152
  thus ?thesis
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   153
    by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   154
qed
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   155
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   156
instance real :: ab_group_add
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   157
proof
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   158
  fix x y z :: real
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   159
  show "(x + y) + z = x + (y + z)"
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   160
    by (cases x, cases y, cases z, simp add: real_add add_assoc)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   161
  show "x + y = y + x"
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   162
    by (cases x, cases y, simp add: real_add add_commute)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   163
  show "0 + x = x"
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   164
    by (cases x, simp add: real_add real_zero_def add_ac)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   165
  show "- x + x = 0"
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   166
    by (cases x, simp add: real_minus real_add real_zero_def add_commute)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   167
  show "x - y = x + - y"
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   168
    by (simp add: real_diff_def)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   169
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   170
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   171
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   172
subsection {* Multiplication *}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   173
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   174
lemma real_mult_congruent2_lemma:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   175
     "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   176
          x * x1 + y * y1 + (x * y2 + y * x2) =
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   177
          x * x2 + y * y2 + (x * y1 + y * x1)"
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   178
apply (simp add: add_left_commute add_assoc [symmetric])
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   179
apply (simp add: add_assoc right_distrib [symmetric])
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   180
apply (simp add: add_commute)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   181
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   182
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   183
lemma real_mult_congruent2:
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   184
    "(%p1 p2.
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   185
        (%(x1,y1). (%(x2,y2). 
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   186
          { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   187
     respects2 realrel"
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14497
diff changeset
   188
apply (rule congruent2_commuteI [OF equiv_realrel], clarify)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   189
apply (simp add: mult_commute add_commute)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   190
apply (auto simp add: real_mult_congruent2_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   191
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   192
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   193
lemma real_mult:
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   194
      "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   195
       Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   196
by (simp add: real_mult_def UN_UN_split_split_eq
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14497
diff changeset
   197
         UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   198
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   199
lemma real_mult_commute: "(z::real) * w = w * z"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   200
by (cases z, cases w, simp add: real_mult add_ac mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   201
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   202
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   203
apply (cases z1, cases z2, cases z3)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   204
apply (simp add: real_mult algebra_simps)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   205
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   206
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   207
lemma real_mult_1: "(1::real) * z = z"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   208
apply (cases z)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   209
apply (simp add: real_mult real_one_def algebra_simps)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   210
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   211
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   212
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   213
apply (cases z1, cases z2, cases w)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   214
apply (simp add: real_add real_mult algebra_simps)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   215
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   216
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   217
text{*one and zero are distinct*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   218
lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   219
proof -
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   220
  have "(1::preal) < 1 + 1"
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   221
    by (simp add: preal_self_less_add_left)
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   222
  thus ?thesis
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   223
    by (simp add: real_zero_def real_one_def)
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   224
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   225
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   226
instance real :: comm_ring_1
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   227
proof
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   228
  fix x y z :: real
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   229
  show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   230
  show "x * y = y * x" by (rule real_mult_commute)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   231
  show "1 * x = x" by (rule real_mult_1)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   232
  show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   233
  show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   234
qed
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   235
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   236
subsection {* Inverse and Division *}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   237
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   238
lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   239
by (simp add: real_zero_def add_commute)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   240
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   241
text{*Instead of using an existential quantifier and constructing the inverse
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   242
within the proof, we could define the inverse explicitly.*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   243
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   244
lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   245
apply (simp add: real_zero_def real_one_def, cases x)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   246
apply (cut_tac x = xa and y = y in linorder_less_linear)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   247
apply (auto dest!: less_add_left_Ex simp add: real_zero_iff)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   248
apply (rule_tac
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   249
        x = "Abs_Real (realrel``{(1, inverse (D) + 1)})"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   250
       in exI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   251
apply (rule_tac [2]
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   252
        x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})" 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   253
       in exI)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   254
apply (auto simp add: real_mult preal_mult_inverse_right algebra_simps)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   255
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   256
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   257
lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   258
apply (simp add: real_inverse_def)
23287
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   259
apply (drule real_mult_inverse_left_ex, safe)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   260
apply (rule theI, assumption, rename_tac z)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   261
apply (subgoal_tac "(z * x) * y = z * (x * y)")
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   262
apply (simp add: mult_commute)
063039db59dd define (1::preal); clean up instance declarations
huffman
parents: 23031
diff changeset
   263
apply (rule mult_assoc)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   264
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   265
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   266
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   267
subsection{*The Real Numbers form a Field*}
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   268
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   269
instance real :: field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   270
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   271
  fix x y z :: real
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   272
  show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14426
diff changeset
   273
  show "x / y = x * inverse y" by (simp add: real_divide_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   274
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   275
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   276
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   277
text{*Inverse of zero!  Useful to simplify certain equations*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   278
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   279
lemma INVERSE_ZERO: "inverse 0 = (0::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   280
by (simp add: real_inverse_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   281
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   282
instance real :: division_by_zero
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   283
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   284
  show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   285
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   286
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   287
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   288
subsection{*The @{text "\<le>"} Ordering*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   289
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   290
lemma real_le_refl: "w \<le> (w::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   291
by (cases w, force simp add: real_le_def)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   292
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   293
text{*The arithmetic decision procedure is not set up for type preal.
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   294
  This lemma is currently unused, but it could simplify the proofs of the
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   295
  following two lemmas.*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   296
lemma preal_eq_le_imp_le:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   297
  assumes eq: "a+b = c+d" and le: "c \<le> a"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   298
  shows "b \<le> (d::preal)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   299
proof -
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   300
  have "c+d \<le> a+d" by (simp add: prems)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   301
  hence "a+b \<le> a+d" by (simp add: prems)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   302
  thus "b \<le> d" by simp
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   303
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   304
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   305
lemma real_le_lemma:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   306
  assumes l: "u1 + v2 \<le> u2 + v1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   307
      and "x1 + v1 = u1 + y1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   308
      and "x2 + v2 = u2 + y2"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   309
  shows "x1 + y2 \<le> x2 + (y1::preal)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   310
proof -
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   311
  have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   312
  hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: add_ac)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   313
  also have "... \<le> (x2+y1) + (u2+v1)" by (simp add: prems)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   314
  finally show ?thesis by simp
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   315
qed
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   316
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   317
lemma real_le: 
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   318
     "(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) =  
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   319
      (x1 + y2 \<le> x2 + y1)"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   320
apply (simp add: real_le_def)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   321
apply (auto intro: real_le_lemma)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   322
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   323
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   324
lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15229
diff changeset
   325
by (cases z, cases w, simp add: real_le)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   326
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   327
lemma real_trans_lemma:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   328
  assumes "x + v \<le> u + y"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   329
      and "u + v' \<le> u' + v"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   330
      and "x2 + v2 = u2 + y2"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   331
  shows "x + v' \<le> u' + (y::preal)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   332
proof -
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   333
  have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: add_ac)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   334
  also have "... \<le> (u+y) + (u+v')" by (simp add: prems)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   335
  also have "... \<le> (u+y) + (u'+v)" by (simp add: prems)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   336
  also have "... = (u'+y) + (u+v)"  by (simp add: add_ac)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   337
  finally show ?thesis by simp
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15229
diff changeset
   338
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   339
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   340
lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   341
apply (cases i, cases j, cases k)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   342
apply (simp add: real_le)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   343
apply (blast intro: real_trans_lemma)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   344
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   345
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   346
instance real :: order
27682
25aceefd4786 added class preorder
haftmann
parents: 27668
diff changeset
   347
proof
25aceefd4786 added class preorder
haftmann
parents: 27668
diff changeset
   348
  fix u v :: real
25aceefd4786 added class preorder
haftmann
parents: 27668
diff changeset
   349
  show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u" 
25aceefd4786 added class preorder
haftmann
parents: 27668
diff changeset
   350
    by (auto simp add: real_less_def intro: real_le_anti_sym)
25aceefd4786 added class preorder
haftmann
parents: 27668
diff changeset
   351
qed (assumption | rule real_le_refl real_le_trans real_le_anti_sym)+
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   352
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   353
(* Axiom 'linorder_linear' of class 'linorder': *)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   354
lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   355
apply (cases z, cases w)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   356
apply (auto simp add: real_le real_zero_def add_ac)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   357
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   358
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   359
instance real :: linorder
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   360
  by (intro_classes, rule real_le_linear)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   361
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   362
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   363
lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   364
apply (cases x, cases y) 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   365
apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   366
                      add_ac)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   367
apply (simp_all add: add_assoc [symmetric])
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15229
diff changeset
   368
done
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   369
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   370
lemma real_add_left_mono: 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   371
  assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   372
proof -
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   373
  have "z + x - (z + y) = (z + -z) + (x - y)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   374
    by (simp add: algebra_simps) 
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   375
  with le show ?thesis 
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
   376
    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   377
qed
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   378
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   379
lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   380
by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   381
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   382
lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   383
by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   384
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   385
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   386
apply (cases x, cases y)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   387
apply (simp add: linorder_not_le [where 'a = real, symmetric] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   388
                 linorder_not_le [where 'a = preal] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   389
                  real_zero_def real_le real_mult)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   390
  --{*Reduce to the (simpler) @{text "\<le>"} relation *}
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16888
diff changeset
   391
apply (auto dest!: less_add_left_Ex
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   392
     simp add: algebra_simps preal_self_less_add_left)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   393
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   394
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   395
lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   396
apply (rule real_sum_gt_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   397
apply (drule real_less_sum_gt_zero [of x y])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   398
apply (drule real_mult_order, assumption)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   399
apply (simp add: right_distrib)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   400
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   401
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   402
instantiation real :: distrib_lattice
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   403
begin
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   404
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   405
definition
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   406
  "(inf \<Colon> real \<Rightarrow> real \<Rightarrow> real) = min"
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   407
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   408
definition
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   409
  "(sup \<Colon> real \<Rightarrow> real \<Rightarrow> real) = max"
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   410
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   411
instance
22456
6070e48ecb78 added lattice definitions
haftmann
parents: 21404
diff changeset
   412
  by default (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
6070e48ecb78 added lattice definitions
haftmann
parents: 21404
diff changeset
   413
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   414
end
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   415
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   416
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   417
subsection{*The Reals Form an Ordered Field*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   418
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   419
instance real :: ordered_field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   420
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   421
  fix x y z :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   422
  show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono)
22962
4bb05ba38939 remove redundant lemmas
huffman
parents: 22958
diff changeset
   423
  show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2)
4bb05ba38939 remove redundant lemmas
huffman
parents: 22958
diff changeset
   424
  show "\<bar>x\<bar> = (if x < 0 then -x else x)" by (simp only: real_abs_def)
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24198
diff changeset
   425
  show "sgn x = (if x=0 then 0 else if 0<x then 1 else - 1)"
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24198
diff changeset
   426
    by (simp only: real_sgn_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   427
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   428
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 25162
diff changeset
   429
instance real :: lordered_ab_group_add ..
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 25162
diff changeset
   430
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   431
text{*The function @{term real_of_preal} requires many proofs, but it seems
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   432
to be essential for proving completeness of the reals from that of the
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   433
positive reals.*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   434
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   435
lemma real_of_preal_add:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   436
     "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   437
by (simp add: real_of_preal_def real_add algebra_simps)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   438
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   439
lemma real_of_preal_mult:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   440
     "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   441
by (simp add: real_of_preal_def real_mult algebra_simps)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   442
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   443
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   444
text{*Gleason prop 9-4.4 p 127*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   445
lemma real_of_preal_trichotomy:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   446
      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   447
apply (simp add: real_of_preal_def real_zero_def, cases x)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   448
apply (auto simp add: real_minus add_ac)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   449
apply (cut_tac x = x and y = y in linorder_less_linear)
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   450
apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric])
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   451
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   452
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   453
lemma real_of_preal_leD:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   454
      "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   455
by (simp add: real_of_preal_def real_le)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   456
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   457
lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   458
by (auto simp add: real_of_preal_leD linorder_not_le [symmetric])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   459
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   460
lemma real_of_preal_lessD:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   461
      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   462
by (simp add: real_of_preal_def real_le linorder_not_le [symmetric])
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   463
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   464
lemma real_of_preal_less_iff [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   465
     "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   466
by (blast intro: real_of_preal_lessI real_of_preal_lessD)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   467
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   468
lemma real_of_preal_le_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   469
     "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   470
by (simp add: linorder_not_less [symmetric])
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   471
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   472
lemma real_of_preal_zero_less: "0 < real_of_preal m"
23288
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   473
apply (insert preal_self_less_add_left [of 1 m])
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   474
apply (auto simp add: real_zero_def real_of_preal_def
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   475
                      real_less_def real_le_def add_ac)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   476
apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI)
3e45b5464d2b remove references to preal-specific theorems
huffman
parents: 23287
diff changeset
   477
apply (simp add: add_ac)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   478
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   479
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   480
lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   481
by (simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   482
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   483
lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   484
proof -
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   485
  from real_of_preal_minus_less_zero
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   486
  show ?thesis by (blast dest: order_less_trans)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   487
qed
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   488
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   489
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   490
subsection{*Theorems About the Ordering*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   491
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   492
lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   493
apply (auto simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   494
apply (cut_tac x = x in real_of_preal_trichotomy)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   495
apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   496
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   497
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   498
lemma real_gt_preal_preal_Ex:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   499
     "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   500
by (blast dest!: real_of_preal_zero_less [THEN order_less_trans]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   501
             intro: real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   502
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   503
lemma real_ge_preal_preal_Ex:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   504
     "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   505
by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   506
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   507
lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   508
by (auto elim: order_le_imp_less_or_eq [THEN disjE] 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   509
            intro: real_of_preal_zero_less [THEN [2] order_less_trans] 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   510
            simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   511
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   512
lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   513
by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   514
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   515
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   516
subsection{*More Lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   517
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   518
lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   519
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   520
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   521
lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   522
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   523
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   524
lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   525
  by (force elim: order_less_asym
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   526
            simp add: Ring_and_Field.mult_less_cancel_right)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   527
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   528
lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   529
apply (simp add: mult_le_cancel_right)
23289
0cf371d943b1 remove redundant lemmas
huffman
parents: 23288
diff changeset
   530
apply (blast intro: elim: order_less_asym)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   531
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   532
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   533
lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
15923
01d5d0c1c078 fixed lin.arith
nipkow
parents: 15542
diff changeset
   534
by(simp add:mult_commute)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   535
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   536
lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x"
23289
0cf371d943b1 remove redundant lemmas
huffman
parents: 23288
diff changeset
   537
by (simp add: one_less_inverse_iff) (* TODO: generalize/move *)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   538
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   539
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   540
subsection {* Embedding numbers into the Reals *}
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   541
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   542
abbreviation
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   543
  real_of_nat :: "nat \<Rightarrow> real"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   544
where
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   545
  "real_of_nat \<equiv> of_nat"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   546
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   547
abbreviation
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   548
  real_of_int :: "int \<Rightarrow> real"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   549
where
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   550
  "real_of_int \<equiv> of_int"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   551
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   552
abbreviation
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   553
  real_of_rat :: "rat \<Rightarrow> real"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   554
where
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   555
  "real_of_rat \<equiv> of_rat"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   556
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   557
consts
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   558
  (*overloaded constant for injecting other types into "real"*)
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   559
  real :: "'a => real"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   560
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   561
defs (overloaded)
28520
376b9c083b04 tuned code setup
haftmann
parents: 28351
diff changeset
   562
  real_of_nat_def [code unfold]: "real == real_of_nat"
376b9c083b04 tuned code setup
haftmann
parents: 28351
diff changeset
   563
  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
   564
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   565
lemma real_eq_of_nat: "real = of_nat"
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   566
  unfolding real_of_nat_def ..
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   567
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   568
lemma real_eq_of_int: "real = of_int"
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   569
  unfolding real_of_int_def ..
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   570
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   571
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
   572
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   573
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   574
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
   575
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   576
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   577
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
   578
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   579
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   580
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
   581
by (simp add: real_of_int_def) 
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   582
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   583
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
   584
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   585
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   586
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
   587
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   588
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   589
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
   590
  apply (subst real_eq_of_int)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   591
  apply (rule of_int_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   592
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   593
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   594
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
   595
    (PROD x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   596
  apply (subst real_eq_of_int)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   597
  apply (rule of_int_setprod)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   598
done
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   599
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   600
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
   601
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   602
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   603
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
   604
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   605
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   606
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
   607
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   608
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   609
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
   610
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   611
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   612
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
   613
by (simp add: real_of_int_def) 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   614
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   615
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
   616
by (simp add: real_of_int_def) 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   617
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   618
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
   619
by (simp add: real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   620
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27652
diff changeset
   621
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
   622
by (simp add: real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   623
16888
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
   624
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
   625
by (auto simp add: abs_if)
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
   626
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   627
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
   628
  apply (subgoal_tac "real n + 1 = real (n + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   629
  apply (simp del: real_of_int_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   630
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   631
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   632
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   633
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
   634
  apply (subgoal_tac "real m + 1 = real (m + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   635
  apply (simp del: real_of_int_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   636
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   637
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   638
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   639
lemma real_of_int_div_aux: "d ~= 0 ==> (real (x::int)) / (real d) = 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   640
    real (x div d) + (real (x mod d)) / (real d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   641
proof -
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   642
  assume "d ~= 0"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   643
  have "x = (x div d) * d + x mod d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   644
    by auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   645
  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
   646
    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
   647
  then have "real x / real d = ... / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   648
    by simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   649
  then show ?thesis
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   650
    by (auto simp add: add_divide_distrib algebra_simps prems)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   651
qed
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   652
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   653
lemma real_of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   654
    real(n div d) = real n / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   655
  apply (frule real_of_int_div_aux [of d n])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   656
  apply simp
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29667
diff changeset
   657
  apply (simp add: dvd_eq_mod_eq_0)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   658
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   659
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   660
lemma real_of_int_div2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   661
  "0 <= real (n::int) / real (x) - real (n div x)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   662
  apply (case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   663
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   664
  apply (case_tac "0 < x")
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   665
  apply (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   666
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   667
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   668
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   669
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   670
  apply auto
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   671
  apply (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   672
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   673
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   674
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   675
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   676
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   677
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   678
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   679
lemma real_of_int_div3:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   680
  "real (n::int) / real (x) - real (n div x) <= 1"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   681
  apply(case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   682
  apply simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   683
  apply (simp add: algebra_simps)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   684
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   685
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   686
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   687
  apply (subst divide_le_eq)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   688
  apply clarsimp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   689
  apply (rule conjI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   690
  apply (rule impI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   691
  apply (rule order_less_imp_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   692
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   693
  apply (rule impI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   694
  apply (rule order_less_imp_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   695
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   696
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   697
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   698
lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" 
27964
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
   699
by (insert real_of_int_div2 [of n x], simp)
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
   700
1e0303048c0b added const Rational
nipkow
parents: 27833
diff changeset
   701
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   702
subsection{*Embedding the Naturals into the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   703
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   704
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
   705
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   706
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
   707
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
   708
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
   709
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   710
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
   711
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   712
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   713
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
   714
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   715
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   716
(*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
   717
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
   718
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   719
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   720
lemma real_of_nat_less_iff [iff]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   721
     "(real (n::nat) < real m) = (n < m)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   722
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   723
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   724
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
   725
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   726
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   727
lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   728
by (simp add: real_of_nat_def zero_le_imp_of_nat)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   729
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   730
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
   731
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
   732
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   733
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
   734
by (simp add: real_of_nat_def of_nat_mult)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   735
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   736
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
   737
    (SUM x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   738
  apply (subst real_eq_of_nat)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   739
  apply (rule of_nat_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   740
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   741
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   742
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
   743
    (PROD x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   744
  apply (subst real_eq_of_nat)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   745
  apply (rule of_nat_setprod)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   746
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   747
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   748
lemma real_of_card: "real (card A) = setsum (%x.1) A"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   749
  apply (subst card_eq_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   750
  apply (subst real_of_nat_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   751
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   752
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   753
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   754
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
   755
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   756
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   757
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
   758
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   759
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   760
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
   761
by (simp add: add: real_of_nat_def of_nat_diff)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   762
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25140
diff changeset
   763
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
   764
by (auto simp: real_of_nat_def)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   765
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   766
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
   767
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   768
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   769
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
   770
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   771
25140
273772abbea2 More changes from >0 to ~=0::nat
nipkow
parents: 25134
diff changeset
   772
lemma real_of_nat_ge_zero_cancel_iff [simp]: "(0 \<le> real (n::nat))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   773
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   774
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   775
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
   776
  apply (subgoal_tac "real n + 1 = real (Suc n)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   777
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   778
  apply (auto simp add: real_of_nat_Suc)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   779
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   780
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   781
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
   782
  apply (subgoal_tac "real m + 1 = real (Suc m)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   783
  apply (simp add: less_Suc_eq_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   784
  apply (simp add: real_of_nat_Suc)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   785
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   786
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   787
lemma real_of_nat_div_aux: "0 < d ==> (real (x::nat)) / (real d) = 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   788
    real (x div d) + (real (x mod d)) / (real d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   789
proof -
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   790
  assume "0 < d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   791
  have "x = (x div d) * d + x mod d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   792
    by auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   793
  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
   794
    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
   795
  then have "real x / real d = \<dots> / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   796
    by simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   797
  then show ?thesis
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   798
    by (auto simp add: add_divide_distrib algebra_simps prems)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   799
qed
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   800
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   801
lemma real_of_nat_div: "0 < (d::nat) ==> d dvd n ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   802
    real(n div d) = real n / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   803
  apply (frule real_of_nat_div_aux [of d n])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   804
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   805
  apply (subst dvd_eq_mod_eq_0 [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   806
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   807
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   808
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   809
lemma real_of_nat_div2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   810
  "0 <= real (n::nat) / real (x) - real (n div x)"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   811
apply(case_tac "x = 0")
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   812
 apply (simp)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   813
apply (simp add: algebra_simps)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   814
apply (subst real_of_nat_div_aux)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   815
 apply simp
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   816
apply simp
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   817
apply (subst zero_le_divide_iff)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   818
apply simp
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   819
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   820
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   821
lemma real_of_nat_div3:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   822
  "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
   823
apply(case_tac "x = 0")
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   824
apply (simp)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   825
apply (simp add: algebra_simps)
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   826
apply (subst real_of_nat_div_aux)
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   827
 apply simp
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   828
apply simp
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   829
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   830
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   831
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
   832
by (insert real_of_nat_div2 [of n x], simp)
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   833
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   834
lemma real_of_int_real_of_nat: "real (int n) = real n"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   835
by (simp add: real_of_nat_def real_of_int_def int_eq_of_nat)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   836
14426
de890c247b38 fixed bugs in the setup of arithmetic procedures
paulson
parents: 14421
diff changeset
   837
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
   838
by (simp add: real_of_int_def real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   839
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   840
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
   841
  apply (subgoal_tac "real(int(nat x)) = real(nat x)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   842
  apply force
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   843
  apply (simp only: real_of_int_real_of_nat)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   844
done
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   845
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   846
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   847
subsection{* Rationals *}
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   848
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   849
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
   850
by (simp add: real_eq_of_nat)
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   851
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   852
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   853
lemma Rats_eq_int_div_int:
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   854
  "\<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
   855
proof
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   856
  show "\<rat> \<subseteq> ?S"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   857
  proof
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   858
    fix x::real assume "x : \<rat>"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   859
    then obtain r where "x = of_rat r" unfolding Rats_def ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   860
    have "of_rat r : ?S"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   861
      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
   862
    thus "x : ?S" using `x = of_rat r` by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   863
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   864
next
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   865
  show "?S \<subseteq> \<rat>"
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   866
  proof(auto simp:Rats_def)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   867
    fix i j :: int assume "j \<noteq> 0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   868
    hence "real i / real j = of_rat(Fract i j)"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   869
      by (simp add:of_rat_rat real_eq_of_int)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   870
    thus "real i / real j \<in> range of_rat" by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   871
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   872
qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   873
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   874
lemma Rats_eq_int_div_nat:
28091
50f2d6ba024c Streamlined parts of Complex/ex/DenumRat and AFP/Integration/Rats and
nipkow
parents: 28001
diff changeset
   875
  "\<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
   876
proof(auto simp:Rats_eq_int_div_int)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   877
  fix i j::int assume "j \<noteq> 0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   878
  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
   879
  proof cases
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   880
    assume "j>0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   881
    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
   882
      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
   883
    thus ?thesis by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   884
  next
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   885
    assume "~ j>0"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   886
    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
   887
      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
   888
    thus ?thesis by blast
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   889
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   890
next
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   891
  fix i::int and n::nat assume "0 < n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   892
  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
   893
  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
   894
qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   895
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   896
lemma Rats_abs_nat_div_natE:
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   897
  assumes "x \<in> \<rat>"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   898
  obtains m n where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   899
proof -
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   900
  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
   901
    by(auto simp add: Rats_eq_int_div_nat)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   902
  hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   903
  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
   904
  let ?gcd = "gcd m n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   905
  from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by (simp add: gcd_zero)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   906
  let ?k = "m div ?gcd"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   907
  let ?l = "n div ?gcd"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   908
  let ?gcd' = "gcd ?k ?l"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   909
  have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   910
    by (rule dvd_mult_div_cancel)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   911
  have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   912
    by (rule dvd_mult_div_cancel)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   913
  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
   914
  moreover
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   915
  have "\<bar>x\<bar> = real ?k / real ?l"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   916
  proof -
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   917
    from gcd have "real ?k / real ?l =
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   918
        real (?gcd * ?k) / real (?gcd * ?l)" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   919
    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
   920
    also from x_rat have "\<dots> = \<bar>x\<bar>" ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   921
    finally show ?thesis ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   922
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   923
  moreover
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   924
  have "?gcd' = 1"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   925
  proof -
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   926
    have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)"
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   927
      by (rule gcd_mult_distrib2)
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   928
    with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   929
    with gcd show ?thesis by simp
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   930
  qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   931
  ultimately show ?thesis ..
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   932
qed
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   933
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27964
diff changeset
   934
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   935
subsection{*Numerals and Arithmetic*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   936
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   937
instantiation real :: number_ring
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   938
begin
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   939
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   940
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28520
diff changeset
   941
  real_number_of_def [code del]: "number_of w = real_of_int w"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   942
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   943
instance
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   944
  by intro_classes (simp add: real_number_of_def)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   945
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   946
end
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25546
diff changeset
   947
25965
05df64f786a4 improved code theorem setup
haftmann
parents: 25885
diff changeset
   948
lemma [code unfold, symmetric, code post]:
24198
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   949
  "number_of k = real_of_int (number_of k)"
4031da6d8ba3 adaptions for code generation
haftmann
parents: 24075
diff changeset
   950
  unfolding number_of_is_id real_number_of_def ..
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   951
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   952
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   953
text{*Collapse applications of @{term real} to @{term number_of}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   954
lemma real_number_of [simp]: "real (number_of v :: int) = number_of v"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   955
by (simp add:  real_of_int_def of_int_number_of_eq)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   956
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   957
lemma real_of_nat_number_of [simp]:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   958
     "real (number_of v :: nat) =  
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   959
        (if neg (number_of v :: int) then 0  
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   960
         else (number_of v :: real))"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   961
by (simp add: real_of_int_real_of_nat [symmetric] int_nat_number_of)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   962
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
   963
declaration {*
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
   964
  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
   965
    (* 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
   966
  #> 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
   967
    (* 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
   968
  #> 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
   969
      @{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
   970
      @{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
   971
      @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
   972
      @{thm real_of_nat_number_of}, @{thm real_number_of}]
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
   973
  #> Lin_Arith.add_inj_const (@{const_name real}, HOLogic.natT --> HOLogic.realT)
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
   974
  #> Lin_Arith.add_inj_const (@{const_name real}, HOLogic.intT --> HOLogic.realT))
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 30242
diff changeset
   975
*}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   976
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   977
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   978
subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   979
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   980
text{*Needed in this non-standard form by Hyperreal/Transcendental*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   981
lemma real_0_le_divide_iff:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   982
     "((0::real) \<le> x/y) = ((x \<le> 0 | 0 \<le> y) & (0 \<le> x | y \<le> 0))"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   983
by (simp add: real_divide_def zero_le_mult_iff, auto)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   984
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   985
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
   986
by arith
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   987
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   988
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
   989
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   990
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   991
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
   992
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   993
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   994
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
   995
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   996
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   997
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
   998
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   999
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
  1000
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
  1001
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1002
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1003
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1004
(*
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1005
FIXME: we should have this, as for type int, but many proofs would break.
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1006
It replaces x+-y by x-y.
15086
e6a2a98d5ef5 removal of more iff-rules from RealDef.thy
paulson
parents: 15085
diff changeset
  1007
declare real_diff_def [symmetric, simp]
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1008
*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1009
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1010
subsubsection{*Density of the Reals*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1011
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1012
lemma real_lbound_gt_zero:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1013
     "[| (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
  1014
apply (rule_tac x = " (min d1 d2) /2" in exI)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1015
apply (simp add: min_def)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1016
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1017
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1018
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1019
text{*Similar results are proved in @{text Ring_and_Field}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1020
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
  1021
  by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1022
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1023
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
  1024
  by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1025
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1026
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1027
subsection{*Absolute Value Function for the Reals*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1028
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1029
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
  1030
by (simp add: abs_if)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1031
23289
0cf371d943b1 remove redundant lemmas
huffman
parents: 23288
diff changeset
  1032
(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1033
lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
  1034
by (force simp add: OrderedGroup.abs_le_iff)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1035
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1036
lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)"
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14754
diff changeset
  1037
by (simp add: abs_if)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1038
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1039
lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
22958
b3a5569a81e5 cleaned up
huffman
parents: 22456
diff changeset
  1040
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
  1041
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1042
lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
  1043
by simp
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1044
 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1045
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
  1046
by simp
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1047
26732
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1048
instance real :: lordered_ring
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1049
proof
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1050
  fix a::real
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1051
  show "abs a = sup a (-a)"
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1052
    by (auto simp add: real_abs_def sup_real_def)
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1053
qed
6ea9de67e576 constant HOL.eq now qualified
haftmann
parents: 26513
diff changeset
  1054
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1055
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1056
subsection {* Implementation of rational real numbers *}
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1057
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1058
definition Ratreal :: "rat \<Rightarrow> real" where
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1059
  [simp]: "Ratreal = of_rat"
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1060
24623
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1061
code_datatype Ratreal
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1062
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1063
lemma Ratreal_number_collapse [code post]:
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1064
  "Ratreal 0 = 0"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1065
  "Ratreal 1 = 1"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1066
  "Ratreal (number_of k) = number_of k"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1067
by simp_all
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1068
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1069
lemma zero_real_code [code, code unfold]:
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1070
  "0 = Ratreal 0"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1071
by simp
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1072
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1073
lemma one_real_code [code, code unfold]:
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1074
  "1 = Ratreal 1"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1075
by simp
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1076
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1077
lemma number_of_real_code [code unfold]:
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1078
  "number_of k = Ratreal (number_of k)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1079
by simp
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1080
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1081
lemma Ratreal_number_of_quotient [code post]:
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1082
  "Ratreal (number_of r) / Ratreal (number_of s) = number_of r / number_of s"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1083
by simp
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1084
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1085
lemma Ratreal_number_of_quotient2 [code post]:
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1086
  "Ratreal (number_of r / number_of s) = number_of r / number_of s"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1087
unfolding Ratreal_number_of_quotient [symmetric] Ratreal_def of_rat_divide ..
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1088
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1089
instantiation real :: eq
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1090
begin
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1091
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1092
definition "eq_class.eq (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1093
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1094
instance by default (simp add: eq_real_def)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1095
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1096
lemma real_eq_code [code]: "eq_class.eq (Ratreal x) (Ratreal y) \<longleftrightarrow> eq_class.eq x y"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1097
  by (simp add: eq_real_def eq)
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1098
28351
abfc66969d1f non left-linear equations for nbe
haftmann
parents: 28091
diff changeset
  1099
lemma real_eq_refl [code nbe]:
abfc66969d1f non left-linear equations for nbe
haftmann
parents: 28091
diff changeset
  1100
  "eq_class.eq (x::real) x \<longleftrightarrow> True"
abfc66969d1f non left-linear equations for nbe
haftmann
parents: 28091
diff changeset
  1101
  by (rule HOL.eq_refl)
abfc66969d1f non left-linear equations for nbe
haftmann
parents: 28091
diff changeset
  1102
26513
6f306c8c2c54 explicit class "eq" for operational equality
haftmann
parents: 25965
diff changeset
  1103
end
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1104
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1105
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
  1106
  by (simp add: of_rat_less_eq)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1107
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1108
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
  1109
  by (simp add: of_rat_less)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1110
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1111
lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1112
  by (simp add: of_rat_add)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1113
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1114
lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1115
  by (simp add: of_rat_mult)
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1116
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1117
lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1118
  by (simp add: of_rat_minus)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1119
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1120
lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1121
  by (simp add: of_rat_diff)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1122
27544
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1123
lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1124
  by (simp add: of_rat_inverse)
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1125
 
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1126
lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)"
5b293a8cf476 improved code generator setup
haftmann
parents: 27106
diff changeset
  1127
  by (simp add: of_rat_divide)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1128
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1129
definition (in term_syntax)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1130
  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Eval.term)" where
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1131
  [code inline]: "valterm_ratreal k = Code_Eval.valtermify Ratreal {\<cdot>} k"
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1132
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1133
notation fcomp (infixl "o>" 60)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1134
notation scomp (infixl "o\<rightarrow>" 60)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1135
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1136
instantiation real :: random
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1137
begin
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1138
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1139
definition
31641
feea4d3d743d hide constant Quickcheck.random
haftmann
parents: 31203
diff changeset
  1140
  "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
31203
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1141
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1142
instance ..
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1143
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1144
end
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1145
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1146
no_notation fcomp (infixl "o>" 60)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1147
no_notation scomp (infixl "o\<rightarrow>" 60)
5c8fb4fd67e0 moved Code_Index, Random and Quickcheck before Main
haftmann
parents: 31100
diff changeset
  1148
24623
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1149
text {* Setup for SML code generator *}
23031
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1150
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1151
types_code
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1152
  real ("(int */ int)")
23031
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1153
attach (term_of) {*
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1154
fun term_of_real (p, q) =
24623
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1155
  let
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1156
    val rT = HOLogic.realT
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1157
  in
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1158
    if q = 1 orelse p = 0 then HOLogic.mk_number rT p
24623
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1159
    else @{term "op / \<Colon> real \<Rightarrow> real \<Rightarrow> real"} $
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1160
      HOLogic.mk_number rT p $ HOLogic.mk_number rT q
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1161
  end;
23031
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1162
*}
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1163
attach (test) {*
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1164
fun gen_real i =
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1165
  let
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1166
    val p = random_range 0 i;
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1167
    val q = random_range 1 (i + 1);
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1168
    val g = Integer.gcd p q;
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24623
diff changeset
  1169
    val p' = p div g;
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24623
diff changeset
  1170
    val q' = q div g;
25885
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25762
diff changeset
  1171
    val r = (if one_of [true, false] then p' else ~ p',
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25762
diff changeset
  1172
      if p' = 0 then 0 else q')
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1173
  in
25885
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25762
diff changeset
  1174
    (r, fn () => term_of_real r)
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1175
  end;
23031
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1176
*}
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1177
9da9585c816e added code generation based on Isabelle's rat type.
nipkow
parents: 22970
diff changeset
  1178
consts_code
24623
7b2bc73405b8 renamed constructor RealC to Ratreal
haftmann
parents: 24534
diff changeset
  1179
  Ratreal ("(_)")
24534
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1180
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1181
consts_code
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1182
  "of_int :: int \<Rightarrow> real" ("\<module>real'_of'_int")
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1183
attach {*
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1184
fun real_of_int 0 = (0, 0)
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1185
  | real_of_int i = (i, 1);
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1186
*}
09b9a47904b7 New code generator setup (taken from Library/Executable_Real.thy,
berghofe
parents: 24506
diff changeset
  1187
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
  1188
end