| author | blanchet | 
| Tue, 19 Nov 2013 19:42:30 +0100 | |
| changeset 54506 | 8b5caa190054 | 
| parent 54230 | b1d955791529 | 
| child 56410 | a14831ac3023 | 
| permissions | -rw-r--r-- | 
| 
54220
 
0e6645622f22
more convenient place for a theory in solitariness
 
haftmann 
parents: 
50282 
diff
changeset
 | 
1  | 
(* Title: HOL/Decision_Procs/Rat_Pair.thy  | 
| 24197 | 2  | 
Author: Amine Chaieb  | 
3  | 
*)  | 
|
4  | 
||
| 
54220
 
0e6645622f22
more convenient place for a theory in solitariness
 
haftmann 
parents: 
50282 
diff
changeset
 | 
5  | 
header {* Rational numbers as pairs *}
 | 
| 24197 | 6  | 
|
| 
54220
 
0e6645622f22
more convenient place for a theory in solitariness
 
haftmann 
parents: 
50282 
diff
changeset
 | 
7  | 
theory Rat_Pair  | 
| 36411 | 8  | 
imports Complex_Main  | 
| 24197 | 9  | 
begin  | 
10  | 
||
| 42463 | 11  | 
type_synonym Num = "int \<times> int"  | 
| 25005 | 12  | 
|
| 44780 | 13  | 
abbreviation Num0_syn :: Num  ("0\<^sub>N")
 | 
| 44779 | 14  | 
where "0\<^sub>N \<equiv> (0, 0)"  | 
| 25005 | 15  | 
|
| 
50282
 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 
wenzelm 
parents: 
47162 
diff
changeset
 | 
16  | 
abbreviation Numi_syn :: "int \<Rightarrow> Num"  ("'((_)')\<^sub>N")
 | 
| 
 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 
wenzelm 
parents: 
47162 
diff
changeset
 | 
17  | 
where "(i)\<^sub>N \<equiv> (i, 1)"  | 
| 24197 | 18  | 
|
| 44779 | 19  | 
definition isnormNum :: "Num \<Rightarrow> bool" where  | 
| 31706 | 20  | 
"isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1))"  | 
| 24197 | 21  | 
|
| 44779 | 22  | 
definition normNum :: "Num \<Rightarrow> Num" where  | 
23  | 
"normNum = (\<lambda>(a,b).  | 
|
24  | 
(if a=0 \<or> b = 0 then (0,0) else  | 
|
| 44780 | 25  | 
(let g = gcd a b  | 
| 44779 | 26  | 
in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"  | 
| 24197 | 27  | 
|
| 44779 | 28  | 
declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger]  | 
29  | 
||
| 24197 | 30  | 
lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"  | 
31  | 
proof -  | 
|
| 44780 | 32  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
33  | 
  { assume "a=0 \<or> b = 0" hence ?thesis by (simp add: x normNum_def isnormNum_def) }
 | 
|
| 24197 | 34  | 
moreover  | 
| 44779 | 35  | 
  { assume anz: "a \<noteq> 0" and bnz: "b \<noteq> 0"
 | 
| 31706 | 36  | 
let ?g = "gcd a b"  | 
| 24197 | 37  | 
let ?a' = "a div ?g"  | 
38  | 
let ?b' = "b div ?g"  | 
|
| 31706 | 39  | 
let ?g' = "gcd ?a' ?b'"  | 
| 44779 | 40  | 
from anz bnz have "?g \<noteq> 0" by simp with gcd_ge_0_int[of a b]  | 
| 41528 | 41  | 
have gpos: "?g > 0" by arith  | 
| 44779 | 42  | 
have gdvd: "?g dvd a" "?g dvd b" by arith+  | 
| 47162 | 43  | 
from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] anz bnz  | 
| 44780 | 44  | 
have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+  | 
| 44779 | 45  | 
from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31706 
diff
changeset
 | 
46  | 
from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" .  | 
| 24197 | 47  | 
from bnz have "b < 0 \<or> b > 0" by arith  | 
48  | 
moreover  | 
|
| 44779 | 49  | 
    { assume b: "b > 0"
 | 
50  | 
from b have "?b' \<ge> 0"  | 
|
51  | 
by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])  | 
|
| 44780 | 52  | 
with nz' have b': "?b' > 0" by arith  | 
53  | 
from b b' anz bnz nz' gp1 have ?thesis  | 
|
54  | 
by (simp add: x isnormNum_def normNum_def Let_def split_def) }  | 
|
| 44779 | 55  | 
    moreover {
 | 
56  | 
assume b: "b < 0"  | 
|
| 44780 | 57  | 
      { assume b': "?b' \<ge> 0"
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
58  | 
from gpos have th: "?g \<ge> 0" by arith  | 
| 47162 | 59  | 
from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)]  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
60  | 
have False using b by arith }  | 
| 44779 | 61  | 
hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric])  | 
| 44780 | 62  | 
from anz bnz nz' b b' gp1 have ?thesis  | 
63  | 
by (simp add: x isnormNum_def normNum_def Let_def split_def) }  | 
|
| 24197 | 64  | 
ultimately have ?thesis by blast  | 
65  | 
}  | 
|
66  | 
ultimately show ?thesis by blast  | 
|
67  | 
qed  | 
|
68  | 
||
69  | 
text {* Arithmetic over Num *}
 | 
