src/HOL/Real/RealDef.thy
author paulson
Tue, 10 Feb 2004 12:02:11 +0100
changeset 14378 69c4d5997669
parent 14369 c50188fe6366
child 14387 e96d5c42c4b0
permissions -rw-r--r--
generic of_nat and of_int functions, and generalization of iszero and neg
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
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     5
    Description : The reals
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     6
*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     7
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     8
theory RealDef = PReal:
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     9
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    10
constdefs
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    11
  realrel   ::  "((preal * preal) * (preal * preal)) set"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    12
  "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
    13
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    14
typedef (REAL)  real = "UNIV//realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    15
  by (auto simp add: quotient_def)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    16
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    17
instance real :: ord ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    18
instance real :: zero ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    19
instance real :: one ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    20
instance real :: plus ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    21
instance real :: times ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    22
instance real :: minus ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    23
instance real :: inverse ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    24
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    25
consts
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
    26
   (*Overloaded constant denoting the Real subset of enclosing
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    27
     types such as hypreal and complex*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    28
   Reals :: "'a set"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    29
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    30
   (*overloaded constant for injecting other types into "real"*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    31
   real :: "'a => real"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    32
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    33
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    34
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    35
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    36
  real_zero_def:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    37
  "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
    38
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    39
  real_one_def:
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    40
  "1 == Abs_REAL(realrel``
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    41
               {(preal_of_rat 1 + preal_of_rat 1,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    42
		 preal_of_rat 1)})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    43
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    44
  real_minus_def:
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    45
  "- R ==  Abs_REAL(UN (x,y):Rep_REAL(R). realrel``{(y,x)})"
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    46
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    47
  real_diff_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    48
  "R - (S::real) == R + - S"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    49
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    50
  real_inverse_def:
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11701
diff changeset
    51
  "inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    52
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    53
  real_divide_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    54
  "R / (S::real) == R * inverse S"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    55
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    56
constdefs
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    57
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    58
  (** these don't use the overloaded "real" function: users don't see them **)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    59
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    60
  real_of_preal :: "preal => real"
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    61
  "real_of_preal m     ==
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    62
           Abs_REAL(realrel``{(m + preal_of_rat 1, preal_of_rat 1)})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    63
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    64
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    65
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    66
  real_add_def:
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    67
  "P+Q == Abs_REAL(\<Union>p1\<in>Rep_REAL(P). \<Union>p2\<in>Rep_REAL(Q).
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    68
                   (%(x1,y1). (%(x2,y2). realrel``{(x1+x2, y1+y2)}) p2) p1)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    69
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    70
  real_mult_def:
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    71
  "P*Q == Abs_REAL(\<Union>p1\<in>Rep_REAL(P). \<Union>p2\<in>Rep_REAL(Q).
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    72
                   (%(x1,y1). (%(x2,y2). realrel``{(x1*x2+y1*y2,x1*y2+x2*y1)})
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
    73
		   p2) p1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    74
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    75
  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
    76
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    77
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    78
  real_le_def:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    79
  "P \<le> (Q::real) == \<exists>x1 y1 x2 y2. x1 + y2 \<le> x2 + y1 &
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    80
                            (x1,y1) \<in> Rep_REAL(P) & (x2,y2) \<in> Rep_REAL(Q)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    81
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    82
  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
    83
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    84
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 12018
diff changeset
    85
