Theory Rational

Up to index of Isabelle/HOL/HOL-Algebra/example_Bicomplex

theory Rational
imports Abstract_Rat
uses rat_arith.ML
begin

(*  Title: HOL/Library/Rational.thy
    ID:    $Id: Rational.thy,v 1.39 2008/04/22 06:33:17 haftmann Exp $
    Author: Markus Wenzel, TU Muenchen
*)

header {* Rational numbers *}

theory Rational
imports "~~/src/HOL/Library/Abstract_Rat"
uses ("rat_arith.ML")
begin

subsection {* Rational numbers *}

subsubsection {* Equivalence of fractions *}

definition
  fraction :: "(int × int) set" where
  "fraction = {x. snd x ≠ 0}"

definition
  ratrel :: "((int × int) × (int × int)) set" where
  "ratrel = {(x,y). snd x ≠ 0 ∧ snd y ≠ 0 ∧ fst x * snd y = fst y * snd x}"

lemma fraction_iff [simp]: "(x ∈ fraction) = (snd x ≠ 0)"
by (simp add: fraction_def)

lemma ratrel_iff [simp]:
  "((x,y) ∈ ratrel) =
   (snd x ≠ 0 ∧ snd y ≠ 0 ∧ fst x * snd y = fst y * snd x)"
by (simp add: ratrel_def)

lemma refl_ratrel: "refl fraction ratrel"
by (auto simp add: refl_def fraction_def ratrel_def)

lemma sym_ratrel: "sym ratrel"
by (simp add: ratrel_def sym_def)

lemma trans_ratrel_lemma:
  assumes 1: "a * b' = a' * b"
  assumes 2: "a' * b'' = a'' * b'"
  assumes 3: "b' ≠ (0::int)"
  shows "a * b'' = a'' * b"
proof -
  have "b' * (a * b'') = b'' * (a * b')" by simp
  also note 1
  also have "b'' * (a' * b) = b * (a' * b'')" by simp
  also note 2
  also have "b * (a'' * b') = b' * (a'' * b)" by simp
  finally have "b' * (a * b'') = b' * (a'' * b)" .
  with 3 show "a * b'' = a'' * b" by simp
qed

lemma trans_ratrel: "trans ratrel"
by (auto simp add: trans_def elim: trans_ratrel_lemma)

lemma equiv_ratrel: "equiv fraction ratrel"
by (rule equiv.intro [OF refl_ratrel sym_ratrel trans_ratrel])

lemmas equiv_ratrel_iff [iff] = eq_equiv_class_iff [OF equiv_ratrel]

lemma equiv_ratrel_iff2:
  "[|snd x ≠ 0; snd y ≠ 0|]
    ==> (ratrel `` {x} = ratrel `` {y}) = ((x,y) ∈ ratrel)"
by (rule eq_equiv_class_iff [OF equiv_ratrel], simp_all)


subsubsection {* The type of rational numbers *}

typedef (Rat) rat = "fraction//ratrel"
proof
  have "(0,1) ∈ fraction" by (simp add: fraction_def)
  thus "ratrel``{(0,1)} ∈ fraction//ratrel" by (rule quotientI)
qed

lemma ratrel_in_Rat [simp]: "snd x ≠ 0 ==> ratrel``{x} ∈ Rat"
by (simp add: Rat_def quotientI)

declare Abs_Rat_inject [simp]  Abs_Rat_inverse [simp]


definition
  Fract :: "int => int => rat" where
  [code func del]: "Fract a b = Abs_Rat (ratrel``{(a,b)})"

lemma Fract_zero:
  "Fract k 0 = Fract l 0"
  by (simp add: Fract_def ratrel_def)

theorem Rat_cases [case_names Fract, cases type: rat]:
    "(!!a b. q = Fract a b ==> b ≠ 0 ==> C) ==> C"
  by (cases q) (clarsimp simp add: Fract_def Rat_def fraction_def quotient_def)

theorem Rat_induct [case_names Fract, induct type: rat]:
    "(!!a b. b ≠ 0 ==> P (Fract a b)) ==> P q"
  by (cases q) simp


subsubsection {* Congruence lemmas *}

lemma add_congruent2:
     "(λx y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
      respects2 ratrel"
apply (rule equiv_ratrel [THEN congruent2_commuteI])
apply (simp_all add: left_distrib)
done

lemma minus_congruent:
  "(λx. ratrel``{(- fst x, snd x)}) respects ratrel"
by (simp add: congruent_def)

