src/HOL/Real/Rational.thy
author wenzelm
Tue, 18 Sep 2007 16:08:00 +0200
changeset 24630 351a308ab58d
parent 24622 8116eb022282
child 24661 a705b9834590
permissions -rw-r--r--
simplified type int (eliminated IntInf.int, integer);
paulson@14365
     1
(*  Title: HOL/Library/Rational.thy
paulson@14365
     2
    ID:    $Id$
paulson@14365
     3
    Author: Markus Wenzel, TU Muenchen
paulson@14365
     4
*)
paulson@14365
     5
wenzelm@14691
     6
header {* Rational numbers *}
paulson@14365
     7
nipkow@15131
     8
theory Rational
berghofe@24533
     9
imports Abstract_Rat
haftmann@16417
    10
uses ("rat_arith.ML")
nipkow@15131
    11
begin
paulson@14365
    12
huffman@18913
    13
subsection {* Rational numbers *}
paulson@14365
    14
paulson@14365
    15
subsubsection {* Equivalence of fractions *}
paulson@14365
    16
wenzelm@19765
    17
definition
wenzelm@21404
    18
  fraction :: "(int \<times> int) set" where
wenzelm@19765
    19
  "fraction = {x. snd x \<noteq> 0}"
paulson@14365
    20
wenzelm@21404
    21
definition
wenzelm@21404
    22
  ratrel :: "((int \<times> int) \<times> (int \<times> int)) set" where
wenzelm@19765
    23
  "ratrel = {(x,y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
paulson@14365
    24
huffman@18913
    25
lemma fraction_iff [simp]: "(x \<in> fraction) = (snd x \<noteq> 0)"
huffman@18913
    26
by (simp add: fraction_def)
paulson@14365
    27
huffman@18913
    28
lemma ratrel_iff [simp]:
huffman@18913
    29
  "((x,y) \<in> ratrel) =
huffman@18913
    30
   (snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)"
huffman@18913
    31
by (simp add: ratrel_def)
huffman@18913
    32
huffman@18913
    33
lemma refl_ratrel: "refl fraction ratrel"
huffman@18913
    34
by (auto simp add: refl_def fraction_def ratrel_def)
huffman@18913
    35
huffman@18913
    36
lemma sym_ratrel: "sym ratrel"
huffman@18913
    37
by (simp add: ratrel_def sym_def)
huffman@18913
    38
huffman@18913
    39
lemma trans_ratrel_lemma:
huffman@18913
    40
  assumes 1: "a * b' = a' * b"
huffman@18913
    41
  assumes 2: "a' * b'' = a'' * b'"
huffman@18913
    42
  assumes 3: "b' \<noteq> (0::int)"
huffman@18913
    43
  shows "a * b'' = a'' * b"
huffman@18913
    44
proof -
huffman@18913
    45
  have "b' * (a * b'') = b'' * (a * b')" by simp
huffman@18913
    46
  also note 1
huffman@18913
    47
  also have "b'' * (a' * b) = b * (a' * b'')" by simp
huffman@18913
    48
  also note 2
huffman@18913
    49
  also have "b * (a'' * b') = b' * (a'' * b)" by simp
huffman@18913
    50
  finally have "b' * (a * b'') = b' * (a'' * b)" .
huffman@18913
    51
  with 3 show "a * b'' = a'' * b" by simp
paulson@14365
    52
qed
paulson@14365
    53
huffman@18913
    54
lemma trans_ratrel: "trans ratrel"
huffman@18913
    55
by (auto simp add: trans_def elim: trans_ratrel_lemma)
paulson@14365
    56
huffman@18913
    57
lemma equiv_ratrel: "equiv fraction ratrel"
huffman@18913
    58
by (rule equiv.intro [OF refl_ratrel sym_ratrel trans_ratrel])
paulson@14365
    59
huffman@18913
    60
lemmas equiv_ratrel_iff [iff] = eq_equiv_class_iff [OF equiv_ratrel]
paulson@14365
    61
huffman@18913
    62
lemma equiv_ratrel_iff2:
huffman@18913
    63
  "\<lbrakk>snd x \<noteq> 0; snd y \<noteq> 0\<rbrakk>
huffman@18913
    64
    \<Longrightarrow> (ratrel `` {x} = ratrel `` {y}) = ((x,y) \<in> ratrel)"
huffman@18913
    65
by (rule eq_equiv_class_iff [OF equiv_ratrel], simp_all)
paulson@14365
    66
paulson@14365
    67
huffman@18913
    68
subsubsection {* The type of rational numbers *}
paulson@14365
    69
huffman@18913
    70
typedef (Rat) rat = "fraction//ratrel"
huffman@18913
    71
proof
huffman@18913
    72
  have "(0,1) \<in> fraction" by (simp add: fraction_def)
huffman@18913
    73
  thus "ratrel``{(0,1)} \<in> fraction//ratrel" by (rule quotientI)
paulson@14365
    74
qed
paulson@14365
    75
huffman@18913
    76
lemma ratrel_in_Rat [simp]: "snd x \<noteq> 0 \<Longrightarrow> ratrel``{x} \<in> Rat"
huffman@18913
    77
by (simp add: Rat_def quotientI)
paulson@14365
    78
huffman@18913
    79
declare Abs_Rat_inject [simp]  Abs_Rat_inverse [simp]
huffman@18913
    80
huffman@18913
    81
wenzelm@19765
    82
