# HG changeset patch # User wenzelm # Date 1434826531 -7200 # Node ID b9add7665d7a5e8c1f51bc9f7f880c96850c0664 # Parent 5398aa5a4df9b58fb8349e13fefbf3ca516dbd6d tuned proofs; diff -r 5398aa5a4df9 -r b9add7665d7a src/HOL/Decision_Procs/Rat_Pair.thy --- a/src/HOL/Decision_Procs/Rat_Pair.thy Sat Jun 20 20:17:29 2015 +0200 +++ b/src/HOL/Decision_Procs/Rat_Pair.thy Sat Jun 20 20:55:31 2015 +0200 @@ -16,12 +16,14 @@ abbreviation Numi_syn :: "int \ Num" ("'((_)')\<^sub>N") where "(i)\<^sub>N \ (i, 1)" -definition isnormNum :: "Num \ bool" where - "isnormNum = (\(a,b). (if a = 0 then b = 0 else b > 0 \ gcd a b = 1))" +definition isnormNum :: "Num \ bool" + where "isnormNum = (\(a, b). if a = 0 then b = 0 else b > 0 \ gcd a b = 1)" -definition normNum :: "Num \ Num" where +definition normNum :: "Num \ Num" +where "normNum = (\(a,b). - (if a=0 \ b = 0 then (0,0) else + (if a = 0 \ b = 0 then (0, 0) + else (let g = gcd a b in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))" @@ -30,61 +32,72 @@ lemma normNum_isnormNum [simp]: "isnormNum (normNum x)" proof - obtain a b where x: "x = (a, b)" by (cases x) - { assume "a=0 \ b = 0" hence ?thesis by (simp add: x normNum_def isnormNum_def) } - moreover - { assume anz: "a \ 0" and bnz: "b \ 0" + consider "a = 0 \ b = 0" | "a \ 0" "b \ 0" by blast + then show ?thesis + proof cases + case 1 + then show ?thesis + by (simp add: x normNum_def isnormNum_def) + next + case 2 let ?g = "gcd a b" let ?a' = "a div ?g" let ?b' = "b div ?g" let ?g' = "gcd ?a' ?b'" - from anz bnz have "?g \ 0" by simp with gcd_ge_0_int[of a b] + from 2 have "?g \ 0" by simp with gcd_ge_0_int[of a b] have gpos: "?g > 0" by arith have gdvd: "?g dvd a" "?g dvd b" by arith+ - from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] anz bnz + from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] 2 have nz': "?a' \ 0" "?b' \ 0" by - (rule notI, simp)+ - from anz bnz have stupid: "a \ 0 \ b \ 0" by arith + from 2 have stupid: "a \ 0 \ b \ 0" by arith from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" . - from bnz have "b < 0 \ b > 0" by arith - moreover - { assume b: "b > 0" - from b have "?b' \ 0" + from 2 consider "b < 0" | "b > 0" by arith + then show ?thesis + proof cases + case 1 + have False if b': "?b' \ 0" + proof - + from gpos have th: "?g \ 0" by arith + from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)] + show ?thesis using 1 by arith + qed + then have b': "?b' < 0" by (presburger add: linorder_not_le[symmetric]) + from \a \ 0\ nz' 1 b' gp1 show ?thesis + by (simp add: x isnormNum_def normNum_def Let_def split_def) + next + case 2 + then have "?b' \ 0" by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos]) with nz' have b': "?b' > 0" by arith - from b b' anz bnz nz' gp1 have ?thesis - by (simp add: x isnormNum_def normNum_def Let_def split_def) } - moreover { - assume b: "b < 0" - { assume b': "?b' \ 0" - from gpos have th: "?g \ 0" by arith - from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)] - have False using b by arith } - hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric]) - from anz bnz nz' b b' gp1 have ?thesis - by (simp add: x isnormNum_def normNum_def Let_def split_def) } - ultimately have ?thesis by blast - } - ultimately show ?thesis by blast + from 2 b' \a \ 0\ nz' gp1 show ?thesis + by (simp add: x isnormNum_def normNum_def Let_def split_def) + qed + qed qed text \Arithmetic over Num\ -definition Nadd :: "Num \ Num \ Num" (infixl "+\<^sub>N" 60) where - "Nadd = (\(a,b) (a',b'). if a = 0 \ b = 0 then normNum(a',b') - else if a'=0 \ b' = 0 then normNum(a,b) - else normNum(a*b' + b*a', b*b'))" +definition Nadd :: "Num \ Num \ Num" (infixl "+\<^sub>N" 60) +where + "Nadd = (\(a, b) (a', b'). + if a = 0 \ b = 0 then normNum (a', b') + else if a' = 0 \ b' = 0 then normNum (a, b) + else normNum (a * b' + b * a', b * b'))" -definition Nmul :: "Num \ Num \ Num" (infixl "*\<^sub>N" 60) where - "Nmul = (\(a,b) (a',b'). let g = gcd (a*a') (b*b') - in (a*a' div g, b*b' div g))" +definition Nmul :: "Num \ Num \ Num" (infixl "*\<^sub>N" 60) +where + "Nmul = (\(a, b) (a', b'). + let g = gcd (a * a') (b * b') + in (a * a' div g, b * b' div g))" definition Nneg :: "Num \ Num" ("~\<^sub>N") - where "Nneg \ (\(a,b). (-a,b))" + where "Nneg = (\(a, b). (- a, b))" definition Nsub :: "Num \ Num \ Num" (infixl "-\<^sub>N" 60) where "Nsub = (\a b. a +\<^sub>N ~\<^sub>N b)" definition Ninv :: "Num \ Num" - where "Ninv = (\(a,b). if a < 0 then (-b, \a\) else (b,a))" + where "Ninv = (\(a, b). if a < 0 then (- b, \a\) else (b, a))" definition Ndiv :: "Num \ Num \ Num" (infixl "\
\<^sub>N" 60) where "Ndiv = (\a b. a *\<^sub>N Ninv b)" @@ -95,53 +108,59 @@ lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)" by (simp add: Nadd_def split_def) -lemma Nsub_normN[simp]: "\ isnormNum y\ \ isnormNum (x -\<^sub>N y)" +lemma Nsub_normN[simp]: "isnormNum y \ isnormNum (x -\<^sub>N y)" by (simp add: Nsub_def split_def) lemma Nmul_normN[simp]: - assumes xn: "isnormNum x" and yn: "isnormNum y" + assumes xn: "isnormNum x" + and yn: "isnormNum y" shows "isnormNum (x *\<^sub>N y)" proof - obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) - { assume "a = 0" - hence ?thesis using xn x y - by (simp add: isnormNum_def Let_def Nmul_def split_def) } - moreover - { assume "a' = 0" - hence ?thesis using yn x y - by (simp add: isnormNum_def Let_def Nmul_def split_def) } - moreover - { assume a: "a \0" and a': "a'\0" - hence bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def) + consider "a = 0" | "a' = 0" | "a \ 0" "a' \ 0" by blast + then show ?thesis + proof cases + case 1 + then show ?thesis + using xn x y by (simp add: isnormNum_def Let_def Nmul_def split_def) + next + case 2 + then show ?thesis + using yn x y by (simp add: isnormNum_def Let_def Nmul_def split_def) + next + case 3 + then have bp: "b > 0" "b' > 0" + using xn yn x y by (simp_all add: isnormNum_def) from bp have "x *\<^sub>N y = normNum (a * a', b * b')" - using x y a a' bp by (simp add: Nmul_def Let_def split_def normNum_def) - hence ?thesis by simp } - ultimately show ?thesis by blast + using x y 3 bp by (simp add: Nmul_def Let_def split_def normNum_def) + then show ?thesis by simp + qed qed lemma Ninv_normN[simp]: "isnormNum x \ isnormNum (Ninv x)" - by (simp add: Ninv_def isnormNum_def split_def) - (cases "fst x = 0", auto simp add: gcd_commute_int) + apply (simp add: Ninv_def isnormNum_def split_def) + apply (cases "fst x = 0") + apply (auto simp add: gcd_commute_int) + done -lemma isnormNum_int[simp]: - "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \ 0 \ isnormNum (i)\<^sub>N" +lemma isnormNum_int[simp]: "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \ 0 \ isnormNum (i)\<^sub>N" by (simp_all add: isnormNum_def) text \Relations over Num\ definition Nlt0:: "Num \ bool" ("0>\<^sub>N") - where "Nlt0 = (\(a,b). a < 0)" + where "Nlt0 = (\(a, b). a < 0)" definition Nle0:: "Num \ bool" ("0\\<^sub>N") - where "Nle0 = (\(a,b). a \ 0)" + where "Nle0 = (\(a, b). a \ 0)" definition Ngt0:: "Num \ bool" ("0<\<^sub>N") - where "Ngt0 = (\(a,b). a > 0)" + where "Ngt0 = (\(a, b). a > 0)" definition Nge0:: "Num \ bool" ("0\\<^sub>N") - where "Nge0 = (\(a,b). a \ 0)" + where "Nge0 = (\(a, b). a \ 0)" definition Nlt :: "Num \ Num \ bool" (infix "<\<^sub>N" 55) where "Nlt = (\a b. 0>\<^sub>N (a -\<^sub>N b))" @@ -149,27 +168,34 @@ definition Nle :: "Num \ Num \ bool" (infix "\\<^sub>N" 55) where "Nle = (\a b. 0\\<^sub>N (a -\<^sub>N b))" -definition "INum = (\(a,b). of_int a / of_int b)" +definition "INum = (\(a, b). of_int a / of_int b)" -lemma INum_int [simp]: "INum (i)\<^sub>N = ((of_int i) ::'a::field)" "INum 0\<^sub>N = (0::'a::field)" +lemma INum_int [simp]: "INum (i)\<^sub>N = (of_int i ::'a::field)" "INum 0\<^sub>N = (0::'a::field)" by (simp_all add: INum_def) lemma isnormNum_unique[simp]: - assumes na: "isnormNum x" and nb: "isnormNum y" - shows "((INum x ::'a::{field_char_0, field}) = INum y) = (x = y)" (is "?lhs = ?rhs") + assumes na: "isnormNum x" + and nb: "isnormNum y" + shows "(INum x ::'a::{field_char_0,field}) = INum y \ x = y" + (is "?lhs = ?rhs") proof obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) - assume H: ?lhs - { assume "a = 0 \ b = 0 \ a' = 0 \ b' = 0" - hence ?rhs using na nb H - by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) } - moreover - { assume az: "a \ 0" and bz: "b \ 0" and a'z: "a'\0" and b'z: "b'\0" - from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def) - from H bz b'z have eq: "a * b' = a'*b" + consider "a = 0 \ b = 0 \ a' = 0 \ b' = 0" | "a \ 0" "b \ 0" "a' \ 0" "b' \ 0" + by blast + then show ?rhs if H: ?lhs + proof cases + case 1 + then show ?thesis + using na nb H by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) + next + case 2 + with na nb have pos: "b > 0" "b' > 0" + by (simp_all add: x y isnormNum_def) + from H \b \ 0\ \b' \ 0\ have eq: "a * b' = a'*b" by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult) - from az a'z na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1" + from \a \ 0\ \a' \ 0\ na nb + have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1" by (simp_all add: x y isnormNum_def add: gcd_commute_int) from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'" apply - @@ -182,23 +208,22 @@ coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]] have eq1: "b = b'" using pos by arith with eq have "a = a'" using pos by simp - with eq1 have ?rhs by (simp add: x y) } - ultimately show ?rhs by blast -next - assume ?rhs thus ?lhs by simp + with eq1 show ?thesis by (simp add: x y) + qed + show ?lhs if ?rhs + using that by simp qed - -lemma isnormNum0[simp]: - "isnormNum x \ (INum x = (0::'a::{field_char_0, field})) = (x = 0\<^sub>N)" +lemma isnormNum0[simp]: "isnormNum x \ INum x = (0::'a::{field_char_0,field}) \ x = 0\<^sub>N" unfolding INum_int(2)[symmetric] by (rule isnormNum_unique) simp_all -lemma of_int_div_aux: "d ~= 0 ==> ((of_int x)::'a::field_char_0) / (of_int d) = - of_int (x div d) + (of_int (x mod d)) / ((of_int d)::'a)" +lemma of_int_div_aux: + assumes "d \ 0" + shows "(of_int x ::'a::field_char_0) / of_int d = + of_int (x div d) + (of_int (x mod d)) / of_int d" proof - - assume "d ~= 0" - let ?t = "of_int (x div d) * ((of_int d)::'a) + of_int(x mod d)" + let ?t = "of_int (x div d) * (of_int d ::'a) + of_int (x mod d)" let ?f = "\x. x / of_int d" have "x = (x div d) * d + x mod d" by auto @@ -206,30 +231,36 @@ by (simp only: of_int_mult[symmetric] of_int_add [symmetric]) then have "of_int x / of_int d = ?t / of_int d" using cong[OF refl[of ?f] eq] by simp - then show ?thesis by (simp add: add_divide_distrib algebra_simps \d ~= 0\) + then show ?thesis + by (simp add: add_divide_distrib algebra_simps \d \ 0\) qed -lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==> - (of_int(n div d)::'a::field_char_0) = of_int n / of_int d" - using of_int_div_aux [of d n, where ?'a = 'a] by simp +lemma of_int_div: + fixes d :: int + assumes "d \ 0" "d dvd n" + shows "(of_int (n div d) ::'a::field_char_0) = of_int n / of_int d" + using assms of_int_div_aux [of d n, where ?'a = 'a] by simp -lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0, field})" +lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0,field})" proof - obtain a b where x: "x = (a, b)" by (cases x) - { assume "a = 0 \ b = 0" - hence ?thesis by (simp add: x INum_def normNum_def split_def Let_def) } - moreover - { assume a: "a \ 0" and b: "b \ 0" + consider "a = 0 \ b = 0" | "a \ 0" "b \ 0" by blast + then show ?thesis + proof cases + case 1 + then show ?thesis + by (simp add: x INum_def normNum_def split_def Let_def) + next + case 2 let ?g = "gcd a b" - from a b have g: "?g \ 0"by simp - from of_int_div[OF g, where ?'a = 'a] - have ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) } - ultimately show ?thesis by blast + from 2 have g: "?g \ 0"by simp + from of_int_div[OF g, where ?'a = 'a] show ?thesis + by (auto simp add: x INum_def normNum_def split_def Let_def) + qed qed -lemma INum_normNum_iff: - "(INum x ::'a::{field_char_0, field}) = INum y \ normNum x = normNum y" - (is "?lhs = ?rhs") +lemma INum_normNum_iff: "(INum x ::'a::{field_char_0,field}) = INum y \ normNum x = normNum y" + (is "?lhs \ _") proof - have "normNum x = normNum y \ (INum (normNum x) :: 'a) = INum (normNum y)" by (simp del: normNum) @@ -237,178 +268,231 @@ finally show ?thesis by simp qed -lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0, field})" -proof - - let ?z = "0:: 'a" - obtain a b where x: "x = (a, b)" by (cases x) - obtain a' b' where y: "y = (a', b')" by (cases y) - { assume "a=0 \ a'= 0 \ b =0 \ b' = 0" - hence ?thesis - apply (cases "a=0", simp_all add: x y Nadd_def) - apply (cases "b= 0", simp_all add: INum_def) - apply (cases "a'= 0", simp_all) - apply (cases "b'= 0", simp_all) - done } - moreover - { assume aa': "a \ 0" "a'\ 0" and bb': "b \ 0" "b' \ 0" - { assume z: "a * b' + b * a' = 0" - hence "of_int (a*b' + b*a') / (of_int b* of_int b') = ?z" by simp - hence "of_int b' * of_int a / (of_int b * of_int b') + - of_int b * of_int a' / (of_int b * of_int b') = ?z" - by (simp add:add_divide_distrib) - hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa' - by simp - from z aa' bb' have ?thesis - by (simp add: x y th Nadd_def normNum_def INum_def split_def) } - moreover { - assume z: "a * b' + b * a' \ 0" - let ?g = "gcd (a * b' + b * a') (b * b')" - have gz: "?g \ 0" using z by simp - have ?thesis using aa' bb' z gz - of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * b' + b * a'" "b * b'"]] - of_int_div [where ?'a = 'a, OF gz gcd_dvd2 [of "a * b' + b * a'" "b * b'"]] - by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps) - } - ultimately have ?thesis using aa' bb' - by (simp add: x y Nadd_def INum_def normNum_def Let_def) } - ultimately show ?thesis by blast -qed - -lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {field_char_0, field})" +lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0,field})" proof - let ?z = "0::'a" obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) - { assume "a=0 \ a'= 0 \ b = 0 \ b' = 0" - hence ?thesis - apply (cases "a=0", simp_all add: x y Nmul_def INum_def Let_def) - apply (cases "b=0", simp_all) - apply (cases "a'=0", simp_all) - done } - moreover - { assume z: "a \ 0" "a' \ 0" "b \ 0" "b' \ 0" - let ?g="gcd (a*a') (b*b')" - have gz: "?g \ 0" using z by simp - from z of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * a'" "b * b'"]] - of_int_div [where ?'a = 'a , OF gz gcd_dvd2 [of "a * a'" "b * b'"]] - have ?thesis by (simp add: Nmul_def x y Let_def INum_def) } - ultimately show ?thesis by blast + consider "a = 0 \ a'= 0 \ b =0 \ b' = 0" | "a \ 0" "a'\ 0" "b \ 0" "b' \ 0" + by blast + then show ?thesis + proof cases + case 1 + then show ?thesis + apply (cases "a = 0") + apply (simp_all add: x y Nadd_def) + apply (cases "b = 0") + apply (simp_all add: INum_def) + apply (cases "a'= 0") + apply simp_all + apply (cases "b'= 0") + apply simp_all + done + next + case 2 + show ?thesis + proof (cases "a * b' + b * a' = 0") + case True + then have "of_int (a * b' + b * a') / (of_int b * of_int b') = ?z" + by simp + then have "of_int b' * of_int a / (of_int b * of_int b') + + of_int b * of_int a' / (of_int b * of_int b') = ?z" + by (simp add: add_divide_distrib) + then have th: "of_int a / of_int b + of_int a' / of_int b' = ?z" + using 2 by simp + from True 2 show ?thesis + by (simp add: x y th Nadd_def normNum_def INum_def split_def) + next + case False + let ?g = "gcd (a * b' + b * a') (b * b')" + have gz: "?g \ 0" + using False by simp + show ?thesis + using 2 False gz + of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * b' + b * a'" "b * b'"]] + of_int_div [where ?'a = 'a, OF gz gcd_dvd2 [of "a * b' + b * a'" "b * b'"]] + by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps) + qed + qed qed -lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)" +lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a::{field_char_0,field})" +proof - + let ?z = "0::'a" + obtain a b where x: "x = (a, b)" by (cases x) + obtain a' b' where y: "y = (a', b')" by (cases y) + consider "a = 0 \ a' = 0 \ b = 0 \ b' = 0" | "a \ 0" "a' \ 0" "b \ 0" "b' \ 0" + by blast + then show ?thesis + proof cases + case 1 + then show ?thesis + apply (cases "a = 0") + apply (simp_all add: x y Nmul_def INum_def Let_def) + apply (cases "b = 0") + apply simp_all + apply (cases "a' = 0") + apply simp_all + done + next + case 2 + let ?g = "gcd (a * a') (b * b')" + have gz: "?g \ 0" + using 2 by simp + from 2 of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * a'" "b * b'"]] + of_int_div [where ?'a = 'a , OF gz gcd_dvd2 [of "a * a'" "b * b'"]] + show ?thesis + by (simp add: Nmul_def x y Let_def INum_def) + qed +qed + +lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a::field)" by (simp add: Nneg_def split_def INum_def) -lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0, field})" +lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0,field})" by (simp add: Nsub_def split_def) lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field) / (INum x)" by (simp add: Ninv_def INum_def split_def) -lemma Ndiv[simp]: "INum (x \
\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0, field})" +lemma Ndiv[simp]: "INum (x \
\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0,field})" by (simp add: Ndiv_def) lemma Nlt0_iff[simp]: assumes nx: "isnormNum x" - shows "((INum x :: 'a :: {field_char_0, linordered_field})< 0) = 0>\<^sub>N x" + shows "((INum x :: 'a::{field_char_0,linordered_field})< 0) = 0>\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) - { assume "a = 0" hence ?thesis by (simp add: x Nlt0_def INum_def) } - moreover - { assume a: "a \ 0" hence b: "(of_int b::'a) > 0" + show ?thesis + proof (cases "a = 0") + case True + then show ?thesis + by (simp add: x Nlt0_def INum_def) + next + case False + then have b: "(of_int b::'a) > 0" using nx by (simp add: x isnormNum_def) from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"] - have ?thesis by (simp add: x Nlt0_def INum_def) } - ultimately show ?thesis by blast + show ?thesis + by (simp add: x Nlt0_def INum_def) + qed qed lemma Nle0_iff[simp]: assumes nx: "isnormNum x" - shows "((INum x :: 'a :: {field_char_0, linordered_field}) \ 0) = 0\\<^sub>N x" + shows "((INum x :: 'a::{field_char_0,linordered_field}) \ 0) = 0\\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) - { assume "a = 0" hence ?thesis by (simp add: x Nle0_def INum_def) } - moreover - { assume a: "a \ 0" hence b: "(of_int b :: 'a) > 0" + show ?thesis + proof (cases "a = 0") + case True + then show ?thesis + by (simp add: x Nle0_def INum_def) + next + case False + then have b: "(of_int b :: 'a) > 0" using nx by (simp add: x isnormNum_def) from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"] - have ?thesis by (simp add: x Nle0_def INum_def) } - ultimately show ?thesis by blast + show ?thesis + by (simp add: x Nle0_def INum_def) + qed qed lemma Ngt0_iff[simp]: assumes nx: "isnormNum x" - shows "((INum x :: 'a :: {field_char_0, linordered_field})> 0) = 0<\<^sub>N x" + shows "((INum x :: 'a::{field_char_0,linordered_field})> 0) = 0<\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) - { assume "a = 0" hence ?thesis by (simp add: x Ngt0_def INum_def) } - moreover - { assume a: "a \ 0" hence b: "(of_int b::'a) > 0" using nx - by (simp add: x isnormNum_def) + show ?thesis + proof (cases "a = 0") + case True + then show ?thesis + by (simp add: x Ngt0_def INum_def) + next + case False + then have b: "(of_int b::'a) > 0" + using nx by (simp add: x isnormNum_def) from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"] - have ?thesis by (simp add: x Ngt0_def INum_def) } - ultimately show ?thesis by blast + show ?thesis + by (simp add: x Ngt0_def INum_def) + qed qed lemma Nge0_iff[simp]: assumes nx: "isnormNum x" - shows "((INum x :: 'a :: {field_char_0, linordered_field}) \ 0) = 0\\<^sub>N x" + shows "((INum x :: 'a::{field_char_0,linordered_field}) \ 0) = 0\\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) - { assume "a = 0" hence ?thesis by (simp add: x Nge0_def INum_def) } - moreover - { assume "a \ 0" hence b: "(of_int b::'a) > 0" using nx - by (simp add: x isnormNum_def) + show ?thesis + proof (cases "a = 0") + case True + then show ?thesis + by (simp add: x Nge0_def INum_def) + next + case False + then have b: "(of_int b::'a) > 0" + using nx by (simp add: x isnormNum_def) from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"] - have ?thesis by (simp add: x Nge0_def INum_def) } - ultimately show ?thesis by blast + show ?thesis + by (simp add: x Nge0_def INum_def) + qed qed lemma Nlt_iff[simp]: - assumes nx: "isnormNum x" and ny: "isnormNum y" - shows "((INum x :: 'a :: {field_char_0, linordered_field}) < INum y) = (x <\<^sub>N y)" + assumes nx: "isnormNum x" + and ny: "isnormNum y" + shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) = (x <\<^sub>N y)" proof - let ?z = "0::'a" have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)" using nx ny by simp also have "\ = (0>\<^sub>N (x -\<^sub>N y))" using Nlt0_iff[OF Nsub_normN[OF ny]] by simp - finally show ?thesis by (simp add: Nlt_def) + finally show ?thesis + by (simp add: Nlt_def) qed lemma Nle_iff[simp]: - assumes nx: "isnormNum x" and ny: "isnormNum y" - shows "((INum x :: 'a :: {field_char_0, linordered_field})\ INum y) = (x \\<^sub>N y)" + assumes nx: "isnormNum x" + and ny: "isnormNum y" + shows "((INum x :: 'a::{field_char_0,linordered_field})\ INum y) = (x \\<^sub>N y)" proof - have "((INum x ::'a) \ INum y) = (INum (x -\<^sub>N y) \ (0::'a))" using nx ny by simp also have "\ = (0\\<^sub>N (x -\<^sub>N y))" using Nle0_iff[OF Nsub_normN[OF ny]] by simp - finally show ?thesis by (simp add: Nle_def) + finally show ?thesis + by (simp add: Nle_def) qed lemma Nadd_commute: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" shows "x +\<^sub>N y = y +\<^sub>N x" proof - - have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all - have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp - with isnormNum_unique[OF n] show ?thesis by simp + have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" + by simp_all + have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" + by simp + with isnormNum_unique[OF n] show ?thesis + by simp qed lemma [simp]: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" shows "(0, b) +\<^sub>N y = normNum y" and "(a, 0) +\<^sub>N y = normNum y" and "x +\<^sub>N (0, b) = normNum x" and "x +\<^sub>N (a, 0) = normNum x" apply (simp add: Nadd_def split_def) apply (simp add: Nadd_def split_def) - apply (subst Nadd_commute, simp add: Nadd_def split_def) - apply (subst Nadd_commute, simp add: Nadd_def split_def) + apply (subst Nadd_commute) + apply (simp add: Nadd_def split_def) + apply (subst Nadd_commute) + apply (simp add: Nadd_def split_def) done lemma normNum_nilpotent_aux[simp]: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" assumes nx: "isnormNum x" shows "normNum x = x" proof - @@ -419,7 +503,7 @@ qed lemma normNum_nilpotent[simp]: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" shows "normNum (normNum x) = normNum x" by simp @@ -427,11 +511,12 @@ by (simp_all add: normNum_def) lemma normNum_Nadd: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" - shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" + shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" + by simp lemma Nadd_normNum1[simp]: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" shows "normNum x +\<^sub>N y = x +\<^sub>N y" proof - have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all @@ -441,7 +526,7 @@ qed lemma Nadd_normNum2[simp]: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" shows "x +\<^sub>N normNum y = x +\<^sub>N y" proof - have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all @@ -451,7 +536,7 @@ qed lemma Nadd_assoc: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)" proof - have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all @@ -463,8 +548,10 @@ by (simp add: Nmul_def split_def Let_def gcd_commute_int mult.commute) lemma Nmul_assoc: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" - assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" + assumes nx: "isnormNum x" + and ny: "isnormNum y" + and nz: "isnormNum z" shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)" proof - from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))" @@ -474,11 +561,11 @@ qed lemma Nsub0: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" - assumes x: "isnormNum x" and y: "isnormNum y" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" + assumes x: "isnormNum x" + and y: "isnormNum y" shows "x -\<^sub>N y = 0\<^sub>N \ x = y" proof - - fix h :: 'a from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"] have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp also have "\ = (INum x = (INum y :: 'a))" by simp @@ -490,11 +577,11 @@ by (simp_all add: Nmul_def Let_def split_def) lemma Nmul_eq0[simp]: - assumes "SORT_CONSTRAINT('a::{field_char_0, field})" - assumes nx: "isnormNum x" and ny: "isnormNum y" + assumes "SORT_CONSTRAINT('a::{field_char_0,field})" + assumes nx: "isnormNum x" + and ny: "isnormNum y" shows "x*\<^sub>N y = 0\<^sub>N \ x = 0\<^sub>N \ y = 0\<^sub>N" proof - - fix h :: 'a obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) have n0: "isnormNum 0\<^sub>N" by simp @@ -508,9 +595,7 @@ lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c" by (simp add: Nneg_def split_def) -lemma Nmul1[simp]: - "isnormNum c \ (1)\<^sub>N *\<^sub>N c = c" - "isnormNum c \ c *\<^sub>N (1)\<^sub>N = c" +lemma Nmul1[simp]: "isnormNum c \ (1)\<^sub>N *\<^sub>N c = c" "isnormNum c \ c *\<^sub>N (1)\<^sub>N = c" apply (simp_all add: Nmul_def Let_def split_def isnormNum_def) apply (cases "fst c = 0", simp_all, cases c, simp_all)+ done