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