| author | wenzelm | 
| Tue, 07 Nov 2006 11:46:49 +0100 | |
| changeset 21207 | cef082634be9 | 
| parent 20522 | 05072ae0d435 | 
| child 21404 | eb85850d3eb7 | 
| permissions | -rw-r--r-- | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 1 | (* Title: HOL/Library/Rational.thy | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 4 | *) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 5 | |
| 14691 | 6 | header {* Rational numbers *}
 | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 7 | |
| 15131 | 8 | theory Rational | 
| 18913 | 9 | imports Main | 
| 16417 | 10 | uses ("rat_arith.ML")
 | 
| 15131 | 11 | begin | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 12 | |
| 18913 | 13 | subsection {* Rational numbers *}
 | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 14 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 15 | subsubsection {* Equivalence of fractions *}
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 16 | |
| 19765 | 17 | definition | 
| 18913 | 18 | fraction :: "(int \<times> int) set" | 
| 19765 | 19 |   "fraction = {x. snd x \<noteq> 0}"
 | 
| 18913 | 20 | |
| 21 | ratrel :: "((int \<times> int) \<times> (int \<times> int)) set" | |
| 19765 | 22 |   "ratrel = {(x,y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
 | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 23 | |
| 18913 | 24 | lemma fraction_iff [simp]: "(x \<in> fraction) = (snd x \<noteq> 0)" | 
| 25 | by (simp add: fraction_def) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 26 | |
| 18913 | 27 | lemma ratrel_iff [simp]: | 
| 28 | "((x,y) \<in> ratrel) = | |
| 29 | (snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)" | |
| 30 | by (simp add: ratrel_def) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 31 | |
| 18913 | 32 | lemma refl_ratrel: "refl fraction ratrel" | 
| 33 | by (auto simp add: refl_def fraction_def ratrel_def) | |
| 34 | ||
| 35 | lemma sym_ratrel: "sym ratrel" | |
| 36 | by (simp add: ratrel_def sym_def) | |
| 37 | ||
| 38 | lemma trans_ratrel_lemma: | |
| 39 | assumes 1: "a * b' = a' * b" | |
| 40 | assumes 2: "a' * b'' = a'' * b'" | |
| 41 | assumes 3: "b' \<noteq> (0::int)" | |
| 42 | shows "a * b'' = a'' * b" | |
| 43 | proof - | |
| 44 | have "b' * (a * b'') = b'' * (a * b')" by simp | |
| 45 | also note 1 | |
| 46 | also have "b'' * (a' * b) = b * (a' * b'')" by simp | |
| 47 | also note 2 | |
| 48 | also have "b * (a'' * b') = b' * (a'' * b)" by simp | |
| 49 | finally have "b' * (a * b'') = b' * (a'' * b)" . | |
| 50 | with 3 show "a * b'' = a'' * b" by simp | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 51 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 52 | |
| 18913 | 53 | lemma trans_ratrel: "trans ratrel" | 
| 54 | by (auto simp add: trans_def elim: trans_ratrel_lemma) | |
| 55 | ||
| 56 | lemma equiv_ratrel: "equiv fraction ratrel" | |
| 57 | by (rule equiv.intro [OF refl_ratrel sym_ratrel trans_ratrel]) | |
| 58 | ||
| 59 | lemmas equiv_ratrel_iff [iff] = eq_equiv_class_iff [OF equiv_ratrel] | |
| 60 | ||
| 61 | lemma equiv_ratrel_iff2: | |
| 62 | "\<lbrakk>snd x \<noteq> 0; snd y \<noteq> 0\<rbrakk> | |
| 63 |     \<Longrightarrow> (ratrel `` {x} = ratrel `` {y}) = ((x,y) \<in> ratrel)"
 | |
| 64 | by (rule eq_equiv_class_iff [OF equiv_ratrel], simp_all) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 65 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 66 | |
| 18913 | 67 | subsubsection {* The type of rational numbers *}
 | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 68 | |
| 18913 | 69 | typedef (Rat) rat = "fraction//ratrel" | 
| 70 | proof | |
| 71 | have "(0,1) \<in> fraction" by (simp add: fraction_def) | |
| 72 |   thus "ratrel``{(0,1)} \<in> fraction//ratrel" by (rule quotientI)
 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 73 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 74 | |
| 18913 | 75 | lemma ratrel_in_Rat [simp]: "snd x \<noteq> 0 \<Longrightarrow> ratrel``{x} \<in> Rat"
 | 