definition
wenzelm@21404
    83
  Fract :: "int \<Rightarrow> int \<Rightarrow> rat" where
haftmann@24198
    84
  [code func del]: "Fract a b = Abs_Rat (ratrel``{(a,b)})"
haftmann@24198
    85
haftmann@24198
    86
lemma Fract_zero:
haftmann@24198
    87
  "Fract k 0 = Fract l 0"
haftmann@24198
    88
  by (simp add: Fract_def ratrel_def)
huffman@18913
    89
huffman@18913
    90
theorem Rat_cases [case_names Fract, cases type: rat]:
wenzelm@21404
    91
    "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
wenzelm@21404
    92
  by (cases q) (clarsimp simp add: Fract_def Rat_def fraction_def quotient_def)
huffman@18913
    93
huffman@18913
    94
theorem Rat_induct [case_names Fract, induct type: rat]:
huffman@18913
    95
    "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
huffman@18913
    96
  by (cases q) simp
huffman@18913
    97
huffman@18913
    98
huffman@18913
    99
subsubsection {* Congruence lemmas *}
huffman@18913
   100
huffman@18913
   101
lemma add_congruent2:
huffman@18913
   102
     "(\<lambda>x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
huffman@18913
   103
      respects2 ratrel"
huffman@18913
   104
apply (rule equiv_ratrel [THEN congruent2_commuteI])
huffman@18913
   105
apply (simp_all add: left_distrib)
huffman@18913
   106
done
huffman@18913
   107
huffman@18913
   108
lemma minus_congruent:
huffman@18913
   109
  "(\<lambda>x. ratrel``{(- fst x, snd x)}) respects ratrel"
huffman@18913
   110
by (simp add: congruent_def)
huffman@18913
   111
huffman@18913
   112
lemma mult_congruent2:
huffman@18913
   113
  "(\<lambda>x y. ratrel``{(fst x * fst y, snd x * snd y)}) respects2 ratrel"
huffman@18913
   114
by (rule equiv_ratrel [THEN congruent2_commuteI], simp_all)
huffman@18913
   115
huffman@18913
   116
lemma inverse_congruent:
huffman@18913
   117
  "(\<lambda>x. ratrel``{if fst x=0 then (0,1) else (snd x, fst x)}) respects ratrel"
huffman@18913
   118
by (auto simp add: congruent_def mult_commute)
huffman@18913
   119
huffman@18913
   120
lemma le_congruent2:
huffman@18982
   121
  "(\<lambda>x y. {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})
huffman@18913
   122
   respects2 ratrel"
huffman@18913
   123
proof (clarsimp simp add: congruent2_def)
huffman@18913
   124
  fix a b a' b' c d c' d'::int
paulson@14365
   125
  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
huffman@18913
   126
  assume eq1: "a * b' = a' * b"
huffman@18913
   127
  assume eq2: "c * d' = c' * d"
paulson@14365
   128
paulson@14365
   129
  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
paulson@14365
   130
  {
paulson@14365
   131
    fix a b c d x :: int assume x: "x \<noteq> 0"
paulson@14365
   132
    have "?le a b c d = ?le (a * x) (b * x) c d"
paulson@14365
   133
    proof -
paulson@14365
   134
      from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
paulson@14365
   135
      hence "?le a b c d =
paulson@14365
   136
          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
paulson@14365
   137
        by (simp add: mult_le_cancel_right)
paulson@14365
   138
      also have "... = ?le (a * x) (b * x) c d"
paulson@14365
   139
        by (simp add: mult_ac)
paulson@14365
   140
      finally show ?thesis .
paulson@14365
   141
    qed
paulson@14365
   142
  } note le_factor = this
paulson@14365
   143
paulson@14365
   144
  let ?D = "b * d" and ?D' = "b' * d'"
paulson@14365
   145
  from neq have D: "?D \<noteq> 0" by simp
paulson@14365
   146
  from neq have "?D' \<noteq> 0" by simp
paulson@14365
   147
  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
paulson@14365
   148
    by (rule le_factor)
paulson@14365
   149
  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
paulson@14365
   150
    by (simp add: mult_ac)
paulson@14365
   151
  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
paulson@14365
   152
    by (simp only: eq1 eq2)
paulson@14365
   153
  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
paulson@14365
   154
    by (simp add: mult_ac)
paulson@14365
   155
  also from D have "... = ?le a' b' c' d'"
paulson@14365
   156
    by (rule le_factor [symmetric])
huffman@18913
   157
  finally show "?le a b c d = ?le a' b' c' d'" .
paulson@14365
   158
qed
paulson@14365
   159
huffman@18913
   160
lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel]
huffman@18913
   161
lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel]
paulson@14365
   162
paulson@14365
   163
paulson@14365
   164
subsubsection {* Standard operations on rational numbers *}
paulson@14365
   165
haftmann@23879
   166
instance rat :: zero
haftmann@23879
   167
  Zero_rat_def: "0 == Fract 0 1" ..
haftmann@24198
   168
lemmas [code func del] = Zero_rat_def
paulson@14365
   169
haftmann@23879
   170
instance rat :: one
haftmann@23879
   171
  One_rat_def: "1 == Fract 1 1" ..
haftmann@24198
   172
lemmas [code func del] = One_rat_def
huffman@18913
   173
haftmann@23879
   174
instance rat :: plus
huffman@18913
   175
  add_rat_def:
huffman@18913
   176
   "q + r ==
huffman@18913
   177
       Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
haftmann@23879
   178
           ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})" ..
