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