|
70  | 
||
| 44780 | 71  | 
definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "+\<^sub>N" 60) where  | 
| 44779 | 72  | 
"Nadd = (\<lambda>(a,b) (a',b'). if a = 0 \<or> b = 0 then normNum(a',b')  | 
| 44780 | 73  | 
else if a'=0 \<or> b' = 0 then normNum(a,b)  | 
| 24197 | 74  | 
else normNum(a*b' + b*a', b*b'))"  | 
75  | 
||
| 44780 | 76  | 
definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "*\<^sub>N" 60) where  | 
77  | 
"Nmul = (\<lambda>(a,b) (a',b'). let g = gcd (a*a') (b*b')  | 
|
| 24197 | 78  | 
in (a*a' div g, b*b' div g))"  | 
79  | 
||
| 44779 | 80  | 
definition Nneg :: "Num \<Rightarrow> Num" ("~\<^sub>N")
 | 
81  | 
where "Nneg \<equiv> (\<lambda>(a,b). (-a,b))"  | 
|
| 24197 | 82  | 
|
| 44780 | 83  | 
definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "-\<^sub>N" 60)  | 
| 44779 | 84  | 
where "Nsub = (\<lambda>a b. a +\<^sub>N ~\<^sub>N b)"  | 
| 24197 | 85  | 
|
| 44779 | 86  | 
definition Ninv :: "Num \<Rightarrow> Num"  | 
87  | 
where "Ninv = (\<lambda>(a,b). if a < 0 then (-b, \<bar>a\<bar>) else (b,a))"  | 
|
| 24197 | 88  | 
|
| 44780 | 89  | 
definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num" (infixl "\<div>\<^sub>N" 60)  | 
| 44779 | 90  | 
where "Ndiv = (\<lambda>a b. a *\<^sub>N Ninv b)"  | 
| 24197 | 91  | 
|
92  | 
lemma Nneg_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (~\<^sub>N x)"  | 
|
| 44779 | 93  | 
by (simp add: isnormNum_def Nneg_def split_def)  | 
94  | 
||
| 24197 | 95  | 
lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)"  | 
96  | 
by (simp add: Nadd_def split_def)  | 
|
| 44779 | 97  | 
|
| 24197 | 98  | 
lemma Nsub_normN[simp]: "\<lbrakk> isnormNum y\<rbrakk> \<Longrightarrow> isnormNum (x -\<^sub>N y)"  | 
99  | 
by (simp add: Nsub_def split_def)  | 
|
| 44779 | 100  | 
|
101  | 
lemma Nmul_normN[simp]:  | 
|
| 44780 | 102  | 
assumes xn: "isnormNum x" and yn: "isnormNum y"  | 
| 24197 | 103  | 
shows "isnormNum (x *\<^sub>N y)"  | 
| 44779 | 104  | 
proof -  | 
| 44780 | 105  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
106  | 
obtain a' b' where y: "y = (a', b')" by (cases y)  | 
|
107  | 
  { assume "a = 0"
 | 
|
108  | 
hence ?thesis using xn x y  | 
|
109  | 
by (simp add: isnormNum_def Let_def Nmul_def split_def) }  | 
|
| 24197 | 110  | 
moreover  | 
| 44780 | 111  | 
  { assume "a' = 0"
 | 
112  | 
hence ?thesis using yn x y  | 
|
113  | 
by (simp add: isnormNum_def Let_def Nmul_def split_def) }  | 
|
| 24197 | 114  | 
moreover  | 
| 44780 | 115  | 
  { assume a: "a \<noteq>0" and a': "a'\<noteq>0"
 | 
116  | 
hence bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def)  | 
|
117  | 
from mult_pos_pos[OF bp] have "x *\<^sub>N y = normNum (a * a', b * b')"  | 
|
118  | 
using x y a a' bp by (simp add: Nmul_def Let_def split_def normNum_def)  | 
|
119  | 
hence ?thesis by simp }  | 
|
| 24197 | 120  | 
ultimately show ?thesis by blast  | 
121  | 
qed  | 
|
122  | 
||
123  | 
lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"  | 
|
| 25005 | 124  | 
by (simp add: Ninv_def isnormNum_def split_def)  | 
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31706 
diff
changeset
 | 