haftmann@23879
   179
lemmas [code func del] = add_rat_def
huffman@18913
   180
haftmann@23879
   181
instance rat :: minus
huffman@18913
   182
  minus_rat_def:
huffman@18913
   183
    "- q == Abs_Rat (\<Union>x \<in> Rep_Rat q. ratrel``{(- fst x, snd x)})"
haftmann@23879
   184
  diff_rat_def:  "q - r == q + - (r::rat)" ..
haftmann@24198
   185
lemmas [code func del] = minus_rat_def diff_rat_def
huffman@18913
   186
haftmann@23879
   187
instance rat :: times
huffman@18913
   188
  mult_rat_def:
huffman@18913
   189
   "q * r ==
huffman@18913
   190
       Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
haftmann@23879
   191
           ratrel``{(fst x * fst y, snd x * snd y)})" ..
haftmann@23879
   192
lemmas [code func del] = mult_rat_def
huffman@18913
   193
haftmann@23879
   194
instance rat :: inverse
huffman@18913
   195
  inverse_rat_def:
huffman@18913
   196
    "inverse q ==
huffman@18913
   197
        Abs_Rat (\<Union>x \<in> Rep_Rat q.
huffman@18913
   198
            ratrel``{if fst x=0 then (0,1) else (snd x, fst x)})"
haftmann@23879
   199
  divide_rat_def:  "q / r == q * inverse (r::rat)" ..
haftmann@24198
   200
lemmas [code func del] = inverse_rat_def divide_rat_def
huffman@18913
   201
haftmann@23879
   202
instance rat :: ord
huffman@18913
   203
  le_rat_def:
huffman@18982
   204
   "q \<le> r == contents (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
huffman@18982
   205
      {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})"
haftmann@23879
   206
  less_rat_def: "(z < (w::rat)) == (z \<le> w & z \<noteq> w)" ..
haftmann@23879
   207
lemmas [code func del] = le_rat_def less_rat_def
huffman@18913
   208
haftmann@23879
   209
instance rat :: abs
haftmann@23879
   210
  abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)" ..
huffman@18913
   211
nipkow@24506
   212
instance rat :: sgn
nipkow@24506
   213
  sgn_rat_def: "sgn(q::rat) == (if q=0 then 0 else if 0<q then 1 else - 1)" ..
nipkow@24506
   214
haftmann@23879
   215
instance rat :: power ..
paulson@14365
   216
huffman@20522
   217
primrec (rat)
huffman@20522
   218
  rat_power_0:   "q ^ 0       = 1"
huffman@20522
   219
  rat_power_Suc: "q ^ (Suc n) = (q::rat) * (q ^ n)"
huffman@20522
   220
huffman@18913
   221
theorem eq_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
huffman@18913
   222
  (Fract a b = Fract c d) = (a * d = c * b)"
huffman@18913
   223
by (simp add: Fract_def)
paulson@14365
   224
paulson@14365
   225
theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
paulson@14365
   226
  Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
huffman@18913
   227
by (simp add: Fract_def add_rat_def add_congruent2 UN_ratrel2)
paulson@14365
   228
paulson@14365
   229
theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b"
huffman@18913
   230
by (simp add: Fract_def minus_rat_def minus_congruent UN_ratrel)
paulson@14365
   231
paulson@14365
   232
theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
paulson@14365
   233
    Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
huffman@18913
   234
by (simp add: diff_rat_def add_rat minus_rat)
paulson@14365
   235
paulson@14365
   236
theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
paulson@14365
   237
  Fract a b * Fract c d = Fract (a * c) (b * d)"
huffman@18913
   238
by (simp add: Fract_def mult_rat_def mult_congruent2 UN_ratrel2)
paulson@14365
   239
huffman@18913
   240
theorem inverse_rat: "a \<noteq> 0 ==> b \<noteq> 0 ==>
paulson@14365
   241
  inverse (Fract a b) = Fract b a"
huffman@18913
   242
by (simp add: Fract_def inverse_rat_def inverse_congruent UN_ratrel)
paulson@14365
   243
huffman@18913
   244
theorem divide_rat: "c \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==>
paulson@14365
   245
  Fract a b / Fract c d = Fract (a * d) (b * c)"
huffman@18913
   246
by (simp add: divide_rat_def inverse_rat mult_rat)
paulson@14365
   247
paulson@14365
   248
theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
paulson@14365
   249
  (Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
huffman@18982
   250
by (simp add: Fract_def le_rat_def le_congruent2 UN_ratrel2)
paulson@14365
   251
paulson@14365
   252
theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
paulson@14365
   253
    (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
huffman@18913
   254
by (simp add: less_rat_def le_rat eq_rat order_less_le)
paulson@14365
   255
paulson@14365
   256
theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
haftmann@23879
   257
  by (simp add: abs_rat_def minus_rat Zero_rat_def less_rat eq_rat)
wenzelm@14691
   258
     (auto simp add: mult_less_0_iff zero_less_mult_iff order_le_less
paulson@14365
   259
                split: abs_split)
paulson@14365
   260
paulson@14365
   261
paulson@14365
   262
subsubsection {* The ordered field of rational numbers *}
paulson@14365
   263
paulson@14365
   264
instance rat :: field
paulson@14365
   265
proof
paulson@14365
   266
  fix q r s :: rat
paulson@14365
   267
  show "(q + r) + s = q + (r + s)"
huffman@18913
   268
    by (induct q, induct r, induct s)
huffman@18913
   269
       (simp add: add_rat add_ac mult_ac int_distrib)
paulson@14365
   270
  show "q + r = r + q"
paulson@14365
   271
    by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
paulson@14365
   272
  show "0 + q = q"
haftmann@23879
   273
    by (induct q) (simp add: Zero_rat_def add_rat)
paulson@14365
   274
  show "(-q) + q = 0"
haftmann@23879
   275
    by (induct q) (simp add: Zero_rat_def minus_rat add_rat eq_rat)
paulson@14365
   276
  show "q - r = q + (-r)"
paulson@14365
   277
    by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
paulson@14365
   278
  show "(q * r) * s = q * (r * s)"
paulson@14365
   279
    by (induct q, induct r, induct s) (simp add: mult_rat mult_ac)
paulson@14365
   280
  show "q * r = r * q"
paulson@14365
   281
    by (induct q, induct r) (simp add: mult_rat mult_ac)
paulson@14365
   282
  show "1 * q = q"
haftmann@23879
   283
    by (induct q) (simp add: One_rat_def mult_rat)
paulson@14365
   284
  show "(q + r) * s = q * s + r * s"
wenzelm@14691
   285
    by (induct q, induct r, induct s)
paulson@14365
   286
       (simp add: add_rat mult_rat eq_rat int_distrib)
paulson@14365
   287
  show "q \<noteq> 0 ==> inverse q * q = 1"
haftmann@23879
   288
    by (induct q) (simp add: inverse_rat mult_rat One_rat_def Zero_rat_def eq_rat)
paulson@14430
   289
  show "q / r = q * inverse r"
wenzelm@14691
   290
    by (simp add: divide_rat_def)
paulson@14365
   291
  show "0 \<noteq> (1::rat)"
haftmann@23879
   292
    by (simp add: Zero_rat_def One_rat_def eq_rat)
paulson@14365
   293
qed
paulson@14365
   294
paulson@14365
   295
instance rat :: linorder
paulson@14365
   296
proof
paulson@14365
   297
  fix q r s :: rat
paulson@14365
   298
  {
paulson@14365
   299
    assume "q \<le> r" and "r \<le> s"
paulson@14365
   300
    show "q \<le> s"
paulson@14365
   301
    proof (insert prems, induct q, induct r, induct s)
paulson@14365
   302
      fix a b c d e f :: int
paulson@14365
   303
      assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
paulson@14365
   304
      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
paulson@14365
   305
      show "Fract a b \<le> Fract e f"
paulson@14365
   306
      proof -
paulson@14365
   307
        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
paulson@14365
   308
          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
paulson@14365
   309
        have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
paulson@14365
   310
        proof -
paulson@14365
   311
          from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
paulson@14365
   312
            by (simp add: le_rat)
paulson@14365
   313
          with ff show ?thesis by (simp add: mult_le_cancel_right)
paulson@14365
   314
        qed
paulson@14365
   315
        also have "... = (c * f) * (d * f) * (b * b)"
paulson@14365
   316
          by (simp only: mult_ac)
paulson@14365
   317
        also have "... \<le> (e * d) * (d * f) * (b * b)"
paulson@14365
   318
        proof -
paulson@14365
   319
          from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
paulson@14365
   320
            by (simp add: le_rat)
paulson@14365
   321
          with bb show ?thesis by (simp add: mult_le_cancel_right)
paulson@14365
   322
        qed
paulson@14365
   323
        finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
paulson@14365
   324
          by (simp only: mult_ac)
paulson@14365
   325
        with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
paulson@14365
   326
          by (simp add: mult_le_cancel_right)
paulson@14365
   327
        with neq show ?thesis by (simp add: le_rat)
paulson@14365
   328
      qed
paulson@14365
   329
    qed
paulson@14365
   330
  next
paulson@14365
   331
    assume "q \<le> r" and "r \<le> q"
paulson@14365
   332
    show "q = r"
paulson@14365
   333
    proof (insert prems, induct q, induct r)
paulson@14365
   334
      fix a b c d :: int
paulson@14365
   335
      assume neq: "b \<noteq> 0"  "d \<noteq> 0"
paulson@14365
   336
      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
paulson@14365
   337
      show "Fract a b = Fract c d"
paulson@14365
   338
      proof -
paulson@14365
   339
        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
paulson@14365
   340
          by (simp add: le_rat)
paulson@14365
   341
        also have "... \<le> (a * d) * (b * d)"
paulson@14365
   342
        proof -
paulson@14365
   343
          from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
paulson@14365
   344
            by (simp add: le_rat)
paulson@14365
   345
          thus ?thesis by (simp only: mult_ac)
paulson@14365
   346
        qed
paulson@14365
   347
        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
paulson@14365
   348
        moreover from neq have "b * d \<noteq> 0" by simp
paulson@14365
   349
        ultimately have "a * d = c * b" by simp
paulson@14365
   350
        with neq show ?thesis by (simp add: eq_rat)
paulson@14365
   351
      qed
paulson@14365
   352
    qed
paulson@14365
   353
  next
paulson@14365
   354
    show "q \<le> q"
paulson@14365
   355
      by (induct q) (simp add: le_rat)
paulson@14365
   356
    show "(q < r) = (q \<le> r \<and> q \<noteq> r)"
paulson@14365
   357
      by (simp only: less_rat_def)
paulson@14365
   358
    show "q \<le> r \<or> r \<le> q"
huffman@18913
   359
      by (induct q, induct r)
huffman@18913
   360
         (simp add: le_rat mult_commute, rule linorder_linear)
paulson@14365
   361
  }
paulson@14365
   362
qed
paulson@14365
   363
haftmann@22456
   364
instance rat :: distrib_lattice
haftmann@22456
   365
  "inf r s \<equiv> min r s"
haftmann@22456
   366
  "sup r s \<equiv> max r s"
haftmann@22456
   367
  by default (auto simp add: min_max.sup_inf_distrib1 inf_rat_def sup_rat_def)
haftmann@22456
   368
paulson@14365
   369
instance rat :: ordered_field
paulson@14365
   370
proof
paulson@14365
   371
  fix q r s :: rat
paulson@14365
   372
  show "q \<le> r ==> s + q \<le> s + r"
paulson@14365
   373
  proof (induct q, induct r, induct s)
paulson@14365
   374
    fix a b c d e f :: int
paulson@14365
   375
    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
paulson@14365
   376
    assume le: "Fract a b \<le> Fract c d"
paulson@14365
   377
    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
paulson@14365
   378
    proof -
paulson@14365
   379
      let ?F = "f * f" from neq have F: "0 < ?F"
paulson@14365
   380
        by (auto simp add: zero_less_mult_iff)
paulson@14365
   381
      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
paulson@14365
   382
        by (simp add: le_rat)
paulson@14365
   383
      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
paulson@14365
   384
        by (simp add: mult_le_cancel_right)
paulson@14365
   385
      with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib)
paulson@14365
   386
    qed
paulson@14365
   387
  qed
paulson@14365
   388
  show "q < r ==> 0 < s ==> s * q < s * r"
paulson@14365
   389
  proof (induct q, induct r, induct s)
paulson@14365
   390
    fix a b c d e f :: int
paulson@14365
   391
    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
paulson@14365
   392
    assume le: "Fract a b < Fract c d"
paulson@14365
   393
    assume gt: "0 < Fract e f"
paulson@14365
   394
    show "Fract e f * Fract a b < Fract e f * Fract c d"
paulson@14365
   395
    proof -
paulson@14365
   396
      let ?E = "e * f" and ?F = "f * f"
paulson@14365
   397
      from neq gt have "0 < ?E"
haftmann@23879
   398
        by (auto simp add: Zero_rat_def less_rat le_rat order_less_le eq_rat)
paulson@14365
   399
      moreover from neq have "0 < ?F"
paulson@14365
   400
        by (auto simp add: zero_less_mult_iff)
paulson@14365
   401
      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
paulson@14365
   402
        by (simp add: less_rat)
paulson@14365
   403
      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
paulson@14365
   404
        by (simp add: mult_less_cancel_right)
paulson@14365
   405
      with neq show ?thesis
paulson@14365
   406
        by (simp add: less_rat mult_rat mult_ac)
paulson@14365
   407
    qed
paulson@14365
   408
  qed
paulson@14365
   409
  show "\<bar>q\<bar> = (if q < 0 then -q else q)"
paulson@14365
   410
    by (simp only: abs_rat_def)
nipkow@24506
   411
qed (auto simp: sgn_rat_def)
paulson@14365
   412
paulson@14365
   413
instance rat :: division_by_zero
paulson@14365
   414
proof
huffman@18913
   415
  show "inverse 0 = (0::rat)"
haftmann@23879
   416
    by (simp add: Zero_rat_def Fract_def inverse_rat_def
huffman@18913
   417
                  inverse_congruent UN_ratrel)
paulson@14365
   418
qed
paulson@14365
   419
huffman@20522
   420
instance rat :: recpower
huffman@20522
   421
proof
huffman@20522
   422
  fix q :: rat
huffman@20522
   423
  fix n :: nat
huffman@20522
   424
  show "q ^ 0 = 1" by simp
huffman@20522
   425
  show "q ^ (Suc n) = q * (q ^ n)" by simp
huffman@20522
   426
qed
huffman@20522
   427
paulson@14365
   428
paulson@14365
   429
subsection {* Various Other Results *}
paulson@14365
   430
paulson@14365
   431
lemma minus_rat_cancel [simp]: "b \<noteq> 0 ==> Fract (-a) (-b) = Fract a b"
huffman@18913
   432
by (simp add: eq_rat)
paulson@14365
   433
paulson@14365
   434
theorem Rat_induct_pos [case_names Fract, induct type: rat]:
paulson@14365
   435
  assumes step: "!!a b. 0 < b ==> P (Fract a b)"
paulson@14365
   436
    shows "P q"
paulson@14365
   437
proof (cases q)
paulson@14365
   438
  have step': "!!a b. b < 0 ==> P (Fract a b)"
paulson@14365
   439
  proof -
paulson@14365
   440
    fix a::int and b::int
paulson@14365
   441
    assume b: "b < 0"
paulson@14365
   442
    hence "0 < -b" by simp
paulson@14365
   443
    hence "P (Fract (-a) (-b))" by (rule step)
paulson@14365
   444
    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
paulson@14365
   445
  qed
paulson@14365
   446
  case (Fract a b)
paulson@14365
   447
  thus "P q" by (force simp add: linorder_neq_iff step step')
paulson@14365
   448
qed
paulson@14365
   449
paulson@14365
   450
lemma zero_less_Fract_iff:
paulson@14365
   451
     "0 < b ==> (0 < Fract a b) = (0 < a)"
haftmann@23879
   452
by (simp add: Zero_rat_def less_rat order_less_imp_not_eq2 zero_less_mult_iff)
paulson@14365
   453
paulson@14378
   454
lemma Fract_add_one: "n \<noteq> 0 ==> Fract (m + n) n = Fract m n + 1"
paulson@14378
   455
apply (insert add_rat [of concl: m n 1 1])
haftmann@23879
   456
apply (simp add: One_rat_def [symmetric])
paulson@14378
   457
done
paulson@14378
   458
huffman@23429
   459
lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1"
haftmann@23879
   460
by (induct k) (simp_all add: Zero_rat_def One_rat_def add_rat)
huffman@23429
   461
huffman@23429
   462
lemma of_int_rat: "of_int k = Fract k 1"
huffman@23429
   463
by (cases k rule: int_diff_cases, simp add: of_nat_rat diff_rat)
huffman@23429
   464
paulson@14378
   465
lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
huffman@23429
   466
by (rule of_nat_rat [symmetric])
paulson@14378
   467
paulson@14378
   468
lemma Fract_of_int_eq: "Fract k 1 = of_int k"
huffman@23429
   469
by (rule of_int_rat [symmetric])
paulson@14378
   470
haftmann@24198
   471
lemma Fract_of_int_quotient: "Fract k l = (if l = 0 then Fract 1 0 else of_int k / of_int l)"
haftmann@24198
   472
by (auto simp add: Fract_zero Fract_of_int_eq [symmetric] divide_rat)
haftmann@24198
   473
paulson@14378
   474
wenzelm@14691
   475
subsection {* Numerals and Arithmetic *}
paulson@14387
   476
haftmann@22456
   477
instance rat :: number
haftmann@22456
   478
  rat_number_of_def: "(number_of w :: rat) \<equiv> of_int w" ..
paulson@14387
   479
paulson@14387
   480
instance rat :: number_ring
wenzelm@19765
   481
  by default (simp add: rat_number_of_def) 
paulson@14387
   482
paulson@14387
   483
use "rat_arith.ML"
wenzelm@24075
   484
declaration {* K rat_arith_setup *}
paulson@14387
   485
huffman@23342
   486
huffman@23342
   487
subsection {* Embedding from Rationals to other Fields *}
huffman@23342
   488
haftmann@24198
   489
class field_char_0 = field + ring_char_0
huffman@23342
   490
huffman@23342
   491
instance ordered_field < field_char_0 ..
huffman@23342
   492
huffman@23342
   493
definition
huffman@23342
   494
  of_rat :: "rat \<Rightarrow> 'a::field_char_0"
huffman@23342
   495
where
haftmann@24198
   496
  [code func del]: "of_rat q = contents (\<Union>(a,b) \<in> Rep_Rat q. {of_int a / of_int b})"
huffman@23342
   497
huffman@23342
   498
lemma of_rat_congruent:
huffman@23342
   499
  "(\<lambda>(a, b). {of_int a / of_int b::'a::field_char_0}) respects ratrel"
huffman@23342
   500
apply (rule congruent.intro)
huffman@23342
   501
apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
huffman@23342
   502
apply (simp only: of_int_mult [symmetric])
huffman@23342
   503
done
huffman@23342
   504
huffman@23342
   505
lemma of_rat_rat:
huffman@23342
   506
  "b \<noteq> 0 \<Longrightarrow> of_rat (Fract a b) = of_int a / of_int b"
huffman@23342
   507
unfolding Fract_def of_rat_def
huffman@23342
   508
by (simp add: UN_ratrel of_rat_congruent)
huffman@23342
   509
huffman@23342
   510
lemma of_rat_0 [simp]: "of_rat 0 = 0"
huffman@23342
   511
by (simp add: Zero_rat_def of_rat_rat)
huffman@23342
   512
huffman@23342
   513
lemma of_rat_1 [simp]: "of_rat 1 = 1"
huffman@23342
   514
by (simp add: One_rat_def of_rat_rat)
huffman@23342
   515
huffman@23342
   516
lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b"
huffman@23342
   517
by (induct a, induct b, simp add: add_rat of_rat_rat add_frac_eq)
huffman@23342
   518
huffman@23343
   519
lemma of_rat_minus: "of_rat (- a) = - of_rat a"
huffman@23343
   520
by (induct a, simp add: minus_rat of_rat_rat)
huffman@23343
   521
huffman@23343
   522
lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
huffman@23343
   523
by (simp only: diff_minus of_rat_add of_rat_minus)
huffman@23343
   524
huffman@23342
   525
lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
huffman@23342
   526
apply (induct a, induct b, simp add: mult_rat of_rat_rat)
huffman@23342
   527
apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac)
huffman@23342
   528
done
huffman@23342
   529
huffman@23342
   530
lemma nonzero_of_rat_inverse:
huffman@23342
   531
  "a \<noteq> 0 \<Longrightarrow> of_rat (inverse a) = inverse (of_rat a)"
huffman@23343
   532
apply (rule inverse_unique [symmetric])
huffman@23343
   533
apply (simp add: of_rat_mult [symmetric])
huffman@23342
   534
done
huffman@23342
   535
huffman@23342
   536
lemma of_rat_inverse:
huffman@23342
   537
  "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
huffman@23342
   538
   inverse (of_rat a)"
huffman@23342
   539
by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)
huffman@23342
   540
