| author | wenzelm |
| Tue, 16 Oct 2001 17:56:12 +0200 | |
| changeset 11808 | c724a9093ebe |
| parent 11701 | 3d51fbf81c17 |
| child 11913 | 673d7bc6b9db |
| permissions | -rw-r--r-- |
| 10681 | 1 |
(* Title: HOL/Library/Rational_Numbers.thy |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
5 |
*) |
|
| 10614 | 6 |
|
7 |
header {*
|
|
8 |
\title{Rational numbers}
|
|
9 |
\author{Markus Wenzel}
|
|
10 |
*} |
|
11 |
||
12 |
theory Rational_Numbers = Quotient + Ring_and_Field: |
|
13 |
||
14 |
subsection {* Fractions *}
|
|
15 |
||
16 |
subsubsection {* The type of fractions *}
|
|
17 |
||
18 |
typedef fraction = "{(a, b) :: int \<times> int | a b. b \<noteq> 0}"
|
|
19 |
proof |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
20 |
show "(0, Numeral1) \<in> ?fraction" by simp |
| 10614 | 21 |
qed |
22 |
||
23 |
constdefs |
|
24 |
fract :: "int => int => fraction" |
|
25 |
"fract a b == Abs_fraction (a, b)" |
|
26 |
num :: "fraction => int" |
|
27 |
"num Q == fst (Rep_fraction Q)" |
|
28 |
den :: "fraction => int" |
|
29 |
"den Q == snd (Rep_fraction Q)" |
|
30 |
||
31 |
lemma fract_num [simp]: "b \<noteq> 0 ==> num (fract a b) = a" |
|
32 |
by (simp add: fract_def num_def fraction_def Abs_fraction_inverse) |
|
33 |
||
34 |
lemma fract_den [simp]: "b \<noteq> 0 ==> den (fract a b) = b" |
|
35 |
by (simp add: fract_def den_def fraction_def Abs_fraction_inverse) |
|
36 |
||
37 |
lemma fraction_cases [case_names fract, cases type: fraction]: |
|
38 |
"(!!a b. Q = fract a b ==> b \<noteq> 0 ==> C) ==> C" |
|
39 |
proof - |
|
40 |
assume r: "!!a b. Q = fract a b ==> b \<noteq> 0 ==> C" |
|
41 |
obtain a b where "Q = fract a b" and "b \<noteq> 0" |
|
42 |
by (cases Q) (auto simp add: fract_def fraction_def) |
|
43 |
thus C by (rule r) |
|
44 |
qed |
|
45 |
||
46 |
lemma fraction_induct [case_names fract, induct type: fraction]: |
|
47 |
"(!!a b. b \<noteq> 0 ==> P (fract a b)) ==> P Q" |
|
48 |
by (cases Q) simp |
|
49 |
||
50 |
||
51 |
subsubsection {* Equivalence of fractions *}
|
|
52 |
||
53 |
instance fraction :: eqv .. |
|
54 |
||
55 |
defs (overloaded) |
|
56 |
equiv_fraction_def: "Q \<sim> R == num Q * den R = num R * den Q" |
|
57 |
||
58 |
lemma equiv_fraction_iff: |
|
59 |
"b \<noteq> 0 ==> b' \<noteq> 0 ==> (fract a b \<sim> fract a' b') = (a * b' = a' * b)" |
|
60 |
by (simp add: equiv_fraction_def) |
|
61 |
||
62 |
lemma equiv_fractionI [intro]: |
|
63 |
"a * b' = a' * b ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> fract a b \<sim> fract a' b'" |
|
64 |
by (insert equiv_fraction_iff) blast |
|
65 |
||
66 |
lemma equiv_fractionD [dest]: |
|
67 |
"fract a b \<sim> fract a' b' ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> a * b' = a' * b" |
|
68 |
by (insert equiv_fraction_iff) blast |
|
69 |
||
70 |
instance fraction :: equiv |
|
71 |
proof |
|
72 |
fix Q R S :: fraction |
|
73 |
{
|
|
74 |
show "Q \<sim> Q" |
|
75 |
proof (induct Q) |
|
76 |
fix a b :: int |
|
77 |
assume "b \<noteq> 0" and "b \<noteq> 0" |
|
78 |
with refl show "fract a b \<sim> fract a b" .. |
|
79 |
qed |
|
80 |
next |
|
81 |
assume "Q \<sim> R" and "R \<sim> S" |
|
82 |
show "Q \<sim> S" |
|
83 |
proof (insert prems, induct Q, induct R, induct S) |
|
84 |
fix a b a' b' a'' b'' :: int |
|
85 |
assume b: "b \<noteq> 0" and b': "b' \<noteq> 0" and b'': "b'' \<noteq> 0" |
|
86 |
assume "fract a b \<sim> fract a' b'" hence eq1: "a * b' = a' * b" .. |
|
87 |
assume "fract a' b' \<sim> fract a'' b''" hence eq2: "a' * b'' = a'' * b'" .. |
|
88 |
have "a * b'' = a'' * b" |
|
89 |
proof cases |
|
90 |
assume "a' = 0" |
|
91 |
with b' eq1 eq2 have "a = 0 \<and> a'' = 0" by auto |
|
92 |
thus ?thesis by simp |
|
93 |
next |
|
94 |
assume a': "a' \<noteq> 0" |
|
95 |
from eq1 eq2 have "(a * b') * (a' * b'') = (a' * b) * (a'' * b')" by simp |
|
96 |
hence "(a * b'') * (a' * b') = (a'' * b) * (a' * b')" by (simp only: zmult_ac) |
|
97 |
with a' b' show ?thesis by simp |
|
98 |
qed |
|
99 |
thus "fract a b \<sim> fract a'' b''" .. |
|
100 |
qed |
|
101 |
next |
|
102 |
show "Q \<sim> R ==> R \<sim> Q" |
|
103 |
proof (induct Q, induct R) |
|
104 |
fix a b a' b' :: int |
|
105 |
assume b: "b \<noteq> 0" and b': "b' \<noteq> 0" |
|
106 |
assume "fract a b \<sim> fract a' b'" |
|
107 |
hence "a * b' = a' * b" .. |
|
108 |
hence "a' * b = a * b'" .. |
|
109 |
thus "fract a' b' \<sim> fract a b" .. |
|
110 |
qed |
|
111 |
} |
|
112 |
qed |
|
113 |
||
114 |
lemma eq_fraction_iff: |
|
115 |
"b \<noteq> 0 ==> b' \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>) = (a * b' = a' * b)" |
|
116 |
by (simp add: equiv_fraction_iff quot_equality) |
|
117 |
||
118 |
lemma eq_fractionI [intro]: |
|
119 |
"a * b' = a' * b ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" |
|
120 |
by (insert eq_fraction_iff) blast |
|
121 |
||
122 |
lemma eq_fractionD [dest]: |
|
123 |
"\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> a * b' = a' * b" |
|
124 |
by (insert eq_fraction_iff) blast |
|
125 |
||
126 |
||
127 |
subsubsection {* Operations on fractions *}
|
|
128 |
||
129 |
text {*
|
|
130 |
We define the basic arithmetic operations on fractions and |
|
131 |
demonstrate their ``well-definedness'', i.e.\ congruence with respect |
|
132 |
to equivalence of fractions. |
|
133 |
*} |
|
134 |
||
135 |
instance fraction :: zero .. |
|
136 |
instance fraction :: plus .. |
|
137 |
instance fraction :: minus .. |
|
138 |
instance fraction :: times .. |
|
139 |
instance fraction :: inverse .. |
|
140 |
instance fraction :: ord .. |
|
141 |
||
142 |
defs (overloaded) |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
143 |
zero_fraction_def: "0 == fract 0 Numeral1" |
| 10614 | 144 |
add_fraction_def: "Q + R == |
145 |
fract (num Q * den R + num R * den Q) (den Q * den R)" |
|
146 |
minus_fraction_def: "-Q == fract (-(num Q)) (den Q)" |
|
147 |
mult_fraction_def: "Q * R == fract (num Q * num R) (den Q * den R)" |
|
148 |
inverse_fraction_def: "inverse Q == fract (den Q) (num Q)" |
|
149 |
le_fraction_def: "Q \<le> R == |
|
150 |
(num Q * den R) * (den Q * den R) \<le> (num R * den Q) * (den Q * den R)" |
|
151 |
||
152 |
lemma is_zero_fraction_iff: "b \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>0\<rfloor>) = (a = 0)" |
|
153 |
by (simp add: zero_fraction_def eq_fraction_iff) |
|
154 |
||
155 |
theorem add_fraction_cong: |
|
156 |
"\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> |
|
157 |
==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 |
|
158 |
==> \<lfloor>fract a b + fract c d\<rfloor> = \<lfloor>fract a' b' + fract c' d'\<rfloor>" |
|
159 |
proof - |
|
160 |
assume neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0" |
|
161 |
assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" .. |
|
162 |
assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" .. |
|
163 |
have "\<lfloor>fract (a * d + c * b) (b * d)\<rfloor> = \<lfloor>fract (a' * d' + c' * b') (b' * d')\<rfloor>" |
|
164 |
proof |
|
165 |
show "(a * d + c * b) * (b' * d') = (a' * d' + c' * b') * (b * d)" |
|
166 |
(is "?lhs = ?rhs") |
|
167 |
proof - |
|
168 |
have "?lhs = (a * b') * (d * d') + (c * d') * (b * b')" |
|
169 |
by (simp add: int_distrib zmult_ac) |
|
170 |
also have "... = (a' * b) * (d * d') + (c' * d) * (b * b')" |
|
171 |
by (simp only: eq1 eq2) |
|
172 |
also have "... = ?rhs" |
|
173 |
by (simp add: int_distrib zmult_ac) |
|
174 |
finally show "?lhs = ?rhs" . |
|
175 |
qed |
|
176 |
from neq show "b * d \<noteq> 0" by simp |
|
177 |
from neq show "b' * d' \<noteq> 0" by simp |
|
178 |
qed |
|
179 |
with neq show ?thesis by (simp add: add_fraction_def) |
|
180 |
qed |
|
181 |
||
182 |
theorem minus_fraction_cong: |
|
183 |
"\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> b \<noteq> 0 ==> b' \<noteq> 0 |
|
184 |
==> \<lfloor>-(fract a b)\<rfloor> = \<lfloor>-(fract a' b')\<rfloor>" |
|
185 |
proof - |
|
186 |
assume neq: "b \<noteq> 0" "b' \<noteq> 0" |
|
187 |
assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" |
|
188 |
hence "a * b' = a' * b" .. |
|
189 |
hence "-a * b' = -a' * b" by simp |
|
190 |
hence "\<lfloor>fract (-a) b\<rfloor> = \<lfloor>fract (-a') b'\<rfloor>" .. |
|
191 |
with neq show ?thesis by (simp add: minus_fraction_def) |
|
192 |
qed |
|
193 |
||
194 |
theorem mult_fraction_cong: |
|
195 |
"\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> |
|
196 |
==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 |
|
197 |
==> \<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>" |
|
198 |
proof - |
|
199 |
assume neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0" |
|
200 |
assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" .. |
|
201 |
assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" .. |
|
202 |
have "\<lfloor>fract (a * c) (b * d)\<rfloor> = \<lfloor>fract (a' * c') (b' * d')\<rfloor>" |
|
203 |
proof |
|
204 |
from eq1 eq2 have "(a * b') * (c * d') = (a' * b) * (c' * d)" by simp |
|
205 |
thus "(a * c) * (b' * d') = (a' * c') * (b * d)" by (simp add: zmult_ac) |
|
206 |
from neq show "b * d \<noteq> 0" by simp |
|
207 |
from neq show "b' * d' \<noteq> 0" by simp |
|
208 |
qed |
|
209 |
with neq show "\<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>" |
|
210 |
by (simp add: mult_fraction_def) |
|
211 |
qed |
|
212 |
||
213 |
theorem inverse_fraction_cong: |
|
214 |
"\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor> ==> \<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor> |
|
215 |
==> b \<noteq> 0 ==> b' \<noteq> 0 |
|
216 |
==> \<lfloor>inverse (fract a b)\<rfloor> = \<lfloor>inverse (fract a' b')\<rfloor>" |
|
217 |
proof - |
|
218 |
assume neq: "b \<noteq> 0" "b' \<noteq> 0" |
|
219 |
assume "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>" and "\<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>" |
|
220 |
with neq obtain "a \<noteq> 0" and "a' \<noteq> 0" by (simp add: is_zero_fraction_iff) |
|
221 |
assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" |
|
222 |
hence "a * b' = a' * b" .. |
|
223 |
hence "b * a' = b' * a" by (simp only: zmult_ac) |
|
224 |
hence "\<lfloor>fract b a\<rfloor> = \<lfloor>fract b' a'\<rfloor>" .. |
|
225 |
with neq show ?thesis by (simp add: inverse_fraction_def) |
|
226 |
qed |
|
227 |
||
228 |
theorem le_fraction_cong: |
|
229 |
"\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> |
|
230 |
==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 |
|
231 |
==> (fract a b \<le> fract c d) = (fract a' b' \<le> fract c' d')" |
|
232 |
proof - |
|
233 |
assume neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0" |
|
234 |
assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" .. |
|
235 |
assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" .. |
|
236 |
||
237 |
let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))" |
|
238 |
{
|
|
239 |
fix a b c d x :: int assume x: "x \<noteq> 0" |
|
240 |
have "?le a b c d = ?le (a * x) (b * x) c d" |
|
241 |
proof - |
|
242 |
from x have "0 < x * x" by (auto simp add: int_less_le) |
|
243 |
hence "?le a b c d = |
|
244 |
((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))" |
|
245 |
by (simp add: zmult_zle_cancel2) |
|
246 |
also have "... = ?le (a * x) (b * x) c d" |
|
247 |
by (simp add: zmult_ac) |
|
248 |
finally show ?thesis . |
|
249 |
qed |
|
250 |
} note le_factor = this |
|
251 |
||
252 |
let ?D = "b * d" and ?D' = "b' * d'" |
|
253 |
from neq have D: "?D \<noteq> 0" by simp |
|
254 |
from neq have "?D' \<noteq> 0" by simp |
|
255 |
hence "?le a b c d = ?le (a * ?D') (b * ?D') c d" |
|
256 |
by (rule le_factor) |
|
257 |
also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')" |
|
258 |
by (simp add: zmult_ac) |
|
259 |
also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')" |
|
260 |
by (simp only: eq1 eq2) |
|
261 |
also have "... = ?le (a' * ?D) (b' * ?D) c' d'" |
|
262 |
by (simp add: zmult_ac) |
|
263 |
also from D have "... = ?le a' b' c' d'" |
|
264 |
by (rule le_factor [symmetric]) |
|
265 |
finally have "?le a b c d = ?le a' b' c' d'" . |
|
266 |
with neq show ?thesis by (simp add: le_fraction_def) |
|
267 |
qed |
|
268 |
||
269 |
||
270 |
subsection {* Rational numbers *}
|
|
271 |
||
272 |
subsubsection {* The type of rational numbers *}
|
|
273 |
||
274 |
typedef (Rat) |
|
275 |
rat = "UNIV :: fraction quot set" .. |
|
276 |
||
277 |
lemma RatI [intro, simp]: "Q \<in> Rat" |
|
278 |
by (simp add: Rat_def) |
|
279 |
||
280 |
constdefs |
|
281 |
fraction_of :: "rat => fraction" |
|
282 |
"fraction_of q == pick (Rep_Rat q)" |
|
283 |
rat_of :: "fraction => rat" |
|
284 |
"rat_of Q == Abs_Rat \<lfloor>Q\<rfloor>" |
|
285 |
||
286 |
theorem rat_of_equality [iff?]: "(rat_of Q = rat_of Q') = (\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor>)" |
|
287 |
by (simp add: rat_of_def Abs_Rat_inject) |
|
288 |
||
289 |
lemma rat_of: "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> rat_of Q = rat_of Q'" .. |
|
290 |
||
291 |
constdefs |
|
292 |
Fract :: "int => int => rat" |
|
293 |
"Fract a b == rat_of (fract a b)" |
|
294 |
||
295 |
theorem Fract_inverse: "\<lfloor>fraction_of (Fract a b)\<rfloor> = \<lfloor>fract a b\<rfloor>" |
|
296 |
by (simp add: fraction_of_def rat_of_def Fract_def Abs_Rat_inverse pick_inverse) |
|
297 |
||
298 |
theorem Fract_equality [iff?]: |
|
299 |
"(Fract a b = Fract c d) = (\<lfloor>fract a b\<rfloor> = \<lfloor>fract c d\<rfloor>)" |
|
300 |
by (simp add: Fract_def rat_of_equality) |
|
301 |
||
302 |
theorem eq_rat: |
|
303 |
"b \<noteq> 0 ==> d \<noteq> 0 ==> (Fract a b = Fract c d) = (a * d = c * b)" |
|
304 |
by (simp add: Fract_equality eq_fraction_iff) |
|
305 |
||
306 |
theorem Rat_cases [case_names Fract, cases type: rat]: |
|
307 |
"(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C" |
|
308 |
proof - |
|
309 |
assume r: "!!a b. q = Fract a b ==> b \<noteq> 0 ==> C" |
|
310 |
obtain x where "q = Abs_Rat x" by (cases q) |
|
311 |
moreover obtain Q where "x = \<lfloor>Q\<rfloor>" by (cases x) |
|
312 |
moreover obtain a b where "Q = fract a b" and "b \<noteq> 0" by (cases Q) |
|
313 |
ultimately have "q = Fract a b" by (simp only: Fract_def rat_of_def) |
|
314 |
thus ?thesis by (rule r) |
|
315 |
qed |
|
316 |
||
317 |
theorem Rat_induct [case_names Fract, induct type: rat]: |
|
318 |
"(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q" |
|
319 |
by (cases q) simp |
|
320 |
||
321 |
||
322 |
subsubsection {* Canonical function definitions *}
|
|
323 |
||
324 |
text {*
|
|
325 |
Note that the unconditional version below is much easier to read. |
|
326 |
*} |
|
327 |
||
328 |
theorem rat_cond_function: |
|
329 |
"(!!q r. P \<lfloor>fraction_of q\<rfloor> \<lfloor>fraction_of r\<rfloor> ==> |
|
330 |
f q r == g (fraction_of q) (fraction_of r)) ==> |
|
331 |
(!!a b a' b' c d c' d'. |
|
332 |
\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==> |
|
333 |
P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==> P \<lfloor>fract a' b'\<rfloor> \<lfloor>fract c' d'\<rfloor> ==> |
|
334 |
b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==> |
|
335 |
g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==> |
|
336 |
P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==> |
|
337 |
f (Fract a b) (Fract c d) = g (fract a b) (fract c d)" |
|
338 |
(is "PROP ?eq ==> PROP ?cong ==> ?P ==> _") |
|
339 |
proof - |
|
340 |
assume eq: "PROP ?eq" and cong: "PROP ?cong" and P: ?P |
|
341 |
have "f (Abs_Rat \<lfloor>fract a b\<rfloor>) (Abs_Rat \<lfloor>fract c d\<rfloor>) = g (fract a b) (fract c d)" |
|
342 |
proof (rule quot_cond_function) |
|
343 |
fix X Y assume "P X Y" |
|
344 |
with eq show "f (Abs_Rat X) (Abs_Rat Y) == g (pick X) (pick Y)" |
|
345 |
by (simp add: fraction_of_def pick_inverse Abs_Rat_inverse) |
|
346 |
next |
|
347 |
fix Q Q' R R' :: fraction |
|
348 |
show "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> \<lfloor>R\<rfloor> = \<lfloor>R'\<rfloor> ==> |
|
349 |
P \<lfloor>Q\<rfloor> \<lfloor>R\<rfloor> ==> P \<lfloor>Q'\<rfloor> \<lfloor>R'\<rfloor> ==> g Q R = g Q' R'" |
|
350 |
by (induct Q, induct Q', induct R, induct R') (rule cong) |
|
351 |
qed |
|
352 |
thus ?thesis by (unfold Fract_def rat_of_def) |
|
353 |
qed |
|
354 |
||
355 |
theorem rat_function: |
|
356 |
"(!!q r. f q r == g (fraction_of q) (fraction_of r)) ==> |
|
357 |
(!!a b a' b' c d c' d'. |
|
358 |
\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==> |
|
359 |
b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==> |
|
360 |
g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==> |
|
361 |
f (Fract a b) (Fract c d) = g (fract a b) (fract c d)" |
|
362 |
proof - |
|
| 11549 | 363 |
case rule_context from this TrueI |
| 10614 | 364 |
show ?thesis by (rule rat_cond_function) |
365 |
qed |
|
366 |
||
367 |
||
368 |
subsubsection {* Standard operations on rational numbers *}
|
|
369 |
||
370 |
instance rat :: zero .. |
|
371 |
instance rat :: plus .. |
|
372 |
instance rat :: minus .. |
|
373 |
instance rat :: times .. |
|
374 |
instance rat :: inverse .. |
|
375 |
instance rat :: ord .. |
|
376 |
instance rat :: number .. |
|
377 |
||
378 |
defs (overloaded) |
|
379 |
zero_rat_def: "0 == rat_of 0" |
|
380 |
add_rat_def: "q + r == rat_of (fraction_of q + fraction_of r)" |
|
381 |
minus_rat_def: "-q == rat_of (-(fraction_of q))" |
|
382 |
diff_rat_def: "q - r == q + (-(r::rat))" |
|
383 |
mult_rat_def: "q * r == rat_of (fraction_of q * fraction_of r)" |
|
384 |
inverse_rat_def: "q \<noteq> 0 ==> inverse q == rat_of (inverse (fraction_of q))" |
|
385 |
divide_rat_def: "r \<noteq> 0 ==> q / r == q * inverse (r::rat)" |
|
386 |
le_rat_def: "q \<le> r == fraction_of q \<le> fraction_of r" |
|
387 |
less_rat_def: "q < r == q \<le> r \<and> q \<noteq> (r::rat)" |
|
388 |
abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)" |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
389 |
number_of_rat_def: "number_of b == Fract (number_of b) Numeral1" |
| 10614 | 390 |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
391 |
theorem zero_rat: "0 = Fract 0 Numeral1" |
| 10614 | 392 |
by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def) |
393 |
||
394 |
theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> |
|
395 |
Fract a b + Fract c d = Fract (a * d + c * b) (b * d)" |
|
396 |
proof - |
|
397 |
have "Fract a b + Fract c d = rat_of (fract a b + fract c d)" |
|
398 |
by (rule rat_function, rule add_rat_def, rule rat_of, rule add_fraction_cong) |
|
399 |
also |
|
400 |
assume "b \<noteq> 0" "d \<noteq> 0" |
|
401 |
hence "fract a b + fract c d = fract (a * d + c * b) (b * d)" |
|
402 |
by (simp add: add_fraction_def) |
|
403 |
finally show ?thesis by (unfold Fract_def) |
|
404 |
qed |
|
405 |
||
406 |
theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b" |
|
407 |
proof - |
|
408 |
have "-(Fract a b) = rat_of (-(fract a b))" |
|
409 |
by (rule rat_function, rule minus_rat_def, rule rat_of, rule minus_fraction_cong) |
|
410 |
also assume "b \<noteq> 0" hence "-(fract a b) = fract (-a) b" |
|
411 |
by (simp add: minus_fraction_def) |
|
412 |
finally show ?thesis by (unfold Fract_def) |
|
413 |
qed |
|
414 |
||
415 |
theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> |
|
416 |
Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" |
|
417 |
by (simp add: diff_rat_def add_rat minus_rat) |
|
418 |
||
419 |
theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> |
|
420 |
Fract a b * Fract c d = Fract (a * c) (b * d)" |
|
421 |
proof - |
|
422 |
have "Fract a b * Fract c d = rat_of (fract a b * fract c d)" |
|
423 |
by (rule rat_function, rule mult_rat_def, rule rat_of, rule mult_fraction_cong) |
|
424 |
also |
|
425 |
assume "b \<noteq> 0" "d \<noteq> 0" |
|
426 |
hence "fract a b * fract c d = fract (a * c) (b * d)" |
|
427 |
by (simp add: mult_fraction_def) |
|
428 |
finally show ?thesis by (unfold Fract_def) |
|
429 |
qed |
|
430 |
||
431 |
theorem inverse_rat: "Fract a b \<noteq> 0 ==> b \<noteq> 0 ==> |
|
432 |
inverse (Fract a b) = Fract b a" |
|
433 |
proof - |
|
434 |
assume neq: "b \<noteq> 0" and nonzero: "Fract a b \<noteq> 0" |
|
435 |
hence "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>" |
|
436 |
by (simp add: zero_rat eq_rat is_zero_fraction_iff) |
|
437 |
with _ inverse_fraction_cong [THEN rat_of] |
|
438 |
have "inverse (Fract a b) = rat_of (inverse (fract a b))" |
|
439 |
proof (rule rat_cond_function) |
|
440 |
fix q assume cond: "\<lfloor>fraction_of q\<rfloor> \<noteq> \<lfloor>0\<rfloor>" |
|
441 |
have "q \<noteq> 0" |
|
442 |
proof (cases q) |
|
443 |
fix a b assume "b \<noteq> 0" and "q = Fract a b" |
|
444 |
from this cond show ?thesis |
|
445 |
by (simp add: Fract_inverse is_zero_fraction_iff zero_rat eq_rat) |
|
446 |
qed |
|
447 |
thus "inverse q == rat_of (inverse (fraction_of q))" |
|
448 |
by (rule inverse_rat_def) |
|
449 |
qed |
|
450 |
also from neq nonzero have "inverse (fract a b) = fract b a" |
|
451 |
by (simp add: inverse_fraction_def) |
|
452 |
finally show ?thesis by (unfold Fract_def) |
|
453 |
qed |
|
454 |
||
455 |
theorem divide_rat: "Fract c d \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==> |
|
456 |
Fract a b / Fract c d = Fract (a * d) (b * c)" |
|
457 |
proof - |
|
458 |
assume neq: "b \<noteq> 0" "d \<noteq> 0" and nonzero: "Fract c d \<noteq> 0" |
|
459 |
hence "c \<noteq> 0" by (simp add: zero_rat eq_rat) |
|
460 |
with neq nonzero show ?thesis |
|
461 |
by (simp add: divide_rat_def inverse_rat mult_rat) |
|
462 |
qed |
|
463 |
||
464 |
theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> |
|
465 |
(Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))" |
|
466 |
proof - |
|
467 |
have "(Fract a b \<le> Fract c d) = (fract a b \<le> fract c d)" |
|
468 |
by (rule rat_function, rule le_rat_def, rule le_fraction_cong) |
|
469 |
also |
|
470 |
assume "b \<noteq> 0" "d \<noteq> 0" |
|
471 |
hence "(fract a b \<le> fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))" |
|
472 |
by (simp add: le_fraction_def) |
|
473 |
finally show ?thesis . |
|
474 |
qed |
|
475 |
||
476 |
theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> |
|
477 |
(Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))" |
|
478 |
by (simp add: less_rat_def le_rat eq_rat int_less_le) |
|
479 |
||
480 |
theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>" |
|
481 |
by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat) |
|
482 |
(auto simp add: zmult_less_0_iff int_0_less_mult_iff int_le_less split: zabs_split) |
|
483 |
||
484 |
||
485 |
subsubsection {* The ordered field of rational numbers *}
|
|
486 |
||
487 |
instance rat :: field |
|
488 |
proof |
|
489 |
fix q r s :: rat |
|
490 |
show "(q + r) + s = q + (r + s)" |
|
491 |
by (induct q, induct r, induct s) (simp add: add_rat zadd_ac zmult_ac int_distrib) |
|
492 |
show "q + r = r + q" |
|
493 |
by (induct q, induct r) (simp add: add_rat zadd_ac zmult_ac) |
|
494 |
show "0 + q = q" |
|
495 |
by (induct q) (simp add: zero_rat add_rat) |
|
| 10621 | 496 |
show "(-q) + q = 0" |
| 10614 | 497 |
by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat) |
498 |
show "q - r = q + (-r)" |
|
499 |
by (induct q, induct r) (simp add: add_rat minus_rat diff_rat) |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
500 |
show "(0::rat) = Numeral0" |
| 10614 | 501 |
by (simp add: zero_rat number_of_rat_def) |
502 |
show "(q * r) * s = q * (r * s)" |
|
503 |
by (induct q, induct r, induct s) (simp add: mult_rat zmult_ac) |
|
504 |
show "q * r = r * q" |
|
505 |
by (induct q, induct r) (simp add: mult_rat zmult_ac) |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
506 |
show "Numeral1 * q = q" |
| 10614 | 507 |
by (induct q) (simp add: number_of_rat_def mult_rat) |
508 |
show "(q + r) * s = q * s + r * s" |
|
509 |
by (induct q, induct r, induct s) (simp add: add_rat mult_rat eq_rat int_distrib) |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
510 |
show "q \<noteq> 0 ==> inverse q * q = Numeral1" |
| 10614 | 511 |
by (induct q) (simp add: inverse_rat mult_rat number_of_rat_def zero_rat eq_rat) |
512 |
show "r \<noteq> 0 ==> q / r = q * inverse r" |
|
513 |
by (induct q, induct r) (simp add: mult_rat divide_rat inverse_rat zero_rat eq_rat) |
|
514 |
qed |
|
515 |
||
516 |
instance rat :: linorder |
|
517 |
proof |
|
518 |
fix q r s :: rat |
|
519 |
{
|
|
520 |
assume "q \<le> r" and "r \<le> s" |
|
521 |
show "q \<le> s" |
|
522 |
proof (insert prems, induct q, induct r, induct s) |
|
523 |
fix a b c d e f :: int |
|
524 |
assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" |
|
525 |
assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f" |
|
526 |
show "Fract a b \<le> Fract e f" |
|
527 |
proof - |
|
528 |
from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f" |
|
529 |
by (auto simp add: int_less_le) |
|
530 |
have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)" |
|
531 |
proof - |
|
532 |
from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" |
|
533 |
by (simp add: le_rat) |
|
534 |
with ff show ?thesis by (simp add: zmult_zle_cancel2) |
|
535 |
qed |
|
536 |
also have "... = (c * f) * (d * f) * (b * b)" |
|
537 |
by (simp only: zmult_ac) |
|
538 |
also have "... \<le> (e * d) * (d * f) * (b * b)" |
|
539 |
proof - |
|
540 |
from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)" |
|
541 |
by (simp add: le_rat) |
|
542 |
with bb show ?thesis by (simp add: zmult_zle_cancel2) |
|
543 |
qed |
|
544 |
finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)" |
|
545 |
by (simp only: zmult_ac) |
|
546 |
with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)" |
|
547 |
by (simp add: zmult_zle_cancel2) |
|
548 |
with neq show ?thesis by (simp add: le_rat) |
|
549 |
qed |
|
550 |
qed |
|
551 |
next |
|
552 |
assume "q \<le> r" and "r \<le> q" |
|
553 |
show "q = r" |
|
554 |
proof (insert prems, induct q, induct r) |
|
555 |
fix a b c d :: int |
|
556 |
assume neq: "b \<noteq> 0" "d \<noteq> 0" |
|
557 |
assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b" |
|
558 |
show "Fract a b = Fract c d" |
|
559 |
proof - |
|
560 |
from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" |
|
561 |
by (simp add: le_rat) |
|
562 |
also have "... \<le> (a * d) * (b * d)" |
|
563 |
proof - |
|
564 |
from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)" |
|
565 |
by (simp add: le_rat) |
|
566 |
thus ?thesis by (simp only: zmult_ac) |
|
567 |
qed |
|
568 |
finally have "(a * d) * (b * d) = (c * b) * (b * d)" . |
|
569 |
moreover from neq have "b * d \<noteq> 0" by simp |
|
570 |
ultimately have "a * d = c * b" by simp |
|
571 |
with neq show ?thesis by (simp add: eq_rat) |
|
572 |
qed |
|
573 |
qed |
|
574 |
next |
|
575 |
show "q \<le> q" |
|
576 |
by (induct q) (simp add: le_rat) |
|
577 |
show "(q < r) = (q \<le> r \<and> q \<noteq> r)" |
|
578 |
by (simp only: less_rat_def) |
|
579 |
show "q \<le> r \<or> r \<le> q" |
|
580 |
by (induct q, induct r) (simp add: le_rat zmult_ac, arith) |
|
581 |
} |
|
582 |
qed |
|
583 |
||
584 |
instance rat :: ordered_field |
|
585 |
proof |
|
586 |
fix q r s :: rat |
|
587 |
show "q \<le> r ==> s + q \<le> s + r" |
|
588 |
proof (induct q, induct r, induct s) |
|
589 |
fix a b c d e f :: int |
|
590 |
assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" |
|
591 |
assume le: "Fract a b \<le> Fract c d" |
|
592 |
show "Fract e f + Fract a b \<le> Fract e f + Fract c d" |
|
593 |
proof - |
|
594 |
let ?F = "f * f" from neq have F: "0 < ?F" |
|
595 |
by (auto simp add: int_less_le) |
|
596 |
from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)" |
|
597 |
by (simp add: le_rat) |
|
598 |
with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F" |
|
599 |
by (simp add: zmult_zle_cancel2) |
|
600 |
with neq show ?thesis by (simp add: add_rat le_rat zmult_ac int_distrib) |
|
601 |
qed |
|
602 |
qed |
|
603 |
show "q < r ==> 0 < s ==> s * q < s * r" |
|
604 |
proof (induct q, induct r, induct s) |
|
605 |
fix a b c d e f :: int |
|
606 |
assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" |
|
607 |
assume le: "Fract a b < Fract c d" |
|
608 |
assume gt: "0 < Fract e f" |
|
609 |
show "Fract e f * Fract a b < Fract e f * Fract c d" |
|
610 |
proof - |
|
611 |
let ?E = "e * f" and ?F = "f * f" |
|
612 |
from neq gt have "0 < ?E" |
|
613 |
by (auto simp add: zero_rat less_rat le_rat int_less_le eq_rat) |
|
614 |
moreover from neq have "0 < ?F" |
|
615 |
by (auto simp add: int_less_le) |
|
616 |
moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)" |
|
617 |
by (simp add: less_rat) |
|
618 |
ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F" |
|
619 |
by (simp add: zmult_zless_cancel2) |
|
620 |
with neq show ?thesis |
|
621 |
by (simp add: less_rat mult_rat zmult_ac) |
|
622 |
qed |
|
623 |
qed |
|
624 |
show "\<bar>q\<bar> = (if q < 0 then -q else q)" |
|
625 |
by (simp only: abs_rat_def) |
|
626 |
qed |
|
627 |
||
628 |
||
629 |
subsection {* Embedding integers *}
|
|
630 |
||
| 10665 | 631 |
constdefs |
632 |
rat :: "int => rat" (* FIXME generalize int to any numeric subtype *) |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
633 |
"rat z == Fract z Numeral1" |
| 10665 | 634 |
int_set :: "rat set" ("\<int>") (* FIXME generalize rat to any numeric supertype *)
|
| 10614 | 635 |
"\<int> == range rat" |
636 |
||
637 |
lemma rat_inject: "(rat z = rat w) = (z = w)" |
|
638 |
proof |
|
639 |
assume "rat z = rat w" |
|
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
640 |
hence "Fract z Numeral1 = Fract w Numeral1" by (unfold rat_def) |
|
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11549
diff
changeset
|
641 |
hence "\<lfloor>fract z Numeral1\<rfloor> = \<lfloor>fract w Numeral1\<rfloor>" .. |
| 10614 | 642 |
thus "z = w" by auto |
643 |
next |
|
644 |
assume "z = w" |
|
645 |
thus "rat z = rat w" by simp |
|
646 |
qed |
|
647 |
||
648 |
lemma int_set_cases [case_names rat, cases set: int_set]: |
|
649 |
"q \<in> \<int> ==> (!!z. q = rat z ==> C) ==> C" |
|
650 |
proof (unfold int_set_def) |
|
651 |
assume "!!z. q = rat z ==> C" |
|
652 |
assume "q \<in> range rat" thus C .. |
|
653 |
qed |
|
654 |
||
655 |
lemma int_set_induct [case_names rat, induct set: int_set]: |
|
656 |
"q \<in> \<int> ==> (!!z. P (rat z)) ==> P q" |
|
657 |
by (rule int_set_cases) auto |
|
658 |
||
659 |
theorem number_of_rat: "number_of b = rat (number_of b)" |
|
660 |
by (simp only: number_of_rat_def rat_def) |
|
661 |
||
662 |
end |