125  | 
(cases "fst x = 0", auto simp add: gcd_commute_int)  | 
| 24197 | 126  | 
|
| 44780 | 127  | 
lemma isnormNum_int[simp]:  | 
| 
50282
 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 
wenzelm 
parents: 
47162 
diff
changeset
 | 
128  | 
"isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N"  | 
| 31706 | 129  | 
by (simp_all add: isnormNum_def)  | 
| 24197 | 130  | 
|
131  | 
||
132  | 
text {* Relations over Num *}
 | 
|
133  | 
||
| 44780 | 134  | 
definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
 | 
| 44779 | 135  | 
where "Nlt0 = (\<lambda>(a,b). a < 0)"  | 
| 24197 | 136  | 
|
| 44780 | 137  | 
definition Nle0:: "Num \<Rightarrow> bool"  ("0\<ge>\<^sub>N")
 | 
| 44779 | 138  | 
where "Nle0 = (\<lambda>(a,b). a \<le> 0)"  | 
| 24197 | 139  | 
|
| 44780 | 140  | 
definition Ngt0:: "Num \<Rightarrow> bool"  ("0<\<^sub>N")
 | 
| 44779 | 141  | 
where "Ngt0 = (\<lambda>(a,b). a > 0)"  | 
| 24197 | 142  | 
|
| 44780 | 143  | 
definition Nge0:: "Num \<Rightarrow> bool"  ("0\<le>\<^sub>N")
 | 
| 44779 | 144  | 
where "Nge0 = (\<lambda>(a,b). a \<ge> 0)"  | 
| 24197 | 145  | 
|
| 44780 | 146  | 
definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix "<\<^sub>N" 55)  | 
| 44779 | 147  | 
where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))"  | 
| 24197 | 148  | 
|
| 44779 | 149  | 
definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool" (infix "\<le>\<^sub>N" 55)  | 
150  | 
where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))"  | 
|
| 24197 | 151  | 
|
| 44779 | 152  | 
definition "INum = (\<lambda>(a,b). of_int a / of_int b)"  | 
| 24197 | 153  | 
|
| 
50282
 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 
wenzelm 
parents: 
47162 
diff
changeset
 | 
154  | 
lemma INum_int [simp]: "INum (i)\<^sub>N = ((of_int i) ::'a::field)" "INum 0\<^sub>N = (0::'a::field)"  | 
| 24197 | 155  | 
by (simp_all add: INum_def)  | 
156  | 
||
| 44780 | 157  | 
lemma isnormNum_unique[simp]:  | 
158  | 
assumes na: "isnormNum x" and nb: "isnormNum y"  | 
|
| 36409 | 159  | 
  shows "((INum x ::'a::{field_char_0, field_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
 | 
| 24197 | 160  | 
proof  | 
| 44780 | 161  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
162  | 
obtain a' b' where y: "y = (a', b')" by (cases y)  | 
|
163  | 
assume H: ?lhs  | 
|
| 44779 | 164  | 
  { assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0"
 | 
| 
32456
 
341c83339aeb
tuned the simp rules for Int involving insert and intervals.
 
nipkow 
parents: 
31967 
diff
changeset
 | 
165  | 
hence ?rhs using na nb H  | 
| 44780 | 166  | 
by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) }  | 
| 24197 | 167  | 
moreover  | 
168  | 
  { assume az: "a \<noteq> 0" and bz: "b \<noteq> 0" and a'z: "a'\<noteq>0" and b'z: "b'\<noteq>0"
 | 
|
| 44780 | 169  | 
from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def)  | 
170  | 
from H bz b'z have eq: "a * b' = a'*b"  | 
|
171  | 
by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)  | 
|
172  | 
from az a'z na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"  | 
|
173  | 
by (simp_all add: x y isnormNum_def add: gcd_commute_int)  | 
|
174  | 
from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'"  | 
|
175  | 
apply -  | 
|
| 27668 | 176  | 
apply algebra  | 
177  | 
apply algebra  | 
|
178  | 
apply simp  | 
|
179  | 
apply algebra  | 
|
| 24197 | 180  | 
done  | 
| 33657 | 181  | 
from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)]  | 
| 44780 | 182  | 
coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]]  | 
| 41528 | 183  | 
have eq1: "b = b'" using pos by arith  | 
| 24197 | 184  | 
with eq have "a = a'" using pos by simp  | 
| 44780 | 185  | 
with eq1 have ?rhs by (simp add: x y) }  | 
| 24197 | 186  | 
ultimately show ?rhs by blast  | 
187  | 
next  | 
|
188  | 
assume ?rhs thus ?lhs by simp  | 
|
189  | 
qed  | 
|
190  | 
||
191  | 
||
| 44779 | 192  | 
lemma isnormNum0[simp]:  | 
193  | 
    "isnormNum x \<Longrightarrow> (INum x = (0::'a::{field_char_0, field_inverse_zero})) = (x = 0\<^sub>N)"
 | 