lemma mult_congruent2:
  "(λx y. ratrel``{(fst x * fst y, snd x * snd y)}) respects2 ratrel"
by (rule equiv_ratrel [THEN congruent2_commuteI], simp_all)

lemma inverse_congruent:
  "(λx. ratrel``{if fst x=0 then (0,1) else (snd x, fst x)}) respects ratrel"
by (auto simp add: congruent_def mult_commute)

lemma le_congruent2:
  "(λx y. {(fst x * snd y)*(snd x * snd y) ≤ (fst y * snd x)*(snd x * snd y)})
   respects2 ratrel"
proof (clarsimp simp add: congruent2_def)
  fix a b a' b' c d c' d'::int
  assume neq: "b ≠ 0"  "b' ≠ 0"  "d ≠ 0"  "d' ≠ 0"
  assume eq1: "a * b' = a' * b"
  assume eq2: "c * d' = c' * d"

  let ?le = "λa b c d. ((a * d) * (b * d) ≤ (c * b) * (b * d))"
  {
    fix a b c d x :: int assume x: "x ≠ 0"
    have "?le a b c d = ?le (a * x) (b * x) c d"
    proof -
      from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
      hence "?le a b c d =
          ((a * d) * (b * d) * (x * x) ≤ (c * b) * (b * d) * (x * x))"
        by (simp add: mult_le_cancel_right)
      also have "... = ?le (a * x) (b * x) c d"
        by (simp add: mult_ac)
      finally show ?thesis .
    qed
  } note le_factor = this

  let ?D = "b * d" and ?D' = "b' * d'"
  from neq have D: "?D ≠ 0" by simp
  from neq have "?D' ≠ 0" by simp
  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
    by (rule le_factor)
  also have "... = ((a * b') * ?D * ?D' * d * d' ≤ (c * d') * ?D * ?D' * b * b')"
    by (simp add: mult_ac)
  also have "... = ((a' * b) * ?D * ?D' * d * d' ≤ (c' * d) * ?D * ?D' * b * b')"
    by (simp only: eq1 eq2)
  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
    by (simp add: mult_ac)
  also from D have "... = ?le a' b' c' d'"
    by (rule le_factor [symmetric])
  finally show "?le a b c d = ?le a' b' c' d'" .
qed

lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel]
lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel]


subsubsection {* Standard operations on rational numbers *}

instantiation rat :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
begin

definition
  Zero_rat_def [code func del]: "0 = Fract 0 1"

definition
  One_rat_def [code func del]: "1 = Fract 1 1"

definition
  add_rat_def [code func del]:
   "q + r =
       Abs_Rat (\<Union>x ∈ Rep_Rat q. \<Union>y ∈ Rep_Rat r.
           ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})"

definition
  minus_rat_def [code func del]:
    "- q = Abs_Rat (\<Union>x ∈ Rep_Rat q. ratrel``{(- fst x, snd x)})"

definition
  diff_rat_def [code func del]: "q - r = q + - (r::rat)"

definition
  mult_rat_def [code func del]:
   "q * r =
       Abs_Rat (\<Union>x ∈ Rep_Rat q. \<Union>y ∈ Rep_Rat r.
           ratrel``{(fst x * fst y, snd x * snd y)})"

definition
  inverse_rat_def [code func del]:
    "inverse q =
        Abs_Rat (\<Union>x ∈ Rep_Rat q.
            ratrel``{if fst x=0 then (0,1) else (snd x, fst x)})"

definition
  divide_rat_def [code func del]: "q / r = q * inverse (r::rat)"

definition
  le_rat_def [code func del]:
   "q ≤ r <-> contents (\<Union>x ∈ Rep_Rat q. \<Union>y ∈ Rep_Rat r.
      {(fst x * snd y)*(snd x * snd y) ≤ (fst y * snd x)*(snd x * snd y)})"

definition
  less_rat_def [code func del]: "z < (w::rat) <-> z ≤ w ∧ z ≠ w"

definition
  abs_rat_def: "¦q¦ = (if q < 0 then -q else (q::rat))"

definition
  sgn_rat_def: "sgn (q::rat) = (if q=0 then 0 else if 0<q then 1 else - 1)"

instance ..

end

instantiation rat :: power
begin

primrec power_rat
where
  rat_power_0:     "q ^ 0       = (1::rat)"
  | rat_power_Suc: "q ^ (Suc n) = (q::rat) * (q ^ n)"

instance ..

end

theorem eq_rat: "b ≠ 0 ==> d ≠ 0 ==>
  (Fract a b = Fract c d) = (a * d = c * b)"