huffman@23342
   541
lemma nonzero_of_rat_divide:
huffman@23342
   542
  "b \<noteq> 0 \<Longrightarrow> of_rat (a / b) = of_rat a / of_rat b"
huffman@23342
   543
by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
huffman@23342
   544
huffman@23342
   545
lemma of_rat_divide:
huffman@23342
   546
  "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
huffman@23342
   547
   = of_rat a / of_rat b"
huffman@23342
   548
by (cases "b = 0", simp_all add: nonzero_of_rat_divide)
huffman@23342
   549
huffman@23343
   550
lemma of_rat_power:
huffman@23343
   551
  "(of_rat (a ^ n)::'a::{field_char_0,recpower}) = of_rat a ^ n"
huffman@23343
   552
by (induct n) (simp_all add: of_rat_mult power_Suc)
huffman@23343
   553
huffman@23343
   554
lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)"
huffman@23343
   555
apply (induct a, induct b)
huffman@23343
   556
apply (simp add: of_rat_rat eq_rat)
huffman@23343
   557
apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
huffman@23343
   558
apply (simp only: of_int_mult [symmetric] of_int_eq_iff)
huffman@23343
   559
done
huffman@23343
   560
huffman@23343
   561
lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified]
huffman@23343
   562
huffman@23343
   563