|
| 24197 | 194  | 
unfolding INum_int(2)[symmetric]  | 
| 44779 | 195  | 
by (rule isnormNum_unique) simp_all  | 
| 24197 | 196  | 
|
| 44780 | 197  | 
lemma of_int_div_aux: "d ~= 0 ==> ((of_int x)::'a::field_char_0) / (of_int d) =  | 
| 24197 | 198  | 
of_int (x div d) + (of_int (x mod d)) / ((of_int d)::'a)"  | 
199  | 
proof -  | 
|
200  | 
assume "d ~= 0"  | 
|
201  | 
let ?t = "of_int (x div d) * ((of_int d)::'a) + of_int(x mod d)"  | 
|
202  | 
let ?f = "\<lambda>x. x / of_int d"  | 
|
203  | 
have "x = (x div d) * d + x mod d"  | 
|
204  | 
by auto  | 
|
205  | 
then have eq: "of_int x = ?t"  | 
|
206  | 
by (simp only: of_int_mult[symmetric] of_int_add [symmetric])  | 
|
| 44780 | 207  | 
then have "of_int x / of_int d = ?t / of_int d"  | 
| 24197 | 208  | 
using cong[OF refl[of ?f] eq] by simp  | 
| 41528 | 209  | 
then show ?thesis by (simp add: add_divide_distrib algebra_simps `d ~= 0`)  | 
| 24197 | 210  | 
qed  | 
211  | 
||
212  | 
lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>  | 
|
| 36409 | 213  | 
(of_int(n div d)::'a::field_char_0) = of_int n / of_int d"  | 
| 24197 | 214  | 
apply (frule of_int_div_aux [of d n, where ?'a = 'a])  | 
215  | 
apply simp  | 
|
| 30042 | 216  | 
apply (simp add: dvd_eq_mod_eq_0)  | 
| 44779 | 217  | 
done  | 
| 24197 | 218  | 
|
219  | 
||
| 36409 | 220  | 
lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0, field_inverse_zero})"
 | 
| 44779 | 221  | 
proof -  | 
| 44780 | 222  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
223  | 
  { assume "a = 0 \<or> b = 0"
 | 
|
224  | 
hence ?thesis by (simp add: x INum_def normNum_def split_def Let_def) }  | 
|
225  | 
moreover  | 
|
226  | 
  { assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
 | 
|
| 31706 | 227  | 
let ?g = "gcd a b"  | 
| 24197 | 228  | 
from a b have g: "?g \<noteq> 0"by simp  | 
229  | 
from of_int_div[OF g, where ?'a = 'a]  | 
|
| 44779 | 230  | 
have ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) }  | 
| 24197 | 231  | 
ultimately show ?thesis by blast  | 
232  | 
qed  | 
|
233  | 
||
| 44779 | 234  | 
lemma INum_normNum_iff:  | 
235  | 
  "(INum x ::'a::{field_char_0, field_inverse_zero}) = INum y \<longleftrightarrow> normNum x = normNum y"
 | 
|
236  | 
(is "?lhs = ?rhs")  | 
|
| 24197 | 237  | 
proof -  | 
238  | 
have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"  | 
|
239  | 
by (simp del: normNum)  | 
|
240  | 
also have "\<dots> = ?lhs" by simp  | 
|
241  | 
finally show ?thesis by simp  | 
|
242  | 
qed  | 
|
243  | 
||
| 36409 | 244  | 
lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0, field_inverse_zero})"
 | 