by (simp add: Fract_def)

theorem add_rat: "b ≠ 0 ==> d ≠ 0 ==>
  Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
by (simp add: Fract_def add_rat_def add_congruent2 UN_ratrel2)

theorem minus_rat: "b ≠ 0 ==> -(Fract a b) = Fract (-a) b"
by (simp add: Fract_def minus_rat_def minus_congruent UN_ratrel)

theorem diff_rat: "b ≠ 0 ==> d ≠ 0 ==>
    Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
by (simp add: diff_rat_def add_rat minus_rat)

theorem mult_rat: "b ≠ 0 ==> d ≠ 0 ==>
  Fract a b * Fract c d = Fract (a * c) (b * d)"
by (simp add: Fract_def mult_rat_def mult_congruent2 UN_ratrel2)

theorem inverse_rat: "a ≠ 0 ==> b ≠ 0 ==>
  inverse (Fract a b) = Fract b a"
by (simp add: Fract_def inverse_rat_def inverse_congruent UN_ratrel)

theorem divide_rat: "c ≠ 0 ==> b ≠ 0 ==> d ≠ 0 ==>
  Fract a b / Fract c d = Fract (a * d) (b * c)"
by (simp add: divide_rat_def inverse_rat mult_rat)

theorem le_rat: "b ≠ 0 ==> d ≠ 0 ==>
  (Fract a b ≤ Fract c d) = ((a * d) * (b * d) ≤ (c * b) * (b * d))"
by (simp add: Fract_def le_rat_def le_congruent2 UN_ratrel2)