syntax (xsymbols)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    86
  Reals     :: "'a set"                   ("\<real>")
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    87
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    88
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    89
subsection{*Proving that realrel is an equivalence relation*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    90
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    91
lemma preal_trans_lemma:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    92
  assumes "x + y1 = x1 + y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    93
      and "x + y2 = x2 + y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    94
  shows "x1 + y2 = x2 + (y1::preal)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    95
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
    96
  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
    97
  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
    98
  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
    99
  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
   100
  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
   101
  finally have "(x1 + y2) + x = (x2 + y1) + x" .
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   102
  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
   103
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   104
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   105
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   106
lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   107
by (unfold realrel_def, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   108
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   109
lemma realrel_refl: "(x,x): realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   110
apply (case_tac "x")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   111
apply (simp add: realrel_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   112
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   113
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   114
lemma equiv_realrel: "equiv UNIV realrel"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   115
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
   116
apply (blast dest: preal_trans_lemma) 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   117
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   118
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   119
(* (realrel `` {x} = realrel `` {y}) = ((x,y) : realrel) *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   120
lemmas equiv_realrel_iff = 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   121
       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
   122
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   123
declare equiv_realrel_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   124
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   125
lemma realrel_in_real [simp]: "realrel``{(x,y)}: REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   126
by (unfold REAL_def realrel_def quotient_def, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   127
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   128
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   129
lemma inj_on_Abs_REAL: "inj_on Abs_REAL REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   130
apply (rule inj_on_inverseI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   131
apply (erule Abs_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   132
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   133
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   134
declare inj_on_Abs_REAL [THEN inj_on_iff, simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   135
declare Abs_REAL_inverse [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   136
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   137
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   138
lemmas eq_realrelD = equiv_realrel [THEN [2] eq_equiv_class]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   139
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   140
lemma inj_Rep_REAL: "inj Rep_REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   141
apply (rule inj_on_inverseI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   142
apply (rule Rep_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   143
done
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
(** real_of_preal: the injection from preal to real **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   146
lemma inj_real_of_preal: "inj(real_of_preal)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   147
apply (rule inj_onI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   148
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   149
apply (drule inj_on_Abs_REAL [THEN inj_onD])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   150
apply (rule realrel_in_real)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   151
apply (drule eq_equiv_class)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   152
apply (rule equiv_realrel, blast)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   153
apply (simp add: realrel_def preal_add_right_cancel_iff)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   154
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   155
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   156
lemma eq_Abs_REAL: 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   157
    "(!!x y. z = Abs_REAL(realrel``{(x,y)}) ==> P) ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   158
apply (rule_tac x1 = z in Rep_REAL [unfolded REAL_def, THEN quotientE])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   159
apply (drule_tac f = Abs_REAL in arg_cong)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   160
apply (case_tac "x")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   161
apply (simp add: Rep_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   162
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   163
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   164
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   165
subsection{*Congruence property for addition*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   166
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   167
lemma real_add_congruent2_lemma:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   168
     "[|a + ba = aa + b; ab + bc = ac + bb|]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   169
      ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   170
apply (simp add: preal_add_assoc) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   171
apply (rule preal_add_left_commute [of ab, THEN ssubst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   172
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   173
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   174
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   175
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   176
lemma real_add:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   177
  "Abs_REAL(realrel``{(x1,y1)}) + Abs_REAL(realrel``{(x2,y2)}) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   178
   Abs_REAL(realrel``{(x1+x2, y1+y2)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   179
apply (simp add: real_add_def UN_UN_split_split_eq)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   180
apply (subst equiv_realrel [THEN UN_equiv_class2])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   181
apply (auto simp add: congruent2_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   182
apply (blast intro: real_add_congruent2_lemma) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   183
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   184
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   185
lemma real_add_commute: "(z::real) + w = w + z"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   186
apply (rule eq_Abs_REAL [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   187
apply (rule eq_Abs_REAL [of w])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   188
apply (simp add: preal_add_ac real_add)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   189
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   190
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   191
lemma real_add_assoc: "((z1::real) + z2) + z3 = z1 + (z2 + z3)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   192
apply (rule eq_Abs_REAL [of z1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   193
apply (rule eq_Abs_REAL [of z2])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   194
apply (rule eq_Abs_REAL [of z3])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   195
apply (simp add: real_add preal_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   196
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   197
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   198
lemma real_add_zero_left: "(0::real) + z = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   199
apply (unfold real_of_preal_def real_zero_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   200
apply (rule eq_Abs_REAL [of z])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   201
apply (simp add: real_add preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   202
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   203
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   204
lemma real_add_zero_right: "z + (0::real) = z"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   205
by (simp add: real_add_zero_left real_add_commute)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   206
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   207
instance real :: plus_ac0
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   208
  by (intro_classes,
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   209
      (assumption | 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   210
       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
   211
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   212
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   213
subsection{*Additive Inverse on real*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   214
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   215
lemma real_minus_congruent:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   216
  "congruent realrel (%p. (%(x,y). realrel``{(y,x)}) p)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   217
apply (unfold congruent_def, clarify)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   218
apply (simp add: preal_add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   219
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   220
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   221
lemma real_minus:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   222
      "- (Abs_REAL(realrel``{(x,y)})) = Abs_REAL(realrel `` {(y,x)})"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   223
apply (unfold real_minus_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   224
apply (rule_tac f = Abs_REAL in arg_cong)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   225
apply (simp add: realrel_in_real [THEN Abs_REAL_inverse] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   226
            UN_equiv_class [OF equiv_realrel real_minus_congruent])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   227
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   228
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   229
lemma real_add_minus_left: "(-z) + z = (0::real)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   230
apply (unfold real_zero_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   231
apply (rule eq_Abs_REAL [of z])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   232
apply (simp add: real_minus real_add preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   233
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   234
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   235
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   236
subsection{*Congruence property for multiplication*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   237
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   238
lemma real_mult_congruent2_lemma:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   239
     "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   240
          x * x1 + y * y1 + (x * y2 + x2 * y) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   241
          x * x2 + y * y2 + (x * y1 + x1 * y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   242
apply (simp add: preal_add_left_commute preal_add_assoc [symmetric] preal_add_mult_distrib2 [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   243
apply (rule preal_mult_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   244
apply (rule_tac y1 = x2 in preal_mult_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   245
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
   246
apply (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   247
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   248
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   249
lemma real_mult_congruent2:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   250
    "congruent2 realrel (%p1 p2.
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   251
        (%(x1,y1). (%(x2,y2). realrel``{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   252
apply (rule equiv_realrel [THEN congruent2_commuteI], clarify)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   253
apply (unfold split_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   254
apply (simp add: preal_mult_commute preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   255
apply (auto simp add: real_mult_congruent2_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   256
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   257
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   258
lemma real_mult:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   259
   "Abs_REAL((realrel``{(x1,y1)})) * Abs_REAL((realrel``{(x2,y2)})) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   260
    Abs_REAL(realrel `` {(x1*x2+y1*y2,x1*y2+x2*y1)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   261
apply (unfold real_mult_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   262
apply (simp add: equiv_realrel [THEN UN_equiv_class2] real_mult_congruent2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   263
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   264
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   265
lemma real_mult_commute: "(z::real) * w = w * z"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   266
apply (rule eq_Abs_REAL [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   267
apply (rule eq_Abs_REAL [of w])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   268
apply (simp add: real_mult preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   269
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   270
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   271
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   272
apply (rule eq_Abs_REAL [of z1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   273
apply (rule eq_Abs_REAL [of z2])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   274
apply (rule eq_Abs_REAL [of z3])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   275
apply (simp add: preal_add_mult_distrib2 real_mult preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   276
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   277
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   278
lemma real_mult_1: "(1::real) * z = z"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   279
apply (unfold real_one_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   280
apply (rule eq_Abs_REAL [of z])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   281
apply (simp add: real_mult preal_add_mult_distrib2 preal_mult_1_right
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   282
                 preal_mult_ac preal_add_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   283
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   284
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   285
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   286
apply (rule eq_Abs_REAL [of z1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   287
apply (rule eq_Abs_REAL [of z2])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   288
apply (rule eq_Abs_REAL [of w])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   289
apply (simp add: preal_add_mult_distrib2 real_add real_mult preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   290
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   291
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   292
text{*one and zero are distinct*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   293
lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   294
apply (subgoal_tac "preal_of_rat 1 < preal_of_rat 1 + preal_of_rat 1")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   295
 prefer 2 apply (simp add: preal_self_less_add_left) 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   296
apply (unfold real_zero_def real_one_def)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   297
apply (auto simp add: preal_add_right_cancel_iff)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   298
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   299
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   300
subsection{*existence of inverse*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   301
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   302
lemma real_zero_iff: "Abs_REAL (realrel `` {(x, x)}) = 0"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   303
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   304
apply (auto simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   305
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   306
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   307
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
   308
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
   309
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   310
lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   311
apply (unfold real_zero_def real_one_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   312
apply (rule eq_Abs_REAL [of x])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   313
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
   314
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
   315
apply (rule_tac
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   316
        x = "Abs_REAL (realrel `` { (preal_of_rat 1, 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   317
                            inverse (D) + preal_of_rat 1)}) " 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   318
       in exI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   319
apply (rule_tac [2]
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   320
        x = "Abs_REAL (realrel `` { (inverse (D) + preal_of_rat 1,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   321
                   preal_of_rat 1)})" 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   322
       in exI)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   323
apply (auto simp add: real_mult preal_mult_1_right
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   324
              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
   325
              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
   326
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   327
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   328
lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   329
apply (unfold real_inverse_def)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   330
apply (frule real_mult_inverse_left_ex, safe)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   331
apply (rule someI2, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   332
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   333
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   334
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   335
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
   336
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   337
instance real :: field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   338
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   339
  fix x y z :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   340
  show "(x + y) + z = x + (y + z)" by (rule real_add_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   341
  show "x + y = y + x" by (rule real_add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   342
  show "0 + x = x" by simp
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   343
  show "- x + x = 0" by (rule real_add_minus_left)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   344
  show "x - y = x + (-y)" by (simp add: real_diff_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   345
  show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   346
  show "x * y = y * x" by (rule real_mult_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   347
  show "1 * x = x" by (rule real_mult_1)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   348
  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
   349
  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
   350
  show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   351
  show "y \<noteq> 0 ==> x / y = x * inverse y" by (simp add: real_divide_def)
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   352
  assume eq: "z+x = z+y" 
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   353
    hence "(-z + z) + x = (-z + z) + y" by (simp only: eq real_add_assoc)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   354
    thus "x = y" by (simp add: real_add_minus_left)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   355
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   356
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   357
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   358
text{*Inverse of zero!  Useful to simplify certain equations*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   359
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   360
lemma INVERSE_ZERO: "inverse 0 = (0::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   361
apply (unfold real_inverse_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   362
apply (rule someI2)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   363
apply (auto simp add: zero_neq_one)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   364
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   365
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   366
lemma DIVISION_BY_ZERO: "a / (0::real) = 0"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   367
  by (simp add: real_divide_def INVERSE_ZERO)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   368
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   369
instance real :: division_by_zero
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   370
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   371
  fix x :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   372
  show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   373
  show "x/0 = 0" by (rule DIVISION_BY_ZERO) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   374
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   375
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   376
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   377
(*Pull negations out*)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   378
declare minus_mult_right [symmetric, simp] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   379
        minus_mult_left [symmetric, simp]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   380
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   381
lemma real_mult_1_right: "z * (1::real) = z"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   382
  by (rule Ring_and_Field.mult_1_right)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   383
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   384
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   385
subsection{*The @{text "\<le>"} Ordering*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   386
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   387
lemma real_le_refl: "w \<le> (w::real)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   388
apply (rule eq_Abs_REAL [of w])
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   389
apply (force simp add: real_le_def)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   390
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   391
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   392
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
   393
  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
   394
  following two lemmas.*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   395
lemma preal_eq_le_imp_le:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   396
  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
   397
  shows "b \<le> (d::preal)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   398
proof -
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   399
  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
   400
  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
   401
  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
   402
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   403
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   404
lemma real_le_lemma:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   405
  assumes l: "u1 + v2 \<le> u2 + v1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   406
      and "x1 + v1 = u1 + y1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   407
      and "x2 + v2 = u2 + y2"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   408
  shows "x1 + y2 \<le> x2 + (y1::preal)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   409
proof -
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   410
  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
   411
  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
   412
  also have "... \<le> (x2+y1) + (u2+v1)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   413
         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
   414
  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
   415
qed						 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   416
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   417
lemma real_le: 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   418
  "(Abs_REAL(realrel``{(x1,y1)}) \<le> Abs_REAL(realrel``{(x2,y2)})) =  
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   419
   (x1 + y2 \<le> x2 + y1)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   420
apply (simp add: real_le_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   421
apply (auto intro: real_le_lemma);
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   422
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   423
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   424
lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   425
apply (rule eq_Abs_REAL [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   426
apply (rule eq_Abs_REAL [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   427
apply (simp add: real_le order_antisym) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   428
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   429
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   430
lemma real_trans_lemma:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   431
  assumes "x + v \<le> u + y"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   432
      and "u + v' \<le> u' + v"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   433
      and "x2 + v2 = u2 + y2"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   434
  shows "x + v' \<le> u' + (y::preal)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   435
proof -
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   436
  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
   437
  also have "... \<le> (u+y) + (u+v')" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   438
    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
   439
  also have "... \<le> (u+y) + (u'+v)" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   440
    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
   441
  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
   442
  finally show ?thesis by (simp add: preal_add_le_cancel_right)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   443
qed						 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   444
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   445
lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   446
apply (rule eq_Abs_REAL [of i])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   447
apply (rule eq_Abs_REAL [of j])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   448
apply (rule eq_Abs_REAL [of k])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   449
apply (simp add: real_le) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   450
apply (blast intro: real_trans_lemma) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   451
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   452
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   453
(* Axiom 'order_less_le' of class 'order': *)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   454
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
   455
by (simp add: real_less_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   456
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   457
instance real :: order
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   458
proof qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   459
 (assumption |
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   460
  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
   461
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   462
(* Axiom 'linorder_linear' of class 'linorder': *)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   463
lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   464
apply (rule eq_Abs_REAL [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   465
apply (rule eq_Abs_REAL [of w]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   466
apply (auto simp add: real_le real_zero_def preal_add_ac preal_cancels)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   467
apply (cut_tac x="x+ya" and y="xa+y" in linorder_linear) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   468
apply (auto ); 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   469
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   470
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   471
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   472
instance real :: linorder
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   473
  by (intro_classes, rule real_le_linear)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   474
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   475
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   476
lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   477
apply (rule eq_Abs_REAL [of x])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   478
apply (rule eq_Abs_REAL [of y]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   479
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
   480
                      preal_add_ac)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   481
apply (simp_all add: preal_add_assoc [symmetric] preal_cancels)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   482
done 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   483
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   484
lemma real_add_left_mono: "x \<le> y ==> z + x \<le> z + (y::real)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   485
apply (auto simp add: real_le_eq_diff [of x] real_le_eq_diff [of "z+x"])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   486
apply (subgoal_tac "z + x - (z + y) = (z + -z) + (x - y)")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   487
 prefer 2 apply (simp add: diff_minus add_ac, simp) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   488
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   489
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   490
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
   491
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
   492
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   493
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
   494
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
   495
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   496
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   497
apply (rule eq_Abs_REAL [of x])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   498
apply (rule eq_Abs_REAL [of y])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   499
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
   500
                 linorder_not_le [where 'a = preal] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   501
                  real_zero_def real_le real_mult)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   502
  --{*Reduce to the (simpler) @{text "\<le>"} relation *}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   503
apply (auto  dest!: less_add_left_Ex 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   504
     simp add: preal_add_ac preal_mult_ac 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   505
          preal_add_mult_distrib2 preal_cancels preal_self_less_add_right)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   506
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   507
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   508
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
   509
apply (rule real_sum_gt_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   510
apply (drule real_less_sum_gt_zero [of x y])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   511
apply (drule real_mult_order, assumption)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   512
apply (simp add: right_distrib)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   513
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   514
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   515
text{*lemma for proving @{term "0<(1::real)"}*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   516
lemma real_zero_le_one: "0 \<le> (1::real)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   517
apply (simp add: real_zero_def real_one_def real_le 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   518
                 preal_self_less_add_left order_less_imp_le)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   519
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   520
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   521
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   522
subsection{*The Reals Form an Ordered Field*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   523
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   524
instance real :: ordered_field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   525
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   526
  fix x y z :: real
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   527
  show "0 < (1::real)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   528
    by (simp add: real_less_def real_zero_le_one real_zero_not_eq_one)  
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   529
  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
   530
  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
   531
  show "\<bar>x\<bar> = (if x < 0 then -x else x)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   532
    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
   533
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   534
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   535
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   536
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   537
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
   538
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
   539
positive reals.*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   540
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   541
lemma real_of_preal_add:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   542
     "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
   543
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
   544
              preal_add_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   545
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   546
lemma real_of_preal_mult:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   547
     "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
   548
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
   549
              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
   550
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   551
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   552
text{*Gleason prop 9-4.4 p 127*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   553
lemma real_of_preal_trichotomy:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   554
      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   555
apply (unfold real_of_preal_def real_zero_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   556
apply (rule eq_Abs_REAL [of x])
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   557
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
   558
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
   559
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
   560
apply (auto simp add: preal_add_commute)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   561
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   562
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   563
lemma real_of_preal_leD:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   564
      "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
   565
apply (unfold real_of_preal_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   566
apply (auto simp add: real_le_def preal_add_ac)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   567
apply (auto simp add: preal_add_assoc [symmetric] preal_add_right_cancel_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   568
apply (auto simp add: preal_add_ac preal_add_le_cancel_left)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   569
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   570
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   571
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
   572
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
   573
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   574
lemma real_of_preal_lessD:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   575
      "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
   576
apply (auto simp add: real_less_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   577
apply (drule real_of_preal_leD) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   578
apply (auto simp add: order_le_less) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   579
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   580
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   581
lemma real_of_preal_less_iff [simp]:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   582
     "(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
   583
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
   584
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   585
lemma real_of_preal_le_iff:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   586
     "(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
   587
by (simp add: linorder_not_less [symmetric]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   588
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   589
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
   590
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
   591
            preal_add_ac preal_cancels)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   592
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
   593
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
   594
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
   595
apply (simp add: preal_add_ac) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   596
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   597
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   598
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
   599
by (simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   600
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   601
lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   602
apply (cut_tac real_of_preal_minus_less_zero)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   603
apply (fast dest: order_less_trans)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   604
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   605
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   606
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   607
subsection{*Theorems About the Ordering*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   608
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   609
text{*obsolete but used a lot*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   610
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   611
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
   612
by blast 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   613
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   614
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
   615
by (simp add: order_le_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   616
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   617
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
   618
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
   619
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
   620
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
   621
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   622
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   623
lemma real_gt_preal_preal_Ex:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   624
     "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
   625
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
   626
             intro: real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   627
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   628
lemma real_ge_preal_preal_Ex:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   629
     "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
   630
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
   631
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   632
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
   633
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
   634
            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
   635
            simp add: real_of_preal_zero_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   636
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   637
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
   638
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
   639
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   640
lemma real_add_less_le_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::real)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   641
  by (rule Ring_and_Field.add_less_le_mono)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   642
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   643
lemma real_add_le_less_mono:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   644
     "!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   645
  by (rule Ring_and_Field.add_le_less_mono)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   646
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   647
lemma real_zero_less_one: "0 < (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   648
  by (rule Ring_and_Field.zero_less_one)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   649
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   650
lemma real_le_square [simp]: "(0::real) \<le> x*x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   651
 by (rule Ring_and_Field.zero_le_square)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   652
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   653
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   654
subsection{*More Lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   655
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   656
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
   657
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   658
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   659
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
   660
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   661
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   662
text{*The precondition could be weakened to @{term "0\<le>x"}*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   663
lemma real_mult_less_mono:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   664
     "[| 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
   665
 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
   666
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   667
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
   668
  by (force elim: order_less_asym
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   669
            simp add: Ring_and_Field.mult_less_cancel_right)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   670
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   671
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
   672
apply (simp add: mult_le_cancel_right)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   673
apply (blast intro: elim: order_less_asym) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   674
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   675
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   676
lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   677
  by (force elim: order_less_asym
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   678
            simp add: Ring_and_Field.mult_le_cancel_left)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   679
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   680
text{*Only two uses?*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   681
lemma real_mult_less_mono':
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   682
     "[| 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
   683
 by (rule Ring_and_Field.mult_strict_mono')
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   684
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   685
text{*FIXME: delete or at least combine the next two lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   686
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   687
apply (drule Ring_and_Field.equals_zero_I [THEN sym])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   688
apply (cut_tac x = y in real_le_square) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   689
apply (auto, drule real_le_anti_sym, auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   690
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   691
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   692
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
   693
apply (rule_tac y = x in real_sum_squares_cancel)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   694
apply (simp add: real_add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   695
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   696
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   697
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
   698
by (drule add_strict_mono [of concl: 0 0], assumption, simp)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   699
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   700
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
   701
apply (drule order_le_imp_less_or_eq)+
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   702
apply (auto intro: real_add_order order_less_imp_le)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   703
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   704
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   705
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
   706
apply (case_tac "x \<noteq> 0")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   707
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
   708
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   709
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   710
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
   711
by (auto dest: less_imp_inverse_less)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   712
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   713
lemma real_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
   714
proof -
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   715
  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
   716
  thus ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   717
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   718
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   719
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   720
subsection{*Embedding the Integers into the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   721
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   722
defs (overloaded)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   723
  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
   724
  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
   725
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   726
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
   727
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   728
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   729
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
   730
by (simp add: real_of_int_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   731
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   732
lemma real_of_int_add: "real (x::int) + real y = real (x + y)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   733
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   734
declare real_of_int_add [symmetric, simp]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   735
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   736
lemma real_of_int_minus: "-real (x::int) = real (-x)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   737
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   738
declare real_of_int_minus [symmetric, simp]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   739
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   740
lemma real_of_int_diff: "real (x::int) - real y = real (x - y)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   741
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   742
declare real_of_int_diff [symmetric, simp]
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   743
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   744
lemma real_of_int_mult: "real (x::int) * real y = real (x * y)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   745
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   746
declare real_of_int_mult [symmetric, simp]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   747
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   748
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
   749
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   750
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   751
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
   752
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   753
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   754
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
   755
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   756
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   757
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
   758
by (simp add: real_of_int_def) 
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   759
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   760
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   761
subsection{*Embedding the Naturals into the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   762
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   763
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
   764
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   765
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   766
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
   767
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   768
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   769
lemma 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
   770
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   771
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   772
(*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
   773
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
   774
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   775
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   776
lemma real_of_nat_less_iff [iff]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   777
     "(real (n::nat) < real m) = (n < m)"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   778
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   779
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   780
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
   781
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   782
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   783
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
   784
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
   785
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   786
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
   787
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
   788
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   789
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
   790
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   791
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   792
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
   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_zero_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
   796
by (simp add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   797
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   798
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
   799
by (simp add: add: real_of_nat_def) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   800
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   801
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
   802
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
   803
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   804
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
   805
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   806
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   807
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
   808
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   809
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   810
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
   811
by (simp add: add: real_of_nat_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   812
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   813
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
   814
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
   815
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   816
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   817
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   818
text{*Still needed for binary arith*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   819
lemma real_of_nat_real_of_int: "~neg z ==> real (nat z) = real z"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   820
proof (simp add: not_neg_eq_ge_0 real_of_nat_def real_of_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   821
  assume "0 \<le> z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   822
  hence eq: "of_nat (nat z) = z" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   823
    by (simp add: nat_0_le int_eq_of_nat[symmetric]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   824
  have "of_nat (nat z) = of_int (of_nat (nat z))" by simp
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   825
  also have "... = of_int z" by (simp add: eq)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   826
  finally show "of_nat (nat z) = of_int z" .
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14369
diff changeset
   827
qed
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   828
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   829
ML
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   830
{*
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   831
val real_abs_def = thm "real_abs_def";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   832
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   833
val real_le_def = thm "real_le_def";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   834
val real_diff_def = thm "real_diff_def";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   835
val real_divide_def = thm "real_divide_def";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   836
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   837
val realrel_iff = thm"realrel_iff";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   838
val realrel_refl = thm"realrel_refl";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   839
val equiv_realrel = thm"equiv_realrel";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   840
val equiv_realrel_iff = thm"equiv_realrel_iff";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   841
val realrel_in_real = thm"realrel_in_real";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   842
val inj_on_Abs_REAL = thm"inj_on_Abs_REAL";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   843
val eq_realrelD = thm"eq_realrelD";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   844
val inj_Rep_REAL = thm"inj_Rep_REAL";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   845
val inj_real_of_preal = thm"inj_real_of_preal";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   846
val eq_Abs_REAL = thm"eq_Abs_REAL";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   847
val real_minus_congruent = thm"real_minus_congruent";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   848
val real_minus = thm"real_minus";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   849
val real_add = thm"real_add";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   850
val real_add_commute = thm"real_add_commute";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   851
val real_add_assoc = thm"real_add_assoc";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   852
val real_add_zero_left = thm"real_add_zero_left";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   853
val real_add_zero_right = thm"real_add_zero_right";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   854
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   855
val real_mult = thm"real_mult";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   856
val real_mult_commute = thm"real_mult_commute";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   857
val real_mult_assoc = thm"real_mult_assoc";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   858
val real_mult_1 = thm"real_mult_1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   859
val real_mult_1_right = thm"real_mult_1_right";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   860
val preal_le_linear = thm"preal_le_linear";
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   861
val real_mult_inverse_left = thm"real_mult_inverse_left";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   862
val real_not_refl2 = thm"real_not_refl2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   863
val real_of_preal_add = thm"real_of_preal_add";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   864
val real_of_preal_mult = thm"real_of_preal_mult";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   865
val real_of_preal_trichotomy = thm"real_of_preal_trichotomy";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   866
val real_of_preal_minus_less_zero = thm"real_of_preal_minus_less_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   867
val real_of_preal_not_minus_gt_zero = thm"real_of_preal_not_minus_gt_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   868
val real_of_preal_zero_less = thm"real_of_preal_zero_less";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   869
val real_le_imp_less_or_eq = thm"real_le_imp_less_or_eq";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   870
val real_le_refl = thm"real_le_refl";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   871
val real_le_linear = thm"real_le_linear";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   872
val real_le_trans = thm"real_le_trans";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   873
val real_le_anti_sym = thm"real_le_anti_sym";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   874
val real_less_le = thm"real_less_le";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   875
val real_less_sum_gt_zero = thm"real_less_sum_gt_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   876
val real_gt_zero_preal_Ex = thm "real_gt_zero_preal_Ex";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   877
val real_gt_preal_preal_Ex = thm "real_gt_preal_preal_Ex";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   878
val real_ge_preal_preal_Ex = thm "real_ge_preal_preal_Ex";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   879
val real_less_all_preal = thm "real_less_all_preal";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   880
val real_less_all_real2 = thm "real_less_all_real2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   881
val real_of_preal_le_iff = thm "real_of_preal_le_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   882
val real_mult_order = thm "real_mult_order";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   883
val real_zero_less_one = thm "real_zero_less_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   884
val real_add_less_le_mono = thm "real_add_less_le_mono";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   885
val real_add_le_less_mono = thm "real_add_le_less_mono";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   886
val real_add_order = thm "real_add_order";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   887
val real_le_add_order = thm "real_le_add_order";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   888
val real_le_square = thm "real_le_square";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   889
val real_mult_less_mono2 = thm "real_mult_less_mono2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   890
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   891
val real_mult_less_iff1 = thm "real_mult_less_iff1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   892
val real_mult_le_cancel_iff1 = thm "real_mult_le_cancel_iff1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   893
val real_mult_le_cancel_iff2 = thm "real_mult_le_cancel_iff2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   894
val real_mult_less_mono = thm "real_mult_less_mono";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   895
val real_mult_less_mono' = thm "real_mult_less_mono'";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   896
val real_sum_squares_cancel = thm "real_sum_squares_cancel";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   897
val real_sum_squares_cancel2 = thm "real_sum_squares_cancel2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   898
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   899
val real_mult_left_cancel = thm"real_mult_left_cancel";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   900
val real_mult_right_cancel = thm"real_mult_right_cancel";
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   901
val real_inverse_unique = thm "real_inverse_unique";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   902
val real_inverse_gt_one = thm "real_inverse_gt_one";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   903
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   904
val real_of_int_zero = thm"real_of_int_zero";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   905
val real_of_one = thm"real_of_one";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   906
val real_of_int_add = thm"real_of_int_add";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   907
val real_of_int_minus = thm"real_of_int_minus";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   908
val real_of_int_diff = thm"real_of_int_diff";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   909
val real_of_int_mult = thm"real_of_int_mult";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   910
val real_of_int_real_of_nat = thm"real_of_int_real_of_nat";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   911
val real_of_int_inject = thm"real_of_int_inject";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   912
val real_of_int_less_iff = thm"real_of_int_less_iff";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   913
val real_of_int_le_iff = thm"real_of_int_le_iff";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   914
val real_of_nat_zero = thm "real_of_nat_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   915
val real_of_nat_one = thm "real_of_nat_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   916
val real_of_nat_add = thm "real_of_nat_add";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   917
val real_of_nat_Suc = thm "real_of_nat_Suc";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   918
val real_of_nat_less_iff = thm "real_of_nat_less_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   919
val real_of_nat_le_iff = thm "real_of_nat_le_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   920
val real_of_nat_ge_zero = thm "real_of_nat_ge_zero";
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14348
diff changeset
   921
val real_of_nat_Suc_gt_zero = thm "real_of_nat_Suc_gt_zero";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   922
val real_of_nat_mult = thm "real_of_nat_mult";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   923
val real_of_nat_inject = thm "real_of_nat_inject";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   924
val real_of_nat_diff = thm "real_of_nat_diff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   925
val real_of_nat_zero_iff = thm "real_of_nat_zero_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   926
val real_of_nat_gt_zero_cancel_iff = thm "real_of_nat_gt_zero_cancel_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   927
val real_of_nat_le_zero_cancel_iff = thm "real_of_nat_le_zero_cancel_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   928
val not_real_of_nat_less_zero = thm "not_real_of_nat_less_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   929
val real_of_nat_ge_zero_cancel_iff = thm "real_of_nat_ge_zero_cancel_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   930
*}
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
   931
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
   932
end