lemma of_rat_eq_id [simp]: "of_rat = (id :: rat \<Rightarrow> rat)"
huffman@23343
   564
proof
huffman@23343
   565
  fix a
huffman@23343
   566
  show "of_rat a = id a"
huffman@23343
   567
  by (induct a)
huffman@23343
   568
     (simp add: of_rat_rat divide_rat Fract_of_int_eq [symmetric])
huffman@23343
   569
qed
huffman@23343
   570
huffman@23343
   571
text{*Collapse nested embeddings*}
huffman@23343
   572
lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n"
huffman@23343
   573
by (induct n) (simp_all add: of_rat_add)
huffman@23343
   574
huffman@23343
   575
lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z"
huffman@23365
   576
by (cases z rule: int_diff_cases, simp add: of_rat_diff)
huffman@23343
   577
huffman@23343
   578
lemma of_rat_number_of_eq [simp]:
huffman@23343
   579
  "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})"
huffman@23343
   580
by (simp add: number_of_eq)
huffman@23343
   581
haftmann@23879
   582
lemmas zero_rat = Zero_rat_def
haftmann@23879
   583
lemmas one_rat = One_rat_def
haftmann@23879
   584
haftmann@24198
   585
abbreviation
haftmann@24198
   586
  rat_of_nat :: "nat \<Rightarrow> rat"