| 44779 | 245  | 
proof -  | 
246  | 
let ?z = "0:: 'a"  | 
|
| 44780 | 247  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
248  | 
obtain a' b' where y: "y = (a', b')" by (cases y)  | 
|
| 44779 | 249  | 
  { assume "a=0 \<or> a'= 0 \<or> b =0 \<or> b' = 0"
 | 
| 44780 | 250  | 
hence ?thesis  | 
251  | 
apply (cases "a=0", simp_all add: x y Nadd_def)  | 
|
| 44779 | 252  | 
apply (cases "b= 0", simp_all add: INum_def)  | 
253  | 
apply (cases "a'= 0", simp_all)  | 
|
254  | 
apply (cases "b'= 0", simp_all)  | 
|
| 24197 | 255  | 
done }  | 
| 44780 | 256  | 
moreover  | 
257  | 
  { assume aa': "a \<noteq> 0" "a'\<noteq> 0" and bb': "b \<noteq> 0" "b' \<noteq> 0"
 | 
|
| 44779 | 258  | 
    { assume z: "a * b' + b * a' = 0"
 | 
| 24197 | 259  | 
hence "of_int (a*b' + b*a') / (of_int b* of_int b') = ?z" by simp  | 
| 44780 | 260  | 
hence "of_int b' * of_int a / (of_int b * of_int b') +  | 
261  | 
of_int b * of_int a' / (of_int b * of_int b') = ?z"  | 
|
262  | 
by (simp add:add_divide_distrib)  | 
|
| 44779 | 263  | 
hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa'  | 
| 44780 | 264  | 
by simp  | 
265  | 
from z aa' bb' have ?thesis  | 
|
266  | 
by (simp add: x y th Nadd_def normNum_def INum_def split_def) }  | 
|
| 44779 | 267  | 
    moreover {
 | 
268  | 
assume z: "a * b' + b * a' \<noteq> 0"  | 
|
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
54220 
diff
changeset
 | 
269  | 
let ?g = "gcd (a * b' + b * a') (b * b')"  | 
| 24197 | 270  | 
have gz: "?g \<noteq> 0" using z by simp  | 
271  | 
have ?thesis using aa' bb' z gz  | 
|
| 44779 | 272  | 
of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]  | 
273  | 
of_int_div[where ?'a = 'a, OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]  | 
|
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
54220 
diff
changeset
 | 
274  | 
by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps)  | 
| 
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
54220 
diff
changeset
 | 
275  | 
}  | 
| 44779 | 276  | 
ultimately have ?thesis using aa' bb'  | 
| 44780 | 277  | 
by (simp add: x y Nadd_def INum_def normNum_def Let_def) }  | 
| 24197 | 278  | 
ultimately show ?thesis by blast  | 
279  | 
qed  | 
|
280  | 
||
| 44779 | 281  | 
lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {field_char_0, field_inverse_zero})"
 | 
282  | 
proof -  | 
|
| 24197 | 283  | 
let ?z = "0::'a"  | 
| 44780 | 284  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
285  | 
obtain a' b' where y: "y = (a', b')" by (cases y)  | 
|
| 44779 | 286  | 
  { assume "a=0 \<or> a'= 0 \<or> b = 0 \<or> b' = 0"
 | 
| 44780 | 287  | 
hence ?thesis  | 
| 44779 | 288  | 
apply (cases "a=0", simp_all add: x y Nmul_def INum_def Let_def)  | 
289  | 
apply (cases "b=0", simp_all)  | 
|
| 44780 | 290  | 
apply (cases "a'=0", simp_all)  | 
| 24197 | 291  | 
done }  | 
292  | 
moreover  | 
|
| 44779 | 293  | 
  { assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
 | 
| 31706 | 294  | 
let ?g="gcd (a*a') (b*b')"  | 
| 24197 | 295  | 
have gz: "?g \<noteq> 0" using z by simp  | 
| 44779 | 296  | 
from z of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a*a'" and y="b*b'"]]  | 
| 44780 | 297  | 
of_int_div[where ?'a = 'a , OF gz gcd_dvd2_int[where x="a*a'" and y="b*b'"]]  | 
| 44779 | 298  | 
have ?thesis by (simp add: Nmul_def x y Let_def INum_def) }  | 
| 24197 | 299  | 
ultimately show ?thesis by blast  | 
300  | 
qed  | 
|
301  | 
||
302  | 
lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"  | 
|
303  | 
by (simp add: Nneg_def split_def INum_def)  | 
|
304  | 
||
| 44779 | 305  | 
lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0, field_inverse_zero})"
 | 
