src/HOL/Library/Fraction_Field.thy
author haftmann
Fri, 23 Apr 2010 16:45:53 +0200
changeset 36312 26eea417ccc4
parent 35372 ca158c7b1144
child 36331 6b9e487450ba
permissions -rw-r--r--
separated instantiation of division_by_zero
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35372
ca158c7b1144 renamed theory Rational to Rat
haftmann
parents: 31998
diff changeset
     1
(*  Title:      HOL/Library/Fraction_Field.thy
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
     3
*)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
     4
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
     5
header{* A formalization of the fraction field of any integral domain 
35372
ca158c7b1144 renamed theory Rational to Rat
haftmann
parents: 31998
diff changeset
     6
         A generalization of Rat.thy from int to any integral domain *}
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
     7
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
     8
theory Fraction_Field
35372
ca158c7b1144 renamed theory Rational to Rat
haftmann
parents: 31998
diff changeset
     9
imports Main
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    10
begin
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    11
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    12
subsection {* General fractions construction *}
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    13
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    14
subsubsection {* Construction of the type of fractions *}
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    15
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    16
definition fractrel :: "(('a::idom * 'a ) * ('a * 'a)) set" where
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    17
  "fractrel == {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    18
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    19
lemma fractrel_iff [simp]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    20
  "(x, y) \<in> fractrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    21
  by (simp add: fractrel_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    22
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    23
lemma refl_fractrel: "refl_on {x. snd x \<noteq> 0} fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    24
  by (auto simp add: refl_on_def fractrel_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    25
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    26
lemma sym_fractrel: "sym fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    27
  by (simp add: fractrel_def sym_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    28
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    29
lemma trans_fractrel: "trans fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    30
proof (rule transI, unfold split_paired_all)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    31
  fix a b a' b' a'' b'' :: 'a
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    32
  assume A: "((a, b), (a', b')) \<in> fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    33
  assume B: "((a', b'), (a'', b'')) \<in> fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    34
  have "b' * (a * b'') = b'' * (a * b')" by (simp add: mult_ac)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    35
  also from A have "a * b' = a' * b" by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    36
  also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: mult_ac)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    37
  also from B have "a' * b'' = a'' * b'" by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    38
  also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: mult_ac)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    39
  finally have "b' * (a * b'') = b' * (a'' * b)" .
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    40
  moreover from B have "b' \<noteq> 0" by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    41
  ultimately have "a * b'' = a'' * b" by simp
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    42
  with A B show "((a, b), (a'', b'')) \<in> fractrel" by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    43
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    44
  
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    45
lemma equiv_fractrel: "equiv {x. snd x \<noteq> 0} fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    46
  by (rule equiv.intro [OF refl_fractrel sym_fractrel trans_fractrel])
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    47
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    48
lemmas UN_fractrel = UN_equiv_class [OF equiv_fractrel]
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    49
lemmas UN_fractrel2 = UN_equiv_class2 [OF equiv_fractrel equiv_fractrel]
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    50
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    51
lemma equiv_fractrel_iff [iff]: 
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    52
  assumes "snd x \<noteq> 0" and "snd y \<noteq> 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    53
  shows "fractrel `` {x} = fractrel `` {y} \<longleftrightarrow> (x, y) \<in> fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    54
  by (rule eq_equiv_class_iff, rule equiv_fractrel) (auto simp add: assms)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    55
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    56
typedef 'a fract = "{(x::'a\<times>'a). snd x \<noteq> (0::'a::idom)} // fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    57
proof
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    58
  have "(0::'a, 1::'a) \<in> {x. snd x \<noteq> 0}" by simp
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    59
  then show "fractrel `` {(0::'a, 1)} \<in> {x. snd x \<noteq> 0} // fractrel" by (rule quotientI)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    60
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    61
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    62
lemma fractrel_in_fract [simp]: "snd x \<noteq> 0 \<Longrightarrow> fractrel `` {x} \<in> fract"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    63
  by (simp add: fract_def quotientI)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    64
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    65
declare Abs_fract_inject [simp] Abs_fract_inverse [simp]
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    66
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    67
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    68
subsubsection {* Representation and basic operations *}
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    69
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    70
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    71
  Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" where
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    72
  [code del]: "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    73
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    74
code_datatype Fract
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    75
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    76
lemma Fract_cases [case_names Fract, cases type: fract]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    77
  assumes "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> C"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    78
  shows C
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    79
  using assms by (cases q) (clarsimp simp add: Fract_def fract_def quotient_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    80
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    81
lemma Fract_induct [case_names Fract, induct type: fract]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    82
  assumes "\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    83
  shows "P q"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    84
  using assms by (cases q) simp
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    85
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    86
lemma eq_fract:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    87
  shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    88
  and "\<And>a. Fract a 0 = Fract 0 1"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    89
  and "\<And>a c. Fract 0 a = Fract 0 c"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    90
  by (simp_all add: Fract_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    91
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    92
instantiation fract :: (idom) "{comm_ring_1, power}"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    93
begin
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    94
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    95
definition
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31761
diff changeset
    96
  Zero_fract_def [code, code_unfold]: "0 = Fract 0 1"
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    97
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
    98
definition
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31761
diff changeset
    99
  One_fract_def [code, code_unfold]: "1 = Fract 1 1"
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   100
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   101
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   102
  add_fract_def [code del]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   103
  "q + r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   104
    fractrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   105
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   106
lemma add_fract [simp]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   107
  assumes "b \<noteq> (0::'a::idom)" and "d \<noteq> 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   108
  shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   109
proof -
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   110
  have "(\<lambda>x y. fractrel``{(fst x * snd y + fst y * snd x, snd x * snd y :: 'a)})
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   111
    respects2 fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   112
  apply (rule equiv_fractrel [THEN congruent2_commuteI]) apply (auto simp add: algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   113
  unfolding mult_assoc[symmetric] .
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   114
  with assms show ?thesis by (simp add: Fract_def add_fract_def UN_fractrel2)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   115
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   116
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   117
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   118
  minus_fract_def [code del]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   119
  "- q = Abs_fract (\<Union>x \<in> Rep_fract q. fractrel `` {(- fst x, snd x)})"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   120
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   121
lemma minus_fract [simp, code]: "- Fract a b = Fract (- a) (b::'a::idom)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   122
proof -
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   123
  have "(\<lambda>x. fractrel `` {(- fst x, snd x :: 'a)}) respects fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   124
    by (simp add: congruent_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   125
  then show ?thesis by (simp add: Fract_def minus_fract_def UN_fractrel)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   126
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   127
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   128
lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   129
  by (cases "b = 0") (simp_all add: eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   130
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   131
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   132
  diff_fract_def [code del]: "q - r = q + - (r::'a fract)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   133
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   134
lemma diff_fract [simp]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   135
  assumes "b \<noteq> 0" and "d \<noteq> 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   136
  shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   137
  using assms by (simp add: diff_fract_def diff_minus)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   138
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   139
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   140
  mult_fract_def [code del]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   141
  "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   142
    fractrel``{(fst x * fst y, snd x * snd y)})"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   143
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   144
lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   145
proof -
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   146
  have "(\<lambda>x y. fractrel `` {(fst x * fst y, snd x * snd y :: 'a)}) respects2 fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   147
    apply (rule equiv_fractrel [THEN congruent2_commuteI]) apply (auto simp add: algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   148
    unfolding mult_assoc[symmetric] .
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   149
  then show ?thesis by (simp add: Fract_def mult_fract_def UN_fractrel2)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   150
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   151
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   152
lemma mult_fract_cancel:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   153
  assumes "c \<noteq> 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   154
  shows "Fract (c * a) (c * b) = Fract a b"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   155
proof -
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   156
  from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   157
  then show ?thesis by (simp add: mult_fract [symmetric])
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   158
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   159
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   160
instance proof
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   161
  fix q r s :: "'a fract" show "(q * r) * s = q * (r * s)" 
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   162
    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   163
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   164
  fix q r :: "'a fract" show "q * r = r * q"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   165
    by (cases q, cases r) (simp add: eq_fract algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   166
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   167
  fix q :: "'a fract" show "1 * q = q"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   168
    by (cases q) (simp add: One_fract_def eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   169
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   170
  fix q r s :: "'a fract" show "(q + r) + s = q + (r + s)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   171
    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   172
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   173
  fix q r :: "'a fract" show "q + r = r + q"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   174
    by (cases q, cases r) (simp add: eq_fract algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   175
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   176
  fix q :: "'a fract" show "0 + q = q"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   177
    by (cases q) (simp add: Zero_fract_def eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   178
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   179
  fix q :: "'a fract" show "- q + q = 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   180
    by (cases q) (simp add: Zero_fract_def eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   181
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   182
  fix q r :: "'a fract" show "q - r = q + - r"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   183
    by (cases q, cases r) (simp add: eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   184
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   185
  fix q r s :: "'a fract" show "(q + r) * s = q * s + r * s"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   186
    by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   187
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   188
  show "(0::'a fract) \<noteq> 1" by (simp add: Zero_fract_def One_fract_def eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   189
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   190
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   191
end
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   192
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   193
lemma of_nat_fract: "of_nat k = Fract (of_nat k) 1"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   194
  by (induct k) (simp_all add: Zero_fract_def One_fract_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   195
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   196
lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   197
  by (rule of_nat_fract [symmetric])
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   198
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31761
diff changeset
   199
lemma fract_collapse [code_post]:
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   200
  "Fract 0 k = 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   201
  "Fract 1 1 = 1"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   202
  "Fract k 0 = 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   203
  by (cases "k = 0")
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   204
    (simp_all add: Zero_fract_def One_fract_def eq_fract Fract_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   205
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31761
diff changeset
   206
lemma fract_expand [code_unfold]:
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   207
  "0 = Fract 0 1"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   208
  "1 = Fract 1 1"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   209
  by (simp_all add: fract_collapse)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   210
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   211
lemma Fract_cases_nonzero [case_names Fract 0]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   212
  assumes Fract: "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> C"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   213
  assumes 0: "q = 0 \<Longrightarrow> C"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   214
  shows C
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   215
proof (cases "q = 0")
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   216
  case True then show C using 0 by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   217
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   218
  case False
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   219
  then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   220
  moreover with False have "0 \<noteq> Fract a b" by simp
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   221
  with `b \<noteq> 0` have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   222
  with Fract `q = Fract a b` `b \<noteq> 0` show C by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   223
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   224
  
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   225
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   226
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   227
subsubsection {* The field of rational numbers *}
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   228
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   229
context idom
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   230
begin
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   231
subclass ring_no_zero_divisors ..
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   232
thm mult_eq_0_iff
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   233
end
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   234
36312
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   235
instantiation fract :: (idom) field
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   236
begin
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   237
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   238
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   239
  inverse_fract_def [code del]:
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   240
  "inverse q = Abs_fract (\<Union>x \<in> Rep_fract q.
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   241
     fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   242
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   243
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   244
lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   245
proof -
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   246
  have stupid: "\<And>x. (0::'a) = x \<longleftrightarrow> x = 0" by auto
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   247
  have "(\<lambda>x. fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x :: 'a)}) respects fractrel"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   248
    by (auto simp add: congruent_def stupid algebra_simps)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   249
  then show ?thesis by (simp add: Fract_def inverse_fract_def UN_fractrel)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   250
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   251
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   252
definition
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   253
  divide_fract_def [code del]: "q / r = q * inverse (r:: 'a fract)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   254
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   255
lemma divide_fract [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   256
  by (simp add: divide_fract_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   257
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   258
instance proof
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   259
  fix q :: "'a fract"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   260
  assume "q \<noteq> 0"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   261
  then show "inverse q * q = 1" apply (cases q rule: Fract_cases_nonzero)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   262
    by (simp_all add: mult_fract  inverse_fract fract_expand eq_fract mult_commute)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   263
next
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   264
  fix q r :: "'a fract"
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   265
  show "q / r = q * inverse r" by (simp add: divide_fract_def)
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   266
qed
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   267
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   268
end
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   269
36312
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   270
instance fract :: (idom) division_by_zero
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   271
proof
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   272
  show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   273
    (simp add: fract_collapse)
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   274
qed
26eea417ccc4 separated instantiation of division_by_zero
haftmann
parents: 35372
diff changeset
   275
31761
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   276
3585bebe49a8 Added Library/Fraction_Field.thy: The fraction field of any integral
chaieb
parents:
diff changeset
   277
end