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