diff -r 58ccbc73a172 -r ac0a3b9c6dae src/HOL/Rat.thy --- a/src/HOL/Rat.thy Thu Jul 14 12:21:12 2016 +0200 +++ b/src/HOL/Rat.thy Fri Jul 15 11:07:51 2016 +0200 @@ -1,11 +1,11 @@ -(* Title: HOL/Rat.thy - Author: Markus Wenzel, TU Muenchen +(* Title: HOL/Rat.thy + Author: Markus Wenzel, TU Muenchen *) section \Rational numbers\ theory Rat -imports GCD Archimedean_Field + imports GCD Archimedean_Field begin subsection \Rational numbers as quotient\ @@ -27,17 +27,17 @@ lemma transp_ratrel: "transp ratrel" proof (rule transpI, unfold split_paired_all) fix a b a' b' a'' b'' :: int - assume A: "ratrel (a, b) (a', b')" - assume B: "ratrel (a', b') (a'', b'')" + assume *: "ratrel (a, b) (a', b')" + assume **: "ratrel (a', b') (a'', b'')" have "b' * (a * b'') = b'' * (a * b')" by simp - also from A have "a * b' = a' * b" by auto + also from * have "a * b' = a' * b" by auto also have "b'' * (a' * b) = b * (a' * b'')" by simp - also from B have "a' * b'' = a'' * b'" by auto + also from ** have "a' * b'' = a'' * b'" by auto also have "b * (a'' * b') = b' * (a'' * b)" by simp finally have "b' * (a * b'') = b' * (a'' * b)" . - moreover from B have "b' \ 0" by auto + moreover from ** have "b' \ 0" by auto ultimately have "a * b'' = a'' * b" by simp - with A B show "ratrel (a, b) (a'', b'')" by auto + with * ** show "ratrel (a, b) (a'', b'')" by auto qed lemma part_equivp_ratrel: "part_equivp ratrel" @@ -120,7 +120,7 @@ lift_definition plus_rat :: "rat \ rat \ rat" is "\x y. (fst x * snd y + fst y * snd x, snd x * snd y)" - by (clarsimp, simp add: distrib_right, simp add: ac_simps) + by (auto simp: distrib_right) (simp add: ac_simps) lemma add_rat [simp]: assumes "b \ 0" and "d \ 0" @@ -139,9 +139,8 @@ definition diff_rat_def: "q - r = q + - r" for q r :: rat lemma diff_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" - using assms by (simp add: diff_rat_def) + "b \ 0 \ d \ 0 \ Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" + by (simp add: diff_rat_def) lift_definition times_rat :: "rat \ rat \ rat" is "\x y. (fst x * fst y, snd x * snd y)" @@ -150,10 +149,8 @@ lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)" by transfer simp -lemma mult_rat_cancel: - assumes "c \ 0" - shows "Fract (c * a) (c * b) = Fract a b" - using assms by transfer simp +lemma mult_rat_cancel: "c \ 0 \ Fract (c * a) (c * b) = Fract a b" + by transfer simp lift_definition inverse_rat :: "rat \ rat" is "\x. if fst x = 0 then (0, 1) else (snd x, fst x)" @@ -220,7 +217,7 @@ "Fract (- 1) 1 = - 1" "Fract k 0 = 0" using Fract_of_int_eq [of "numeral w"] - using Fract_of_int_eq [of "- numeral w"] + and Fract_of_int_eq [of "- numeral w"] by (simp_all add: Zero_rat_def One_rat_def eq_rat) lemma rat_number_expand: @@ -255,7 +252,8 @@ lemma Fract_coprime: "Fract (a div gcd a b) (b div gcd a b) = Fract a b" proof (cases "b = 0") case True - then show ?thesis by (simp add: eq_rat) + then show ?thesis + by (simp add: eq_rat) next case False moreover have "b div gcd a b * gcd a b = b" @@ -282,28 +280,27 @@ have *: "p * s = q * r" if "p * gcd r s = sgn (q * s) * r * gcd p q" and "q * gcd r s = sgn (q * s) * s * gcd p q" proof - - from that - have "(p * gcd r s) * (sgn (q * s) * s * gcd p q) = (q * gcd r s) * (sgn (q * s) * r * gcd p q)" + from that have "(p * gcd r s) * (sgn (q * s) * s * gcd p q) = + (q * gcd r s) * (sgn (q * s) * r * gcd p q)" by simp with assms show ?thesis by (auto simp add: ac_simps sgn_times sgn_0_0) qed from assms show ?thesis - by (auto simp add: normalize_def Let_def dvd_div_div_eq_mult mult.commute sgn_times + by (auto simp: normalize_def Let_def dvd_div_div_eq_mult mult.commute sgn_times split: if_splits intro: *) qed lemma normalize_eq: "normalize (a, b) = (p, q) \ Fract p q = Fract a b" - by (auto simp add: normalize_def Let_def Fract_coprime dvd_div_neg rat_number_collapse + by (auto simp: normalize_def Let_def Fract_coprime dvd_div_neg rat_number_collapse split: if_split_asm) lemma normalize_denom_pos: "normalize r = (p, q) \ q > 0" - by (auto simp add: normalize_def Let_def dvd_div_neg pos_imp_zdiv_neg_iff nonneg1_imp_zdiv_pos_iff + by (auto simp: normalize_def Let_def dvd_div_neg pos_imp_zdiv_neg_iff nonneg1_imp_zdiv_pos_iff split: if_split_asm) lemma normalize_coprime: "normalize r = (p, q) \ coprime p q" - by (auto simp add: normalize_def Let_def dvd_div_neg div_gcd_coprime - split: if_split_asm) + by (auto simp: normalize_def Let_def dvd_div_neg div_gcd_coprime split: if_split_asm) lemma normalize_stable [simp]: "q > 0 \ coprime p q \ normalize (p, q) = (p, q)" by (simp add: normalize_def) @@ -325,7 +322,8 @@ lemma quotient_of_unique: "\!p. r = Fract (fst p) (snd p) \ snd p > 0 \ coprime (fst p) (snd p)" proof (cases r) case (Fract a b) - then have "r = Fract (fst (a, b)) (snd (a, b)) \ snd (a, b) > 0 \ coprime (fst (a, b)) (snd (a, b))" + then have "r = Fract (fst (a, b)) (snd (a, b)) \ + snd (a, b) > 0 \ coprime (fst (a, b)) (snd (a, b))" by auto then show ?thesis proof (rule ex1I) @@ -453,7 +451,8 @@ lemma positive_add: "positive x \ positive y \ positive (x + y)" apply transfer apply (simp add: zero_less_mult_iff) - apply (elim disjE, simp_all add: add_pos_pos add_neg_neg mult_pos_neg mult_neg_pos mult_neg_neg) + apply (elim disjE) + apply (simp_all add: add_pos_pos add_neg_neg mult_pos_neg mult_neg_pos mult_neg_neg) done lemma positive_mult: "positive x \ positive y \ positive (x * y)" @@ -484,8 +483,8 @@ show "a < b \ a \ b \ \ b \ a" unfolding less_eq_rat_def less_rat_def apply auto - apply (drule (1) positive_add) - apply (simp_all add: positive_zero) + apply (drule (1) positive_add) + apply (simp_all add: positive_zero) done show "a \ a" unfolding less_eq_rat_def by simp @@ -533,15 +532,12 @@ by transfer simp lemma less_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b < Fract c d \ (a * d) * (b * d) < (c * b) * (b * d)" - using assms unfolding less_rat_def - by (simp add: positive_rat algebra_simps) + "b \ 0 \ d \ 0 \ Fract a b < Fract c d \ (a * d) * (b * d) < (c * b) * (b * d)" + by (simp add: less_rat_def positive_rat algebra_simps) lemma le_rat [simp]: - assumes "b \ 0" and "d \ 0" - shows "Fract a b \ Fract c d \ (a * d) * (b * d) \ (c * b) * (b * d)" - using assms unfolding le_less by (simp add: eq_rat) + "b \ 0 \ d \ 0 \ Fract a b \ Fract c d \ (a * d) * (b * d) \ (c * b) * (b * d)" + by (simp add: le_less eq_rat) lemma abs_rat [simp, code]: "\Fract a b\ = Fract \a\ \b\" by (auto simp add: abs_rat_def zabs_def Zero_rat_def not_less le_less eq_rat zero_less_mult_iff) @@ -565,7 +561,8 @@ then show "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b]) qed - from Fract show "P q" by (auto simp add: linorder_neq_iff step step') + from Fract show "P q" + by (auto simp add: linorder_neq_iff step step') qed lemma zero_less_Fract_iff: "0 < b \ 0 < Fract a b \ 0 < a" @@ -662,9 +659,7 @@ lift_definition of_rat :: "rat \ 'a" is "\x. of_int (fst x) / of_int (snd x)" - apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq) - apply (simp only: of_int_mult [symmetric]) - done + by (auto simp: nonzero_divide_eq_eq nonzero_eq_divide_eq) (simp only: of_int_mult [symmetric]) end @@ -779,7 +774,7 @@ by (induct a) (simp add: of_rat_rat Fract_of_int_eq [symmetric]) qed -text \Collapse nested embeddings\ +text \Collapse nested embeddings.\ lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n" by (induct n) (simp_all add: of_rat_add) @@ -854,31 +849,36 @@ apply (rule of_rat_mult [symmetric]) done -lemma nonzero_Rats_inverse: "a \ \ \ a \ 0 \ inverse a \ \" for a :: "'a::field_char_0" +lemma nonzero_Rats_inverse: "a \ \ \ a \ 0 \ inverse a \ \" + for a :: "'a::field_char_0" apply (auto simp add: Rats_def) apply (rule range_eqI) apply (erule nonzero_of_rat_inverse [symmetric]) done -lemma Rats_inverse [simp]: "a \ \ \ inverse a \ \" for a :: "'a::{field_char_0,field}" +lemma Rats_inverse [simp]: "a \ \ \ inverse a \ \" + for a :: "'a::{field_char_0,field}" apply (auto simp add: Rats_def) apply (rule range_eqI) apply (rule of_rat_inverse [symmetric]) done -lemma nonzero_Rats_divide: "a \ \ \ b \ \ \ b \ 0 \ a / b \ \" for a b :: "'a::field_char_0" +lemma nonzero_Rats_divide: "a \ \ \ b \ \ \ b \ 0 \ a / b \ \" + for a b :: "'a::field_char_0" apply (auto simp add: Rats_def) apply (rule range_eqI) apply (erule nonzero_of_rat_divide [symmetric]) done -lemma Rats_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" for a b :: "'a::{field_char_0, field}" +lemma Rats_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" + for a b :: "'a::{field_char_0, field}" apply (auto simp add: Rats_def) apply (rule range_eqI) apply (rule of_rat_divide [symmetric]) done -lemma Rats_power [simp]: "a \ \ \ a ^ n \ \" for a :: "'a::field_char_0" +lemma Rats_power [simp]: "a \ \ \ a ^ n \ \" + for a :: "'a::field_char_0" apply (auto simp add: Rats_def) apply (rule range_eqI) apply (rule of_rat_power [symmetric]) @@ -888,7 +888,8 @@ assumes "q \ \" obtains (of_rat) r where "q = of_rat r" proof - - from \q \ \\ have "q \ range of_rat" unfolding Rats_def . + from \q \ \\ have "q \ range of_rat" + by (simp only: Rats_def) then obtain r where "q = of_rat r" .. then show thesis .. qed @@ -1028,7 +1029,8 @@ text \Quickcheck\ definition (in term_syntax) - valterm_fract :: "int \ (unit \ Code_Evaluation.term) \ int \ (unit \ Code_Evaluation.term) \ + valterm_fract :: "int \ (unit \ Code_Evaluation.term) \ + int \ (unit \ Code_Evaluation.term) \ rat \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\} k {\} l"