theorem less_rat: "b ≠ 0 ==> d ≠ 0 ==>
    (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
by (simp add: less_rat_def le_rat eq_rat order_less_le)

theorem abs_rat: "b ≠ 0 ==> ¦Fract a b¦ = Fract ¦a¦ ¦b¦"
  by (simp add: abs_rat_def minus_rat Zero_rat_def less_rat eq_rat)
     (auto simp add: mult_less_0_iff zero_less_mult_iff order_le_less
                split: abs_split)


subsubsection {* The ordered field of rational numbers *}

instance rat :: field
proof
  fix q r s :: rat
  show "(q + r) + s = q + (r + s)"
    by (induct q, induct r, induct s)
       (simp add: add_rat add_ac mult_ac int_distrib)
  show "q + r = r + q"
    by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
  show "0 + q = q"
    by (induct q) (simp add: Zero_rat_def add_rat)
  show "(-q) + q = 0"
    by (induct q) (simp add: Zero_rat_def minus_rat add_rat eq_rat)
  show "q - r = q + (-r)"
    by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
  show "(q * r) * s = q * (r * s)"
    by (induct q, induct r, induct s) (simp add: mult_rat mult_ac)
  show "q * r = r * q"
    by (induct q, induct r) (simp add: mult_rat mult_ac)
  show "1 * q = q"
    by (induct q) (simp add: One_rat_def mult_rat)
  show "(q + r) * s = q * s + r * s"
    by (induct q, induct r, induct s)
       (simp add: add_rat mult_rat eq_rat int_distrib)
  show "q ≠ 0 ==> inverse q * q = 1"
    by (induct q) (simp add: inverse_rat mult_rat One_rat_def Zero_rat_def eq_rat)
  show "q / r = q * inverse r"
    by (simp add: divide_rat_def)
  show "0 ≠ (1::rat)"
    by (simp add: Zero_rat_def One_rat_def eq_rat)
qed

instance rat :: linorder
proof
  fix q r s :: rat
  {
    assume "q ≤ r" and "r ≤ s"
    show "q ≤ s"
    proof (insert prems, induct q, induct r, induct s)
      fix a b c d e f :: int
      assume neq: "b ≠ 0"  "d ≠ 0"  "f ≠ 0"
      assume 1: "Fract a b ≤ Fract c d" and 2: "Fract c d ≤ Fract e f"
      show "Fract a b ≤ Fract e f"
      proof -
        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
        have "(a * d) * (b * d) * (f * f) ≤ (c * b) * (b * d) * (f * f)"
        proof -
          from neq 1 have "(a * d) * (b * d) ≤ (c * b) * (b * d)"
            by (simp add: le_rat)
          with ff show ?thesis by (simp add: mult_le_cancel_right)
        qed
        also have "... = (c * f) * (d * f) * (b * b)"
          by (simp only: mult_ac)
        also have "... ≤ (e * d) * (d * f) * (b * b)"
        proof -
          from neq 2 have "(c * f) * (d * f) ≤ (e * d) * (d * f)"
            by (simp add: le_rat)
          with bb show ?thesis by (simp add: mult_le_cancel_right)
        qed
        finally have "(a * f) * (b * f) * (d * d) ≤ e * b * (b * f) * (d * d)"
          by (simp only: mult_ac)
        with dd have "(a * f) * (b * f) ≤ (e * b) * (b * f)"
          by (simp add: mult_le_cancel_right)
        with neq show ?thesis by (simp add: le_rat)
      qed
    qed
  next
    assume "q ≤ r" and "r ≤ q"
    show "q = r"
    proof (insert prems, induct q, induct r)
      fix a b c d :: int
      assume neq: "b ≠ 0"  "d ≠ 0"
      assume 1: "Fract a b ≤ Fract c d" and 2: "Fract c d ≤ Fract a b"
      show "Fract a b = Fract c d"
      proof -
        from neq 1 have "(a * d) * (b * d) ≤ (c * b) * (b * d)"
          by (simp add: le_rat)
        also have "... ≤ (a * d) * (b * d)"
        proof -
          from neq 2 have "(c * b) * (d * b) ≤ (a * d) * (d * b)"
            by (simp add: le_rat)
          thus ?thesis by (simp only: mult_ac)
        qed
        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
        moreover from neq have "b * d ≠ 0" by simp
        ultimately have "a * d = c * b" by simp
        with neq show ?thesis by (simp add: eq_rat)
      qed
    qed
  next
    show "q ≤ q"
      by (induct q) (simp add: le_rat)
    show "(q < r) = (q ≤ r ∧ q ≠ r)"
      by (simp only: less_rat_def)
    show "q ≤ r ∨ r ≤ q"
      by (induct q, induct r)
         (simp add: le_rat mult_commute, rule linorder_linear)
  }
qed

instantiation rat :: distrib_lattice
begin

definition
  "(inf :: rat => rat => rat) = min"

definition
  "(sup :: rat => rat => rat) = max"

instance
  by default (auto simp add: min_max.sup_inf_distrib1 inf_rat_def sup_rat_def)

end

instance rat :: ordered_field
proof
  fix q r s :: rat
  show "q ≤ r ==> s + q ≤ s + r"
  proof (induct q, induct r, induct s)
    fix a b c d e f :: int
    assume neq: "b ≠ 0"  "d ≠ 0"  "f ≠ 0"
    assume le: "Fract a b ≤ Fract c d"
    show "Fract e f + Fract a b ≤ Fract e f + Fract c d"
    proof -
      let ?F = "f * f" from neq have F: "0 < ?F"
        by (auto simp add: zero_less_mult_iff)
      from neq le have "(a * d) * (b * d) ≤ (c * b) * (b * d)"
        by (simp add: le_rat)
      with F have "(a * d) * (b * d) * ?F * ?F ≤ (c * b) * (b * d) * ?F * ?F"
        by (simp add: mult_le_cancel_right)
      with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib)
    qed
  qed
  show "q < r ==> 0 < s ==> s * q < s * r"
  proof (induct q, induct r, induct s)
    fix a b c d e f :: int
    assume neq: "b ≠ 0"  "d ≠ 0"  "f ≠ 0"
    assume le: "Fract a b < Fract c d"
    assume gt: "0 < Fract e f"
    show "Fract e f * Fract a b < Fract e f * Fract c d"
    proof -
      let ?E = "e * f" and ?F = "f * f"
      from neq gt have "0 < ?E"
        by (auto simp add: Zero_rat_def less_rat le_rat order_less_le eq_rat)
      moreover from neq have "0 < ?F"
        by (auto simp add: zero_less_mult_iff)
      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
        by (simp add: less_rat)
      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
        by (simp add: mult_less_cancel_right)
      with neq show ?thesis
        by (simp add: less_rat mult_rat mult_ac)
    qed
  qed
  show "¦q¦ = (if q < 0 then -q else q)"
    by (simp only: abs_rat_def)
qed (auto simp: sgn_rat_def)

instance rat :: division_by_zero
proof
  show "inverse 0 = (0::rat)"
    by (simp add: Zero_rat_def Fract_def inverse_rat_def
                  inverse_congruent UN_ratrel)
qed

instance rat :: recpower
proof
  fix q :: rat
  fix n :: nat
  show "q ^ 0 = 1" by simp
  show "q ^ (Suc n) = q * (q ^ n)" by simp