| 76 | by (simp add: Rat_def quotientI) | |
| 77 | ||
| 78 | declare Abs_Rat_inject [simp] Abs_Rat_inverse [simp] | |
| 79 | ||
| 80 | ||
| 19765 | 81 | definition | 
| 18913 | 82 | Fract :: "int \<Rightarrow> int \<Rightarrow> rat" | 
| 19765 | 83 |   "Fract a b = Abs_Rat (ratrel``{(a,b)})"
 | 
| 18913 | 84 | |
| 85 | theorem Rat_cases [case_names Fract, cases type: rat]: | |
| 86 | "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C" | |
| 87 | by (cases q, clarsimp simp add: Fract_def Rat_def fraction_def quotient_def) | |
| 88 | ||
| 89 | theorem Rat_induct [case_names Fract, induct type: rat]: | |
| 90 | "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q" | |
| 91 | by (cases q) simp | |
| 92 | ||
| 93 | ||
| 94 | subsubsection {* Congruence lemmas *}
 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 95 | |
| 18913 | 96 | lemma add_congruent2: | 
| 97 |      "(\<lambda>x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
 | |
| 98 | respects2 ratrel" | |
| 99 | apply (rule equiv_ratrel [THEN congruent2_commuteI]) | |
| 100 | apply (simp_all add: left_distrib) | |
| 101 | done | |
| 102 | ||
| 103 | lemma minus_congruent: | |
| 104 |   "(\<lambda>x. ratrel``{(- fst x, snd x)}) respects ratrel"
 | |
| 105 | by (simp add: congruent_def) | |
| 106 | ||
| 107 | lemma mult_congruent2: | |
| 108 |   "(\<lambda>x y. ratrel``{(fst x * fst y, snd x * snd y)}) respects2 ratrel"
 | |
| 109 | by (rule equiv_ratrel [THEN congruent2_commuteI], simp_all) | |
| 110 | ||
| 111 | lemma inverse_congruent: | |
| 112 |   "(\<lambda>x. ratrel``{if fst x=0 then (0,1) else (snd x, fst x)}) respects ratrel"
 | |
| 113 | by (auto simp add: congruent_def mult_commute) | |
| 114 | ||
| 115 | lemma le_congruent2: | |
| 18982 | 116 |   "(\<lambda>x y. {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})
 | 
| 18913 | 117 | respects2 ratrel" | 
| 118 | proof (clarsimp simp add: congruent2_def) | |
| 119 | fix a b a' b' c d c' d'::int | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 120 | assume neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0" | 
| 18913 | 121 | assume eq1: "a * b' = a' * b" | 
| 122 | assume eq2: "c * d' = c' * d" | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 123 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 124 | let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 125 |   {
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 126 | fix a b c d x :: int assume x: "x \<noteq> 0" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 127 | have "?le a b c d = ?le (a * x) (b * x) c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 128 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 129 | from x have "0 < x * x" by (auto simp add: zero_less_mult_iff) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 130 | hence "?le a b c d = | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 131 | ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 132 | by (simp add: mult_le_cancel_right) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 133 | also have "... = ?le (a * x) (b * x) c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 134 | by (simp add: mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 135 | finally show ?thesis . | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 136 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 137 | } note le_factor = this | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 138 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 139 | let ?D = "b * d" and ?D' = "b' * d'" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 140 | from neq have D: "?D \<noteq> 0" by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 141 | from neq have "?D' \<noteq> 0" by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 142 | hence "?le a b c d = ?le (a * ?D') (b * ?D') c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 143 | by (rule le_factor) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 144 | also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 145 | by (simp add: mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 146 | also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 147 | by (simp only: eq1 eq2) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 148 | also have "... = ?le (a' * ?D) (b' * ?D) c' d'" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 149 | by (simp add: mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 150 | also from D have "... = ?le a' b' c' d'" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 151 | by (rule le_factor [symmetric]) | 
| 18913 | 152 | finally show "?le a b c d = ?le a' b' c' d'" . | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 153 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 154 | |
| 18913 | 155 | lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel] | 
| 156 | lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel] | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 157 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 158 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 159 | subsubsection {* Standard operations on rational numbers *}
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 160 | |
| 20522 | 161 | instance rat :: "{ord, zero, one, plus, times, minus, inverse, power}" ..
 | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 162 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 163 | defs (overloaded) | 
| 18913 | 164 | Zero_rat_def: "0 == Fract 0 1" | 
| 165 | One_rat_def: "1 == Fract 1 1" | |
| 166 | ||
| 167 | add_rat_def: | |
| 168 | "q + r == | |
| 169 | Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r. | |
| 170 |            ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})"
 | |
