| author | wenzelm | 
| Tue, 28 Jul 2015 21:47:03 +0200 | |
| changeset 60825 | bacfb7c45d81 | 
| parent 60500 | 903bb1495239 | 
| child 61076 | bdc1e2f0a86a | 
| child 61106 | 5bafa612ede4 | 
| permissions | -rw-r--r-- | 
| 35372 | 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 | |
| 60500 | 5 | section\<open>A formalization of the fraction field of any integral domain; | 
| 6 | generalization of theory Rat from int to any integral domain\<close> | |
| 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 | 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 | |
| 60500 | 12 | subsection \<open>General fractions construction\<close> | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 13 | |
| 60500 | 14 | subsubsection \<open>Construction of the type of fractions\<close> | 
| 31761 
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
 | 
| 46573 | 17 |   "fractrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
 | 
| 31761 
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" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 34 | have "b' * (a * b'') = b'' * (a * b')" by (simp add: ac_simps) | 
| 31761 
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 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 36 | also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: ac_simps) | 
| 31761 
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 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 38 | also have "b * (a'' * b') = b' * (a'' * b)" by (simp add: ac_simps) | 
| 31761 
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 | 
| 54463 | 44 | |
| 31761 
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"
 | 
| 40815 | 46 | by (rule equivI [OF refl_fractrel sym_fractrel trans_fractrel]) | 
| 31761 
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 | |
| 54463 | 51 | lemma equiv_fractrel_iff [iff]: | 
| 31761 
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 | |
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
40822diff
changeset | 56 | definition "fract = {(x::'a\<times>'a). snd x \<noteq> (0::'a::idom)} // fractrel"
 | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
40822diff
changeset | 57 | |
| 49834 | 58 | typedef 'a fract = "fract :: ('a * 'a::idom) set set"
 | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
40822diff
changeset | 59 | unfolding fract_def | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 60 | proof | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 61 |   have "(0::'a, 1::'a) \<in> {x. snd x \<noteq> 0}" by simp
 | 
| 54463 | 62 |   then show "fractrel `` {(0::'a, 1)} \<in> {x. snd x \<noteq> 0} // fractrel"
 | 
| 63 | by (rule quotientI) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 64 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 65 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 66 | 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 | 67 | by (simp add: fract_def quotientI) | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 68 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 69 | 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 | 70 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 71 | |
| 60500 | 72 | subsubsection \<open>Representation and basic operations\<close> | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 73 | |
| 54463 | 74 | definition Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract" | 
| 75 |   where "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
 | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 76 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 77 | code_datatype Fract | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 78 | |
| 53196 | 79 | lemma Fract_cases [cases type: fract]: | 
| 80 | obtains (Fract) a b where "q = Fract a b" "b \<noteq> 0" | |
| 81 | by (cases q) (clarsimp simp add: Fract_def fract_def quotient_def) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 82 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 83 | lemma Fract_induct [case_names Fract, induct type: fract]: | 
| 54463 | 84 | "(\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)) \<Longrightarrow> P q" | 
| 53196 | 85 | by (cases q) simp | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 86 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 87 | lemma eq_fract: | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 88 | 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" | 
| 53196 | 89 | and "\<And>a. Fract a 0 = Fract 0 1" | 
| 90 | and "\<And>a c. Fract 0 a = Fract 0 c" | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 91 | by (simp_all add: Fract_def) | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 92 | |
| 53196 | 93 | instantiation fract :: (idom) "{comm_ring_1,power}"
 | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 94 | begin | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 95 | |
| 46573 | 96 | definition Zero_fract_def [code_unfold]: "0 = Fract 0 1" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 97 | |
| 46573 | 98 | definition One_fract_def [code_unfold]: "1 = Fract 1 1" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 99 | |
| 46573 | 100 | definition add_fract_def: | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 101 | "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 | 102 |     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 | 103 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 104 | lemma add_fract [simp]: | 
| 53196 | 105 | assumes "b \<noteq> (0::'a::idom)" | 
| 106 | and "d \<noteq> 0" | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 107 | 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 | 108 | proof - | 
| 54463 | 109 |   have "(\<lambda>x y. fractrel``{(fst x * snd y + fst y * snd x, snd x * snd y :: 'a)}) respects2 fractrel"
 | 
| 110 | by (rule equiv_fractrel [THEN congruent2_commuteI]) (simp_all add: algebra_simps) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 111 | 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 | 112 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 113 | |
| 46573 | 114 | definition minus_fract_def: | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 115 |   "- 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 | 116 | |
| 54463 | 117 | lemma minus_fract [simp, code]: | 
| 118 | fixes a b :: "'a::idom" | |
| 119 | shows "- Fract a b = Fract (- a) b" | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 120 | proof - | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 121 |   have "(\<lambda>x. fractrel `` {(- fst x, snd x :: 'a)}) respects fractrel"
 | 