306  | 
by (simp add: Nsub_def split_def)  | 
|
| 24197 | 307  | 
|
| 36409 | 308  | 
lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field_inverse_zero) / (INum x)"  | 
| 24197 | 309  | 
by (simp add: Ninv_def INum_def split_def)  | 
310  | 
||
| 44779 | 311  | 
lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0, field_inverse_zero})"
 | 
312  | 
by (simp add: Ndiv_def)  | 
|
| 24197 | 313  | 
|
| 44779 | 314  | 
lemma Nlt0_iff[simp]:  | 
| 44780 | 315  | 
assumes nx: "isnormNum x"  | 
| 44779 | 316  | 
  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})< 0) = 0>\<^sub>N x"
 | 
317  | 
proof -  | 
|
| 44780 | 318  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
319  | 
  { assume "a = 0" hence ?thesis by (simp add: x Nlt0_def INum_def) }
 | 
|
| 24197 | 320  | 
moreover  | 
| 44780 | 321  | 
  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0"
 | 
322  | 
using nx by (simp add: x isnormNum_def)  | 
|
| 24197 | 323  | 
from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"]  | 
| 44780 | 324  | 
have ?thesis by (simp add: x Nlt0_def INum_def) }  | 
| 24197 | 325  | 
ultimately show ?thesis by blast  | 
326  | 
qed  | 
|
327  | 
||
| 44779 | 328  | 
lemma Nle0_iff[simp]:  | 
329  | 
assumes nx: "isnormNum x"  | 
|
| 36409 | 330  | 
  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<le> 0) = 0\<ge>\<^sub>N x"
 | 
| 44779 | 331  | 
proof -  | 
| 44780 | 332  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
333  | 
  { assume "a = 0" hence ?thesis by (simp add: x Nle0_def INum_def) }
 | 
|
| 24197 | 334  | 
moreover  | 
| 44780 | 335  | 
  { assume a: "a \<noteq> 0" hence b: "(of_int b :: 'a) > 0"
 | 
336  | 
using nx by (simp add: x isnormNum_def)  | 
|
| 24197 | 337  | 
from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"]  | 
| 44780 | 338  | 
have ?thesis by (simp add: x Nle0_def INum_def) }  | 
| 24197 | 339  | 
ultimately show ?thesis by blast  | 
340  | 
qed  | 
|
341  | 
||
| 44779 | 342  | 
lemma Ngt0_iff[simp]:  | 
343  | 
assumes nx: "isnormNum x"  | 
|
344  | 
  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})> 0) = 0<\<^sub>N x"
 | 
|
345  | 
proof -  | 
|
| 44780 | 346  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
347  | 
  { assume "a = 0" hence ?thesis by (simp add: x Ngt0_def INum_def) }
 | 
|
| 24197 | 348  | 
moreover  | 
| 44780 | 349  | 
  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
 | 
350  | 
by (simp add: x isnormNum_def)  | 
|
| 24197 | 351  | 
from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"]  | 
| 44780 | 352  | 
have ?thesis by (simp add: x Ngt0_def INum_def) }  | 
| 24197 | 353  | 
ultimately show ?thesis by blast  | 
354  | 
qed  | 
|
355  | 
||
| 44779 | 356  | 
lemma Nge0_iff[simp]:  | 
357  | 
assumes nx: "isnormNum x"  | 
|
358  | 
  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<ge> 0) = 0\<le>\<^sub>N x"
 | 
|
359  | 
proof -  | 
|
| 44780 | 360  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
361  | 
  { assume "a = 0" hence ?thesis by (simp add: x Nge0_def INum_def) }
 | 
|
| 44779 | 362  | 
moreover  | 
363  | 
  { assume "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
 | 
|
| 44780 | 364  | 
by (simp add: x isnormNum_def)  | 
| 44779 | 365  | 
from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"]  | 
| 44780 | 366  | 
have ?thesis by (simp add: x Nge0_def INum_def) }  | 
| 44779 | 367  | 
ultimately show ?thesis by blast  | 
368  | 
qed  | 
|
369  | 
||
370  | 
lemma Nlt_iff[simp]:  | 
|
371  | 
assumes nx: "isnormNum x" and ny: "isnormNum y"  | 
|
| 36409 | 372  | 
  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) < INum y) = (x <\<^sub>N y)"
 | 