haftmann@24198
   587
where
haftmann@24198
   588
  "rat_of_nat \<equiv> of_nat"
haftmann@24198
   589
haftmann@24198
   590
abbreviation
haftmann@24198
   591
  rat_of_int :: "int \<Rightarrow> rat"
haftmann@24198
   592
where
haftmann@24198
   593
  "rat_of_int \<equiv> of_int"
haftmann@24198
   594
berghofe@24533
   595
berghofe@24533
   596
subsection {* Implementation of rational numbers as pairs of integers *}
berghofe@24533
   597
berghofe@24533
   598
definition
haftmann@24622
   599
  Rational :: "int \<times> int \<Rightarrow> rat"
berghofe@24533
   600
where
haftmann@24622
   601
  "Rational = INum"
berghofe@24533
   602
haftmann@24622
   603
code_datatype Rational
berghofe@24533
   604
haftmann@24622
   605
lemma Rational_simp:
haftmann@24622
   606
  "Rational (k, l) = rat_of_int k / rat_of_int l"
haftmann@24622
   607
  unfolding Rational_def INum_def by simp
berghofe@24533
   608
haftmann@24622
   609
lemma Rational_zero [simp]: "Rational 0\<^sub>N = 0"
haftmann@24622
   610
  by (simp add: Rational_simp)
berghofe@24533
   611
haftmann@24622
   612
lemma Rational_lit [simp]: "Rational i\<^sub>N = rat_of_int i"
haftmann@24622
   613
  by (simp add: Rational_simp)
berghofe@24533
   614
berghofe@24533
   615
lemma zero_rat_code [code, code unfold]:
haftmann@24622
   616
  "0 = Rational 0\<^sub>N" by simp
berghofe@24533
   617
berghofe@24533
   618
lemma zero_rat_code [code, code unfold]:
haftmann@24622
   619
  "1 = Rational 1\<^sub>N" by simp
berghofe@24533
   620
berghofe@24533
   621
lemma [code, code unfold]:
berghofe@24533
   622
  "number_of k = rat_of_int (number_of k)"
berghofe@24533
   623
  by (simp add: number_of_is_id rat_number_of_def)
berghofe@24533
   624
berghofe@24533
   625
definition
berghofe@24533
   626
  [code func del]: "Fract' (b\<Colon>bool) k l = Fract k l"
berghofe@24533
   627
berghofe@24533
   628
lemma [code]:
berghofe@24533
   629
  "Fract k l = Fract' (l \<noteq> 0) k l"
berghofe@24533
   630
  unfolding Fract'_def ..
berghofe@24533
   631
berghofe@24533
   632
lemma [code]:
haftmann@24622
   633
  "Fract' True k l = (if l \<noteq> 0 then Rational (k, l) else Fract 1 0)"
