src/HOL/Real/RealDef.thy
author wenzelm
Tue, 16 May 2006 13:01:22 +0200
changeset 19640 40ec89317425
parent 19023 5652a536b7e8
child 19765 dfe940911617
permissions -rw-r--r--
added Ferrante and Rackoff Algorithm -- by Amine Chaieb;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     1
(*  Title       : Real/RealDef.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7127
diff changeset
     2
    ID          : $Id$
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
     5
    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
     6
    Additional contributions by Jeremy Avigad
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     7
*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     8
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
     9
header{*Defining the Reals from the Positive Reals*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    10
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15086
diff changeset
    11
theory RealDef
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    12
imports PReal
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15923
diff changeset
    13
uses ("real_arith.ML")
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15086
diff changeset
    14
begin
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    15
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    16
constdefs
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    17
  realrel   ::  "((preal * preal) * (preal * preal)) set"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    18
  "realrel == {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    19
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    20
typedef (Real)  real = "UNIV//realrel"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    21
  by (auto simp add: quotient_def)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    22
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
    23
instance real :: "{ord, zero, one, plus, times, minus, inverse}" ..
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    24
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    25
constdefs
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    26
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    27
  (** these don't use the overloaded "real" function: users don't see them **)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    28
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    29
  real_of_preal :: "preal => real"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    30
  "real_of_preal m     ==
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    31
           Abs_Real(realrel``{(m + preal_of_rat 1, preal_of_rat 1)})"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    32
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    33
consts
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
    34
   (*Overloaded constant denoting the Real subset of enclosing
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    35
     types such as hypreal and complex*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    36
   Reals :: "'a set"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    37
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    38
   (*overloaded constant for injecting other types into "real"*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    39
   real :: "'a => real"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    40
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
    41
syntax (xsymbols)
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
    42
  Reals     :: "'a set"                   ("\<real>")
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
    43
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    44
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    45
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    46
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    47
  real_zero_def:
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    48
  "0 == Abs_Real(realrel``{(preal_of_rat 1, preal_of_rat 1)})"
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    49
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    50
  real_one_def:
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    51
  "1 == Abs_Real(realrel``
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    52
               {(preal_of_rat 1 + preal_of_rat 1,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    53
		 preal_of_rat 1)})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    54
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    55
  real_minus_def:
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    56
  "- r ==  contents (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    57
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    58
  real_add_def:
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    59
   "z + w ==
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    60
       contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    61
		 { Abs_Real(realrel``{(x+u, y+v)}) })"
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    62
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    63
  real_diff_def:
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    64
   "r - (s::real) == r + - s"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    65
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    66
  real_mult_def:
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    67
    "z * w ==
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    68
       contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    69
		 { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    70
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    71
  real_inverse_def:
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11701
diff changeset
    72
  "inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    73
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    74
  real_divide_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    75
  "R / (S::real) == R * inverse S"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    76
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    77
  real_le_def:
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    78
   "z \<le> (w::real) == 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
    79
    \<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    80
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    81
  real_less_def: "(x < (y::real)) == (x \<le> y & x \<noteq> y)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    82
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    83
  real_abs_def:  "abs (r::real) == (if 0 \<le> r then r else -r)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    84
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    85
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    86
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    87
subsection{*Proving that realrel is an equivalence relation*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    88
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    89
lemma preal_trans_lemma:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    90
  assumes "x + y1 = x1 + y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    91
      and "x + y2 = x2 + y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    92
  shows "x1 + y2 = x2 + (y1::preal)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    93
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    94
  have "(x1 + y2) + x = (x + y2) + x1" by (simp add: preal_add_ac) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    95
  also have "... = (x2 + y) + x1"  by (simp add: prems)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    96
  also have "... = x2 + (x1 + y)"  by (simp add: preal_add_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    97
  also have "... = x2 + (x + y1)"  by (simp add: prems)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    98
  also have "... = (x2 + y1) + x"  by (simp add: preal_add_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    99
  finally have "(x1 + y2) + x = (x2 + y1) + x" .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   100
  thus ?thesis by (simp add: preal_add_right_cancel_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   101
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   102
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   103
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   104
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
   105
by (simp add: realrel_def)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   106
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   107
lemma equiv_realrel: "equiv UNIV realrel"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   108
apply (auto simp add: equiv_def refl_def sym_def trans_def realrel_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   109
apply (blast dest: preal_trans_lemma) 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   110
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   111
14497
paulson
parents: 14484
diff changeset
   112
text{*Reduces equality of equivalence classes to the @{term realrel} relation:
paulson
parents: 14484
diff changeset
   113
  @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   114
lemmas equiv_realrel_iff = 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   115
       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
   116
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   117
declare equiv_realrel_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   118
14497
paulson
parents: 14484
diff changeset
   119
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   120
lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   121
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
   122
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   123
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   124
lemma inj_on_Abs_Real: "inj_on Abs_Real Real"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   125
apply (rule inj_on_inverseI)
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   126
apply (erule Abs_Real_inverse)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   127
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   128
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   129
declare inj_on_Abs_Real [THEN inj_on_iff, simp]
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   130
declare Abs_Real_inverse [simp]
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   131
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   132
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   133
text{*Case analysis on the representation of a real number as an equivalence
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   134
      class of pairs of positive reals.*}
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   135
lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   136
     "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   137
apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE])
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   138
apply (drule arg_cong [where f=Abs_Real])
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   139
apply (auto simp add: Rep_Real_inverse)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   140
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   141
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   142
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   143
subsection{*Congruence property for addition*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   144
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   145
lemma real_add_congruent2_lemma:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   146
     "[|a + ba = aa + b; ab + bc = ac + bb|]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   147
      ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   148
apply (simp add: preal_add_assoc) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   149
apply (rule preal_add_left_commute [of ab, THEN ssubst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   150
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   151
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   152
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   153
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   154
lemma real_add:
14497
paulson
parents: 14484
diff changeset
   155
     "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
paulson
parents: 14484
diff changeset
   156
      Abs_Real (realrel``{(x+u, y+v)})"
paulson
parents: 14484
diff changeset
   157
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   158
  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
   159
        respects2 realrel"
14497
paulson
parents: 14484
diff changeset
   160
    by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma) 
paulson
parents: 14484
diff changeset
   161
  thus ?thesis
paulson
parents: 14484
diff changeset
   162
    by (simp add: real_add_def UN_UN_split_split_eq
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14497
diff changeset
   163
                  UN_equiv_class2 [OF equiv_realrel equiv_realrel])
14497
paulson
parents: 14484
diff changeset
   164
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   165
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   166
lemma real_add_commute: "(z::real) + w = w + z"
14497
paulson
parents: 14484
diff changeset
   167
by (cases z, cases w, simp add: real_add preal_add_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   168
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   169
lemma real_add_assoc: "((z1::real) + z2) + z3 = z1 + (z2 + z3)"
14497
paulson
parents: 14484
diff changeset
   170
by (cases z1, cases z2, cases z3, simp add: real_add preal_add_assoc)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   171
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   172
lemma real_add_zero_left: "(0::real) + z = z"
14497
paulson
parents: 14484
diff changeset
   173
by (cases z, simp add: real_add real_zero_def preal_add_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   174
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   175
instance real :: comm_monoid_add
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   176
  by (intro_classes,
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   177
      (assumption | 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   178
       rule real_add_commute real_add_assoc real_add_zero_left)+)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   179
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   180
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   181
subsection{*Additive Inverse on real*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   182
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   183
lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   184
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   185
  have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   186
    by (simp add: congruent_def preal_add_commute) 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   187
  thus ?thesis
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   188
    by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   189
qed
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   190
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   191
lemma real_add_minus_left: "(-z) + z = (0::real)"
14497
paulson
parents: 14484
diff changeset
   192
by (cases z, simp add: real_minus real_add real_zero_def preal_add_commute)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   193
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   194
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   195
subsection{*Congruence property for multiplication*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   196
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   197
lemma real_mult_congruent2_lemma:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   198
     "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   199
          x * x1 + y * y1 + (x * y2 + y * x2) =
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   200
          x * x2 + y * y2 + (x * y1 + y * x1)"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   201
apply (simp add: preal_add_left_commute preal_add_assoc [symmetric])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   202
apply (simp add: preal_add_assoc preal_add_mult_distrib2 [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   203
apply (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   204
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   205
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   206
lemma real_mult_congruent2:
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   207
    "(%p1 p2.
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   208
        (%(x1,y1). (%(x2,y2). 
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   209
          { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   210
     respects2 realrel"
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14497
diff changeset
   211
apply (rule congruent2_commuteI [OF equiv_realrel], clarify)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   212
apply (simp add: preal_mult_commute preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   213
apply (auto simp add: real_mult_congruent2_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   214
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   215
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   216
lemma real_mult:
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   217
      "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   218
       Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   219
by (simp add: real_mult_def UN_UN_split_split_eq
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14497
diff changeset
   220
         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
   221
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   222
lemma real_mult_commute: "(z::real) * w = w * z"
14497
paulson
parents: 14484
diff changeset
   223
by (cases z, cases w, simp add: real_mult preal_add_ac preal_mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   224
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   225
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   226
apply (cases z1, cases z2, cases z3)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   227
apply (simp add: real_mult preal_add_mult_distrib2 preal_add_ac preal_mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   228
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   229
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   230
lemma real_mult_1: "(1::real) * z = z"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   231
apply (cases z)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   232
apply (simp add: real_mult real_one_def preal_add_mult_distrib2
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   233
                 preal_mult_1_right preal_mult_ac preal_add_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   234
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   235
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   236
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
   237
apply (cases z1, cases z2, cases w)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   238
apply (simp add: real_add real_mult preal_add_mult_distrib2 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   239
                 preal_add_ac preal_mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   240
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   241
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   242
text{*one and zero are distinct*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   243
lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   244
proof -
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   245
  have "preal_of_rat 1 < preal_of_rat 1 + preal_of_rat 1"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   246
    by (simp add: preal_self_less_add_left) 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   247
  thus ?thesis
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   248
    by (simp add: real_zero_def real_one_def preal_add_right_cancel_iff)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   249
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   250
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   251
subsection{*existence of inverse*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   252
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   253
lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
14497
paulson
parents: 14484
diff changeset
   254
by (simp add: real_zero_def preal_add_commute)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   255
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   256
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
   257
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
   258
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   259
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
   260
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
   261
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
   262
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
   263
apply (rule_tac
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   264
        x = "Abs_Real (realrel `` { (preal_of_rat 1, 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   265
                            inverse (D) + preal_of_rat 1)}) " 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   266
       in exI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   267
apply (rule_tac [2]
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   268
        x = "Abs_Real (realrel `` { (inverse (D) + preal_of_rat 1,
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   269
                   preal_of_rat 1)})" 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   270
       in exI)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   271
apply (auto simp add: real_mult preal_mult_1_right
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   272
              preal_add_mult_distrib2 preal_add_mult_distrib preal_mult_1
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   273
              preal_mult_inverse_right preal_add_ac preal_mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   274
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   275
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   276
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
   277
apply (simp add: real_inverse_def)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   278
apply (frule real_mult_inverse_left_ex, safe)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   279
apply (rule someI2, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   280
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   281
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   282
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   283
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
   284
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   285
instance real :: field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   286
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   287
  fix x y z :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   288
  show "- x + x = 0" by (rule real_add_minus_left)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   289
  show "x - y = x + (-y)" by (simp add: real_diff_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   290
  show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   291
  show "x * y = y * x" by (rule real_mult_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   292
  show "1 * x = x" by (rule real_mult_1)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   293
  show "(x + y) * z = x * z + y * z" by (simp add: real_add_mult_distrib)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   294
  show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   295
  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
   296
  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
   297
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   298
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   299
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   300
text{*Inverse of zero!  Useful to simplify certain equations*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   301
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   302
lemma INVERSE_ZERO: "inverse 0 = (0::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   303
by (simp add: real_inverse_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   304
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   305
instance real :: division_by_zero
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   306
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   307
  show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   308
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   309
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   310
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   311
(*Pull negations out*)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   312
declare minus_mult_right [symmetric, simp] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   313
        minus_mult_left [symmetric, simp]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   314
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   315
lemma real_mult_1_right: "z * (1::real) = z"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   316
  by (rule OrderedGroup.mult_1_right)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   317
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   318
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   319
subsection{*The @{text "\<le>"} Ordering*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   320
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   321
lemma real_le_refl: "w \<le> (w::real)"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   322
by (cases w, force simp add: real_le_def)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   323
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   324
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
   325
  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
   326
  following two lemmas.*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   327
lemma preal_eq_le_imp_le:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   328
  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
   329
  shows "b \<le> (d::preal)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   330
proof -
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   331
  have "c+d \<le> a+d" by (simp add: prems preal_cancels)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   332
  hence "a+b \<le> a+d" by (simp add: prems)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   333
  thus "b \<le> d" by (simp add: preal_cancels)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   334
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   335
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   336
lemma real_le_lemma:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   337
  assumes l: "u1 + v2 \<le> u2 + v1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   338
      and "x1 + v1 = u1 + y1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   339
      and "x2 + v2 = u2 + y2"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   340
  shows "x1 + y2 \<le> x2 + (y1::preal)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   341
proof -
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   342
  have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   343
  hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: preal_add_ac)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   344
  also have "... \<le> (x2+y1) + (u2+v1)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   345
         by (simp add: prems preal_add_le_cancel_left)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   346
  finally show ?thesis by (simp add: preal_add_le_cancel_right)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   347
qed						 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   348
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   349
lemma real_le: 
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   350
     "(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) =  
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   351
      (x1 + y2 \<le> x2 + y1)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   352
apply (simp add: real_le_def) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   353
apply (auto intro: real_le_lemma)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   354
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   355
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   356
lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15229
diff changeset
   357
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
   358
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   359
lemma real_trans_lemma:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   360
  assumes "x + v \<le> u + y"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   361
      and "u + v' \<le> u' + v"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   362
      and "x2 + v2 = u2 + y2"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   363
  shows "x + v' \<le> u' + (y::preal)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   364
proof -
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   365
  have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: preal_add_ac)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   366
  also have "... \<le> (u+y) + (u+v')" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   367
    by (simp add: preal_add_le_cancel_right prems) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   368
  also have "... \<le> (u+y) + (u'+v)" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   369
    by (simp add: preal_add_le_cancel_left prems) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   370
  also have "... = (u'+y) + (u+v)"  by (simp add: preal_add_ac)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   371
  finally show ?thesis by (simp add: preal_add_le_cancel_right)
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15229
diff changeset
   372
qed
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   373
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   374
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
   375
apply (cases i, cases j, cases k)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   376
apply (simp add: real_le)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   377
apply (blast intro: real_trans_lemma) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   378
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   379
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   380
(* Axiom 'order_less_le' of class 'order': *)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   381
lemma real_less_le: "((w::real) < z) = (w \<le> z & w \<noteq> z)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   382
by (simp add: real_less_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   383
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   384
instance real :: order
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   385
proof qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   386
 (assumption |
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   387
  rule real_le_refl real_le_trans real_le_anti_sym real_less_le)+
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   388
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   389
(* Axiom 'linorder_linear' of class 'linorder': *)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   390
lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   391
apply (cases z, cases w) 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   392
apply (auto simp add: real_le real_zero_def preal_add_ac preal_cancels)
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
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   396
instance real :: linorder
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   397
  by (intro_classes, rule real_le_linear)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   398
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   399
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   400
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
   401
apply (cases x, cases y) 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   402
apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   403
                      preal_add_ac)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   404
apply (simp_all add: preal_add_assoc [symmetric] preal_cancels)
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15229
diff changeset
   405
done
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   406
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   407
lemma real_add_left_mono: 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   408
  assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   409
proof -
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   410
  have "z + x - (z + y) = (z + -z) + (x - y)"
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   411
    by (simp add: diff_minus add_ac) 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   412
  with le show ?thesis 
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
   413
    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
   414
qed
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   415
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   416
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
   417
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
   418
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   419
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
   420
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
   421
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   422
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   423
apply (cases x, cases y)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   424
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
   425
                 linorder_not_le [where 'a = preal] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   426
                  real_zero_def real_le real_mult)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   427
  --{*Reduce to the (simpler) @{text "\<le>"} relation *}
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16888
diff changeset
   428
apply (auto dest!: less_add_left_Ex
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   429
     simp add: preal_add_ac preal_mult_ac 
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16888
diff changeset
   430
          preal_add_mult_distrib2 preal_cancels preal_self_less_add_left)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   431
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   432
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   433
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
   434
apply (rule real_sum_gt_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   435
apply (drule real_less_sum_gt_zero [of x y])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   436
apply (drule real_mult_order, assumption)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   437
apply (simp add: right_distrib)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   438
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   439
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   440
text{*lemma for proving @{term "0<(1::real)"}*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   441
lemma real_zero_le_one: "0 \<le> (1::real)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   442
by (simp add: real_zero_def real_one_def real_le 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   443
                 preal_self_less_add_left order_less_imp_le)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   444
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   445
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   446
subsection{*The Reals Form an Ordered Field*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   447
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   448
instance real :: ordered_field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   449
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   450
  fix x y z :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   451
  show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   452
  show "x < y ==> 0 < z ==> z * x < z * y" by (simp add: real_mult_less_mono2)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   453
  show "\<bar>x\<bar> = (if x < 0 then -x else x)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   454
    by (auto dest: order_le_less_trans simp add: real_abs_def linorder_not_le)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   455
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   456
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   457
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   458
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   459
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
   460
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
   461
positive reals.*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   462
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   463
lemma real_of_preal_add:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   464
     "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   465
by (simp add: real_of_preal_def real_add preal_add_mult_distrib preal_mult_1 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   466
              preal_add_ac)
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_mult:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   469
     "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   470
by (simp add: real_of_preal_def real_mult preal_add_mult_distrib2
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   471
              preal_mult_1 preal_mult_1_right preal_add_ac preal_mult_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   472
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   473
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   474
text{*Gleason prop 9-4.4 p 127*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   475
lemma real_of_preal_trichotomy:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   476
      "\<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
   477
apply (simp add: real_of_preal_def real_zero_def, cases x)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   478
apply (auto simp add: real_minus preal_add_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   479
apply (cut_tac x = x and y = y in linorder_less_linear)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   480
apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc [symmetric])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   481
done
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_leD:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   484
      "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   485
by (simp add: real_of_preal_def real_le preal_cancels)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   486
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   487
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
   488
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
   489
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   490
lemma real_of_preal_lessD:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   491
      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   492
by (simp add: real_of_preal_def real_le linorder_not_le [symmetric] 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   493
              preal_cancels) 
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   494
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   495
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   496
lemma real_of_preal_less_iff [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   497
     "(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
   498
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
   499
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   500
lemma real_of_preal_le_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   501
     "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   502
by (simp add: linorder_not_less [symmetric]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   503
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   504
lemma real_of_preal_zero_less: "0 < real_of_preal m"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   505
apply (auto simp add: real_zero_def real_of_preal_def real_less_def real_le_def
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   506
            preal_add_ac preal_cancels)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   507
apply (simp_all add: preal_add_assoc [symmetric] preal_cancels)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   508
apply (blast intro: preal_self_less_add_left order_less_imp_le)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   509
apply (insert preal_not_eq_self [of "preal_of_rat 1" m]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   510
apply (simp add: preal_add_ac) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   511
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   512
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   513
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
   514
by (simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   515
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   516
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
   517
proof -
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   518
  from real_of_preal_minus_less_zero
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   519
  show ?thesis by (blast dest: order_less_trans)
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
   520
qed
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   521
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   522
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   523
subsection{*Theorems About the Ordering*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   524
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   525
text{*obsolete but used a lot*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   526
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   527
lemma real_not_refl2: "x < y ==> x \<noteq> (y::real)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   528
by blast 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   529
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   530
lemma real_le_imp_less_or_eq: "!!(x::real). x \<le> y ==> x < y | x = y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   531
by (simp add: order_le_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   532
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   533
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
   534
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
   535
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
   536
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
   537
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   538
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   539
lemma real_gt_preal_preal_Ex:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   540
     "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
   541
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
   542
             intro: real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   543
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   544
lemma real_ge_preal_preal_Ex:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   545
     "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
   546
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
   547
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   548
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
   549
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
   550
            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
   551
            simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   552
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   553
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
   554
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
   555
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   556
lemma real_add_less_le_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::real)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   557
  by (rule OrderedGroup.add_less_le_mono)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   558
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   559
lemma real_add_le_less_mono:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   560
     "!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   561
  by (rule OrderedGroup.add_le_less_mono)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   562
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   563
lemma real_le_square [simp]: "(0::real) \<le> x*x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   564
 by (rule Ring_and_Field.zero_le_square)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   565
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   566
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   567
subsection{*More Lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   568
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   569
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
   570
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   571
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   572
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
   573
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   574
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   575
text{*The precondition could be weakened to @{term "0\<le>x"}*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   576
lemma real_mult_less_mono:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   577
     "[| u<v;  x<y;  (0::real) < v;  0 < x |] ==> u*x < v* y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   578
 by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   579
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   580
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
   581
  by (force elim: order_less_asym
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   582
            simp add: Ring_and_Field.mult_less_cancel_right)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   583
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   584
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
   585
apply (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   586
apply (blast intro: elim: order_less_asym) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   587
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   588
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   589
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
   590
by(simp add:mult_commute)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   591
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   592
text{*Only two uses?*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   593
lemma real_mult_less_mono':
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   594
     "[| x < y;  r1 < r2;  (0::real) \<le> r1;  0 \<le> x|] ==> r1 * x < r2 * y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   595
 by (rule Ring_and_Field.mult_strict_mono')
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   596
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   597
text{*FIXME: delete or at least combine the next two lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   598
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   599
apply (drule OrderedGroup.equals_zero_I [THEN sym])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   600
apply (cut_tac x = y in real_le_square) 
14476
758e7acdea2f removed redundant thms
paulson
parents: 14443
diff changeset
   601
apply (auto, drule order_antisym, auto)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   602
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   603
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   604
lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   605
apply (rule_tac y = x in real_sum_squares_cancel)
14476
758e7acdea2f removed redundant thms
paulson
parents: 14443
diff changeset
   606
apply (simp add: add_commute)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   607
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   608
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   609
lemma real_add_order: "[| 0 < x; 0 < y |] ==> (0::real) < x + y"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   610
by (drule add_strict_mono [of concl: 0 0], assumption, simp)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   611
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   612
lemma real_le_add_order: "[| 0 \<le> x; 0 \<le> y |] ==> (0::real) \<le> x + y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   613
apply (drule order_le_imp_less_or_eq)+
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   614
apply (auto intro: real_add_order order_less_imp_le)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   615
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   616
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   617
lemma real_inverse_unique: "x*y = (1::real) ==> y = inverse x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   618
apply (case_tac "x \<noteq> 0")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   619
apply (rule_tac c1 = x in real_mult_left_cancel [THEN iffD1], auto)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   620
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   621
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   622
lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   623
by (auto dest: less_imp_inverse_less)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   624
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   625
lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   626
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   627
  have "0 + 0 \<le> x*x + y*y" by (blast intro: add_mono zero_le_square)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   628
  thus ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   629
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   630
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   631
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   632
subsection{*Embedding the Integers into the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   633
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   634
defs (overloaded)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   635
  real_of_nat_def: "real z == of_nat z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   636
  real_of_int_def: "real z == of_int z"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   637
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   638
lemma real_eq_of_nat: "real = of_nat"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   639
  apply (rule ext)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   640
  apply (unfold real_of_nat_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   641
  apply (rule refl)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   642
  done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   643
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   644
lemma real_eq_of_int: "real = of_int"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   645
  apply (rule ext)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   646
  apply (unfold real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   647
  apply (rule refl)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   648
  done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   649
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   650
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
   651
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   652
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   653
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
   654
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   655
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   656
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
   657
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   658
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   659
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
   660
by (simp add: real_of_int_def) 
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   661
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   662
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
   663
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   664
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   665
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
   666
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   667
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   668
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
   669
  apply (subst real_eq_of_int)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   670
  apply (rule of_int_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   671
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   672
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   673
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
   674
    (PROD x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   675
  apply (subst real_eq_of_int)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   676
  apply (rule of_int_setprod)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   677
done
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   678
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   679
lemma real_of_int_zero_cancel [simp]: "(real x = 0) = (x = (0::int))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   680
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   681
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   682
lemma real_of_int_inject [iff]: "(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
   683
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   684
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   685
lemma real_of_int_less_iff [iff]: "(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
   686
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   687
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   688
lemma real_of_int_le_iff [simp]: "(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
   689
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   690
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   691
lemma real_of_int_gt_zero_cancel_iff [simp]: "(0 < real (n::int)) = (0 < n)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   692
by (simp add: real_of_int_def) 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   693
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   694
lemma real_of_int_ge_zero_cancel_iff [simp]: "(0 <= real (n::int)) = (0 <= n)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   695
by (simp add: real_of_int_def) 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   696
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   697
lemma real_of_int_lt_zero_cancel_iff [simp]: "(real (n::int) < 0) = (n < 0)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   698
by (simp add: real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   699
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   700
lemma real_of_int_le_zero_cancel_iff [simp]: "(real (n::int) <= 0) = (n <= 0)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   701
by (simp add: real_of_int_def)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   702
16888
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
   703
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
   704
by (auto simp add: abs_if)
7cb4bcfa058e added list of theorem changes to NEWS
avigad
parents: 16819
diff changeset
   705
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   706
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
   707
  apply (subgoal_tac "real n + 1 = real (n + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   708
  apply (simp del: real_of_int_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   709
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   710
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   711
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   712
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
   713
  apply (subgoal_tac "real m + 1 = real (m + 1)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   714
  apply (simp del: real_of_int_add)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   715
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   716
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   717
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   718
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
   719
    real (x div d) + (real (x mod d)) / (real d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   720
proof -
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   721
  assume "d ~= 0"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   722
  have "x = (x div d) * d + x mod d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   723
    by auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   724
  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
   725
    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
   726
  then have "real x / real d = ... / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   727
    by simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   728
  then show ?thesis
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   729
    by (auto simp add: add_divide_distrib ring_eq_simps prems)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   730
qed
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   731
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   732
lemma real_of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   733
    real(n div d) = real n / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   734
  apply (frule real_of_int_div_aux [of d n])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   735
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   736
  apply (simp add: zdvd_iff_zmod_eq_0)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   737
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   738
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   739
lemma real_of_int_div2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   740
  "0 <= real (n::int) / real (x) - real (n div x)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   741
  apply (case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   742
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   743
  apply (case_tac "0 < x")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   744
  apply (simp add: compare_rls)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   745
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   746
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   747
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   748
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   749
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   750
  apply (simp add: compare_rls)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   751
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   752
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   753
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   754
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   755
  apply auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   756
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   757
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   758
lemma real_of_int_div3:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   759
  "real (n::int) / real (x) - real (n div x) <= 1"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   760
  apply(case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   761
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   762
  apply (simp add: compare_rls)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   763
  apply (subst real_of_int_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   764
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   765
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   766
  apply (subst divide_le_eq)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   767
  apply clarsimp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   768
  apply (rule conjI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   769
  apply (rule impI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   770
  apply (rule order_less_imp_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   771
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   772
  apply (rule impI)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   773
  apply (rule order_less_imp_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   774
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   775
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   776
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   777
lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   778
  by (insert real_of_int_div2 [of n x], simp)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   779
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   780
subsection{*Embedding the Naturals into the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   781
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   782
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
   783
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   784
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   785
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
   786
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   787
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   788
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
   789
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   790
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   791
(*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
   792
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
   793
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   794
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   795
lemma real_of_nat_less_iff [iff]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   796
     "(real (n::nat) < real m) = (n < m)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   797
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   798
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   799
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
   800
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   801
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   802
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
   803
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
   804
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   805
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
   806
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
   807
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   808
lemma real_of_nat_mult [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
   809
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   810
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   811
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
   812
    (SUM x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   813
  apply (subst real_eq_of_nat)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   814
  apply (rule of_nat_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   815
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   816
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   817
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
   818
    (PROD x:A. real(f x))"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   819
  apply (subst real_eq_of_nat)+
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   820
  apply (rule of_nat_setprod)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   821
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   822
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   823
lemma real_of_card: "real (card A) = setsum (%x.1) A"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   824
  apply (subst card_eq_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   825
  apply (subst real_of_nat_setsum)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   826
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   827
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   828
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   829
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
   830
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   831
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   832
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
   833
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   834
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   835
lemma real_of_nat_diff: "n \<le> m ==> 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
   836
by (simp add: add: real_of_nat_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   837
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   838
lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   839
by (simp add: add: real_of_nat_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   840
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   841
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
   842
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   843
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   844
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
   845
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   846
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   847
lemma real_of_nat_ge_zero_cancel_iff [simp]: "(0 \<le> real (n::nat)) = (0 \<le> n)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   848
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   849
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   850
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
   851
  apply (subgoal_tac "real n + 1 = real (Suc n)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   852
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   853
  apply (auto simp add: real_of_nat_Suc)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   854
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   855
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   856
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
   857
  apply (subgoal_tac "real m + 1 = real (Suc m)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   858
  apply (simp add: less_Suc_eq_le)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   859
  apply (simp add: real_of_nat_Suc)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   860
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   861
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   862
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
   863
    real (x div d) + (real (x mod d)) / (real d)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   864
proof -
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   865
  assume "0 < d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   866
  have "x = (x div d) * d + x mod d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   867
    by auto
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   868
  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
   869
    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
   870
  then have "real x / real d = \<dots> / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   871
    by simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   872
  then show ?thesis
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   873
    by (auto simp add: add_divide_distrib ring_eq_simps prems)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   874
qed
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   875
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   876
lemma real_of_nat_div: "0 < (d::nat) ==> d dvd n ==>
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   877
    real(n div d) = real n / real d"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   878
  apply (frule real_of_nat_div_aux [of d n])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   879
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   880
  apply (subst dvd_eq_mod_eq_0 [THEN sym])
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   881
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   882
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   883
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   884
lemma real_of_nat_div2:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   885
  "0 <= real (n::nat) / real (x) - real (n div x)"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   886
  apply(case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   887
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   888
  apply (simp add: compare_rls)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   889
  apply (subst real_of_nat_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   890
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   891
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   892
  apply (subst zero_le_divide_iff)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   893
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   894
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   895
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   896
lemma real_of_nat_div3:
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   897
  "real (n::nat) / real (x) - real (n div x) <= 1"
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   898
  apply(case_tac "x = 0")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   899
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   900
  apply (simp add: compare_rls)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   901
  apply (subst real_of_nat_div_aux)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   902
  apply assumption
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   903
  apply simp
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   904
done
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   905
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   906
lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" 
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   907
  by (insert real_of_nat_div2 [of n x], simp)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   908
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   909
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
   910
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
   911
14426
de890c247b38 fixed bugs in the setup of arithmetic procedures
paulson
parents: 14421
diff changeset
   912
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
   913
by (simp add: real_of_int_def real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   914
16819
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   915
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
   916
  apply (subgoal_tac "real(int(nat x)) = real(nat x)")
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   917
  apply force
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   918
  apply (simp only: real_of_int_real_of_nat)
00d8f9300d13 Additions to the Real (and Hyperreal) libraries:
avigad
parents: 16417
diff changeset
   919
done
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   920
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   921
subsection{*Numerals and Arithmetic*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   922
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   923
instance real :: number ..
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   924
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   925
defs (overloaded)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   926
  real_number_of_def: "(number_of w :: real) == of_int (Rep_Bin w)"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   927
    --{*the type constraint is essential!*}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   928
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   929
instance real :: number_ring
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   930
by (intro_classes, simp add: real_number_of_def) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   931
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   932
text{*Collapse applications of @{term real} to @{term number_of}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   933
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
   934
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
   935
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   936
lemma real_of_nat_number_of [simp]:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   937
     "real (number_of v :: nat) =  
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   938
        (if neg (number_of v :: int) then 0  
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   939
         else (number_of v :: real))"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   940
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
   941
 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   942
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   943
use "real_arith.ML"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   944
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   945
setup real_arith_setup
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   946
19023
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   947
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   948
lemma real_diff_mult_distrib:
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   949
  fixes a::real
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   950
  shows "a * (b - c) = a * b - a * c" 
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   951
proof -
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   952
  have "a * (b - c) = a * (b + -c)" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   953
  also have "\<dots> = (b + -c) * a" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   954
  also have "\<dots> = b*a + (-c)*a" by (rule real_add_mult_distrib)
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   955
  also have "\<dots> = a*b - a*c" by simp
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   956
  finally show ?thesis .
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   957
qed
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   958
5652a536b7e8 * include generalised MVT in HyperReal (contributed by Benjamin Porter)
kleing
parents: 16973
diff changeset
   959
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   960
subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   961
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   962
text{*Needed in this non-standard form by Hyperreal/Transcendental*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   963
lemma real_0_le_divide_iff:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   964
     "((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
   965
by (simp add: real_divide_def zero_le_mult_iff, auto)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   966
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   967
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
   968
by arith
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   969
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   970
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
   971
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   972
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   973
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
   974
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   975
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   976
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
   977
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   978
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   979
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
   980
by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   981
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15077
diff changeset
   982
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
   983
by 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
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   986
(*
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   987
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
   988
It replaces x+-y by x-y.
15086
e6a2a98d5ef5 removal of more iff-rules from RealDef.thy
paulson
parents: 15085
diff changeset
   989
declare real_diff_def [symmetric, simp]
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   990
*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   991
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   992
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   993
subsubsection{*Density of the Reals*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   994
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   995
lemma real_lbound_gt_zero:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   996
     "[| (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
   997
apply (rule_tac x = " (min d1 d2) /2" in exI)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   998
apply (simp add: min_def)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   999
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1000
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1001
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1002
text{*Similar results are proved in @{text Ring_and_Field}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1003
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
  1004
  by auto
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1005
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1006
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
  1007
  by auto
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
subsection{*Absolute Value Function for 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 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
  1013
by (simp add: abs_if)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1014
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1015
lemma abs_interval_iff: "(abs x < r) = (-r < x & x < (r::real))"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1016
by (force simp add: Ring_and_Field.abs_less_iff)
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
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
  1019
by (force simp add: OrderedGroup.abs_le_iff)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1020
14484
ef8c7c5eb01b new treatment of equivalence classes
paulson
parents: 14476
diff changeset
  1021
(*FIXME: used only once, in SEQ.ML*)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1022
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
  1023
by (simp add: abs_if)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1024
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1025
lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15169
diff changeset
  1026
by (simp add: real_of_nat_ge_zero)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1027
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1028
lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1029
apply (simp add: linorder_not_less)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1030
apply (auto intro: abs_ge_self [THEN order_trans])
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1031
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1032
 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1033
text{*Used only in Hyperreal/Lim.ML*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1034
lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1035
apply (simp add: real_add_assoc)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1036
apply (rule_tac a1 = y in add_left_commute [THEN ssubst])
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1037
apply (rule real_add_assoc [THEN subst])
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1038
apply (rule abs_triangle_ineq)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1039
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1040
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1041
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1042
ML
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1043
{*
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1044
val real_lbound_gt_zero = thm"real_lbound_gt_zero";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1045
val real_less_half_sum = thm"real_less_half_sum";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1046
val real_gt_half_sum = thm"real_gt_half_sum";
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1047
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1048
val abs_interval_iff = thm"abs_interval_iff";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1049
val abs_le_interval_iff = thm"abs_le_interval_iff";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1050
val abs_add_one_gt_zero = thm"abs_add_one_gt_zero";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1051
val abs_add_one_not_less_self = thm"abs_add_one_not_less_self";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1052
val abs_sum_triangle_ineq = thm"abs_sum_triangle_ineq";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1053
*}
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
  1054
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1055
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
  1056
end