src/HOL/Rational.thy
author haftmann
Tue, 19 May 2009 13:57:32 +0200
changeset 31203 5c8fb4fd67e0
parent 31100 6a2e67fe4488
child 31205 98370b26c2ce
permissions -rw-r--r--
moved Code_Index, Random and Quickcheck before Main
     1 (*  Title:  HOL/Rational.thy
     2     Author: Markus Wenzel, TU Muenchen
     3 *)
     4 
     5 header {* Rational numbers *}
     6 
     7 theory Rational
     8 imports GCD Archimedean_Field
     9 begin
    10 
    11 subsection {* Rational numbers as quotient *}
    12 
    13 subsubsection {* Construction of the type of rational numbers *}
    14 
    15 definition
    16   ratrel :: "((int \<times> int) \<times> (int \<times> int)) set" where
    17   "ratrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
    18 
    19 lemma ratrel_iff [simp]:
    20   "(x, y) \<in> ratrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
    21   by (simp add: ratrel_def)
    22 
    23 lemma refl_on_ratrel: "refl_on {x. snd x \<noteq> 0} ratrel"
    24   by (auto simp add: refl_on_def ratrel_def)
    25 
    26 lemma sym_ratrel: "sym ratrel"
    27   by (simp add: ratrel_def sym_def)
    28 
    29 lemma trans_ratrel: "trans ratrel"
    30 proof (rule transI, unfold split_paired_all)
    31   fix a b a' b' a'' b'' :: int
    32   assume A: "((a, b), (a', b')) \<in> ratrel"
    33   assume B: "((a', b'), (a'', b'')) \<in> ratrel"
    34   have "b' * (a * b'') = b'' * (a * b')" by simp
    35   also from A have "a * b' = a' * b" by auto
    36   also have "b'' * (a' * b) = b * (a' * b'')" by simp
    37   also from B have "a' * b'' = a'' * b'" by auto
    38   also have "b * (a'' * b') = b' * (a'' * b)" by simp
    39   finally have "b' * (a * b'') = b' * (a'' * b)" .
    40   moreover from B have "b' \<noteq> 0" by auto
    41   ultimately have "a * b'' = a'' * b" by simp
    42   with A B show "((a, b), (a'', b'')) \<in> ratrel" by auto
    43 qed
    44   
    45 lemma equiv_ratrel: "equiv {x. snd x \<noteq> 0} ratrel"
    46   by (rule equiv.intro [OF refl_on_ratrel sym_ratrel trans_ratrel])
    47 
    48 lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel]
    49 lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel]
    50 
    51 lemma equiv_ratrel_iff [iff]: 
    52   assumes "snd x \<noteq> 0" and "snd y \<noteq> 0"
    53   shows "ratrel `` {x} = ratrel `` {y} \<longleftrightarrow> (x, y) \<in> ratrel"
    54   by (rule eq_equiv_class_iff, rule equiv_ratrel) (auto simp add: assms)
    55 
    56 typedef (Rat) rat = "{x. snd x \<noteq> 0} // ratrel"
    57 proof
    58   have "(0::int, 1::int) \<in> {x. snd x \<noteq> 0}" by simp
    59   then show "ratrel `` {(0, 1)} \<in> {x. snd x \<noteq> 0} // ratrel" by (rule quotientI)
    60 qed
    61 
    62 lemma ratrel_in_Rat [simp]: "snd x \<noteq> 0 \<Longrightarrow> ratrel `` {x} \<in> Rat"
    63   by (simp add: Rat_def quotientI)
    64 
    65 declare Abs_Rat_inject [simp] Abs_Rat_inverse [simp]
    66 
    67 
    68 subsubsection {* Representation and basic operations *}
    69 
    70 definition
    71   Fract :: "int \<Rightarrow> int \<Rightarrow> rat" where
    72   [code del]: "Fract a b = Abs_Rat (ratrel `` {if b = 0 then (0, 1) else (a, b)})"
    73 
    74 code_datatype Fract
    75 
    76 lemma Rat_cases [case_names Fract, cases type: rat]:
    77   assumes "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> C"
    78   shows C
    79   using assms by (cases q) (clarsimp simp add: Fract_def Rat_def quotient_def)
    80 
    81 lemma Rat_induct [case_names Fract, induct type: rat]:
    82   assumes "\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)"
    83   shows "P q"
    84   using assms by (cases q) simp
    85 
    86 lemma eq_rat:
    87   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"
    88   and "\<And>a. Fract a 0 = Fract 0 1"
    89   and "\<And>a c. Fract 0 a = Fract 0 c"
    90   by (simp_all add: Fract_def)
    91 
    92 instantiation rat :: comm_ring_1
    93 begin
    94 
    95 definition
    96   Zero_rat_def [code, code unfold]: "0 = Fract 0 1"
    97 
    98 definition
    99   One_rat_def [code, code unfold]: "1 = Fract 1 1"
   100 
   101 definition
   102   add_rat_def [code del]:
   103   "q + r = Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
   104     ratrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
   105 
   106 lemma add_rat [simp]:
   107   assumes "b \<noteq> 0" and "d \<noteq> 0"
   108   shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
   109 proof -
   110   have "(\<lambda>x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
   111     respects2 ratrel"
   112   by (rule equiv_ratrel [THEN congruent2_commuteI]) (simp_all add: left_distrib)
   113   with assms show ?thesis by (simp add: Fract_def add_rat_def UN_ratrel2)
   114 qed
   115 
   116 definition
   117   minus_rat_def [code del]:
   118   "- q = Abs_Rat (\<Union>x \<in> Rep_Rat q. ratrel `` {(- fst x, snd x)})"
   119 
   120 lemma minus_rat [simp, code]: "- Fract a b = Fract (- a) b"
   121 proof -
   122   have "(\<lambda>x. ratrel `` {(- fst x, snd x)}) respects ratrel"
   123     by (simp add: congruent_def)
   124   then show ?thesis by (simp add: Fract_def minus_rat_def UN_ratrel)
   125 qed
   126 
   127 lemma minus_rat_cancel [simp]: "Fract (- a) (- b) = Fract a b"
   128   by (cases "b = 0") (simp_all add: eq_rat)
   129 
   130 definition
   131   diff_rat_def [code del]: "q - r = q + - (r::rat)"
   132 
   133 lemma diff_rat [simp]:
   134   assumes "b \<noteq> 0" and "d \<noteq> 0"
   135   shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
   136   using assms by (simp add: diff_rat_def)
   137 
   138 definition
   139   mult_rat_def [code del]:
   140   "q * r = Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
   141     ratrel``{(fst x * fst y, snd x * snd y)})"
   142 
   143 lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)"
   144 proof -
   145   have "(\<lambda>x y. ratrel `` {(fst x * fst y, snd x * snd y)}) respects2 ratrel"
   146     by (rule equiv_ratrel [THEN congruent2_commuteI]) simp_all
   147   then show ?thesis by (simp add: Fract_def mult_rat_def UN_ratrel2)
   148 qed
   149 
   150 lemma mult_rat_cancel:
   151   assumes "c \<noteq> 0"
   152   shows "Fract (c * a) (c * b) = Fract a b"
   153 proof -
   154   from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def)
   155   then show ?thesis by (simp add: mult_rat [symmetric])
   156 qed
   157 
   158 instance proof
   159   fix q r s :: rat show "(q * r) * s = q * (r * s)" 
   160     by (cases q, cases r, cases s) (simp add: eq_rat)
   161 next
   162   fix q r :: rat show "q * r = r * q"
   163     by (cases q, cases r) (simp add: eq_rat)
   164 next
   165   fix q :: rat show "1 * q = q"
   166     by (cases q) (simp add: One_rat_def eq_rat)
   167 next
   168   fix q r s :: rat show "(q + r) + s = q + (r + s)"
   169     by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
   170 next
   171   fix q r :: rat show "q + r = r + q"
   172     by (cases q, cases r) (simp add: eq_rat)
   173 next
   174   fix q :: rat show "0 + q = q"
   175     by (cases q) (simp add: Zero_rat_def eq_rat)
   176 next
   177   fix q :: rat show "- q + q = 0"
   178     by (cases q) (simp add: Zero_rat_def eq_rat)
   179 next
   180   fix q r :: rat show "q - r = q + - r"
   181     by (cases q, cases r) (simp add: eq_rat)
   182 next
   183   fix q r s :: rat show "(q + r) * s = q * s + r * s"
   184     by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
   185 next
   186   show "(0::rat) \<noteq> 1" by (simp add: Zero_rat_def One_rat_def eq_rat)
   187 qed
   188 
   189 end
   190 
   191 lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1"
   192   by (induct k) (simp_all add: Zero_rat_def One_rat_def)
   193 
   194 lemma of_int_rat: "of_int k = Fract k 1"
   195   by (cases k rule: int_diff_cases) (simp add: of_nat_rat)
   196 
   197 lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
   198   by (rule of_nat_rat [symmetric])
   199 
   200 lemma Fract_of_int_eq: "Fract k 1 = of_int k"
   201   by (rule of_int_rat [symmetric])
   202 
   203 instantiation rat :: number_ring
   204 begin
   205 
   206 definition
   207   rat_number_of_def [code del]: "number_of w = Fract w 1"
   208 
   209 instance proof
   210 qed (simp add: rat_number_of_def of_int_rat)
   211 
   212 end
   213 
   214 lemma rat_number_collapse [code post]:
   215   "Fract 0 k = 0"
   216   "Fract 1 1 = 1"
   217   "Fract (number_of k) 1 = number_of k"
   218   "Fract k 0 = 0"
   219   by (cases "k = 0")
   220     (simp_all add: Zero_rat_def One_rat_def number_of_is_id number_of_eq of_int_rat eq_rat Fract_def)
   221 
   222 lemma rat_number_expand [code unfold]:
   223   "0 = Fract 0 1"
   224   "1 = Fract 1 1"
   225   "number_of k = Fract (number_of k) 1"
   226   by (simp_all add: rat_number_collapse)
   227 
   228 lemma iszero_rat [simp]:
   229   "iszero (number_of k :: rat) \<longleftrightarrow> iszero (number_of k :: int)"
   230   by (simp add: iszero_def rat_number_expand number_of_is_id eq_rat)
   231 
   232 lemma Rat_cases_nonzero [case_names Fract 0]:
   233   assumes Fract: "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> C"
   234   assumes 0: "q = 0 \<Longrightarrow> C"
   235   shows C
   236 proof (cases "q = 0")
   237   case True then show C using 0 by auto
   238 next
   239   case False
   240   then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
   241   moreover with False have "0 \<noteq> Fract a b" by simp
   242   with `b \<noteq> 0` have "a \<noteq> 0" by (simp add: Zero_rat_def eq_rat)
   243   with Fract `q = Fract a b` `b \<noteq> 0` show C by auto
   244 qed
   245 
   246 
   247 subsubsection {* The field of rational numbers *}
   248 
   249 instantiation rat :: "{field, division_by_zero}"
   250 begin
   251 
   252 definition
   253   inverse_rat_def [code del]:
   254   "inverse q = Abs_Rat (\<Union>x \<in> Rep_Rat q.
   255      ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
   256 
   257 lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a"
   258 proof -
   259   have "(\<lambda>x. ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)}) respects ratrel"
   260     by (auto simp add: congruent_def mult_commute)
   261   then show ?thesis by (simp add: Fract_def inverse_rat_def UN_ratrel)
   262 qed
   263 
   264 definition
   265   divide_rat_def [code del]: "q / r = q * inverse (r::rat)"
   266 
   267 lemma divide_rat [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
   268   by (simp add: divide_rat_def)
   269 
   270 instance proof
   271   show "inverse 0 = (0::rat)" by (simp add: rat_number_expand)
   272     (simp add: rat_number_collapse)
   273 next
   274   fix q :: rat
   275   assume "q \<noteq> 0"
   276   then show "inverse q * q = 1" by (cases q rule: Rat_cases_nonzero)
   277    (simp_all add: mult_rat  inverse_rat rat_number_expand eq_rat)
   278 next
   279   fix q r :: rat
   280   show "q / r = q * inverse r" by (simp add: divide_rat_def)
   281 qed
   282 
   283 end
   284 
   285 
   286 subsubsection {* Various *}
   287 
   288 lemma Fract_add_one: "n \<noteq> 0 ==> Fract (m + n) n = Fract m n + 1"
   289   by (simp add: rat_number_expand)
   290 
   291 lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l"
   292   by (simp add: Fract_of_int_eq [symmetric])
   293 
   294 lemma Fract_number_of_quotient [code post]:
   295   "Fract (number_of k) (number_of l) = number_of k / number_of l"
   296   unfolding Fract_of_int_quotient number_of_is_id number_of_eq ..
   297 
   298 lemma Fract_1_number_of [code post]:
   299   "Fract 1 (number_of k) = 1 / number_of k"
   300   unfolding Fract_of_int_quotient number_of_eq by simp
   301 
   302 subsubsection {* The ordered field of rational numbers *}
   303 
   304 instantiation rat :: linorder
   305 begin
   306 
   307 definition
   308   le_rat_def [code del]:
   309    "q \<le> r \<longleftrightarrow> contents (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
   310       {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})"
   311 
   312 lemma le_rat [simp]:
   313   assumes "b \<noteq> 0" and "d \<noteq> 0"
   314   shows "Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
   315 proof -
   316   have "(\<lambda>x y. {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})
   317     respects2 ratrel"
   318   proof (clarsimp simp add: congruent2_def)
   319     fix a b a' b' c d c' d'::int
   320     assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
   321     assume eq1: "a * b' = a' * b"
   322     assume eq2: "c * d' = c' * d"
   323 
   324     let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
   325     {
   326       fix a b c d x :: int assume x: "x \<noteq> 0"
   327       have "?le a b c d = ?le (a * x) (b * x) c d"
   328       proof -
   329         from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
   330         hence "?le a b c d =
   331             ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
   332           by (simp add: mult_le_cancel_right)
   333         also have "... = ?le (a * x) (b * x) c d"
   334           by (simp add: mult_ac)
   335         finally show ?thesis .
   336       qed
   337     } note le_factor = this
   338 
   339     let ?D = "b * d" and ?D' = "b' * d'"
   340     from neq have D: "?D \<noteq> 0" by simp
   341     from neq have "?D' \<noteq> 0" by simp
   342     hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
   343       by (rule le_factor)
   344     also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')" 
   345       by (simp add: mult_ac)
   346     also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
   347       by (simp only: eq1 eq2)
   348     also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
   349       by (simp add: mult_ac)
   350     also from D have "... = ?le a' b' c' d'"
   351       by (rule le_factor [symmetric])
   352     finally show "?le a b c d = ?le a' b' c' d'" .
   353   qed
   354   with assms show ?thesis by (simp add: Fract_def le_rat_def UN_ratrel2)
   355 qed
   356 
   357 definition
   358   less_rat_def [code del]: "z < (w::rat) \<longleftrightarrow> z \<le> w \<and> z \<noteq> w"
   359 
   360 lemma less_rat [simp]:
   361   assumes "b \<noteq> 0" and "d \<noteq> 0"
   362   shows "Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
   363   using assms by (simp add: less_rat_def eq_rat order_less_le)
   364 
   365 instance proof
   366   fix q r s :: rat
   367   {
   368     assume "q \<le> r" and "r \<le> s"
   369     show "q \<le> s"
   370     proof (insert prems, induct q, induct r, induct s)
   371       fix a b c d e f :: int
   372       assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
   373       assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
   374       show "Fract a b \<le> Fract e f"
   375       proof -
   376         from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
   377           by (auto simp add: zero_less_mult_iff linorder_neq_iff)
   378         have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
   379         proof -
   380           from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
   381             by simp
   382           with ff show ?thesis by (simp add: mult_le_cancel_right)
   383         qed
   384         also have "... = (c * f) * (d * f) * (b * b)" by algebra
   385         also have "... \<le> (e * d) * (d * f) * (b * b)"
   386         proof -
   387           from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
   388             by simp
   389           with bb show ?thesis by (simp add: mult_le_cancel_right)
   390         qed
   391         finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
   392           by (simp only: mult_ac)
   393         with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
   394           by (simp add: mult_le_cancel_right)
   395         with neq show ?thesis by simp
   396       qed
   397     qed
   398   next
   399     assume "q \<le> r" and "r \<le> q"
   400     show "q = r"
   401     proof (insert prems, induct q, induct r)
   402       fix a b c d :: int
   403       assume neq: "b \<noteq> 0"  "d \<noteq> 0"
   404       assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
   405       show "Fract a b = Fract c d"
   406       proof -
   407         from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
   408           by simp
   409         also have "... \<le> (a * d) * (b * d)"
   410         proof -
   411           from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
   412             by simp
   413           thus ?thesis by (simp only: mult_ac)
   414         qed
   415         finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
   416         moreover from neq have "b * d \<noteq> 0" by simp
   417         ultimately have "a * d = c * b" by simp
   418         with neq show ?thesis by (simp add: eq_rat)
   419       qed
   420     qed
   421   next
   422     show "q \<le> q"
   423       by (induct q) simp
   424     show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"
   425       by (induct q, induct r) (auto simp add: le_less mult_commute)
   426     show "q \<le> r \<or> r \<le> q"
   427       by (induct q, induct r)
   428          (simp add: mult_commute, rule linorder_linear)
   429   }
   430 qed
   431 
   432 end
   433 
   434 instantiation rat :: "{distrib_lattice, abs_if, sgn_if}"
   435 begin
   436 
   437 definition
   438   abs_rat_def [code del]: "\<bar>q\<bar> = (if q < 0 then -q else (q::rat))"
   439 
   440 lemma abs_rat [simp, code]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
   441   by (auto simp add: abs_rat_def zabs_def Zero_rat_def less_rat not_less le_less minus_rat eq_rat zero_compare_simps)
   442 
   443 definition
   444   sgn_rat_def [code del]: "sgn (q::rat) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"
   445 
   446 lemma sgn_rat [simp, code]: "sgn (Fract a b) = of_int (sgn a * sgn b)"
   447   unfolding Fract_of_int_eq
   448   by (auto simp: zsgn_def sgn_rat_def Zero_rat_def eq_rat)
   449     (auto simp: rat_number_collapse not_less le_less zero_less_mult_iff)
   450 
   451 definition
   452   "(inf \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat) = min"
   453 
   454 definition
   455   "(sup \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat) = max"
   456 
   457 instance by intro_classes
   458   (auto simp add: abs_rat_def sgn_rat_def min_max.sup_inf_distrib1 inf_rat_def sup_rat_def)
   459 
   460 end
   461 
   462 instance rat :: ordered_field
   463 proof
   464   fix q r s :: rat
   465   show "q \<le> r ==> s + q \<le> s + r"
   466   proof (induct q, induct r, induct s)
   467     fix a b c d e f :: int
   468     assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
   469     assume le: "Fract a b \<le> Fract c d"
   470     show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
   471     proof -
   472       let ?F = "f * f" from neq have F: "0 < ?F"
   473         by (auto simp add: zero_less_mult_iff)
   474       from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
   475         by simp
   476       with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
   477         by (simp add: mult_le_cancel_right)
   478       with neq show ?thesis by (simp add: mult_ac int_distrib)
   479     qed
   480   qed
   481   show "q < r ==> 0 < s ==> s * q < s * r"
   482   proof (induct q, induct r, induct s)
   483     fix a b c d e f :: int
   484     assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
   485     assume le: "Fract a b < Fract c d"
   486     assume gt: "0 < Fract e f"
   487     show "Fract e f * Fract a b < Fract e f * Fract c d"
   488     proof -
   489       let ?E = "e * f" and ?F = "f * f"
   490       from neq gt have "0 < ?E"
   491         by (auto simp add: Zero_rat_def order_less_le eq_rat)
   492       moreover from neq have "0 < ?F"
   493         by (auto simp add: zero_less_mult_iff)
   494       moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
   495         by simp
   496       ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
   497         by (simp add: mult_less_cancel_right)
   498       with neq show ?thesis
   499         by (simp add: mult_ac)
   500     qed
   501   qed
   502 qed auto
   503 
   504 lemma Rat_induct_pos [case_names Fract, induct type: rat]:
   505   assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"
   506   shows "P q"
   507 proof (cases q)
   508   have step': "\<And>a b. b < 0 \<Longrightarrow> P (Fract a b)"
   509   proof -
   510     fix a::int and b::int
   511     assume b: "b < 0"
   512     hence "0 < -b" by simp
   513     hence "P (Fract (-a) (-b))" by (rule step)
   514     thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
   515   qed
   516   case (Fract a b)
   517   thus "P q" by (force simp add: linorder_neq_iff step step')
   518 qed
   519 
   520 lemma zero_less_Fract_iff:
   521   "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"
   522   by (simp add: Zero_rat_def zero_less_mult_iff)
   523 
   524 lemma Fract_less_zero_iff:
   525   "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"
   526   by (simp add: Zero_rat_def mult_less_0_iff)
   527 
   528 lemma zero_le_Fract_iff:
   529   "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"
   530   by (simp add: Zero_rat_def zero_le_mult_iff)
   531 
   532 lemma Fract_le_zero_iff:
   533   "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"
   534   by (simp add: Zero_rat_def mult_le_0_iff)
   535 
   536 lemma one_less_Fract_iff:
   537   "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"
   538   by (simp add: One_rat_def mult_less_cancel_right_disj)
   539 
   540 lemma Fract_less_one_iff:
   541   "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"
   542   by (simp add: One_rat_def mult_less_cancel_right_disj)
   543 
   544 lemma one_le_Fract_iff:
   545   "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"
   546   by (simp add: One_rat_def mult_le_cancel_right)
   547 
   548 lemma Fract_le_one_iff:
   549   "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"
   550   by (simp add: One_rat_def mult_le_cancel_right)
   551 
   552 
   553 subsubsection {* Rationals are an Archimedean field *}
   554 
   555 lemma rat_floor_lemma:
   556   assumes "0 < b"
   557   shows "of_int (a div b) \<le> Fract a b \<and> Fract a b < of_int (a div b + 1)"
   558 proof -
   559   have "Fract a b = of_int (a div b) + Fract (a mod b) b"
   560     using `0 < b` by (simp add: of_int_rat)
   561   moreover have "0 \<le> Fract (a mod b) b \<and> Fract (a mod b) b < 1"
   562     using `0 < b` by (simp add: zero_le_Fract_iff Fract_less_one_iff)
   563   ultimately show ?thesis by simp
   564 qed
   565 
   566 instance rat :: archimedean_field
   567 proof
   568   fix r :: rat
   569   show "\<exists>z. r \<le> of_int z"
   570   proof (induct r)
   571     case (Fract a b)
   572     then have "Fract a b \<le> of_int (a div b + 1)"
   573       using rat_floor_lemma [of b a] by simp
   574     then show "\<exists>z. Fract a b \<le> of_int z" ..
   575   qed
   576 qed
   577 
   578 lemma floor_Fract:
   579   assumes "0 < b" shows "floor (Fract a b) = a div b"
   580   using rat_floor_lemma [OF `0 < b`, of a]
   581   by (simp add: floor_unique)
   582 
   583 
   584 subsection {* Linear arithmetic setup *}
   585 
   586 declaration {*
   587   K (Lin_Arith.add_inj_thms [@{thm of_nat_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2]
   588     (* not needed because x < (y::nat) can be rewritten as Suc x <= y: of_nat_less_iff RS iffD2 *)
   589   #> Lin_Arith.add_inj_thms [@{thm of_int_le_iff} RS iffD2, @{thm of_int_eq_iff} RS iffD2]
   590     (* not needed because x < (y::int) can be rewritten as x + 1 <= y: of_int_less_iff RS iffD2 *)
   591   #> Lin_Arith.add_simps [@{thm neg_less_iff_less},
   592       @{thm True_implies_equals},
   593       read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib},
   594       @{thm divide_1}, @{thm divide_zero_left},
   595       @{thm times_divide_eq_right}, @{thm times_divide_eq_left},
   596       @{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym,
   597       @{thm of_int_minus}, @{thm of_int_diff},
   598       @{thm of_int_of_nat_eq}]
   599   #> Lin_Arith.add_simprocs Numeral_Simprocs.field_cancel_numeral_factors
   600   #> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat => rat"})
   601   #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int => rat"}))
   602 *}
   603 
   604 
   605 subsection {* Embedding from Rationals to other Fields *}
   606 
   607 class field_char_0 = field + ring_char_0
   608 
   609 subclass (in ordered_field) field_char_0 ..
   610 
   611 context field_char_0
   612 begin
   613 
   614 definition of_rat :: "rat \<Rightarrow> 'a" where
   615   [code del]: "of_rat q = contents (\<Union>(a,b) \<in> Rep_Rat q. {of_int a / of_int b})"
   616 
   617 end
   618 
   619 lemma of_rat_congruent:
   620   "(\<lambda>(a, b). {of_int a / of_int b :: 'a::field_char_0}) respects ratrel"
   621 apply (rule congruent.intro)
   622 apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
   623 apply (simp only: of_int_mult [symmetric])
   624 done
   625 
   626 lemma of_rat_rat: "b \<noteq> 0 \<Longrightarrow> of_rat (Fract a b) = of_int a / of_int b"
   627   unfolding Fract_def of_rat_def by (simp add: UN_ratrel of_rat_congruent)
   628 
   629 lemma of_rat_0 [simp]: "of_rat 0 = 0"
   630 by (simp add: Zero_rat_def of_rat_rat)
   631 
   632 lemma of_rat_1 [simp]: "of_rat 1 = 1"
   633 by (simp add: One_rat_def of_rat_rat)
   634 
   635 lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b"
   636 by (induct a, induct b, simp add: of_rat_rat add_frac_eq)
   637 
   638 lemma of_rat_minus: "of_rat (- a) = - of_rat a"
   639 by (induct a, simp add: of_rat_rat)
   640 
   641 lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
   642 by (simp only: diff_minus of_rat_add of_rat_minus)
   643 
   644 lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
   645 apply (induct a, induct b, simp add: of_rat_rat)
   646 apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac)
   647 done
   648 
   649 lemma nonzero_of_rat_inverse:
   650   "a \<noteq> 0 \<Longrightarrow> of_rat (inverse a) = inverse (of_rat a)"
   651 apply (rule inverse_unique [symmetric])
   652 apply (simp add: of_rat_mult [symmetric])
   653 done
   654 
   655 lemma of_rat_inverse:
   656   "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
   657    inverse (of_rat a)"
   658 by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)
   659 
   660 lemma nonzero_of_rat_divide:
   661   "b \<noteq> 0 \<Longrightarrow> of_rat (a / b) = of_rat a / of_rat b"
   662 by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
   663 
   664 lemma of_rat_divide:
   665   "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
   666    = of_rat a / of_rat b"
   667 by (cases "b = 0") (simp_all add: nonzero_of_rat_divide)
   668 
   669 lemma of_rat_power:
   670   "(of_rat (a ^ n)::'a::field_char_0) = of_rat a ^ n"
   671 by (induct n) (simp_all add: of_rat_mult)
   672 
   673 lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)"
   674 apply (induct a, induct b)
   675 apply (simp add: of_rat_rat eq_rat)
   676 apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
   677 apply (simp only: of_int_mult [symmetric] of_int_eq_iff)
   678 done
   679 
   680 lemma of_rat_less:
   681   "(of_rat r :: 'a::ordered_field) < of_rat s \<longleftrightarrow> r < s"
   682 proof (induct r, induct s)
   683   fix a b c d :: int
   684   assume not_zero: "b > 0" "d > 0"
   685   then have "b * d > 0" by (rule mult_pos_pos)
   686   have of_int_divide_less_eq:
   687     "(of_int a :: 'a) / of_int b < of_int c / of_int d
   688       \<longleftrightarrow> (of_int a :: 'a) * of_int d < of_int c * of_int b"
   689     using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq)
   690   show "(of_rat (Fract a b) :: 'a::ordered_field) < of_rat (Fract c d)
   691     \<longleftrightarrow> Fract a b < Fract c d"
   692     using not_zero `b * d > 0`
   693     by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult)
   694 qed
   695 
   696 lemma of_rat_less_eq:
   697   "(of_rat r :: 'a::ordered_field) \<le> of_rat s \<longleftrightarrow> r \<le> s"
   698   unfolding le_less by (auto simp add: of_rat_less)
   699 
   700 lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified]
   701 
   702 lemma of_rat_eq_id [simp]: "of_rat = id"
   703 proof
   704   fix a
   705   show "of_rat a = id a"
   706   by (induct a)
   707      (simp add: of_rat_rat Fract_of_int_eq [symmetric])
   708 qed
   709 
   710 text{*Collapse nested embeddings*}
   711 lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n"
   712 by (induct n) (simp_all add: of_rat_add)
   713 
   714 lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z"
   715 by (cases z rule: int_diff_cases) (simp add: of_rat_diff)
   716 
   717 lemma of_rat_number_of_eq [simp]:
   718   "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})"
   719 by (simp add: number_of_eq)
   720 
   721 lemmas zero_rat = Zero_rat_def
   722 lemmas one_rat = One_rat_def
   723 
   724 abbreviation
   725   rat_of_nat :: "nat \<Rightarrow> rat"
   726 where
   727   "rat_of_nat \<equiv> of_nat"
   728 
   729 abbreviation
   730   rat_of_int :: "int \<Rightarrow> rat"
   731 where
   732   "rat_of_int \<equiv> of_int"
   733 
   734 subsection {* The Set of Rational Numbers *}
   735 
   736 context field_char_0
   737 begin
   738 
   739 definition
   740   Rats  :: "'a set" where
   741   [code del]: "Rats = range of_rat"
   742 
   743 notation (xsymbols)
   744   Rats  ("\<rat>")
   745 
   746 end
   747 
   748 lemma Rats_of_rat [simp]: "of_rat r \<in> Rats"
   749 by (simp add: Rats_def)
   750 
   751 lemma Rats_of_int [simp]: "of_int z \<in> Rats"
   752 by (subst of_rat_of_int_eq [symmetric], rule Rats_of_rat)
   753 
   754 lemma Rats_of_nat [simp]: "of_nat n \<in> Rats"
   755 by (subst of_rat_of_nat_eq [symmetric], rule Rats_of_rat)
   756 
   757 lemma Rats_number_of [simp]:
   758   "(number_of w::'a::{number_ring,field_char_0}) \<in> Rats"
   759 by (subst of_rat_number_of_eq [symmetric], rule Rats_of_rat)
   760 
   761 lemma Rats_0 [simp]: "0 \<in> Rats"
   762 apply (unfold Rats_def)
   763 apply (rule range_eqI)
   764 apply (rule of_rat_0 [symmetric])
   765 done
   766 
   767 lemma Rats_1 [simp]: "1 \<in> Rats"
   768 apply (unfold Rats_def)
   769 apply (rule range_eqI)
   770 apply (rule of_rat_1 [symmetric])
   771 done
   772 
   773 lemma Rats_add [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a + b \<in> Rats"
   774 apply (auto simp add: Rats_def)
   775 apply (rule range_eqI)
   776 apply (rule of_rat_add [symmetric])
   777 done
   778 
   779 lemma Rats_minus [simp]: "a \<in> Rats \<Longrightarrow> - a \<in> Rats"
   780 apply (auto simp add: Rats_def)
   781 apply (rule range_eqI)
   782 apply (rule of_rat_minus [symmetric])
   783 done
   784 
   785 lemma Rats_diff [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a - b \<in> Rats"
   786 apply (auto simp add: Rats_def)
   787 apply (rule range_eqI)
   788 apply (rule of_rat_diff [symmetric])
   789 done
   790 
   791 lemma Rats_mult [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a * b \<in> Rats"
   792 apply (auto simp add: Rats_def)
   793 apply (rule range_eqI)
   794 apply (rule of_rat_mult [symmetric])
   795 done
   796 
   797 lemma nonzero_Rats_inverse:
   798   fixes a :: "'a::field_char_0"
   799   shows "\<lbrakk>a \<in> Rats; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Rats"
   800 apply (auto simp add: Rats_def)
   801 apply (rule range_eqI)
   802 apply (erule nonzero_of_rat_inverse [symmetric])
   803 done
   804 
   805 lemma Rats_inverse [simp]:
   806   fixes a :: "'a::{field_char_0,division_by_zero}"
   807   shows "a \<in> Rats \<Longrightarrow> inverse a \<in> Rats"
   808 apply (auto simp add: Rats_def)
   809 apply (rule range_eqI)
   810 apply (rule of_rat_inverse [symmetric])
   811 done
   812 
   813 lemma nonzero_Rats_divide:
   814   fixes a b :: "'a::field_char_0"
   815   shows "\<lbrakk>a \<in> Rats; b \<in> Rats; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
   816 apply (auto simp add: Rats_def)
   817 apply (rule range_eqI)
   818 apply (erule nonzero_of_rat_divide [symmetric])
   819 done
   820 
   821 lemma Rats_divide [simp]:
   822   fixes a b :: "'a::{field_char_0,division_by_zero}"
   823   shows "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
   824 apply (auto simp add: Rats_def)
   825 apply (rule range_eqI)
   826 apply (rule of_rat_divide [symmetric])
   827 done
   828 
   829 lemma Rats_power [simp]:
   830   fixes a :: "'a::field_char_0"
   831   shows "a \<in> Rats \<Longrightarrow> a ^ n \<in> Rats"
   832 apply (auto simp add: Rats_def)
   833 apply (rule range_eqI)
   834 apply (rule of_rat_power [symmetric])
   835 done
   836 
   837 lemma Rats_cases [cases set: Rats]:
   838   assumes "q \<in> \<rat>"
   839   obtains (of_rat) r where "q = of_rat r"
   840   unfolding Rats_def
   841 proof -
   842   from `q \<in> \<rat>` have "q \<in> range of_rat" unfolding Rats_def .
   843   then obtain r where "q = of_rat r" ..
   844   then show thesis ..
   845 qed
   846 
   847 lemma Rats_induct [case_names of_rat, induct set: Rats]:
   848   "q \<in> \<rat> \<Longrightarrow> (\<And>r. P (of_rat r)) \<Longrightarrow> P q"
   849   by (rule Rats_cases) auto
   850 
   851 
   852 subsection {* Implementation of rational numbers as pairs of integers *}
   853 
   854 lemma Fract_norm: "Fract (a div zgcd a b) (b div zgcd a b) = Fract a b"
   855 proof (cases "a = 0 \<or> b = 0")
   856   case True then show ?thesis by (auto simp add: eq_rat)
   857 next
   858   let ?c = "zgcd a b"
   859   case False then have "a \<noteq> 0" and "b \<noteq> 0" by auto
   860   then have "?c \<noteq> 0" by simp
   861   then have "Fract ?c ?c = Fract 1 1" by (simp add: eq_rat)
   862   moreover have "Fract (a div ?c * ?c + a mod ?c) (b div ?c * ?c + b mod ?c) = Fract a b"
   863     by (simp add: semiring_div_class.mod_div_equality)
   864   moreover have "a mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric])
   865   moreover have "b mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric])
   866   ultimately show ?thesis
   867     by (simp add: mult_rat [symmetric])
   868 qed
   869 
   870 definition Fract_norm :: "int \<Rightarrow> int \<Rightarrow> rat" where
   871   [simp, code del]: "Fract_norm a b = Fract a b"
   872 
   873 lemma Fract_norm_code [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = zgcd a b in
   874   if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))"
   875   by (simp add: eq_rat Zero_rat_def Let_def Fract_norm)
   876 
   877 lemma [code]:
   878   "of_rat (Fract a b) = (if b \<noteq> 0 then of_int a / of_int b else 0)"
   879   by (cases "b = 0") (simp_all add: rat_number_collapse of_rat_rat)
   880 
   881 instantiation rat :: eq
   882 begin
   883 
   884 definition [code del]: "eq_class.eq (a\<Colon>rat) b \<longleftrightarrow> a - b = 0"
   885 
   886 instance by default (simp add: eq_rat_def)
   887 
   888 lemma rat_eq_code [code]:
   889   "eq_class.eq (Fract a b) (Fract c d) \<longleftrightarrow> (if b = 0
   890        then c = 0 \<or> d = 0
   891      else if d = 0
   892        then a = 0 \<or> b = 0
   893      else a * d = b * c)"
   894   by (auto simp add: eq eq_rat)
   895 
   896 lemma rat_eq_refl [code nbe]:
   897   "eq_class.eq (r::rat) r \<longleftrightarrow> True"
   898   by (rule HOL.eq_refl)
   899 
   900 end
   901 
   902 lemma le_rat':
   903   assumes "b \<noteq> 0"
   904     and "d \<noteq> 0"
   905   shows "Fract a b \<le> Fract c d \<longleftrightarrow> a * \<bar>d\<bar> * sgn b \<le> c * \<bar>b\<bar> * sgn d"
   906 proof -
   907   have abs_sgn: "\<And>k::int. \<bar>k\<bar> = k * sgn k" unfolding abs_if sgn_if by simp
   908   have "a * d * (b * d) \<le> c * b * (b * d) \<longleftrightarrow> a * d * (sgn b * sgn d) \<le> c * b * (sgn b * sgn d)"
   909   proof (cases "b * d > 0")
   910     case True
   911     moreover from True have "sgn b * sgn d = 1"
   912       by (simp add: sgn_times [symmetric] sgn_1_pos)
   913     ultimately show ?thesis by (simp add: mult_le_cancel_right)
   914   next
   915     case False with assms have "b * d < 0" by (simp add: less_le)
   916     moreover from this have "sgn b * sgn d = - 1"
   917       by (simp only: sgn_times [symmetric] sgn_1_neg)
   918     ultimately show ?thesis by (simp add: mult_le_cancel_right)
   919   qed
   920   also have "\<dots> \<longleftrightarrow> a * \<bar>d\<bar> * sgn b \<le> c * \<bar>b\<bar> * sgn d"
   921     by (simp add: abs_sgn mult_ac)
   922   finally show ?thesis using assms by simp
   923 qed
   924 
   925 lemma less_rat': 
   926   assumes "b \<noteq> 0"
   927     and "d \<noteq> 0"
   928   shows "Fract a b < Fract c d \<longleftrightarrow> a * \<bar>d\<bar> * sgn b < c * \<bar>b\<bar> * sgn d"
   929 proof -
   930   have abs_sgn: "\<And>k::int. \<bar>k\<bar> = k * sgn k" unfolding abs_if sgn_if by simp
   931   have "a * d * (b * d) < c * b * (b * d) \<longleftrightarrow> a * d * (sgn b * sgn d) < c * b * (sgn b * sgn d)"
   932   proof (cases "b * d > 0")
   933     case True
   934     moreover from True have "sgn b * sgn d = 1"
   935       by (simp add: sgn_times [symmetric] sgn_1_pos)
   936     ultimately show ?thesis by (simp add: mult_less_cancel_right)
   937   next
   938     case False with assms have "b * d < 0" by (simp add: less_le)
   939     moreover from this have "sgn b * sgn d = - 1"
   940       by (simp only: sgn_times [symmetric] sgn_1_neg)
   941     ultimately show ?thesis by (simp add: mult_less_cancel_right)
   942   qed
   943   also have "\<dots> \<longleftrightarrow> a * \<bar>d\<bar> * sgn b < c * \<bar>b\<bar> * sgn d"
   944     by (simp add: abs_sgn mult_ac)
   945   finally show ?thesis using assms by simp
   946 qed
   947 
   948 lemma (in ordered_idom) sgn_greater [simp]:
   949   "0 < sgn a \<longleftrightarrow> 0 < a"
   950   unfolding sgn_if by auto
   951 
   952 lemma (in ordered_idom) sgn_less [simp]:
   953   "sgn a < 0 \<longleftrightarrow> a < 0"
   954   unfolding sgn_if by auto
   955 
   956 lemma rat_le_eq_code [code]:
   957   "Fract a b < Fract c d \<longleftrightarrow> (if b = 0
   958        then sgn c * sgn d > 0
   959      else if d = 0
   960        then sgn a * sgn b < 0
   961      else a * \<bar>d\<bar> * sgn b < c * \<bar>b\<bar> * sgn d)"
   962   by (auto simp add: sgn_times mult_less_0_iff zero_less_mult_iff less_rat' eq_rat simp del: less_rat)
   963 
   964 lemma rat_less_eq_code [code]:
   965   "Fract a b \<le> Fract c d \<longleftrightarrow> (if b = 0
   966        then sgn c * sgn d \<ge> 0
   967      else if d = 0
   968        then sgn a * sgn b \<le> 0
   969      else a * \<bar>d\<bar> * sgn b \<le> c * \<bar>b\<bar> * sgn d)"
   970   by (auto simp add: sgn_times mult_le_0_iff zero_le_mult_iff le_rat' eq_rat simp del: le_rat)
   971     (auto simp add: le_less not_less sgn_0_0)
   972 
   973 
   974 lemma rat_plus_code [code]:
   975   "Fract a b + Fract c d = (if b = 0
   976      then Fract c d
   977    else if d = 0
   978      then Fract a b
   979    else Fract_norm (a * d + c * b) (b * d))"
   980   by (simp add: eq_rat, simp add: Zero_rat_def)
   981 
   982 lemma rat_times_code [code]:
   983   "Fract a b * Fract c d = Fract_norm (a * c) (b * d)"
   984   by simp
   985 
   986 lemma rat_minus_code [code]:
   987   "Fract a b - Fract c d = (if b = 0
   988      then Fract (- c) d
   989    else if d = 0
   990      then Fract a b
   991    else Fract_norm (a * d - c * b) (b * d))"
   992   by (simp add: eq_rat, simp add: Zero_rat_def)
   993 
   994 lemma rat_inverse_code [code]:
   995   "inverse (Fract a b) = (if b = 0 then Fract 1 0
   996     else if a < 0 then Fract (- b) (- a)
   997     else Fract b a)"
   998   by (simp add: eq_rat)
   999 
  1000 lemma rat_divide_code [code]:
  1001   "Fract a b / Fract c d = Fract_norm (a * d) (b * c)"
  1002   by simp
  1003 
  1004 definition (in term_syntax)
  1005   valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Eval.term)" where
  1006   [code inline]: "valterm_fract k l = Code_Eval.valtermify Fract {\<cdot>} k {\<cdot>} l"
  1007 
  1008 notation fcomp (infixl "o>" 60)
  1009 notation scomp (infixl "o\<rightarrow>" 60)
  1010 
  1011 instantiation rat :: random
  1012 begin
  1013 
  1014 definition
  1015   "random i = random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
  1016      let j = Code_Index.int_of (denom + 1)
  1017      in valterm_fract num (j, \<lambda>u. Code_Eval.term_of j))))"
  1018 
  1019 instance ..
  1020 
  1021 end
  1022 
  1023 no_notation fcomp (infixl "o>" 60)
  1024 no_notation scomp (infixl "o\<rightarrow>" 60)
  1025 
  1026 hide (open) const Fract_norm
  1027 
  1028 text {* Setup for SML code generator *}
  1029 
  1030 types_code
  1031   rat ("(int */ int)")
  1032 attach (term_of) {*
  1033 fun term_of_rat (p, q) =
  1034   let
  1035     val rT = Type ("Rational.rat", [])
  1036   in
  1037     if q = 1 orelse p = 0 then HOLogic.mk_number rT p
  1038     else @{term "op / \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat"} $
  1039       HOLogic.mk_number rT p $ HOLogic.mk_number rT q
  1040   end;
  1041 *}
  1042 attach (test) {*
  1043 fun gen_rat i =
  1044   let
  1045     val p = random_range 0 i;
  1046     val q = random_range 1 (i + 1);
  1047     val g = Integer.gcd p q;
  1048     val p' = p div g;
  1049     val q' = q div g;
  1050     val r = (if one_of [true, false] then p' else ~ p',
  1051       if p' = 0 then 0 else q')
  1052   in
  1053     (r, fn () => term_of_rat r)
  1054   end;
  1055 *}
  1056 
  1057 consts_code
  1058   Fract ("(_,/ _)")
  1059 
  1060 consts_code
  1061   "of_int :: int \<Rightarrow> rat" ("\<module>rat'_of'_int")
  1062 attach {*
  1063 fun rat_of_int 0 = (0, 0)
  1064   | rat_of_int i = (i, 1);
  1065 *}
  1066 
  1067 end