| 44779 | 373  | 
proof -  | 
| 24197 | 374  | 
let ?z = "0::'a"  | 
| 44779 | 375  | 
have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)"  | 
376  | 
using nx ny by simp  | 
|
377  | 
also have "\<dots> = (0>\<^sub>N (x -\<^sub>N y))"  | 
|
378  | 
using Nlt0_iff[OF Nsub_normN[OF ny]] by simp  | 
|
| 24197 | 379  | 
finally show ?thesis by (simp add: Nlt_def)  | 
380  | 
qed  | 
|
381  | 
||
| 44779 | 382  | 
lemma Nle_iff[simp]:  | 
383  | 
assumes nx: "isnormNum x" and ny: "isnormNum y"  | 
|
| 36409 | 384  | 
  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})\<le> INum y) = (x \<le>\<^sub>N y)"
 | 
| 44779 | 385  | 
proof -  | 
386  | 
have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))"  | 
|
387  | 
using nx ny by simp  | 
|
388  | 
also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))"  | 
|
389  | 
using Nle0_iff[OF Nsub_normN[OF ny]] by simp  | 
|
| 24197 | 390  | 
finally show ?thesis by (simp add: Nle_def)  | 
391  | 
qed  | 
|
392  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
393  | 
lemma Nadd_commute:  | 
| 36409 | 394  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
395  | 
shows "x +\<^sub>N y = y +\<^sub>N x"  | 
| 44779 | 396  | 
proof -  | 
| 24197 | 397  | 
have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all  | 
| 31964 | 398  | 
have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp  | 
| 24197 | 399  | 
with isnormNum_unique[OF n] show ?thesis by simp  | 
400  | 
qed  | 
|
401  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
402  | 
lemma [simp]:  | 
| 36409 | 403  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
404  | 
shows "(0, b) +\<^sub>N y = normNum y"  | 
| 44780 | 405  | 
and "(a, 0) +\<^sub>N y = normNum y"  | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
406  | 
and "x +\<^sub>N (0, b) = normNum x"  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
407  | 
and "x +\<^sub>N (a, 0) = normNum x"  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
408  | 
apply (simp add: Nadd_def split_def)  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
409  | 
apply (simp add: Nadd_def split_def)  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
410  | 
apply (subst Nadd_commute, simp add: Nadd_def split_def)  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
411  | 
apply (subst Nadd_commute, simp add: Nadd_def split_def)  | 
| 24197 | 412  | 
done  | 
413  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
414  | 
lemma normNum_nilpotent_aux[simp]:  | 
| 36409 | 415  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 44780 | 416  | 
assumes nx: "isnormNum x"  | 
| 24197 | 417  | 
shows "normNum x = x"  | 
| 44779 | 418  | 
proof -  | 
| 24197 | 419  | 
let ?a = "normNum x"  | 
420  | 
have n: "isnormNum ?a" by simp  | 
|
| 44779 | 421  | 
have th: "INum ?a = (INum x ::'a)" by simp  | 
422  | 
with isnormNum_unique[OF n nx] show ?thesis by simp  | 
|
| 24197 | 423  | 
qed  | 
424  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
425  | 
lemma normNum_nilpotent[simp]:  | 
| 36409 | 426  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
427  | 
shows "normNum (normNum x) = normNum x"  | 
| 24197 | 428  | 
by simp  | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
429  | 
|
| 24197 | 430  | 
lemma normNum0[simp]: "normNum (0,b) = 0\<^sub>N" "normNum (a,0) = 0\<^sub>N"  | 
431  | 
by (simp_all add: normNum_def)  | 
|
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
432  | 
|
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
433  | 
lemma normNum_Nadd:  | 
| 36409 | 434  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
435  | 
shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
436  | 
|
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
437  | 
lemma Nadd_normNum1[simp]:  | 
| 36409 | 438  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
439  | 
shows "normNum x +\<^sub>N y = x +\<^sub>N y"  | 
| 44779 | 440  | 
proof -  | 
| 24197 | 441  | 
have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all  | 
| 31964 | 442  | 
have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp  | 
| 24197 | 443  | 
also have "\<dots> = INum (x +\<^sub>N y)" by simp  | 
444  | 
finally show ?thesis using isnormNum_unique[OF n] by simp  | 
|
445  | 
qed  | 
|
446  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
447  | 
lemma Nadd_normNum2[simp]:  | 
| 36409 | 448  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
449  | 
shows "x +\<^sub>N normNum y = x +\<^sub>N y"  | 
| 44779 | 450  | 
proof -  | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
451  | 
have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all  | 
| 31964 | 452  | 
have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp  | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
453  | 
also have "\<dots> = INum (x +\<^sub>N y)" by simp  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
454  | 
finally show ?thesis using isnormNum_unique[OF n] by simp  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
455  | 
qed  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
456  | 
|
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
457  | 
lemma Nadd_assoc:  | 
| 36409 | 458  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
459  | 
shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"  | 
| 44779 | 460  | 
proof -  | 
| 24197 | 461  | 
have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all  | 
| 31964 | 462  | 
have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp  | 
| 24197 | 463  | 
with isnormNum_unique[OF n] show ?thesis by simp  | 
464  | 
qed  | 
|
465  | 
||
466  | 
lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31706 
diff
changeset
 | 