| 171 | ||
| 172 | minus_rat_def: | |
| 173 |     "- q == Abs_Rat (\<Union>x \<in> Rep_Rat q. ratrel``{(- fst x, snd x)})"
 | |
| 174 | ||
| 175 | diff_rat_def: "q - r == q + - (r::rat)" | |
| 176 | ||
| 177 | mult_rat_def: | |
| 178 | "q * r == | |
| 179 | Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r. | |
| 180 |            ratrel``{(fst x * fst y, snd x * snd y)})"
 | |
| 181 | ||
| 182 | inverse_rat_def: | |
| 183 | "inverse q == | |
| 184 | Abs_Rat (\<Union>x \<in> Rep_Rat q. | |
| 185 |             ratrel``{if fst x=0 then (0,1) else (snd x, fst x)})"
 | |
| 186 | ||
| 187 | divide_rat_def: "q / r == q * inverse (r::rat)" | |
| 188 | ||
| 189 | le_rat_def: | |
| 18982 | 190 | "q \<le> r == contents (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r. | 
| 191 |       {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})"
 | |
| 18913 | 192 | |
| 193 | less_rat_def: "(z < (w::rat)) == (z \<le> w & z \<noteq> w)" | |
| 194 | ||
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 195 | abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 196 | |
| 20522 | 197 | primrec (rat) | 
| 198 | rat_power_0: "q ^ 0 = 1" | |
| 199 | rat_power_Suc: "q ^ (Suc n) = (q::rat) * (q ^ n)" | |
| 200 | ||
| 18913 | 201 | lemma zero_rat: "0 = Fract 0 1" | 
| 202 | by (simp add: Zero_rat_def) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 203 | |
| 18913 | 204 | lemma one_rat: "1 = Fract 1 1" | 
| 205 | by (simp add: One_rat_def) | |
| 206 | ||
| 207 | theorem eq_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> | |
| 208 | (Fract a b = Fract c d) = (a * d = c * b)" | |
| 209 | by (simp add: Fract_def) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 210 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 211 | theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 212 | Fract a b + Fract c d = Fract (a * d + c * b) (b * d)" | 
| 18913 | 213 | by (simp add: Fract_def add_rat_def add_congruent2 UN_ratrel2) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 214 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 215 | theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b" | 
| 18913 | 216 | by (simp add: Fract_def minus_rat_def minus_congruent UN_ratrel) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 217 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 218 | theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 219 | Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" | 
| 18913 | 220 | by (simp add: diff_rat_def add_rat minus_rat) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 221 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 222 | theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 223 | Fract a b * Fract c d = Fract (a * c) (b * d)" | 
| 18913 | 224 | by (simp add: Fract_def mult_rat_def mult_congruent2 UN_ratrel2) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 225 | |
| 18913 | 226 | theorem inverse_rat: "a \<noteq> 0 ==> b \<noteq> 0 ==> | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 227 | inverse (Fract a b) = Fract b a" | 
| 18913 | 228 | by (simp add: Fract_def inverse_rat_def inverse_congruent UN_ratrel) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 229 | |
| 18913 | 230 | theorem divide_rat: "c \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==> | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 231 | Fract a b / Fract c d = Fract (a * d) (b * c)" | 
| 18913 | 232 | by (simp add: divide_rat_def inverse_rat mult_rat) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 233 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 234 | theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 235 | (Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))" | 
| 18982 | 236 | by (simp add: Fract_def le_rat_def le_congruent2 UN_ratrel2) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 237 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 238 | theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==> | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 239 | (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))" | 
| 18913 | 240 | by (simp add: less_rat_def le_rat eq_rat order_less_le) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 241 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 242 | theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 243 | by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat) | 
| 14691 | 244 | (auto simp add: mult_less_0_iff zero_less_mult_iff order_le_less | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 245 | split: abs_split) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 246 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 247 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 248 | subsubsection {* The ordered field of rational numbers *}
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 249 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 250 | instance rat :: field | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 251 | proof | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 252 | fix q r s :: rat | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 253 | show "(q + r) + s = q + (r + s)" | 
| 18913 | 254 | by (induct q, induct r, induct s) | 
| 255 | (simp add: add_rat add_ac mult_ac int_distrib) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 256 | show "q + r = r + q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 257 | by (induct q, induct r) (simp add: add_rat add_ac mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 258 | show "0 + q = q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 259 | by (induct q) (simp add: zero_rat add_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 260 | show "(-q) + q = 0" | 
| 18913 | 261 | by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 262 | show "q - r = q + (-r)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 263 | by (induct q, induct r) (simp add: add_rat minus_rat diff_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 264 | show "(q * r) * s = q * (r * s)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 265 | by (induct q, induct r, induct s) (simp add: mult_rat mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 266 | show "q * r = r * q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 267 | by (induct q, induct r) (simp add: mult_rat mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 268 | show "1 * q = q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 269 | by (induct q) (simp add: one_rat mult_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 270 | show "(q + r) * s = q * s + r * s" | 
| 14691 | 271 | by (induct q, induct r, induct s) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 272 | (simp add: add_rat mult_rat eq_rat int_distrib) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 273 | show "q \<noteq> 0 ==> inverse q * q = 1" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 274 | by (induct q) (simp add: inverse_rat mult_rat one_rat zero_rat eq_rat) | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 275 | show "q / r = q * inverse r" | 
| 14691 | 276 | by (simp add: divide_rat_def) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 277 | show "0 \<noteq> (1::rat)" | 
| 14691 | 278 | by (simp add: zero_rat one_rat eq_rat) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 279 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 280 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 281 | instance rat :: linorder | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 282 | proof | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 283 | fix q r s :: rat | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 284 |   {
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 285 | assume "q \<le> r" and "r \<le> s" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 286 | show "q \<le> s" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 287 | proof (insert prems, induct q, induct r, induct s) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 288 | fix a b c d e f :: int | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 289 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 290 | assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 291 | show "Fract a b \<le> Fract e f" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 292 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 293 | from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 294 | by (auto simp add: zero_less_mult_iff linorder_neq_iff) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 295 | have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 296 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 297 | from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 298 | by (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 299 | with ff show ?thesis by (simp add: mult_le_cancel_right) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 300 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 301 | also have "... = (c * f) * (d * f) * (b * b)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 302 | by (simp only: mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 303 | also have "... \<le> (e * d) * (d * f) * (b * b)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 304 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 305 | from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 306 | by (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 307 | with bb show ?thesis by (simp add: mult_le_cancel_right) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 308 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 309 | finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 310 | by (simp only: mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 311 | with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 312 | by (simp add: mult_le_cancel_right) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 313 | with neq show ?thesis by (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 314 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 315 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 316 | next | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 317 | assume "q \<le> r" and "r \<le> q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 318 | show "q = r" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 319 | proof (insert prems, induct q, induct r) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 320 | fix a b c d :: int | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 321 | assume neq: "b \<noteq> 0" "d \<noteq> 0" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 322 | assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 323 | show "Fract a b = Fract c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 324 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 325 | from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 326 | by (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 327 | also have "... \<le> (a * d) * (b * d)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 328 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 329 | from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 330 | by (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 331 | thus ?thesis by (simp only: mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 332 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 333 | finally have "(a * d) * (b * d) = (c * b) * (b * d)" . | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 334 | moreover from neq have "b * d \<noteq> 0" by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 335 | ultimately have "a * d = c * b" by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 336 | with neq show ?thesis by (simp add: eq_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 337 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 338 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 339 | next | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 340 | show "q \<le> q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 341 | by (induct q) (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 342 | show "(q < r) = (q \<le> r \<and> q \<noteq> r)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 343 | by (simp only: less_rat_def) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 344 | show "q \<le> r \<or> r \<le> q" | 
| 18913 | 345 | by (induct q, induct r) | 
| 346 | (simp add: le_rat mult_commute, rule linorder_linear) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 347 | } | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 348 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 349 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 350 | instance rat :: ordered_field | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 351 | proof | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 352 | fix q r s :: rat | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 353 | show "q \<le> r ==> s + q \<le> s + r" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 354 | proof (induct q, induct r, induct s) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 355 | fix a b c d e f :: int | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 356 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 357 | assume le: "Fract a b \<le> Fract c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 358 | show "Fract e f + Fract a b \<le> Fract e f + Fract c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 359 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 360 | let ?F = "f * f" from neq have F: "0 < ?F" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 361 | by (auto simp add: zero_less_mult_iff) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 362 | from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 363 | by (simp add: le_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 364 | with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 365 | by (simp add: mult_le_cancel_right) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 366 | with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 367 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 368 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 369 | show "q < r ==> 0 < s ==> s * q < s * r" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 370 | proof (induct q, induct r, induct s) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 371 | fix a b c d e f :: int | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 372 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 373 | assume le: "Fract a b < Fract c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 374 | assume gt: "0 < Fract e f" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 375 | show "Fract e f * Fract a b < Fract e f * Fract c d" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 376 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 377 | let ?E = "e * f" and ?F = "f * f" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 378 | from neq gt have "0 < ?E" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 379 | by (auto simp add: zero_rat less_rat le_rat order_less_le eq_rat) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 380 | moreover from neq have "0 < ?F" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 381 | by (auto simp add: zero_less_mult_iff) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 382 | moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 383 | by (simp add: less_rat) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 384 | ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 385 | by (simp add: mult_less_cancel_right) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 386 | with neq show ?thesis | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 387 | by (simp add: less_rat mult_rat mult_ac) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 388 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 389 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 390 | show "\<bar>q\<bar> = (if q < 0 then -q else q)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 391 | by (simp only: abs_rat_def) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 392 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 393 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 394 | instance rat :: division_by_zero | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 395 | proof | 
| 18913 | 396 | show "inverse 0 = (0::rat)" | 
| 397 | by (simp add: zero_rat Fract_def inverse_rat_def | |
| 398 | inverse_congruent UN_ratrel) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 399 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 400 | |
| 20522 | 401 | instance rat :: recpower | 
| 402 | proof | |
| 403 | fix q :: rat | |
| 404 | fix n :: nat | |
| 405 | show "q ^ 0 = 1" by simp | |
| 406 | show "q ^ (Suc n) = q * (q ^ n)" by simp | |
| 407 | qed | |
| 408 | ||
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 409 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 410 | subsection {* Various Other Results *}
 | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 411 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 412 | lemma minus_rat_cancel [simp]: "b \<noteq> 0 ==> Fract (-a) (-b) = Fract a b" | 
| 18913 | 413 | by (simp add: eq_rat) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 414 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 415 | theorem Rat_induct_pos [case_names Fract, induct type: rat]: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 416 | assumes step: "!!a b. 0 < b ==> P (Fract a b)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 417 | shows "P q" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 418 | proof (cases q) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 419 | have step': "!!a b. b < 0 ==> P (Fract a b)" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 420 | proof - | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 421 | fix a::int and b::int | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 422 | assume b: "b < 0" | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 423 | hence "0 < -b" by simp | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 424 | hence "P (Fract (-a) (-b))" by (rule step) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 425 | thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b]) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 426 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 427 | case (Fract a b) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 428 | thus "P q" by (force simp add: linorder_neq_iff step step') | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 429 | qed | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 430 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 431 | lemma zero_less_Fract_iff: | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 432 | "0 < b ==> (0 < Fract a b) = (0 < a)" | 
| 14691 | 433 | by (simp add: zero_rat less_rat order_less_imp_not_eq2 zero_less_mult_iff) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 434 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 435 | lemma Fract_add_one: "n \<noteq> 0 ==> Fract (m + n) n = Fract m n + 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 436 | apply (insert add_rat [of concl: m n 1 1]) | 
| 14691 | 437 | apply (simp add: one_rat [symmetric]) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 438 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 439 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 440 | lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k" | 
| 14691 | 441 | apply (induct k) | 
| 442 | apply (simp add: zero_rat) | |
| 443 | apply (simp add: Fract_add_one) | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 444 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 445 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 446 | lemma Fract_of_int_eq: "Fract k 1 = of_int k" | 
| 14691 | 447 | proof (cases k rule: int_cases) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 448 | case (nonneg n) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 449 | thus ?thesis by (simp add: int_eq_of_nat Fract_of_nat_eq) | 
| 14691 | 450 | next | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 451 | case (neg n) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 452 | hence "Fract k 1 = - (Fract (of_nat (Suc n)) 1)" | 
| 14691 | 453 | by (simp only: minus_rat int_eq_of_nat) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 454 | also have "... = - (of_nat (Suc n))" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 455 | by (simp only: Fract_of_nat_eq) | 
| 14691 | 456 | finally show ?thesis | 
| 457 | by (simp add: only: prems int_eq_of_nat of_int_minus of_int_of_nat_eq) | |
| 458 | qed | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 459 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14365diff
changeset | 460 | |
| 14691 | 461 | subsection {* Numerals and Arithmetic *}
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 462 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 463 | instance rat :: number .. | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 464 | |
| 15013 | 465 | defs (overloaded) | 
| 20485 | 466 | rat_number_of_def: "(number_of w :: rat) == of_int w" | 
| 15013 | 467 |     --{*the type constraint is essential!*}
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 468 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 469 | instance rat :: number_ring | 
| 19765 | 470 | by default (simp add: rat_number_of_def) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 471 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 472 | use "rat_arith.ML" | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 473 | setup rat_arith_setup | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 474 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 475 | end |