| 40822 | 122 | by (simp add: congruent_def split_paired_all) | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 123 | 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 | 124 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 125 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 126 | 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 | 127 | 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 | 128 | |
| 46573 | 129 | definition diff_fract_def: "q - r = q + - (r::'a fract)" | 
| 31761 
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 | lemma diff_fract [simp]: | 
| 54463 | 132 | assumes "b \<noteq> 0" | 
| 133 | and "d \<noteq> 0" | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 134 | shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53374diff
changeset | 135 | using assms by (simp add: diff_fract_def) | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 136 | |
| 46573 | 137 | definition mult_fract_def: | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 138 | "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 | 139 |     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 | 140 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 141 | 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 | 142 | proof - | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 143 |   have "(\<lambda>x y. fractrel `` {(fst x * fst y, snd x * snd y :: 'a)}) respects2 fractrel"
 | 
| 54463 | 144 | by (rule equiv_fractrel [THEN congruent2_commuteI]) (simp_all add: algebra_simps) | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 145 | 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 | 146 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 147 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 148 | lemma mult_fract_cancel: | 
| 47252 | 149 | assumes "c \<noteq> (0::'a)" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 150 | 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 | 151 | proof - | 
| 54463 | 152 | from assms have "Fract c c = Fract 1 1" | 
| 153 | by (simp add: Fract_def) | |
| 154 | then show ?thesis | |
| 155 | by (simp add: mult_fract [symmetric]) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 156 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 157 | |
| 47252 | 158 | instance | 
| 159 | proof | |
| 53196 | 160 | fix q r s :: "'a fract" | 
| 54463 | 161 | show "(q * r) * s = q * (r * s)" | 
| 31761 
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) | 
| 53196 | 163 | show "q * r = r * q" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 164 | by (cases q, cases r) (simp add: eq_fract algebra_simps) | 
| 53196 | 165 | show "1 * q = q" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 166 | by (cases q) (simp add: One_fract_def eq_fract) | 
| 53196 | 167 | show "(q + r) + s = q + (r + s)" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 168 | by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps) | 
| 53196 | 169 | show "q + r = r + q" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 170 | by (cases q, cases r) (simp add: eq_fract algebra_simps) | 
| 53196 | 171 | show "0 + q = q" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 172 | by (cases q) (simp add: Zero_fract_def eq_fract) | 
| 53196 | 173 | show "- q + q = 0" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 174 | by (cases q) (simp add: Zero_fract_def eq_fract) | 
| 53196 | 175 | show "q - r = q + - r" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 176 | by (cases q, cases r) (simp add: eq_fract) | 
| 53196 | 177 | show "(q + r) * s = q * s + r * s" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 178 | by (cases q, cases r, cases s) (simp add: eq_fract algebra_simps) | 
| 53196 | 179 | show "(0::'a fract) \<noteq> 1" | 
| 180 | by (simp add: Zero_fract_def One_fract_def eq_fract) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 181 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 182 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 183 | end | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 184 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 185 | 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 | 186 | 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 | 187 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 188 | 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 | 189 | by (rule of_nat_fract [symmetric]) | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 190 | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31761diff
changeset | 191 | lemma fract_collapse [code_post]: | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 192 | "Fract 0 k = 0" | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 193 | "Fract 1 1 = 1" | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 194 | "Fract k 0 = 0" | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 195 | by (cases "k = 0") | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 196 | (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 | 197 | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31761diff
changeset | 198 | lemma fract_expand [code_unfold]: | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 199 | "0 = Fract 0 1" | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 200 | "1 = Fract 1 1" | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 201 | by (simp_all add: fract_collapse) | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 202 | |
| 53196 | 203 | lemma Fract_cases_nonzero: | 
| 54463 | 204 | obtains (Fract) a b where "q = Fract a b" and "b \<noteq> 0" and "a \<noteq> 0" | 
| 53196 | 205 | | (0) "q = 0" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 206 | proof (cases "q = 0") | 
| 53196 | 207 | case True | 
| 208 | then show thesis using 0 by auto | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 209 | next | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 210 | case False | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 211 | then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53196diff
changeset | 212 | with False have "0 \<noteq> Fract a b" by simp | 
| 60500 | 213 | with \<open>b \<noteq> 0\<close> have "a \<noteq> 0" by (simp add: Zero_fract_def eq_fract) | 
| 214 | with Fract \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> show thesis by auto | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 215 | qed | 
| 54463 | 216 | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 217 | |
| 60500 | 218 | subsubsection \<open>The field of rational numbers\<close> | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 219 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 220 | context idom | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 221 | begin | 
| 53196 | 222 | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 223 | subclass ring_no_zero_divisors .. | 
| 53196 | 224 | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 225 | end | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 226 | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
58881diff
changeset | 227 | instantiation fract :: (idom) field | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 228 | begin | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 229 | |
| 46573 | 230 | definition inverse_fract_def: | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 231 | "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 | 232 |      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 | 233 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 234 | 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 | 235 | proof - | 
| 54463 | 236 | have *: "\<And>x. (0::'a) = x \<longleftrightarrow> x = 0" | 
| 237 | by auto | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 238 |   have "(\<lambda>x. fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x :: 'a)}) respects fractrel"
 | 
| 53196 | 239 | by (auto simp add: congruent_def * algebra_simps) | 
| 54463 | 240 | then show ?thesis | 
| 241 | by (simp add: Fract_def inverse_fract_def UN_fractrel) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 242 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 243 | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 244 | definition divide_fract_def: "q div r = q * inverse (r:: 'a fract)" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 245 | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 246 | lemma divide_fract [simp]: "Fract a b div Fract c d = Fract (a * d) (b * c)" | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 247 | by (simp add: divide_fract_def) | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 248 | |
| 47252 | 249 | instance | 
| 250 | proof | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 251 | fix q :: "'a fract" | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 252 | assume "q \<noteq> 0" | 
| 46573 | 253 | then show "inverse q * q = 1" | 
| 254 | by (cases q rule: Fract_cases_nonzero) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
54863diff
changeset | 255 | (simp_all add: fract_expand eq_fract mult.commute) | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 256 | next | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 257 | fix q r :: "'a fract" | 
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 258 | show "q div r = q * inverse r" by (simp add: divide_fract_def) | 
| 36409 | 259 | next | 
| 46573 | 260 | show "inverse 0 = (0:: 'a fract)" | 
| 261 | by (simp add: fract_expand) (simp add: fract_collapse) | |
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 262 | qed | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 263 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 264 | end | 
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 265 | |
| 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: diff
changeset | 266 | |
| 60500 | 267 | subsubsection \<open>The ordered field of fractions over an ordered idom\<close> | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 268 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 269 | lemma le_congruent2: | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 270 | "(\<lambda>x y::'a \<times> 'a::linordered_idom. | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 271 |     {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})
 | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 272 | respects2 fractrel" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 273 | proof (clarsimp simp add: congruent2_def) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 274 | fix a b a' b' c d c' d' :: 'a | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 275 | assume neq: "b \<noteq> 0" "b' \<noteq> 0" "d \<noteq> 0" "d' \<noteq> 0" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 276 | assume eq1: "a * b' = a' * b" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 277 | assume eq2: "c * d' = c' * d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 278 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 279 | let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 280 |   {
 | 
| 54463 | 281 | fix a b c d x :: 'a | 
| 282 | assume x: "x \<noteq> 0" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 283 | have "?le a b c d = ?le (a * x) (b * x) c d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 284 | proof - | 
| 54463 | 285 | from x have "0 < x * x" | 
| 286 | by (auto simp add: zero_less_mult_iff) | |
| 46573 | 287 | then have "?le a b c d = | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 288 | ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 289 | by (simp add: mult_le_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 290 | also have "... = ?le (a * x) (b * x) c d" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 291 | by (simp add: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 292 | finally show ?thesis . | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 293 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 294 | } note le_factor = this | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 295 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 296 | let ?D = "b * d" and ?D' = "b' * d'" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 297 | from neq have D: "?D \<noteq> 0" by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 298 | from neq have "?D' \<noteq> 0" by simp | 
| 46573 | 299 | then have "?le a b c d = ?le (a * ?D') (b * ?D') c d" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 300 | by (rule le_factor) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 301 | also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 302 | by (simp add: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 303 | also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 304 | by (simp only: eq1 eq2) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 305 | also have "... = ?le (a' * ?D) (b' * ?D) c' d'" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 306 | by (simp add: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 307 | also from D have "... = ?le a' b' c' d'" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 308 | by (rule le_factor [symmetric]) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 309 | finally show "?le a b c d = ?le a' b' c' d'" . | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 310 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 311 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 312 | instantiation fract :: (linordered_idom) linorder | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 313 | begin | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 314 | |
| 46573 | 315 | definition le_fract_def: | 
| 53196 | 316 | "q \<le> r \<longleftrightarrow> the_elem (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r. | 
| 317 |     {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})"
 | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 318 | |
| 46573 | 319 | definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 320 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 321 | lemma le_fract [simp]: | 
| 54463 | 322 | assumes "b \<noteq> 0" | 
| 323 | and "d \<noteq> 0" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 324 | shows "Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 53196 | 325 | by (simp add: Fract_def le_fract_def le_congruent2 UN_fractrel2 assms) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 326 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 327 | lemma less_fract [simp]: | 
| 54463 | 328 | assumes "b \<noteq> 0" | 
| 329 | and "d \<noteq> 0" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 330 | shows "Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 331 | by (simp add: less_fract_def less_le_not_le ac_simps assms) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 332 | |
| 47252 | 333 | instance | 
| 334 | proof | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 335 | fix q r s :: "'a fract" | 
| 54463 | 336 | assume "q \<le> r" and "r \<le> s" | 
| 337 | then show "q \<le> s" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 338 | proof (induct q, induct r, induct s) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 339 | fix a b c d e f :: 'a | 
| 54463 | 340 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | 
| 341 | assume 1: "Fract a b \<le> Fract c d" | |
| 342 | assume 2: "Fract c d \<le> Fract e f" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 343 | show "Fract a b \<le> Fract e f" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 344 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 345 | from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 346 | by (auto simp add: zero_less_mult_iff linorder_neq_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 347 | have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 348 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 349 | from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 350 | by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 351 | with ff show ?thesis by (simp add: mult_le_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 352 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 353 | also have "... = (c * f) * (d * f) * (b * b)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 354 | by (simp only: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 355 | also have "... \<le> (e * d) * (d * f) * (b * b)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 356 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 357 | from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 358 | by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 359 | with bb show ?thesis by (simp add: mult_le_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 360 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 361 | finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 362 | by (simp only: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 363 | with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 364 | by (simp add: mult_le_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 365 | with neq show ?thesis by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 366 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 367 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 368 | next | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 369 | fix q r :: "'a fract" | 
| 54463 | 370 | assume "q \<le> r" and "r \<le> q" | 
| 371 | then show "q = r" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 372 | proof (induct q, induct r) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 373 | fix a b c d :: 'a | 
| 54463 | 374 | assume neq: "b \<noteq> 0" "d \<noteq> 0" | 
| 375 | assume 1: "Fract a b \<le> Fract c d" | |
| 376 | assume 2: "Fract c d \<le> Fract a b" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 377 | show "Fract a b = Fract c d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 378 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 379 | from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 380 | by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 381 | also have "... \<le> (a * d) * (b * d)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 382 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 383 | from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 384 | by simp | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 385 | then show ?thesis by (simp only: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 386 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 387 | finally have "(a * d) * (b * d) = (c * b) * (b * d)" . | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 388 | moreover from neq have "b * d \<noteq> 0" by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 389 | ultimately have "a * d = c * b" by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 390 | with neq show ?thesis by (simp add: eq_fract) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 391 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 392 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 393 | next | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 394 | fix q r :: "'a fract" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 395 | show "q \<le> q" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 396 | by (induct q) simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 397 | show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 398 | by (simp only: less_fract_def) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 399 | show "q \<le> r \<or> r \<le> q" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 400 | by (induct q, induct r) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
54863diff
changeset | 401 | (simp add: mult.commute, rule linorder_linear) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 402 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 403 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 404 | end | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 405 | |
| 54463 | 406 | instantiation fract :: (linordered_idom) "{distrib_lattice,abs_if,sgn_if}"
 | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 407 | begin | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 408 | |
| 46573 | 409 | definition abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 410 | |
| 46573 | 411 | definition sgn_fract_def: | 
| 54463 | 412 | "sgn (q::'a fract) = (if q = 0 then 0 else if 0 < q then 1 else - 1)" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 413 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 414 | theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 415 | by (auto simp add: abs_fract_def Zero_fract_def le_less | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 416 | eq_fract zero_less_mult_iff mult_less_0_iff split: abs_split) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 417 | |
| 46573 | 418 | definition inf_fract_def: | 
| 419 | "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 420 | |
| 46573 | 421 | definition sup_fract_def: | 
| 422 | "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 423 | |
| 46573 | 424 | instance | 
| 425 | by intro_classes | |
| 426 | (auto simp add: abs_fract_def sgn_fract_def | |
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
54463diff
changeset | 427 | max_min_distrib2 inf_fract_def sup_fract_def) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 428 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 429 | end | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 430 | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
58881diff
changeset | 431 | instance fract :: (linordered_idom) linordered_field | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 432 | proof | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 433 | fix q r s :: "'a fract" | 
| 53196 | 434 | assume "q \<le> r" | 
| 435 | then show "s + q \<le> s + r" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 436 | proof (induct q, induct r, induct s) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 437 | fix a b c d e f :: 'a | 
| 53196 | 438 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 439 | assume le: "Fract a b \<le> Fract c d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 440 | show "Fract e f + Fract a b \<le> Fract e f + Fract c d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 441 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 442 | let ?F = "f * f" from neq have F: "0 < ?F" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 443 | by (auto simp add: zero_less_mult_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 444 | from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 445 | by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 446 | with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 447 | by (simp add: mult_le_cancel_right) | 
| 36348 
89c54f51f55a
dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
 haftmann parents: 
36331diff
changeset | 448 | with neq show ?thesis by (simp add: field_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 449 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 450 | qed | 
| 53196 | 451 | next | 
| 452 | fix q r s :: "'a fract" | |
| 453 | assume "q < r" and "0 < s" | |
| 454 | then show "s * q < s * r" | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 455 | proof (induct q, induct r, induct s) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 456 | fix a b c d e f :: 'a | 
| 54463 | 457 | assume neq: "b \<noteq> 0" "d \<noteq> 0" "f \<noteq> 0" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 458 | assume le: "Fract a b < Fract c d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 459 | assume gt: "0 < Fract e f" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 460 | show "Fract e f * Fract a b < Fract e f * Fract c d" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 461 | proof - | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 462 | let ?E = "e * f" and ?F = "f * f" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 463 | from neq gt have "0 < ?E" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 464 | by (auto simp add: Zero_fract_def order_less_le eq_fract) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 465 | moreover from neq have "0 < ?F" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 466 | by (auto simp add: zero_less_mult_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 467 | moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 468 | by simp | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 469 | ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 470 | by (simp add: mult_less_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 471 | with neq show ?thesis | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 472 | by (simp add: ac_simps) | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 473 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 474 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 475 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 476 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 477 | lemma fract_induct_pos [case_names Fract]: | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 478 | fixes P :: "'a::linordered_idom fract \<Rightarrow> bool" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 479 | assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 480 | shows "P q" | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 481 | proof (cases q) | 
| 54463 | 482 | case (Fract a b) | 
| 483 |   {
 | |
| 484 | fix a b :: 'a | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 485 | assume b: "b < 0" | 
| 54463 | 486 | have "P (Fract a b)" | 
| 487 | proof - | |
| 488 | from b have "0 < - b" by simp | |
| 489 | then have "P (Fract (- a) (- b))" | |
| 490 | by (rule step) | |
| 491 | then show "P (Fract a b)" | |
| 492 | by (simp add: order_less_imp_not_eq [OF b]) | |
| 493 | qed | |
| 494 | } | |
| 495 | with Fract show "P q" | |
| 496 | by (auto simp add: linorder_neq_iff step) | |
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 497 | qed | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 498 | |
| 53196 | 499 | lemma zero_less_Fract_iff: "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 500 | by (auto simp add: Zero_fract_def zero_less_mult_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 501 | |
| 53196 | 502 | lemma Fract_less_zero_iff: "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 503 | by (auto simp add: Zero_fract_def mult_less_0_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 504 | |
| 53196 | 505 | lemma zero_le_Fract_iff: "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 506 | by (auto simp add: Zero_fract_def zero_le_mult_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 507 | |
| 53196 | 508 | lemma Fract_le_zero_iff: "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 509 | by (auto simp add: Zero_fract_def mult_le_0_iff) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 510 | |
| 53196 | 511 | lemma one_less_Fract_iff: "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 512 | by (auto simp add: One_fract_def mult_less_cancel_right_disj) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 513 | |
| 53196 | 514 | lemma Fract_less_one_iff: "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 515 | by (auto simp add: One_fract_def mult_less_cancel_right_disj) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 516 | |
| 53196 | 517 | lemma one_le_Fract_iff: "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 518 | by (auto simp add: One_fract_def mult_le_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 519 | |
| 53196 | 520 | lemma Fract_le_one_iff: "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b" | 
| 36331 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 521 | by (auto simp add: One_fract_def mult_le_cancel_right) | 
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 522 | |
| 
6b9e487450ba
Library/Fraction_Field.thy: ordering relations for fractions
 huffman parents: 
36312diff
changeset | 523 | end |