qed


subsection {* Various Other Results *}

lemma minus_rat_cancel [simp]: "b ≠ 0 ==> Fract (-a) (-b) = Fract a b"
by (simp add: eq_rat)

theorem Rat_induct_pos [case_names Fract, induct type: rat]:
  assumes step: "!!a b. 0 < b ==> P (Fract a b)"
    shows "P q"
proof (cases q)
  have step': "!!a b. b < 0 ==> P (Fract a b)"
  proof -
    fix a::int and b::int
    assume b: "b < 0"
    hence "0 < -b" by simp
    hence "P (Fract (-a) (-b))" by (rule step)
    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
  qed
  case (Fract a b)
  thus "P q" by (force simp add: linorder_neq_iff step step')
qed

lemma zero_less_Fract_iff:
     "0 < b ==> (0 < Fract a b) = (0 < a)"
by (simp add: Zero_rat_def less_rat order_less_imp_not_eq2 zero_less_mult_iff)

lemma Fract_add_one: "n ≠ 0 ==> Fract (m + n) n = Fract m n + 1"
apply (insert add_rat [of concl: m n 1 1])
apply (simp add: One_rat_def [symmetric])
done

lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1"
by (induct k) (simp_all add: Zero_rat_def One_rat_def add_rat)

lemma of_int_rat: "of_int k = Fract k 1"
by (cases k rule: int_diff_cases, simp add: of_nat_rat diff_rat)

lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
by (rule of_nat_rat [symmetric])

lemma Fract_of_int_eq: "Fract k 1 = of_int k"
by (rule of_int_rat [symmetric])

lemma Fract_of_int_quotient: "Fract k l = (if l = 0 then Fract 1 0 else of_int k / of_int l)"
by (auto simp add: Fract_zero Fract_of_int_eq [symmetric] divide_rat)


subsection {* Numerals and Arithmetic *}

instantiation rat :: number_ring
begin

definition
  rat_number_of_def [code func del]: "number_of w = (of_int w :: rat)"

instance
  by default (simp add: rat_number_of_def)

end 

use "rat_arith.ML"
declaration {* K rat_arith_setup *}


subsection {* Embedding from Rationals to other Fields *}

class field_char_0 = field + ring_char_0

instance ordered_field < field_char_0 .. 

definition
  of_rat :: "rat => 'a::field_char_0"
where
  [code func del]: "of_rat q = contents (\<Union>(a,b) ∈ Rep_Rat q. {of_int a / of_int b})"

lemma of_rat_congruent:
  "(λ(a, b). {of_int a / of_int b::'a::field_char_0}) respects ratrel"
apply (rule congruent.intro)
apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
apply (simp only: of_int_mult [symmetric])
done

lemma of_rat_rat:
  "b ≠ 0 ==> of_rat (Fract a b) = of_int a / of_int b"
unfolding Fract_def of_rat_def
by (simp add: UN_ratrel of_rat_congruent)

lemma of_rat_0 [simp]: "of_rat 0 = 0"
by (simp add: Zero_rat_def of_rat_rat)

lemma of_rat_1 [simp]: "of_rat 1 = 1"
by (simp add: One_rat_def of_rat_rat)

lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b"
by (induct a, induct b, simp add: add_rat of_rat_rat add_frac_eq)

lemma of_rat_minus: "of_rat (- a) = - of_rat a"
by (induct a, simp add: minus_rat of_rat_rat)

lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
by (simp only: diff_minus of_rat_add of_rat_minus)

lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
apply (induct a, induct b, simp add: mult_rat of_rat_rat)
apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac)
done

lemma nonzero_of_rat_inverse:
  "a ≠ 0 ==> of_rat (inverse a) = inverse (of_rat a)"
apply (rule inverse_unique [symmetric])
apply (simp add: of_rat_mult [symmetric])
done

lemma of_rat_inverse:
  "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
   inverse (of_rat a)"
by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)

lemma nonzero_of_rat_divide:
  "b ≠ 0 ==> of_rat (a / b) = of_rat a / of_rat b"
by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)

lemma of_rat_divide:
  "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
   = of_rat a / of_rat b"
by (cases "b = 0", simp_all add: nonzero_of_rat_divide)

lemma of_rat_power:
  "(of_rat (a ^ n)::'a::{field_char_0,recpower}) = of_rat a ^ n"
by (induct n) (simp_all add: of_rat_mult power_Suc)

lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)"
apply (induct a, induct b)
apply (simp add: of_rat_rat eq_rat)
apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
apply (simp only: of_int_mult [symmetric] of_int_eq_iff)
done

lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified]

lemma of_rat_eq_id [simp]: "of_rat = (id :: rat => rat)"
proof
  fix a
  show "of_rat a = id a"
  by (induct a)
     (simp add: of_rat_rat divide_rat Fract_of_int_eq [symmetric])
qed

text{*Collapse nested embeddings*}
lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n"
by (induct n) (simp_all add: of_rat_add)

lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z"
by (cases z rule: int_diff_cases, simp add: of_rat_diff)

lemma of_rat_number_of_eq [simp]:
  "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})"
by (simp add: number_of_eq)

lemmas zero_rat = Zero_rat_def
lemmas one_rat = One_rat_def

abbreviation
  rat_of_nat :: "nat => rat"
where
  "rat_of_nat ≡ of_nat"

abbreviation
  rat_of_int :: "int => rat"
where
  "rat_of_int ≡ of_int"


subsection {* Implementation of rational numbers as pairs of integers *}

definition
  Rational :: "int × int => rat"
where
  "Rational = INum"

code_datatype Rational

lemma Rational_simp:
  "Rational (k, l) = rat_of_int k / rat_of_int l"
  unfolding Rational_def INum_def by simp

lemma Rational_zero [simp]: "Rational 0N = 0"
  by (simp add: Rational_simp)

lemma Rational_lit [simp]: "Rational iN = rat_of_int i"
  by (simp add: Rational_simp)

lemma zero_rat_code [code, code unfold]:
  "0 = Rational 0N" by simp
declare zero_rat_code [symmetric, code post]

lemma one_rat_code [code, code unfold]:
  "1 = Rational 1N" by simp
declare one_rat_code [symmetric, code post]

lemma [code unfold, symmetric, code post]:
  "number_of k = rat_of_int (number_of k)"
  by (simp add: number_of_is_id rat_number_of_def)

definition
  [code func del]: "Fract' (b::bool) k l = Fract k l"

lemma [code]:
  "Fract k l = Fract' (l ≠ 0) k l"
  unfolding Fract'_def ..

