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