src/HOL/Real/RealDef.thy
author paulson
Fri, 09 Jan 2004 10:46:18 +0100
changeset 14348 744c868ee0b7
parent 14341 a09441bd4f1e
child 14365 3d4df8c166ae
permissions -rw-r--r--
Defining the type class "ringpower" and deleting superseded theorems for types nat, int, real, hypreal
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
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    26
   (*Overloaded constants denoting the Nat and Real subsets of enclosing
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
   Nats  :: "'a set"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    29
   Reals :: "'a set"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    30
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    31
   (*overloaded constant for injecting other types into "real"*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    32
   real :: "'a => real"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    33
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    34
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    35
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    36
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    37
  real_zero_def:
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    38
  "0 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    39
			    preal_of_prat(prat_of_pnat 1))})"
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    40
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    41
  real_one_def:
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    42
  "1 == Abs_REAL(realrel``
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    43
               {(preal_of_prat(prat_of_pnat 1) + preal_of_prat(prat_of_pnat 1),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    44
		 preal_of_prat(prat_of_pnat 1))})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    45
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    46
  real_minus_def:
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    47
  "- R ==  Abs_REAL(UN (x,y):Rep_REAL(R). realrel``{(y,x)})"
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    48
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    49
  real_diff_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    50
  "R - (S::real) == R + - S"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    51
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    52
  real_inverse_def:
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11701
diff changeset
    53
  "inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    54
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    55
  real_divide_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    56
  "R / (S::real) == R * inverse S"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    57
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    58
constdefs
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    59
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    60
  (** 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
    61
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    62
  real_of_preal :: "preal => real"
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    63
  "real_of_preal m     ==
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    64
           Abs_REAL(realrel``{(m + preal_of_prat(prat_of_pnat 1),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    65
                               preal_of_prat(prat_of_pnat 1))})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    66
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    67
  real_of_posnat :: "nat => real"
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    68
  "real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))"
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    69
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    70
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    71
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    72
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    73
  real_of_nat_def:   "real n == real_of_posnat n + (- 1)"
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    74
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    75
  real_add_def:
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    76
  "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
    77
                   (%(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
    78
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    79
  real_mult_def:
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    80
  "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
    81
                   (%(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
    82
		   p2) p1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    83
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    84
  real_less_def:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    85
  "P<Q == \<exists>x1 y1 x2 y2. x1 + y2 < x2 + y1 &
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    86
                            (x1,y1)\<in>Rep_REAL(P) & (x2,y2)\<in>Rep_REAL(Q)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    87
  real_le_def:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    88
  "P \<le> (Q::real) == ~(Q < P)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    89
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    90
  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
    91
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    92
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 12018
diff changeset
    93
syntax (xsymbols)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    94
  Reals     :: "'a set"                   ("\<real>")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    95
  Nats      :: "'a set"                   ("\<nat>")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    96
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    97
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
    98
subsection{*Proving that realrel is an equivalence relation*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    99
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   100
lemma preal_trans_lemma:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   101
     "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |]
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   102
      ==> x1 + y3 = x3 + y1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   103
apply (rule_tac C = y2 in preal_add_right_cancel)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   104
apply (rotate_tac 1, drule sym)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   105
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   106
apply (rule preal_add_left_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   107
apply (rule_tac x1 = x1 in preal_add_assoc [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   108
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   109
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   110
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   111
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
   112
by (unfold realrel_def, blast)
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 realrel_refl: "(x,x): realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   115
apply (case_tac "x")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   116
apply (simp add: realrel_def)
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
lemma equiv_realrel: "equiv UNIV realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   120
apply (unfold equiv_def refl_def sym_def trans_def realrel_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   121
apply (fast elim!: sym preal_trans_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   122
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   123
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   124
(* (realrel `` {x} = realrel `` {y}) = ((x,y) : realrel) *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   125
lemmas equiv_realrel_iff = 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   126
       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
   127
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   128
declare equiv_realrel_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   129
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   130
lemma realrel_in_real [simp]: "realrel``{(x,y)}: REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   131
by (unfold REAL_def realrel_def quotient_def, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   132
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   133
lemma inj_on_Abs_REAL: "inj_on Abs_REAL REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   134
apply (rule inj_on_inverseI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   135
apply (erule Abs_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   136
done
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
declare inj_on_Abs_REAL [THEN inj_on_iff, simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   139
declare Abs_REAL_inverse [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   140
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   141
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   142
lemmas eq_realrelD = equiv_realrel [THEN [2] eq_equiv_class]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   143
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   144
lemma inj_Rep_REAL: "inj Rep_REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   145
apply (rule inj_on_inverseI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   146
apply (rule Rep_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   147
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   148
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   149
(** real_of_preal: the injection from preal to real **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   150
lemma inj_real_of_preal: "inj(real_of_preal)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   151
apply (rule inj_onI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   152
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   153
apply (drule inj_on_Abs_REAL [THEN inj_onD])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   154
apply (rule realrel_in_real)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   155
apply (drule eq_equiv_class)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   156
apply (rule equiv_realrel, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   157
apply (simp add: realrel_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   158
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   159
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   160
lemma eq_Abs_REAL: 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   161
    "(!!x y. z = Abs_REAL(realrel``{(x,y)}) ==> P) ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   162
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
   163
apply (drule_tac f = Abs_REAL in arg_cong)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   164
apply (case_tac "x")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   165
apply (simp add: Rep_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   166
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   167
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   168
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   169
subsection{*Congruence property for addition*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   170
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   171
lemma real_add_congruent2_lemma:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   172
     "[|a + ba = aa + b; ab + bc = ac + bb|]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   173
      ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   174
apply (simp add: preal_add_assoc) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   175
apply (rule preal_add_left_commute [of ab, THEN ssubst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   176
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   177
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   178
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   179
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   180
lemma real_add:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   181
  "Abs_REAL(realrel``{(x1,y1)}) + Abs_REAL(realrel``{(x2,y2)}) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   182
   Abs_REAL(realrel``{(x1+x2, y1+y2)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   183
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
   184
apply (subst equiv_realrel [THEN UN_equiv_class2])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   185
apply (auto simp add: congruent2_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   186
apply (blast intro: real_add_congruent2_lemma) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   187
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   188
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   189
lemma real_add_commute: "(z::real) + w = w + z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   190
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   191
apply (rule_tac z = w in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   192
apply (simp add: preal_add_ac real_add)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   193
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   194
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   195
lemma real_add_assoc: "((z1::real) + z2) + z3 = z1 + (z2 + z3)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   196
apply (rule_tac z = z1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   197
apply (rule_tac z = z2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   198
apply (rule_tac z = z3 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   199
apply (simp add: real_add preal_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   200
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   201
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   202
lemma real_add_zero_left: "(0::real) + z = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   203
apply (unfold real_of_preal_def real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   204
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   205
apply (simp add: real_add preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   206
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   207
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   208
lemma real_add_zero_right: "z + (0::real) = z"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   209
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
   210
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   211
instance real :: plus_ac0
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   212
  by (intro_classes,
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   213
      (assumption | 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   214
       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
   215
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   216
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   217
subsection{*Additive Inverse on real*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   218
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   219
lemma real_minus_congruent:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   220
  "congruent realrel (%p. (%(x,y). realrel``{(y,x)}) p)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   221
apply (unfold congruent_def, clarify)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   222
apply (simp add: preal_add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   223
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   224
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   225
lemma real_minus:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   226
      "- (Abs_REAL(realrel``{(x,y)})) = Abs_REAL(realrel `` {(y,x)})"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   227
apply (unfold real_minus_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   228
apply (rule_tac f = Abs_REAL in arg_cong)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   229
apply (simp add: realrel_in_real [THEN Abs_REAL_inverse] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   230
            UN_equiv_class [OF equiv_realrel real_minus_congruent])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   231
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   232
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   233
lemma real_add_minus_left: "(-z) + z = (0::real)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   234
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   235
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   236
apply (simp add: real_minus real_add preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   237
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   238
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   239
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   240
subsection{*Congruence property for multiplication*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   241
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   242
lemma real_mult_congruent2_lemma:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   243
     "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   244
          x * x1 + y * y1 + (x * y2 + x2 * y) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   245
          x * x2 + y * y2 + (x * y1 + x1 * y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   246
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
   247
apply (rule preal_mult_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   248
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
   249
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
   250
apply (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   251
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   252
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   253
lemma real_mult_congruent2:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   254
    "congruent2 realrel (%p1 p2.
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   255
          (%(x1,y1). (%(x2,y2). realrel``{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   256
apply (rule equiv_realrel [THEN congruent2_commuteI], clarify)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   257
apply (unfold split_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   258
apply (simp add: preal_mult_commute preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   259
apply (auto simp add: real_mult_congruent2_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   260
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   261
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   262
lemma real_mult:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   263
   "Abs_REAL((realrel``{(x1,y1)})) * Abs_REAL((realrel``{(x2,y2)})) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   264
    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
   265
apply (unfold real_mult_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   266
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
   267
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   268
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   269
lemma real_mult_commute: "(z::real) * w = w * z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   270
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   271
apply (rule_tac z = w in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   272
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
   273
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   274
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   275
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   276
apply (rule_tac z = z1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   277
apply (rule_tac z = z2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   278
apply (rule_tac z = z3 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   279
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
   280
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   281
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   282
lemma real_mult_1: "(1::real) * z = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   283
apply (unfold real_one_def pnat_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   284
apply (rule_tac z = z in eq_Abs_REAL)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   285
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
   286
                 preal_mult_ac preal_add_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   287
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   288
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   289
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   290
apply (rule_tac z = z1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   291
apply (rule_tac z = z2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   292
apply (rule_tac z = w in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   293
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
   294
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   295
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   296
text{*one and zero are distinct*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   297
lemma real_zero_not_eq_one: "0 ~= (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   298
apply (unfold real_zero_def real_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   299
apply (auto simp add: preal_self_less_add_left [THEN preal_not_refl2])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   300
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   301
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   302
subsection{*existence of inverse*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   303
(** lemma -- alternative definition of 0 **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   304
lemma real_zero_iff: "0 = Abs_REAL (realrel `` {(x, x)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   305
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   306
apply (auto simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   307
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   308
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   309
lemma real_mult_inv_left_ex: "x ~= 0 ==> \<exists>y. y*x = (1::real)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   310
apply (unfold real_zero_def real_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   311
apply (rule_tac z = x in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   312
apply (cut_tac x = xa and y = y in linorder_less_linear)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   313
apply (auto dest!: preal_less_add_left_Ex simp add: real_zero_iff [symmetric])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   314
apply (rule_tac
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   315
        x = "Abs_REAL (realrel `` { (preal_of_prat (prat_of_pnat 1), 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   316
                            pinv (D) + preal_of_prat (prat_of_pnat 1))}) " 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   317
       in exI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   318
apply (rule_tac [2]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   319
        x = "Abs_REAL (realrel `` { (pinv (D) + preal_of_prat (prat_of_pnat 1),
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   320
                   preal_of_prat (prat_of_pnat 1))})" 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   321
       in exI)
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   322
apply (auto simp add: real_mult pnat_one_def preal_mult_1_right
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   323
              preal_add_mult_distrib2 preal_add_mult_distrib preal_mult_1
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   324
              preal_mult_inv_right preal_add_ac preal_mult_ac)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   325
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   326
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   327
lemma real_mult_inv_left: "x ~= 0 ==> inverse(x)*x = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   328
apply (unfold real_inverse_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   329
apply (frule real_mult_inv_left_ex, safe)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   330
apply (rule someI2, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   331
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   332
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   333
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   334
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
   335
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   336
instance real :: field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   337
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   338
  fix x y z :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   339
  show "(x + y) + z = x + (y + z)" by (rule real_add_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   340
  show "x + y = y + x" by (rule real_add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   341
  show "0 + x = x" by simp
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   342
  show "- x + x = 0" by (rule real_add_minus_left)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   343
  show "x - y = x + (-y)" by (simp add: real_diff_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   344
  show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   345
  show "x * y = y * x" by (rule real_mult_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   346
  show "1 * x = x" by (rule real_mult_1)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   347
  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
   348
  show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   349
  show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inv_left)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   350
  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
   351
  assume eq: "z+x = z+y" 
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   352
    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
   353
    thus "x = y" by (simp add: real_add_minus_left)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   354
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   355
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   356
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   357
text{*Inverse of zero!  Useful to simplify certain equations*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   358
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   359
lemma INVERSE_ZERO: "inverse 0 = (0::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   360
apply (unfold real_inverse_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   361
apply (rule someI2)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   362
apply (auto simp add: zero_neq_one)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   363
done
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   364
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   365
lemma DIVISION_BY_ZERO: "a / (0::real) = 0"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   366
  by (simp add: real_divide_def INVERSE_ZERO)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   367
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   368
instance real :: division_by_zero
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   369
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   370
  fix x :: real
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   371
  show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   372
  show "x/0 = 0" by (rule DIVISION_BY_ZERO) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   373
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   374
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
(*Pull negations out*)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   377
declare minus_mult_right [symmetric, simp] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   378
        minus_mult_left [symmetric, simp]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   379
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   380
text{*Used in RealBin*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   381
lemma real_minus_mult_commute: "(-x) * y = x * (- y :: real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   382
by simp
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   383
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   384
lemma real_mult_1_right: "z * (1::real) = z"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   385
  by (rule Ring_and_Field.mult_1_right)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   386
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   387
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   388
subsection{*Theorems for Ordering*}
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   389
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   390
(* real_less is a strict order: irreflexive *)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   391
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   392
text{*lemmas*}
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   393
lemma preal_lemma_eq_rev_sum:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   394
     "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   395
by (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   396
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   397
lemma preal_add_left_commute_cancel:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   398
     "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   399
by (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   400
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   401
lemma preal_lemma_for_not_refl:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   402
     "!!(x::preal). [| x + y2a = x2a + y;
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   403
                       x + y2b = x2b + y |]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   404
                    ==> x2a + y2b = x2b + y2a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   405
apply (drule preal_lemma_eq_rev_sum, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   406
apply (erule_tac V = "x + y2b = x2b + y" in thin_rl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   407
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   408
apply (drule preal_add_left_commute_cancel)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   409
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   410
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   411
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   412
lemma real_less_not_refl: "~ (R::real) < R"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   413
apply (rule_tac z = R in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   414
apply (auto simp add: real_less_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   415
apply (drule preal_lemma_for_not_refl, assumption, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   416
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   417
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   418
(*** y < y ==> P ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   419
lemmas real_less_irrefl = real_less_not_refl [THEN notE, standard]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   420
declare real_less_irrefl [elim!]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   421
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   422
lemma real_not_refl2: "!!(x::real). x < y ==> x ~= y"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   423
by (auto simp add: real_less_not_refl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   424
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   425
(* lemma re-arranging and eliminating terms *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   426
lemma preal_lemma_trans: "!! (a::preal). [| a + b = c + d;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   427
             x2b + d + (c + y2e) < a + y2b + (x2e + b) |]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   428
          ==> x2b + y2e < x2e + y2b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   429
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   430
apply (rule_tac C = "c+d" in preal_add_left_less_cancel)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   431
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   432
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   433
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   434
(** A MESS!  heavy re-writing involved*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   435
lemma real_less_trans: "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   436
apply (rule_tac z = R1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   437
apply (rule_tac z = R2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   438
apply (rule_tac z = R3 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   439
apply (auto simp add: real_less_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   440
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   441
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   442
 prefer 2 apply blast 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   443
 prefer 2 apply blast 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   444
apply (drule preal_lemma_for_not_refl, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   445
apply (blast dest: preal_add_less_mono intro: preal_lemma_trans)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   446
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   447
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   448
lemma real_less_not_sym: "!! (R1::real). R1 < R2 ==> ~ (R2 < R1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   449
apply (rule notI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   450
apply (drule real_less_trans, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   451
apply (simp add: real_less_not_refl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   452
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   453
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   454
(* [| x < y;  ~P ==> y < x |] ==> P *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   455
lemmas real_less_asym = real_less_not_sym [THEN contrapos_np, standard]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   456
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   457
lemma real_of_preal_add:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   458
     "real_of_preal ((z1::preal) + z2) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   459
      real_of_preal z1 + real_of_preal z2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   460
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   461
apply (simp add: real_add preal_add_mult_distrib preal_mult_1 add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   462
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   463
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   464
lemma real_of_preal_mult:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   465
     "real_of_preal ((z1::preal) * z2) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   466
      real_of_preal z1* real_of_preal z2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   467
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   468
apply (simp (no_asm_use) add: real_mult preal_add_mult_distrib2 preal_mult_1 preal_mult_1_right pnat_one_def preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   469
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   470
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   471
lemma real_of_preal_ExI:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   472
      "!!(x::preal). y < x ==>
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   473
       \<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   474
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   475
apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   476
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   477
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   478
lemma real_of_preal_ExD:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   479
      "!!(x::preal). \<exists>m. Abs_REAL (realrel `` {(x,y)}) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   480
                     real_of_preal m ==> y < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   481
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   482
apply (auto simp add: preal_add_commute preal_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   483
apply (simp add: preal_add_assoc [symmetric] preal_self_less_add_left)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   484
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   485
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   486
lemma real_of_preal_iff:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   487
     "(\<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m) = (y < x)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   488
by (blast intro!: real_of_preal_ExI real_of_preal_ExD)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   489
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   490
text{*Gleason prop 9-4.4 p 127*}
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   491
lemma real_of_preal_trichotomy:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   492
      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   493
apply (unfold real_of_preal_def real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   494
apply (rule_tac z = x in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   495
apply (auto simp add: real_minus preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   496
apply (cut_tac x = x and y = y in linorder_less_linear)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   497
apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   498
apply (auto simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   499
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   500
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   501
lemma real_of_preal_trichotomyE:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   502
     "!!P. [| !!m. x = real_of_preal m ==> P;
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   503
              x = 0 ==> P;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   504
              !!m. x = -(real_of_preal m) ==> P |] ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   505
apply (cut_tac x = x in real_of_preal_trichotomy, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   506
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   507
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   508
lemma real_of_preal_lessD:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   509
      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   510
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   511
apply (auto simp add: real_less_def preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   512
apply (auto simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   513
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   514
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   515
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   516
lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   517
apply (drule preal_less_add_left_Ex)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   518
apply (auto simp add: real_of_preal_add real_of_preal_def real_less_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   519
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   520
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   521
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   522
apply (simp add: preal_self_less_add_left del: preal_add_less_iff2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   523
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   524
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   525
lemma real_of_preal_less_iff1:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   526
     "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   527
by (blast intro: real_of_preal_lessI real_of_preal_lessD)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   528
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   529
declare real_of_preal_less_iff1 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   530
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   531
lemma real_of_preal_minus_less_self: "- real_of_preal m < real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   532
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   533
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   534
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   535
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   536
apply (simp (no_asm_use) add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   537
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   538
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   539
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   540
lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   541
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   542
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   543
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   544
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   545
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   546
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   547
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   548
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   549
lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   550
apply (cut_tac real_of_preal_minus_less_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   551
apply (fast dest: real_less_trans elim: real_less_irrefl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   552
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   553
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   554
lemma real_of_preal_zero_less: "0 < real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   555
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   556
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   557
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   558
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   559
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   560
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   561
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   562
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   563
lemma real_of_preal_not_less_zero: "~ real_of_preal m < 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   564
apply (cut_tac real_of_preal_zero_less)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   565
apply (blast dest: real_less_trans elim: real_less_irrefl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   566
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   567
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   568
lemma real_minus_minus_zero_less: "0 < - (- real_of_preal m)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   569
by (simp add: real_of_preal_zero_less)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   570
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   571
(* another lemma *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   572
lemma real_of_preal_sum_zero_less:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   573
      "0 < real_of_preal m + real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   574
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   575
apply (auto simp add: real_of_preal_def real_less_def real_add)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   576
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   577
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   578
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   579
apply (simp (no_asm_use) add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   580
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   581
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   582
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   583
lemma real_of_preal_minus_less_all: "- real_of_preal m < real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   584
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   585
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   586
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   587
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   588
apply (simp (no_asm_use) add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   589
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   590
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   591
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   592
lemma real_of_preal_not_minus_gt_all: "~ real_of_preal m < - real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   593
apply (cut_tac real_of_preal_minus_less_all)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   594
apply (blast dest: real_less_trans elim: real_less_irrefl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   595
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   596
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   597
lemma real_of_preal_minus_less_rev1:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   598
     "- real_of_preal m1 < - real_of_preal m2
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   599
      ==> real_of_preal m2 < real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   600
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   601
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   602
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   603
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   604
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   605
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   606
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   607
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   608
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   609
lemma real_of_preal_minus_less_rev2:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   610
     "real_of_preal m1 < real_of_preal m2
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   611
      ==> - real_of_preal m2 < - real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   612
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   613
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   614
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   615
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   616
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   617
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   618
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   619
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   620
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   621
lemma real_of_preal_minus_less_rev_iff:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   622
     "(- real_of_preal m1 < - real_of_preal m2) =
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   623
      (real_of_preal m2 < real_of_preal m1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   624
apply (blast intro!: real_of_preal_minus_less_rev1 real_of_preal_minus_less_rev2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   625
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   626
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   627
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   628
subsection{*Linearity of the Ordering*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   629
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   630
lemma real_linear: "(x::real) < y | x = y | y < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   631
apply (rule_tac x = x in real_of_preal_trichotomyE)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   632
apply (rule_tac [!] x = y in real_of_preal_trichotomyE)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   633
apply (auto dest!: preal_le_anti_sym 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   634
            simp add: preal_less_le_iff real_of_preal_minus_less_zero 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   635
                      real_of_preal_zero_less real_of_preal_minus_less_all
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   636
                      real_of_preal_minus_less_rev_iff)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   637
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   638
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   639
lemma real_neq_iff: "!!w::real. (w ~= z) = (w<z | z<w)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   640
by (cut_tac real_linear, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   641
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   642
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   643
lemma real_linear_less2:
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14270
diff changeset
   644
     "!!(R1::real). [| R1 < R2 ==> P;  R1 = R2 ==> P;
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   645
                       R2 < R1 ==> P |] ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   646
apply (cut_tac x = R1 and y = R2 in real_linear, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   647
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   648
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   649
lemma real_minus_zero_less_iff: "(0 < -R) = (R < (0::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   650
apply (rule_tac x = R in real_of_preal_trichotomyE)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   651
apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   652
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   653
declare real_minus_zero_less_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   654
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   655
lemma real_minus_zero_less_iff2: "(-R < 0) = ((0::real) < R)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   656
apply (rule_tac x = R in real_of_preal_trichotomyE)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   657
apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   658
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   659
declare real_minus_zero_less_iff2 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   660
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   661
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   662
subsection{*Properties of Less-Than Or Equals*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   663
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   664
lemma real_le_imp_less_or_eq: "!!(x::real). x \<le> y ==> x < y | x = y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   665
apply (unfold real_le_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   666
apply (cut_tac real_linear)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   667
apply (blast elim: real_less_irrefl real_less_asym)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   668
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   669
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   670
lemma real_less_or_eq_imp_le: "z<w | z=w ==> z \<le>(w::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   671
apply (unfold real_le_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   672
apply (cut_tac real_linear)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   673
apply (fast elim: real_less_irrefl real_less_asym)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   674
done
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_le_less: "(x \<le> (y::real)) = (x < y | x=y)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   677
by (blast intro!: real_less_or_eq_imp_le dest!: real_le_imp_less_or_eq)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   678
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   679
lemma real_le_refl: "w \<le> (w::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   680
by (simp add: real_le_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   681
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   682
lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   683
apply (drule real_le_imp_less_or_eq) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   684
apply (drule real_le_imp_less_or_eq) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   685
apply (rule real_less_or_eq_imp_le) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   686
apply (blast intro: real_less_trans) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   687
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   688
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   689
lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   690
apply (drule real_le_imp_less_or_eq) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   691
apply (drule real_le_imp_less_or_eq) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   692
apply (fast elim: real_less_irrefl real_less_asym)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   693
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   694
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   695
(* Axiom 'order_less_le' of class 'order': *)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   696
lemma real_less_le: "((w::real) < z) = (w \<le> z & w \<noteq> z)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   697
apply (simp add: real_le_def real_neq_iff)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   698
apply (blast elim!: real_less_asym)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   699
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   700
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   701
instance real :: order
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   702
  by (intro_classes,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   703
      (assumption | 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   704
       rule real_le_refl real_le_trans real_le_anti_sym real_less_le)+)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   705
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   706
(* Axiom 'linorder_linear' of class 'linorder': *)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   707
lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   708
apply (simp add: real_le_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   709
apply (cut_tac real_linear, blast)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   710
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   711
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   712
instance real :: linorder
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   713
  by (intro_classes, rule real_le_linear)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   714
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   715
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   716
subsection{*Theorems About the Ordering*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   717
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   718
lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   719
apply (auto simp add: real_of_preal_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   720
apply (cut_tac x = x in real_of_preal_trichotomy)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   721
apply (blast elim!: real_less_irrefl real_of_preal_not_minus_gt_zero [THEN notE])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   722
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   723
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   724
lemma real_gt_preal_preal_Ex:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   725
     "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   726
by (blast dest!: real_of_preal_zero_less [THEN real_less_trans]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   727
             intro: real_gt_zero_preal_Ex [THEN iffD1])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   728
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   729
lemma real_ge_preal_preal_Ex:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   730
     "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   731
by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   732
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   733
lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   734
by (auto elim: order_le_imp_less_or_eq [THEN disjE] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   735
            intro: real_of_preal_zero_less [THEN [2] real_less_trans] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   736
            simp add: real_of_preal_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   737
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   738
lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   739
by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   740
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   741
lemma real_of_preal_le_iff:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   742
     "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   743
by (auto intro!: preal_le_iff_less_or_eq [THEN iffD1]  
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   744
          simp add: linorder_not_less [symmetric])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   745
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   746
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   747
subsection{*Monotonicity of Addition*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   748
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   749
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   750
apply (auto simp add: real_gt_zero_preal_Ex)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   751
apply (rule_tac x = "y*ya" in exI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   752
apply (simp (no_asm_use) add: real_of_preal_mult)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   753
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   754
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   755
(*Alternative definition for real_less*)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   756
lemma real_less_add_positive_left_Ex: "R < S ==> \<exists>T::real. 0 < T & R + T = S"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   757
apply (rule_tac x = R in real_of_preal_trichotomyE)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   758
apply (rule_tac [!] x = S in real_of_preal_trichotomyE)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   759
apply (auto dest!: preal_less_add_left_Ex 
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   760
        simp add: real_of_preal_not_minus_gt_all real_of_preal_add
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   761
                real_of_preal_not_less_zero real_less_not_refl 
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   762
             real_of_preal_not_minus_gt_zero real_of_preal_minus_less_rev_iff)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   763
apply (rule real_of_preal_zero_less) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   764
apply (rule_tac [1] x = "real_of_preal m+real_of_preal ma" in exI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   765
apply (rule_tac [2] x = "real_of_preal D" in exI)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   766
apply (auto simp add: real_of_preal_minus_less_rev_iff real_of_preal_zero_less
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   767
                 real_of_preal_sum_zero_less real_add_assoc)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   768
apply (simp add: real_add_assoc [symmetric])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   769
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   770
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   771
lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   772
apply (drule real_less_add_positive_left_Ex)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   773
apply (auto simp add: add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   774
done
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_lemma_change_eq_subj: "!!S::real. T = S + W ==> S = T + (-W)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   777
by (simp add: add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   778
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   779
(* FIXME: long! *)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   780
lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   781
apply (rule ccontr)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   782
apply (drule linorder_not_less [THEN iffD1, THEN real_le_imp_less_or_eq])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   783
apply (auto simp add: real_less_not_refl)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   784
apply (drule real_less_add_positive_left_Ex, clarify, simp)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   785
apply (drule real_lemma_change_eq_subj, auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   786
apply (drule real_less_sum_gt_zero)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   787
apply (auto elim: real_less_asym simp add: add_left_commute [of W] add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   788
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   789
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   790
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
   791
apply (rule real_sum_gt_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   792
apply (drule real_less_sum_gt_zero [of x y])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   793
apply (drule real_mult_order, assumption)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   794
apply (simp add: right_distrib)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   795
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   796
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   797
lemma real_less_sum_gt_0_iff: "(0 < S + (-W::real)) = (W < S)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   798
by (blast intro: real_less_sum_gt_zero real_sum_gt_zero_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   799
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   800
lemma real_less_eq_diff: "(x<y) = (x-y < (0::real))"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   801
apply (unfold real_diff_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   802
apply (subst real_minus_zero_less_iff [symmetric])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   803
apply (simp add: real_add_commute real_less_sum_gt_0_iff)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   804
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   805
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   806
lemma real_less_eqI: "(x::real) - y = x' - y' ==> (x<y) = (x'<y')"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   807
apply (subst real_less_eq_diff)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   808
apply (rule_tac y1 = y in real_less_eq_diff [THEN ssubst], simp)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   809
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   810
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   811
lemma real_le_eqI: "(x::real) - y = x' - y' ==> (y\<le>x) = (y'\<le>x')"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   812
apply (drule real_less_eqI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   813
apply (simp add: real_le_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   814
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   815
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   816
lemma real_add_left_mono: "x \<le> y ==> z + x \<le> z + (y::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   817
apply (rule real_le_eqI [THEN iffD1]) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   818
 prefer 2 apply assumption
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   819
apply (simp add: real_diff_def add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   820
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   821
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   822
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   823
subsection{*The Reals Form an Ordered Field*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   824
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   825
instance real :: ordered_field
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   826
proof
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   827
  fix x y z :: real
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   828
  show "0 < (1::real)" 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   829
    by (force intro: real_gt_zero_preal_Ex [THEN iffD2]
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   830
              simp add: real_one_def real_of_preal_def)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   831
  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
   832
  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
   833
  show "\<bar>x\<bar> = (if x < 0 then -x else x)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   834
    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
   835
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   836
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   837
text{*These two need to be proved in @{text Ring_and_Field}*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   838
lemma real_add_less_le_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   839
apply (erule add_strict_right_mono [THEN order_less_le_trans])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   840
apply (erule add_left_mono) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   841
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   842
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   843
lemma real_add_le_less_mono:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   844
     "!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   845
apply (erule add_right_mono [THEN order_le_less_trans])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   846
apply (erule add_strict_left_mono) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   847
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   848
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   849
lemma real_zero_less_one: "0 < (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   850
  by (rule Ring_and_Field.zero_less_one)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   851
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   852
lemma real_le_square [simp]: "(0::real) \<le> x*x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   853
 by (rule Ring_and_Field.zero_le_square)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   854
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   855
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   856
subsection{*More Lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   857
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   858
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
   859
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   860
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   861
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
   862
by auto
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   863
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   864
text{*The precondition could be weakened to @{term "0\<le>x"}*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   865
lemma real_mult_less_mono:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   866
     "[| 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
   867
 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
   868
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   869
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
   870
  by (force elim: order_less_asym
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   871
            simp add: Ring_and_Field.mult_less_cancel_right)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   872
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   873
lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   874
by (auto simp add: real_le_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   875
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   876
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
   877
  by (force elim: order_less_asym
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   878
            simp add: Ring_and_Field.mult_le_cancel_left)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   879
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   880
text{*Only two uses?*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   881
lemma real_mult_less_mono':
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   882
     "[| 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
   883
 by (rule Ring_and_Field.mult_strict_mono')
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   884
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   885
text{*FIXME: delete or at least combine the next two lemmas*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   886
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
   887
apply (drule Ring_and_Field.equals_zero_I [THEN sym])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   888
apply (cut_tac x = y in real_le_square) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   889
apply (auto, drule real_le_anti_sym, auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   890
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   891
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   892
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
   893
apply (rule_tac y = x in real_sum_squares_cancel)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   894
apply (simp add: real_add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   895
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   896
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   897
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   898
lemma real_add_order: "[| 0 < x; 0 < y |] ==> (0::real) < x + y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   899
apply (drule add_strict_mono [of concl: 0 0], assumption)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   900
apply simp 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   901
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   902
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   903
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
   904
apply (drule order_le_imp_less_or_eq)+
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   905
apply (auto intro: real_add_order order_less_imp_le)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   906
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   907
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   908
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   909
subsection{*An Embedding of the Naturals in the Reals*}
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   910
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   911
lemma real_of_posnat_one: "real_of_posnat 0 = (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   912
by (simp add: real_of_posnat_def pnat_one_iff [symmetric]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   913
              real_of_preal_def symmetric real_one_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   914
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   915
lemma real_of_posnat_two: "real_of_posnat (Suc 0) = (1::real) + (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   916
by (simp add: real_of_posnat_def real_of_preal_def real_one_def pnat_two_eq
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   917
                 real_add
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   918
            prat_of_pnat_add [symmetric] preal_of_prat_add [symmetric]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   919
            pnat_add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   920
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   921
lemma real_of_posnat_add: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   922
     "real_of_posnat n1 + real_of_posnat n2 = real_of_posnat (n1 + n2) + (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   923
apply (unfold real_of_posnat_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   924
apply (simp (no_asm_use) add: real_of_posnat_one [symmetric] real_of_posnat_def real_of_preal_add [symmetric] preal_of_prat_add [symmetric] prat_of_pnat_add [symmetric] pnat_of_nat_add)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   925
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   926
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   927
lemma real_of_posnat_add_one:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   928
     "real_of_posnat (n + 1) = real_of_posnat n + (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   929
apply (rule_tac a1 = " (1::real) " in add_right_cancel [THEN iffD1])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   930
apply (rule real_of_posnat_add [THEN subst])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   931
apply (simp (no_asm_use) add: real_of_posnat_two real_add_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   932
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   933
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   934
lemma real_of_posnat_Suc:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   935
     "real_of_posnat (Suc n) = real_of_posnat n + (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   936
by (subst real_of_posnat_add_one [symmetric], simp)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   937
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   938
lemma inj_real_of_posnat: "inj(real_of_posnat)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   939
apply (rule inj_onI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   940
apply (unfold real_of_posnat_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   941
apply (drule inj_real_of_preal [THEN injD])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   942
apply (drule inj_preal_of_prat [THEN injD])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   943
apply (drule inj_prat_of_pnat [THEN injD])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   944
apply (erule inj_pnat_of_nat [THEN injD])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   945
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   946
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   947
lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   948
by (simp add: real_of_nat_def real_of_posnat_one)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   949
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   950
lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   951
by (simp add: real_of_nat_def real_of_posnat_two real_add_assoc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   952
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   953
lemma real_of_nat_add [simp]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   954
     "real (m + n) = real (m::nat) + real n"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   955
apply (simp add: real_of_nat_def add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   956
apply (simp add: real_of_posnat_add add_assoc [symmetric])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   957
apply (simp add: add_commute) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   958
apply (simp add: add_assoc [symmetric])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   959
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   960
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   961
(*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
   962
lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   963
by (simp add: real_of_nat_def real_of_posnat_Suc add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   964
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   965
lemma real_of_nat_less_iff [iff]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   966
     "(real (n::nat) < real m) = (n < m)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   967
by (auto simp add: real_of_nat_def real_of_posnat_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   968
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   969
lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   970
by (simp add: linorder_not_less [symmetric])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   971
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   972
lemma inj_real_of_nat: "inj (real :: nat => real)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   973
apply (rule inj_onI)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   974
apply (auto intro!: inj_real_of_posnat [THEN injD]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   975
            simp add: real_of_nat_def add_right_cancel)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   976
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   977
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   978
lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   979
apply (induct_tac "n")
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   980
apply (auto simp add: real_of_nat_Suc)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   981
apply (drule real_add_le_less_mono)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   982
apply (rule zero_less_one)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   983
apply (simp add: order_less_imp_le)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   984
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   985
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   986
lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   987
apply (induct_tac "m")
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   988
apply (auto simp add: real_of_nat_Suc left_distrib add_commute)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   989
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   990
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   991
lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   992
by (auto dest: inj_real_of_nat [THEN injD])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   993
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   994
lemma real_of_nat_diff [rule_format]:
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   995
     "n \<le> m --> real (m - n) = real (m::nat) - real n"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   996
apply (induct_tac "m", simp)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   997
apply (simp add: real_diff_def Suc_diff_le le_Suc_eq real_of_nat_Suc add_ac)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   998
apply (simp add: add_left_commute [of _ "- 1"]) 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   999
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1000
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1001
lemma real_of_nat_zero_iff: "(real (n::nat) = 0) = (n = 0)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1002
  proof 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1003
    assume "real n = 0"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1004
    have "real n = real (0::nat)" by simp
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1005
    then show "n = 0" by (simp only: real_of_nat_inject)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1006
  next
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1007
    show "n = 0 \<Longrightarrow> real n = 0" by simp
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1008
  qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1009
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1010
lemma real_of_nat_neg_int [simp]: "neg z ==> real (nat z) = 0"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1011
by (simp add: neg_nat real_of_nat_zero)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1012
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1013
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1014
lemma real_inverse_unique: "x*y = (1::real) ==> y = inverse x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1015
apply (case_tac "x \<noteq> 0")
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1016
apply (rule_tac c1 = x in real_mult_left_cancel [THEN iffD1], auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1017
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1018
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1019
lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1020
by (auto dest: less_imp_inverse_less)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1021
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1022
lemma real_of_nat_gt_zero_cancel_iff: "(0 < real (n::nat)) = (0 < n)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1023
by (rule real_of_nat_less_iff [THEN subst], auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1024
declare real_of_nat_gt_zero_cancel_iff [simp]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1025
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1026
lemma real_of_nat_le_zero_cancel_iff: "(real (n::nat) <= 0) = (n = 0)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1027
apply (rule real_of_nat_zero [THEN subst])
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1028
apply (subst real_of_nat_le_iff, auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1029
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1030
declare real_of_nat_le_zero_cancel_iff [simp]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1031
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1032
lemma not_real_of_nat_less_zero: "~ real (n::nat) < 0"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1033
apply (simp (no_asm) add: real_le_def [symmetric] real_of_nat_ge_zero)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1034
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1035
declare not_real_of_nat_less_zero [simp]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1036
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1037
lemma real_of_nat_ge_zero_cancel_iff [simp]: 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1038
      "(0 <= real (n::nat)) = (0 <= n)"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1039
apply (simp add: real_le_def le_def)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1040
done
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1041
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1042
lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y"
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1043
proof -
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1044
  have "0 + 0 \<le> x*x + y*y" by (blast intro: add_mono zero_le_square)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1045
  thus ?thesis by simp
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1046
qed
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1047
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1048
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1049
ML
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1050
{*
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1051
val real_abs_def = thm "real_abs_def";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1052
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1053
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
  1054
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
  1055
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
  1056
val real_of_nat_def = thm "real_of_nat_def";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1057
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1058
val preal_trans_lemma = thm"preal_trans_lemma";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1059
val realrel_iff = thm"realrel_iff";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1060
val realrel_refl = thm"realrel_refl";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1061
val equiv_realrel = thm"equiv_realrel";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1062
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
  1063
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
  1064
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
  1065
val eq_realrelD = thm"eq_realrelD";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1066
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
  1067
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
  1068
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
  1069
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
  1070
val real_minus = thm"real_minus";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1071
val real_add = thm"real_add";
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
  1072
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
  1073
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
  1074
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
  1075
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
  1076
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1077
val real_less_eq_diff = thm "real_less_eq_diff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1078
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1079
val real_mult = thm"real_mult";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1080
val real_mult_commute = thm"real_mult_commute";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1081
val real_mult_assoc = thm"real_mult_assoc";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1082
val real_mult_1 = thm"real_mult_1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1083
val real_mult_1_right = thm"real_mult_1_right";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1084
val real_minus_mult_commute = thm"real_minus_mult_commute";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1085
val preal_le_linear = thm"preal_le_linear";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1086
val real_mult_inv_left = thm"real_mult_inv_left";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1087
val real_less_not_refl = thm"real_less_not_refl";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1088
val real_less_irrefl = thm"real_less_irrefl";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1089
val real_not_refl2 = thm"real_not_refl2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1090
val preal_lemma_trans = thm"preal_lemma_trans";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1091
val real_less_trans = thm"real_less_trans";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1092
val real_less_not_sym = thm"real_less_not_sym";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1093
val real_less_asym = thm"real_less_asym";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1094
val real_of_preal_add = thm"real_of_preal_add";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1095
val real_of_preal_mult = thm"real_of_preal_mult";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1096
val real_of_preal_ExI = thm"real_of_preal_ExI";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1097
val real_of_preal_ExD = thm"real_of_preal_ExD";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1098
val real_of_preal_iff = thm"real_of_preal_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1099
val real_of_preal_trichotomy = thm"real_of_preal_trichotomy";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1100
val real_of_preal_trichotomyE = thm"real_of_preal_trichotomyE";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1101
val real_of_preal_lessD = thm"real_of_preal_lessD";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1102
val real_of_preal_lessI = thm"real_of_preal_lessI";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1103
val real_of_preal_less_iff1 = thm"real_of_preal_less_iff1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1104
val real_of_preal_minus_less_self = thm"real_of_preal_minus_less_self";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1105
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
  1106
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
  1107
val real_of_preal_zero_less = thm"real_of_preal_zero_less";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1108
val real_of_preal_not_less_zero = thm"real_of_preal_not_less_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1109
val real_minus_minus_zero_less = thm"real_minus_minus_zero_less";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1110
val real_of_preal_sum_zero_less = thm"real_of_preal_sum_zero_less";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1111
val real_of_preal_minus_less_all = thm"real_of_preal_minus_less_all";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1112
val real_of_preal_not_minus_gt_all = thm"real_of_preal_not_minus_gt_all";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1113
val real_of_preal_minus_less_rev1 = thm"real_of_preal_minus_less_rev1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1114
val real_of_preal_minus_less_rev2 = thm"real_of_preal_minus_less_rev2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1115
val real_linear = thm"real_linear";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1116
val real_neq_iff = thm"real_neq_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1117
val real_linear_less2 = thm"real_linear_less2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1118
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
  1119
val real_less_or_eq_imp_le = thm"real_less_or_eq_imp_le";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1120
val real_le_less = thm"real_le_less";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1121
val real_le_refl = thm"real_le_refl";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1122
val real_le_linear = thm"real_le_linear";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1123
val real_le_trans = thm"real_le_trans";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1124
val real_le_anti_sym = thm"real_le_anti_sym";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1125
val real_less_le = thm"real_less_le";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1126
val real_less_sum_gt_zero = thm"real_less_sum_gt_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1127
val real_sum_gt_zero_less = thm"real_sum_gt_zero_less";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1128
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1129
val real_gt_zero_preal_Ex = thm "real_gt_zero_preal_Ex";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1130
val real_gt_preal_preal_Ex = thm "real_gt_preal_preal_Ex";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1131
val real_ge_preal_preal_Ex = thm "real_ge_preal_preal_Ex";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1132
val real_less_all_preal = thm "real_less_all_preal";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1133
val real_less_all_real2 = thm "real_less_all_real2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1134
val real_of_preal_le_iff = thm "real_of_preal_le_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1135
val real_mult_order = thm "real_mult_order";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1136
val real_zero_less_one = thm "real_zero_less_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1137
val real_add_less_le_mono = thm "real_add_less_le_mono";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1138
val real_add_le_less_mono = thm "real_add_le_less_mono";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1139
val real_add_order = thm "real_add_order";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1140
val real_le_add_order = thm "real_le_add_order";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1141
val real_le_square = thm "real_le_square";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1142
val real_mult_less_mono2 = thm "real_mult_less_mono2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1143
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1144
val real_mult_less_iff1 = thm "real_mult_less_iff1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1145
val real_mult_le_cancel_iff1 = thm "real_mult_le_cancel_iff1";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1146
val real_mult_le_cancel_iff2 = thm "real_mult_le_cancel_iff2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1147
val real_mult_less_mono = thm "real_mult_less_mono";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1148
val real_mult_less_mono' = thm "real_mult_less_mono'";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1149
val real_sum_squares_cancel = thm "real_sum_squares_cancel";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1150
val real_sum_squares_cancel2 = thm "real_sum_squares_cancel2";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1151
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1152
val real_mult_left_cancel = thm"real_mult_left_cancel";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1153
val real_mult_right_cancel = thm"real_mult_right_cancel";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1154
val real_of_posnat_one = thm "real_of_posnat_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1155
val real_of_posnat_two = thm "real_of_posnat_two";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1156
val real_of_posnat_add = thm "real_of_posnat_add";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1157
val real_of_posnat_add_one = thm "real_of_posnat_add_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1158
val real_of_nat_zero = thm "real_of_nat_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1159
val real_of_nat_one = thm "real_of_nat_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1160
val real_of_nat_add = thm "real_of_nat_add";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1161
val real_of_nat_Suc = thm "real_of_nat_Suc";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1162
val real_of_nat_less_iff = thm "real_of_nat_less_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1163
val real_of_nat_le_iff = thm "real_of_nat_le_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1164
val inj_real_of_nat = thm "inj_real_of_nat";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1165
val real_of_nat_ge_zero = thm "real_of_nat_ge_zero";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1166
val real_of_nat_mult = thm "real_of_nat_mult";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1167
val real_of_nat_inject = thm "real_of_nat_inject";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1168
val real_of_nat_diff = thm "real_of_nat_diff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1169
val real_of_nat_zero_iff = thm "real_of_nat_zero_iff";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1170
val real_of_nat_neg_int = thm "real_of_nat_neg_int";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1171
val real_inverse_unique = thm "real_inverse_unique";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1172
val real_inverse_gt_one = thm "real_inverse_gt_one";
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1173
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
  1174
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
  1175
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
  1176
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
  1177
*}
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
  1178
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
  1179
end