lemma [code]:
  "Fract' True k l = (if l ≠ 0 then Rational (k, l) else Fract 1 0)"
  by (simp add: Fract'_def Rational_simp Fract_of_int_quotient [of k l])

lemma [code]:
  "of_rat (Rational (k, l)) = (if l ≠ 0 then of_int k / of_int l else 0)"
  by (cases "l = 0")
    (auto simp add: Rational_simp of_rat_rat [simplified Fract_of_int_quotient [of k l], symmetric])

instantiation rat :: eq
begin

definition [code func del]: "eq_class.eq (r::rat) s <-> r = s"

instance by default (simp add: eq_rat_def)

lemma rat_eq_code [code]: "eq_class.eq (Rational x) (Rational y) <-> eq_class.eq (normNum x) (normNum y)"
  unfolding Rational_def INum_normNum_iff eq ..

end

lemma rat_less_eq_code [code]: "Rational x ≤ Rational y <-> normNum x ≤N normNum y"
proof -
  have "normNum x ≤N normNum y <-> Rational (normNum x) ≤ Rational (normNum y)" 
    by (simp add: Rational_def del: normNum)
  also have "… = (Rational x ≤ Rational y)" by (simp add: Rational_def)
  finally show ?thesis by simp
qed

lemma rat_less_code [code]: "Rational x < Rational y <-> normNum x <N normNum y"
proof -
  have "normNum x <N normNum y <-> Rational (normNum x) < Rational (normNum y)" 
    by (simp add: Rational_def del: normNum)
  also have "… = (Rational x < Rational y)" by (simp add: Rational_def)
  finally show ?thesis by simp
qed

lemma rat_add_code [code]: "Rational x + Rational y = Rational (x +N y)"
  unfolding Rational_def by simp

lemma rat_mul_code [code]: "Rational x * Rational y = Rational (x *N y)"
  unfolding Rational_def by simp

lemma rat_neg_code [code]: "- Rational x = Rational (~N x)"
  unfolding Rational_def by simp

lemma rat_sub_code [code]: "Rational x - Rational y = Rational (x -N y)"
  unfolding Rational_def by simp

lemma rat_inv_code [code]: "inverse (Rational x) = Rational (Ninv x)"
  unfolding Rational_def Ninv divide_rat_def by simp

lemma rat_div_code [code]: "Rational x / Rational y = Rational (x ÷N y)"
  unfolding Rational_def by simp

text {* Setup for SML code generator *}

types_code
  rat ("(int */ int)")
attach (term_of) {*
fun term_of_rat (p, q) =
  let
    val rT = Type ("Rational.rat", [])
  in
    if q = 1 orelse p = 0 then HOLogic.mk_number rT p
    else @{term "op / :: rat => rat => rat"} $
      HOLogic.mk_number rT p $ HOLogic.mk_number rT q
  end;
*}
attach (test) {*
fun gen_rat i =
  let
    val p = random_range 0 i;
    val q = random_range 1 (i + 1);
    val g = Integer.gcd p q;
    val p' = p div g;
    val q' = q div g;
    val r = (if one_of [true, false] then p' else ~ p',
      if p' = 0 then 0 else q')
  in
    (r, fn () => term_of_rat r)
  end;
*}

consts_code
  Rational ("(_)")

consts_code
  "of_int :: int => rat" ("\<module>rat'_of'_int")
attach {*
fun rat_of_int 0 = (0, 0)
  | rat_of_int i = (i, 1);
*}

end

Rational numbers

Equivalence of fractions

lemma fraction_iff:

  (xfraction) = (snd x  0)

lemma ratrel_iff:

  ((x, y) ∈ ratrel) = (snd x  0 ∧ snd y  0 ∧ fst x * snd y = fst y * snd x)

lemma refl_ratrel:

  refl fraction ratrel

lemma sym_ratrel:

  sym ratrel

lemma trans_ratrel_lemma:

  [| a * b' = a' * b; a' * b'' = a'' * b'; b'  0 |] ==> a * b'' = a'' * b

lemma trans_ratrel:

  trans ratrel

lemma equiv_ratrel:

  equiv fraction ratrel

lemma equiv_ratrel_iff:

  [| xfraction; yfraction |]
  ==> (ratrel `` {x} = ratrel `` {y}) = ((x, y) ∈ ratrel)

lemma equiv_ratrel_iff2:

  [| snd x  0; snd y  0 |]
  ==> (ratrel `` {x} = ratrel `` {y}) = ((x, y) ∈ ratrel)

The type of rational numbers

lemma ratrel_in_Rat:

  snd x  0 ==> ratrel `` {x} ∈ Rat

lemma Fract_zero:

  Fract k 0 = Fract l 0

theorem Rat_cases:

  (!!a b. [| q = Fract a b; b  0 |] ==> C) ==> C

theorem Rat_induct:

  (!!a b. b  0 ==> P (Fract a b)) ==> P q

Congruence lemmas

lemma add_congruent2:

  x y. ratrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)}) respects2
  ratrel

lemma minus_congruent:

  x. ratrel `` {(- fst x, snd x)}) respects ratrel

lemma mult_congruent2:

  x y. ratrel `` {(fst x * fst y, snd x * snd y)}) respects2 ratrel

lemma inverse_congruent:

  x. ratrel `` {if fst x = 0 then 0N else (snd x, fst x)}) respects ratrel

lemma le_congruent2:

  x y. {fst x * snd y * (snd x * snd y)
           fst y * snd x * (snd x * snd y)}) respects2
  ratrel

lemma UN_ratrel:

  [| f respects ratrel; afraction |] ==> (UN x:ratrel `` {a}. f x) = f a

lemma UN_ratrel2:

  [| f respects2 ratrel; a1.0fraction; a2.0fraction |]
  ==> (UN x1:ratrel `` {a1.0}. UN x2:ratrel `` {a2.0}. f x1 x2) = f a1.0 a2.0

Standard operations on rational numbers

theorem eq_rat:

  [| b  0; d  0 |] ==> (Fract a b = Fract c d) = (a * d = c * b)

theorem add_rat:

  [| b  0; d  0 |] ==> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)

theorem minus_rat:

  b  0 ==> - Fract a b = Fract (- a) b

theorem diff_rat:

  [| b  0; d  0 |] ==> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)

theorem mult_rat:

  [| b  0; d  0 |] ==> Fract a b * Fract c d = Fract (a * c) (b * d)

theorem inverse_rat:

  [| a  0; b  0 |] ==> inverse (Fract a b) = Fract b a

theorem divide_rat:

  [| c  0; b  0; d  0 |] ==> Fract a b / Fract c d = Fract (a * d) (b * c)

theorem le_rat:

  [| b  0; d  0 |]
  ==> (Fract a b  Fract c d) = (a * d * (b * d)  c * b * (b * d))

