| author | wenzelm | 
| Fri, 14 Mar 2025 23:03:58 +0100 | |
| changeset 82276 | d22e9c5b5dc6 | 
| parent 80061 | 4c1347e172b1 | 
| permissions | -rw-r--r-- | 
| 65435 | 1 | (* Title: HOL/Computational_Algebra/Fraction_Field.thy | 
| 65417 | 2 | Author: Amine Chaieb, University of Cambridge | 
| 3 | *) | |
| 4 | ||
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
65435diff
changeset | 5 | section\<open>The fraction field of any integral domain\<close> | 
| 65417 | 6 | |
| 7 | theory Fraction_Field | |
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 11 | subsection \<open>General fractions construction\<close> | |
| 12 | ||
| 13 | subsubsection \<open>Construction of the type of fractions\<close> | |
| 14 | ||
| 15 | context idom begin | |
| 16 | ||
| 17 | definition fractrel :: "'a \<times> 'a \<Rightarrow> 'a * 'a \<Rightarrow> bool" where | |
| 18 | "fractrel = (\<lambda>x y. snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)" | |
| 19 | ||
| 20 | lemma fractrel_iff [simp]: | |
| 21 | "fractrel x y \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x" | |
| 22 | by (simp add: fractrel_def) | |
| 23 | ||
| 24 | lemma symp_fractrel: "symp fractrel" | |
| 25 | by (simp add: symp_def) | |
| 26 | ||
| 27 | lemma transp_fractrel: "transp fractrel" | |
| 28 | proof (rule transpI, unfold split_paired_all) | |
| 29 | fix a b a' b' a'' b'' :: 'a | |
| 30 | assume A: "fractrel (a, b) (a', b')" | |
| 31 | assume B: "fractrel (a', b') (a'', b'')" | |
| 32 | have "b' * (a * b'') = b'' * (a * b')" by (simp add: ac_simps) | |
| 33 | also from A have "a * b' = a' * b" by auto | |
| 34 | also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: ac_simps) | |
| 35 | also from B have "a' * b'' = a'' * b'" by auto | |
| 36 | also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: ac_simps) | |
| 37 | finally have "b' * (a * b'') = b' * (a'' * b)" . | |
| 38 | moreover from B have "b' \<noteq> 0" by auto | |
| 39 | ultimately have "a * b'' = a'' * b" by simp | |
| 40 | with A B show "fractrel (a, b) (a'', b'')" by auto | |
| 41 | qed | |
| 42 | ||
| 43 | lemma part_equivp_fractrel: "part_equivp fractrel" | |
| 44 | using _ symp_fractrel transp_fractrel | |
| 45 | by(rule part_equivpI)(rule exI[where x="(0, 1)"]; simp) | |
| 46 | ||
| 47 | end | |
| 48 | ||
| 49 | quotient_type (overloaded) 'a fract = "'a :: idom \<times> 'a" / partial: "fractrel" | |
| 50 | by(rule part_equivp_fractrel) | |
| 51 | ||
| 52 | subsubsection \<open>Representation and basic operations\<close> | |
| 53 | ||
| 54 | lift_definition Fract :: "'a :: idom \<Rightarrow> 'a \<Rightarrow> 'a fract" | |
| 55 | is "\<lambda>a b. if b = 0 then (0, 1) else (a, b)" | |
| 56 | by simp | |
| 57 | ||
| 58 | lemma Fract_cases [cases type: fract]: | |
| 59 | obtains (Fract) a b where "q = Fract a b" "b \<noteq> 0" | |
| 60 | by transfer simp | |
| 61 | ||
| 62 | lemma Fract_induct [case_names Fract, induct type: fract]: | |
| 63 | "(\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)) \<Longrightarrow> P q" | |
| 64 | by (cases q) simp | |
| 65 | ||
| 66 | lemma eq_fract: | |
| 67 | shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b" | |
| 68 | and "\<And>a. Fract a 0 = Fract 0 1" | |
| 69 | and "\<And>a c. Fract 0 a = Fract 0 c" | |
| 70 | by(transfer; simp)+ | |
| 71 | ||
| 72 | instantiation fract :: (idom) comm_ring_1 | |
| 73 | begin | |
| 74 | ||
| 75 | lift_definition zero_fract :: "'a fract" is "(0, 1)" by simp | |
| 76 | ||
| 77 | lemma Zero_fract_def: "0 = Fract 0 1" | |
| 78 | by transfer simp | |
| 79 | ||
| 80 | lift_definition one_fract :: "'a fract" is "(1, 1)" by simp | |
| 81 | ||
| 82 | lemma One_fract_def: "1 = Fract 1 1" | |
| 83 | by transfer simp | |
| 84 | ||
| 85 | lift_definition plus_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract" | |
| 86 | is "\<lambda>q r. (fst q * snd r + fst r * snd q, snd q * snd r)" | |
| 87 | by(auto simp add: algebra_simps) | |
| 88 | ||
| 89 | lemma add_fract [simp]: | |
| 90 | "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)" | |
| 91 | by transfer simp | |
| 92 | ||
| 93 | lift_definition uminus_fract :: "'a fract \<Rightarrow> 'a fract" | |
| 94 | is "\<lambda>x. (- fst x, snd x)" | |
| 95 | by simp | |
| 96 | ||
| 97 | lemma minus_fract [simp]: | |
| 98 | fixes a b :: "'a::idom" | |
| 99 | shows "- Fract a b = Fract (- a) b" | |
| 100 | by transfer simp | |
| 101 | ||
| 102 | lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b" | |
| 103 | by (cases "b = 0") (simp_all add: eq_fract) | |
| 104 | ||
| 105 | definition diff_fract_def: "q - r = q + - (r::'a fract)" | |
| 106 | ||
| 107 | lemma diff_fract [simp]: | |
| 108 | "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" | |
| 109 | by (simp add: diff_fract_def) | |
| 110 | ||
| 111 | lift_definition times_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract" | |
| 112 | is "\<lambda>q r. (fst q * fst r, snd q * snd r)" | |
| 113 | by(simp add: algebra_simps) | |
| 114 | ||
| 115 | lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)" | |
| 116 | by transfer simp | |
| 117 | ||
| 118 | lemma mult_fract_cancel: | |
| 119 | "c \<noteq> 0 \<Longrightarrow> Fract (c * a) (c * b) = Fract a b" | |
| 120 | by transfer simp | |
| 121 | ||
| 122 | instance | |
| 123 | proof | |
| 124 | fix q r s :: "'a fract" | |
| 125 | show "(q * r) * s = q * (r * s)" | |
| 126 | by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps) | |
| 127 | show "q * r = r * q" | |
| 128 | by (cases q, cases r) (simp add: eq_fract algebra_simps) | |
| 129 | show "1 * q = q" | |
| 130 | by (cases q) (simp add: One_fract_def eq_fract) | |
| 131 | show "(q + r) + s = q + (r + s)" | |
| 132 | by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps) | |
| 133 | show "q + r = r + q" | |
| 134 | by (cases q, cases r) (simp add: eq_fract algebra_simps) | |
| 135 | show "0 + q = q" | |
| 136 | by (cases q) (simp add: Zero_fract_def eq_fract) | |
| 137 | show "- q + q = 0" | |
| 138 | by (cases q) (simp add: Zero_fract_def eq_fract) | |
| 139 | show "q - r = q + - r" | |
| 140 | by (cases q, cases r) (simp add: eq_fract) | |
| 141 | show "(q + r) * s = q * s + r * s" | |
| 142 | by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps) | |
| 143 | show "(0::'a fract) \<noteq> 1" | |
| 144 | by (simp add: Zero_fract_def One_fract_def eq_fract) | |
| 145 | qed | |
| 146 | ||
| 147 | end | |
| 148 | ||
| 149 | lemma of_nat_fract: "of_nat k = Fract (of_nat k) 1" | |
| 150 | by (induct k) (simp_all add: Zero_fract_def One_fract_def) | |
| 151 | ||
| 152 | lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k" | |
| 153 | by (rule of_nat_fract [symmetric]) | |
| 154 | ||
| 155 | lemma fract_collapse: | |
| 156 | "Fract 0 k = 0" | |
| 157 | "Fract 1 1 = 1" | |
| 158 | "Fract k 0 = 0" | |
| 159 | by(transfer; simp)+ | |
| 160 | ||
| 161 | lemma fract_expand: | |
| 162 | "0 = Fract 0 1" | |
| 163 | "1 = Fract 1 1" | |
| 164 | by (simp_all add: fract_collapse) | |
| 165 | ||
| 166 | lemma Fract_cases_nonzero: | |
| 167 | obtains (Fract) a b where "q = Fract a b" and "b \<noteq> 0" and "a \<noteq> 0" | |
| 168 | | (0) "q = 0" | |
| 169 | proof (cases "q = 0") | |
| 170 | case True | |
| 171 | then show thesis using 0 by auto | |
| 172 | next | |
| 173 | case False | |
| 174 | then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto | |
| 175 | with False have "0 \<noteq> Fract a b" by simp | |
| 176 | with \<open>b \<noteq> 0\<close> have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract) | |
| 177 | with Fract \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> show thesis by auto | |
| 178 | qed | |
| 179 | ||
| 180 | ||
| 181 | subsubsection \<open>The field of rational numbers\<close> | |
| 182 | ||
| 183 | context idom | |
| 184 | begin | |
| 185 | ||
| 186 | subclass ring_no_zero_divisors .. | |
| 187 | ||
| 188 | end | |
| 189 | ||
| 190 | instantiation fract :: (idom) field | |
| 191 | begin | |
| 192 | ||
| 193 | lift_definition inverse_fract :: "'a fract \<Rightarrow> 'a fract" | |
| 194 | is "\<lambda>x. if fst x = 0 then (0, 1) else (snd x, fst x)" | |
| 195 | by(auto simp add: algebra_simps) | |
| 196 | ||
| 197 | lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a" | |
| 198 | by transfer simp | |
| 199 | ||
| 200 | definition divide_fract_def: "q div r = q * inverse (r:: 'a fract)" | |
| 201 | ||
| 202 | lemma divide_fract [simp]: "Fract a b div Fract c d = Fract (a * d) (b * c)" | |
| 203 | by (simp add: divide_fract_def) | |
| 204 | ||
| 205 | instance | |
| 206 | proof | |
| 207 | fix q :: "'a fract" | |
| 208 | assume "q \<noteq> 0" | |
| 209 | then show "inverse q * q = 1" | |
| 210 | by (cases q rule: Fract_cases_nonzero) | |
| 211 | (simp_all add: fract_expand eq_fract mult.commute) | |
| 212 | next | |
| 213 | fix q r :: "'a fract" | |
| 214 | show "q div r = q * inverse r" by (simp add: divide_fract_def) | |
| 215 | next | |
| 216 | show "inverse 0 = (0:: 'a fract)" | |
| 217 | by (simp add: fract_expand) (simp add: fract_collapse) | |
| 218 | qed | |
| 219 | ||
| 220 | end | |
| 221 | ||
| 222 | ||
| 223 | subsubsection \<open>The ordered field of fractions over an ordered idom\<close> | |
| 224 | ||
| 225 | instantiation fract :: (linordered_idom) linorder | |
| 226 | begin | |
| 227 | ||
| 228 | lemma less_eq_fract_respect: | |
| 229 | fixes a b a' b' c d c' d' :: 'a | |
| 230 | assumes neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0" | |
| 231 | assumes eq1: "a * b' = a' * b" | |
| 232 | assumes eq2: "c * d' = c' * d" | |
| 233 | shows "((a * d) * (b * d) \<le> (c * b) * (b * d)) \<longleftrightarrow> ((a' * d') * (b' * d') \<le> (c' * b') * (b' * d'))" | |
| 234 | proof - | |
| 235 | let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))" | |
| 236 |   {
 | |
| 237 | fix a b c d x :: 'a | |
| 238 | assume x: "x \<noteq> 0" | |
| 239 | have "?le a b c d = ?le (a * x) (b * x) c d" | |
| 240 | proof - | |
| 241 | from x have "0 < x * x" | |
| 242 | by (auto simp add: zero_less_mult_iff) | |
| 243 | then have "?le a b c d = | |
| 244 | ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))" | |
| 245 | by (simp add: mult_le_cancel_right) | |
| 246 | also have "... = ?le (a * x) (b * x) c d" | |
| 247 | by (simp add: ac_simps) | |
| 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 | then have "?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: ac_simps) | |
| 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: ac_simps) | |
| 263 | also from D have "... = ?le a' b' c' d'" | |
| 264 | by (rule le_factor [symmetric]) | |
| 265 | finally show "?le a b c d = ?le a' b' c' d'" . | |
| 266 | qed | |
| 267 | ||
| 268 | lift_definition less_eq_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> bool" | |
| 269 | is "\<lambda>q r. (fst q * snd r) * (snd q * snd r) \<le> (fst r * snd q) * (snd q * snd r)" | |
| 270 | by (clarsimp simp add: less_eq_fract_respect) | |
| 271 | ||
| 272 | definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z" | |
| 273 | ||
| 274 | lemma le_fract [simp]: | |
| 275 | "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)" | |
| 276 | by transfer simp | |
| 277 | ||
| 278 | lemma less_fract [simp]: | |
| 279 | "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)" | |
| 280 | by (simp add: less_fract_def less_le_not_le ac_simps) | |
| 281 | ||
| 282 | instance | |
| 283 | proof | |
| 284 | fix q r s :: "'a fract" | |
| 285 | assume "q \<le> r" and "r \<le> s" | |
| 286 | then show "q \<le> s" | |
| 287 | proof (induct q, induct r, induct s) | |
| 288 | fix a b c d e f :: 'a | |
| 289 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | |
| 290 | assume 1: "Fract a b \<le> Fract c d" | |
| 291 | assume 2: "Fract c d \<le> Fract e f" | |
| 292 | show "Fract a b \<le> Fract e f" | |
| 293 | proof - | |
| 294 | from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f" | |
| 295 | by (auto simp add: zero_less_mult_iff linorder_neq_iff) | |
| 296 | have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)" | |
| 297 | proof - | |
| 298 | from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | |
| 299 | by simp | |
| 300 | with ff show ?thesis by (simp add: mult_le_cancel_right) | |
| 301 | qed | |
| 302 | also have "... = (c * f) * (d * f) * (b * b)" | |
| 303 | by (simp only: ac_simps) | |
| 304 | also have "... \<le> (e * d) * (d * f) * (b * b)" | |
| 305 | proof - | |
| 306 | from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)" | |
| 307 | by simp | |
| 308 | with bb show ?thesis by (simp add: mult_le_cancel_right) | |
| 309 | qed | |
| 310 | finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)" | |
| 311 | by (simp only: ac_simps) | |
| 312 | with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)" | |
| 313 | by (simp add: mult_le_cancel_right) | |
| 314 | with neq show ?thesis by simp | |
| 315 | qed | |
| 316 | qed | |
| 317 | next | |
| 318 | fix q r :: "'a fract" | |
| 319 | assume "q \<le> r" and "r \<le> q" | |
| 320 | then show "q = r" | |
| 321 | proof (induct q, induct r) | |
| 322 | fix a b c d :: 'a | |
| 323 | assume neq: "b \<noteq> 0" "d \<noteq> 0" | |
| 324 | assume 1: "Fract a b \<le> Fract c d" | |
| 325 | assume 2: "Fract c d \<le> Fract a b" | |
| 326 | show "Fract a b = Fract c d" | |
| 327 | proof - | |
| 328 | from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | |
| 329 | by simp | |
| 330 | also have "... \<le> (a * d) * (b * d)" | |
| 331 | proof - | |
| 332 | from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)" | |
| 333 | by simp | |
| 334 | then show ?thesis by (simp only: ac_simps) | |
| 335 | qed | |
| 336 | finally have "(a * d) * (b * d) = (c * b) * (b * d)" . | |
| 337 | moreover from neq have "b * d \<noteq> 0" by simp | |
| 338 | ultimately have "a * d = c * b" by simp | |
| 339 | with neq show ?thesis by (simp add: eq_fract) | |
| 340 | qed | |
| 341 | qed | |
| 342 | next | |
| 343 | fix q r :: "'a fract" | |
| 344 | show "q \<le> q" | |
| 345 | by (induct q) simp | |
| 346 | show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)" | |
| 347 | by (simp only: less_fract_def) | |
| 348 | show "q \<le> r \<or> r \<le> q" | |
| 349 | by (induct q, induct r) | |
| 350 | (simp add: mult.commute, rule linorder_linear) | |
| 351 | qed | |
| 352 | ||
| 353 | end | |
| 354 | ||
| 355 | instantiation fract :: (linordered_idom) linordered_field | |
| 356 | begin | |
| 357 | ||
| 358 | definition abs_fract_def2: | |
| 359 | "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))" | |
| 360 | ||
| 361 | definition sgn_fract_def: | |
| 362 | "sgn (q::'a fract) = (if q = 0 then 0 else if 0 < q then 1 else - 1)" | |
| 363 | ||
| 364 | theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>" | |
| 365 | unfolding abs_fract_def2 not_le [symmetric] | |
| 366 | by transfer (auto simp add: zero_less_mult_iff le_less) | |
| 367 | ||
| 368 | instance proof | |
| 369 | fix q r s :: "'a fract" | |
| 370 | assume "q \<le> r" | |
| 371 | then show "s + q \<le> s + r" | |
| 372 | proof (induct q, induct r, induct s) | |
| 373 | fix a b c d e f :: 'a | |
| 374 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | |
| 375 | assume le: "Fract a b \<le> Fract c d" | |
| 376 | show "Fract e f + Fract a b \<le> Fract e f + Fract c d" | |
| 377 | proof - | |
| 378 | let ?F = "f * f" from neq have F: "0 < ?F" | |
| 379 | by (auto simp add: zero_less_mult_iff) | |
| 380 | from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | |
| 381 | by simp | |
| 382 | with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F" | |
| 383 | by (simp add: mult_le_cancel_right) | |
| 384 | with neq show ?thesis by (simp add: field_simps) | |
| 385 | qed | |
| 386 | qed | |
| 387 | next | |
| 388 | fix q r s :: "'a fract" | |
| 389 | assume "q < r" and "0 < s" | |
| 390 | then show "s * q < s * r" | |
| 391 | proof (induct q, induct r, induct s) | |
| 392 | fix a b c d e f :: 'a | |
| 393 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | |
| 394 | assume le: "Fract a b < Fract c d" | |
| 395 | assume gt: "0 < Fract e f" | |
| 396 | show "Fract e f * Fract a b < Fract e f * Fract c d" | |
| 397 | proof - | |
| 398 | let ?E = "e * f" and ?F = "f * f" | |
| 399 | from neq gt have "0 < ?E" | |
| 400 | by (auto simp add: Zero_fract_def order_less_le eq_fract) | |
| 401 | moreover from neq have "0 < ?F" | |
| 402 | by (auto simp add: zero_less_mult_iff) | |
| 403 | moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)" | |
| 404 | by simp | |
| 405 | ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F" | |
| 406 | by (simp add: mult_less_cancel_right) | |
| 407 | with neq show ?thesis | |
| 408 | by (simp add: ac_simps) | |
| 409 | qed | |
| 410 | qed | |
| 411 | qed (fact sgn_fract_def abs_fract_def2)+ | |
| 412 | ||
| 413 | end | |
| 414 | ||
| 415 | instantiation fract :: (linordered_idom) distrib_lattice | |
| 416 | begin | |
| 417 | ||
| 418 | definition inf_fract_def: | |
| 419 | "(inf :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min" | |
| 420 | ||
| 421 | definition sup_fract_def: | |
| 422 | "(sup :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max" | |
| 423 | ||
| 424 | instance | |
| 425 | by standard (simp_all add: inf_fract_def sup_fract_def max_min_distrib2) | |
| 426 | ||
| 427 | end | |
| 428 | ||
| 429 | lemma fract_induct_pos [case_names Fract]: | |
| 430 | fixes P :: "'a::linordered_idom fract \<Rightarrow> bool" | |
| 431 | assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)" | |
| 432 | shows "P q" | |
| 433 | proof (cases q) | |
| 434 | case (Fract a b) | |
| 435 |   {
 | |
| 436 | fix a b :: 'a | |
| 437 | assume b: "b < 0" | |
| 438 | have "P (Fract a b)" | |
| 439 | proof - | |
| 440 | from b have "0 < - b" by simp | |
| 441 | then have "P (Fract (- a) (- b))" | |
| 442 | by (rule step) | |
| 443 | then show "P (Fract a b)" | |
| 444 | by (simp add: order_less_imp_not_eq [OF b]) | |
| 445 | qed | |
| 446 | } | |
| 447 | with Fract show "P q" | |
| 448 | by (auto simp add: linorder_neq_iff step) | |
| 449 | qed | |
| 450 | ||
| 451 | lemma zero_less_Fract_iff: "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a" | |
| 452 | by (auto simp add: Zero_fract_def zero_less_mult_iff) | |
| 453 | ||
| 454 | lemma Fract_less_zero_iff: "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0" | |
| 455 | by (auto simp add: Zero_fract_def mult_less_0_iff) | |
| 456 | ||
| 457 | lemma zero_le_Fract_iff: "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a" | |
| 458 | by (auto simp add: Zero_fract_def zero_le_mult_iff) | |
| 459 | ||
| 460 | lemma Fract_le_zero_iff: "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0" | |
| 461 | by (auto simp add: Zero_fract_def mult_le_0_iff) | |
| 462 | ||
| 463 | lemma one_less_Fract_iff: "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a" | |
| 464 | by (auto simp add: One_fract_def mult_less_cancel_right_disj) | |
| 465 | ||
| 466 | lemma Fract_less_one_iff: "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b" | |
| 467 | by (auto simp add: One_fract_def mult_less_cancel_right_disj) | |
| 468 | ||
| 469 | lemma one_le_Fract_iff: "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a" | |
| 470 | by (auto simp add: One_fract_def mult_le_cancel_right) | |
| 471 | ||
| 472 | lemma Fract_le_one_iff: "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b" | |
| 473 | by (auto simp add: One_fract_def mult_le_cancel_right) | |
| 474 | ||
| 475 | end |