haftmann@24622
   634
  by (simp add: Fract'_def Rational_simp Fract_of_int_quotient [of k l])
berghofe@24533
   635
berghofe@24533
   636
lemma [code]:
haftmann@24622
   637
  "of_rat (Rational (k, l)) = (if l \<noteq> 0 then of_int k / of_int l else 0)"
berghofe@24533
   638
  by (cases "l = 0")
haftmann@24622
   639
    (auto simp add: Rational_simp of_rat_rat [simplified Fract_of_int_quotient [of k l], symmetric])
berghofe@24533
   640
berghofe@24533
   641
instance rat :: eq ..
berghofe@24533
   642
haftmann@24622
   643
lemma rat_eq_code [code]: "Rational x = Rational y \<longleftrightarrow> normNum x = normNum y"
haftmann@24622
   644
  unfolding Rational_def INum_normNum_iff ..
berghofe@24533
   645
haftmann@24622
   646
lemma rat_less_eq_code [code]: "Rational x \<le> Rational y \<longleftrightarrow> normNum x \<le>\<^sub>N normNum y"
berghofe@24533
   647
proof -
haftmann@24622
   648
  have "normNum x \<le>\<^sub>N normNum y \<longleftrightarrow> Rational (normNum x) \<le> Rational (normNum y)" 
haftmann@24622
   649
    by (simp add: Rational_def del: normNum)
haftmann@24622
   650
  also have "\<dots> = (Rational x \<le> Rational y)" by (simp add: Rational_def)
berghofe@24533
   651
  finally show ?thesis by simp
berghofe@24533
   652
qed
berghofe@24533
   653
haftmann@24622
   654
lemma rat_less_code [code]: "Rational x < Rational y \<longleftrightarrow> normNum x <\<^sub>N normNum y"
berghofe@24533
   655
proof -
haftmann@24622
   656
  have "normNum x <\<^sub>N normNum y \<longleftrightarrow> Rational (normNum x) < Rational (normNum y)" 
haftmann@24622
   657
    by (simp add: Rational_def del: normNum)
haftmann@24622
   658
  also have "\<dots> = (Rational x < Rational y)" by (simp add: Rational_def)
berghofe@24533
   659
  finally show ?thesis by simp
berghofe@24533
   660
qed
berghofe@24533
   661
haftmann@24622
   662
lemma rat_add_code [code]: "Rational x + Rational y = Rational (x +\<^sub>N y)"
haftmann@24622
   663
  unfolding Rational_def by simp
berghofe@24533
   664
haftmann@24622
   665
lemma rat_mul_code [code]: "Rational x * Rational y = Rational (x *\<^sub>N y)"
haftmann@24622
   666
  unfolding Rational_def by simp
berghofe@24533
   667
haftmann@24622
   668
lemma rat_neg_code [code]: "- Rational x = Rational (~\<^sub>N x)"
haftmann@24622
   669
  unfolding Rational_def by simp
berghofe@24533
   670
haftmann@24622
   671
lemma rat_sub_code [code]: "Rational x - Rational y = Rational (x -\<^sub>N y)"
haftmann@24622
   672
  unfolding Rational_def by simp
berghofe@24533
   673
haftmann@24622
   674
lemma rat_inv_code [code]: "inverse (Rational x) = Rational (Ninv x)"
haftmann@24622
   675
  unfolding Rational_def Ninv divide_rat_def by simp
berghofe@24533
   676
haftmann@24622
   677
lemma rat_div_code [code]: "Rational x / Rational y = Rational (x \<div>\<^sub>N y)"
haftmann@24622
   678
  unfolding Rational_def by simp
berghofe@24533
   679
haftmann@24622
   680
text {* Setup for SML code generator *}
berghofe@24533
   681
berghofe@24533
   682
types_code
berghofe@24533
   683
  rat ("(int */ int)")
berghofe@24533
   684
attach (term_of) {*
berghofe@24533
   685
fun term_of_rat (p, q) =
haftmann@24622
   686
  let
haftmann@24622
   687
    val rT = @{typ rat}
berghofe@24533
   688
  in
berghofe@24533
   689
    if q = 1 orelse p = 0 then HOLogic.mk_number rT p
haftmann@24622
   690
    else @{term "op / \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat"} $
berghofe@24533
   691
      HOLogic.mk_number rT p $ HOLogic.mk_number rT q
berghofe@24533
   692
  end;
berghofe@24533
   693
*}
berghofe@24533
   694
attach (test) {*
berghofe@24533
   695
fun gen_rat i =
berghofe@24533
   696
  let
berghofe@24533
   697
    val p = random_range 0 i;
berghofe@24533
   698
    val q = random_range 1 (i + 1);
berghofe@24533
   699
    val g = Integer.gcd p q;
wenzelm@24630
   700
    val p' = p div g;
wenzelm@24630
   701
    val q' = q div g;
berghofe@24533
   702
  in
berghofe@24533
   703
    (if one_of [true, false] then p' else ~ p',
berghofe@24533
   704
     if p' = 0 then 0 else q')
berghofe@24533
   705
  end;
berghofe@24533
   706
*}
berghofe@24533
   707
berghofe@24533
   708
consts_code
haftmann@24622
   709
  Rational ("(_)")
berghofe@24533
   710
berghofe@24533
   711
consts_code
berghofe@24533
   712
  "of_int :: int \<Rightarrow> rat" ("\<module>rat'_of'_int")
berghofe@24533
   713
attach {*
berghofe@24533
   714
fun rat_of_int 0 = (0, 0)
berghofe@24533
   715
  | rat_of_int i = (i, 1);
berghofe@24533
   716
*}
berghofe@24533
   717
paulson@14365
   718
end