theorem less_rat:

  [| b  0; d  0 |]
  ==> (Fract a b < Fract c d) = (a * d * (b * d) < c * b * (b * d))

theorem abs_rat:

  b  0 ==> ¦Fract a b¦ = Fract ¦a¦ ¦b¦

The ordered field of rational numbers

Various Other Results

lemma minus_rat_cancel:

  b  0 ==> Fract (- a) (- b) = Fract a b

theorem Rat_induct_pos:

  (!!a b. 0 < b ==> P (Fract a b)) ==> P q

lemma zero_less_Fract_iff:

  0 < b ==> (0 < Fract a b) = (0 < a)

lemma Fract_add_one:

  n  0 ==> Fract (m + n) n = Fract m n + 1

lemma of_nat_rat:

  of_nat k = Fract (int k) 1

lemma of_int_rat:

  of_int k = Fract k 1

lemma Fract_of_nat_eq:

  Fract (int k) 1 = of_nat k

lemma Fract_of_int_eq:

  Fract k 1 = of_int k

lemma Fract_of_int_quotient:

  Fract k l = (if l = 0 then Fract 1 0 else of_int k / of_int l)

Numerals and Arithmetic

Embedding from Rationals to other Fields

lemma of_rat_congruent:

  (λ(a, b). {of_int a / of_int b}) respects ratrel

lemma of_rat_rat:

  b  0 ==> of_rat (Fract a b) = of_int a / of_int b

lemma of_rat_0:

  of_rat 0 = (0::'a)

lemma of_rat_1:

  of_rat 1 = (1::'a)

lemma of_rat_add:

  of_rat (a + b) = of_rat a + of_rat b

lemma of_rat_minus:

  of_rat (- a) = - of_rat a

lemma of_rat_diff:

  of_rat (a - b) = of_rat a - of_rat b

lemma of_rat_mult:

  of_rat (a * b) = of_rat a * of_rat b

lemma nonzero_of_rat_inverse:

  a  0 ==> of_rat (inverse a) = inverse (of_rat a)

lemma of_rat_inverse:

  of_rat (inverse a) = inverse (of_rat a)

lemma nonzero_of_rat_divide:

  b  0 ==> of_rat (a / b) = of_rat a / of_rat b

lemma of_rat_divide:

  of_rat (a / b) = of_rat a / of_rat b

lemma of_rat_power:

  of_rat (a ^ n) = of_rat a ^ n

lemma of_rat_eq_iff:

  (of_rat a = of_rat b) = (a = b)

lemma of_rat_eq_0_iff:

  (of_rat a = (0::'a)) = (a = 0)

lemma of_rat_eq_id:

  of_rat = id

lemma of_rat_of_nat_eq:

  of_rat (of_nat n) = of_nat n

lemma of_rat_of_int_eq:

  of_rat (of_int z) = of_int z

lemma of_rat_number_of_eq:

  of_rat (number_of w) = number_of w

lemma zero_rat:

  0 = Fract 0 1

lemma one_rat:

  1 = Fract 1 1

Implementation of rational numbers as pairs of integers

lemma Rational_simp:

  Rational (k, l) = rat_of_int k / rat_of_int l

lemma Rational_zero:

  Rational 0N = 0

lemma Rational_lit:

  Rational iN = rat_of_int i

lemma zero_rat_code:

  0 = Rational 0N

lemma one_rat_code:

  1 = Rational 1N

lemma

  rat_of_int (number_of k) = number_of k

lemma

  Fract k l = Fract' (l  0) k l

lemma

  Fract' True k l = (if l  0 then Rational (k, l) else Fract 1 0)

lemma

  of_rat (Rational (k, l)) = (if l  0 then of_int k / of_int l else 0::'a)

lemma rat_eq_code:

  eq_class.eq (Rational x) (Rational y) = eq_class.eq (normNum x) (normNum y)

lemma rat_less_eq_code:

  (Rational x  Rational y) = normNum x N normNum y

lemma rat_less_code:

  (Rational x < Rational y) = normNum x <N normNum y

lemma rat_add_code:

  Rational x + Rational y = Rational (x +N y)

lemma rat_mul_code:

  Rational x * Rational y = Rational (x *N y)

lemma rat_neg_code:

  - Rational x = Rational (~N x)

lemma rat_sub_code:

  Rational x - Rational y = Rational (x -N y)

lemma rat_inv_code:

  inverse (Rational x) = Rational (Ninv x)

lemma rat_div_code:

  Rational x / Rational y = Rational (x ÷N y)