467  | 
by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)  | 
| 24197 | 468  | 
|
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
469  | 
lemma Nmul_assoc:  | 
| 36409 | 470  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 44780 | 471  | 
assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z"  | 
| 24197 | 472  | 
shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"  | 
| 44779 | 473  | 
proof -  | 
| 44780 | 474  | 
from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))"  | 
| 24197 | 475  | 
by simp_all  | 
| 31964 | 476  | 
have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp  | 
| 24197 | 477  | 
with isnormNum_unique[OF n] show ?thesis by simp  | 
478  | 
qed  | 
|
479  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
480  | 
lemma Nsub0:  | 
| 36409 | 481  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 44780 | 482  | 
assumes x: "isnormNum x" and y: "isnormNum y"  | 
483  | 
shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y"  | 
|
| 44779 | 484  | 
proof -  | 
485  | 
fix h :: 'a  | 
|
| 44780 | 486  | 
from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"]  | 
| 44779 | 487  | 
have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp  | 
488  | 
also have "\<dots> = (INum x = (INum y :: 'a))" by simp  | 
|
489  | 
also have "\<dots> = (x = y)" using x y by simp  | 
|
490  | 
finally show ?thesis .  | 
|
| 24197 | 491  | 
qed  | 
492  | 
||
493  | 
lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N"  | 
|
494  | 
by (simp_all add: Nmul_def Let_def split_def)  | 
|
495  | 
||
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
496  | 
lemma Nmul_eq0[simp]:  | 
| 36409 | 497  | 
  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
 | 
| 44780 | 498  | 
assumes nx: "isnormNum x" and ny: "isnormNum y"  | 
499  | 
shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N"  | 
|
| 44779 | 500  | 
proof -  | 
501  | 
fix h :: 'a  | 
|
| 44780 | 502  | 
obtain a b where x: "x = (a, b)" by (cases x)  | 
503  | 
obtain a' b' where y: "y = (a', b')" by (cases y)  | 
|
| 44779 | 504  | 
have n0: "isnormNum 0\<^sub>N" by simp  | 
| 44780 | 505  | 
show ?thesis using nx ny  | 
| 44779 | 506  | 
apply (simp only: isnormNum_unique[where ?'a = 'a, OF Nmul_normN[OF nx ny] n0, symmetric]  | 
507  | 
Nmul[where ?'a = 'a])  | 
|
| 44780 | 508  | 
apply (simp add: x y INum_def split_def isnormNum_def split: split_if_asm)  | 
| 44779 | 509  | 
done  | 
| 24197 | 510  | 
qed  | 
| 44779 | 511  | 
|
| 24197 | 512  | 
lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"  | 
513  | 
by (simp add: Nneg_def split_def)  | 
|
514  | 
||
| 44780 | 515  | 
lemma Nmul1[simp]:  | 
| 
50282
 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 
wenzelm 
parents: 
47162 
diff
changeset
 | 
516  | 
"isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c"  | 
| 
 
fe4d4bb9f4c2
more robust syntax that survives collapse of \<^isub> and \<^sub>;
 
wenzelm 
parents: 
47162 
diff
changeset
 | 
517  | 
"isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c"  | 
| 24197 | 518  | 
apply (simp_all add: Nmul_def Let_def split_def isnormNum_def)  | 
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
519  | 
apply (cases "fst c = 0", simp_all, cases c, simp_all)+  | 
| 
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
520  | 
done  | 
| 24197 | 521  | 
|
| 
28615
 
4c8fa015ec7f
explicit SORT_CONSTRAINT for proofs depending implicitly on certain sorts;
 
wenzelm 
parents: 
27668 
diff
changeset
 | 
522  | 
end  |