| author | wenzelm | 
| Fri, 07 Dec 2018 14:58:32 +0100 | |
| changeset 69419 | 8985ee17bfd2 | 
| parent 69182 | 2424301cc73d | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Int.thy  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 41959 | 3  | 
Author: Tobias Nipkow, Florian Haftmann, TU Muenchen  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
5  | 
|
| 60758 | 6  | 
section \<open>The Integers as Equivalence Classes over Pairs of Natural Numbers\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
8  | 
theory Int  | 
| 63652 | 9  | 
imports Equiv_Relations Power Quotient Fun_Def  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
11  | 
|
| 60758 | 12  | 
subsection \<open>Definition of integers as a quotient type\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
13  | 
|
| 63652 | 14  | 
definition intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"  | 
15  | 
where "intrel = (\<lambda>(x, y) (u, v). x + v = u + y)"  | 
|
| 48045 | 16  | 
|
17  | 
lemma intrel_iff [simp]: "intrel (x, y) (u, v) \<longleftrightarrow> x + v = u + y"  | 
|
18  | 
by (simp add: intrel_def)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
|
| 48045 | 20  | 
quotient_type int = "nat \<times> nat" / "intrel"  | 
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
45607 
diff
changeset
 | 
21  | 
morphisms Rep_Integ Abs_Integ  | 
| 48045 | 22  | 
proof (rule equivpI)  | 
| 63652 | 23  | 
show "reflp intrel" by (auto simp: reflp_def)  | 
24  | 
show "symp intrel" by (auto simp: symp_def)  | 
|
25  | 
show "transp intrel" by (auto simp: transp_def)  | 
|
| 48045 | 26  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
27  | 
|
| 48045 | 28  | 
lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:  | 
| 63652 | 29  | 
"(\<And>x y. z = Abs_Integ (x, y) \<Longrightarrow> P) \<Longrightarrow> P"  | 
30  | 
by (induct z) auto  | 
|
31  | 
||
| 48045 | 32  | 
|
| 60758 | 33  | 
subsection \<open>Integers form a commutative ring\<close>  | 
| 48045 | 34  | 
|
35  | 
instantiation int :: comm_ring_1  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
36  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
37  | 
|
| 51994 | 38  | 
lift_definition zero_int :: "int" is "(0, 0)" .  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
39  | 
|
| 51994 | 40  | 
lift_definition one_int :: "int" is "(1, 0)" .  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
41  | 
|
| 48045 | 42  | 
lift_definition plus_int :: "int \<Rightarrow> int \<Rightarrow> int"  | 
43  | 
is "\<lambda>(x, y) (u, v). (x + u, y + v)"  | 
|
44  | 
by clarsimp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
45  | 
|
| 48045 | 46  | 
lift_definition uminus_int :: "int \<Rightarrow> int"  | 
47  | 
is "\<lambda>(x, y). (y, x)"  | 
|
48  | 
by clarsimp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
49  | 
|
| 48045 | 50  | 
lift_definition minus_int :: "int \<Rightarrow> int \<Rightarrow> int"  | 
51  | 
is "\<lambda>(x, y) (u, v). (x + v, y + u)"  | 
|
52  | 
by clarsimp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
53  | 
|
| 48045 | 54  | 
lift_definition times_int :: "int \<Rightarrow> int \<Rightarrow> int"  | 
55  | 
is "\<lambda>(x, y) (u, v). (x*u + y*v, x*v + y*u)"  | 
|
56  | 
proof (clarsimp)  | 
|
57  | 
fix s t u v w x y z :: nat  | 
|
58  | 
assume "s + v = u + t" and "w + z = y + x"  | 
|
| 63652 | 59  | 
then have "(s + v) * w + (u + t) * x + u * (w + z) + v * (y + x) =  | 
60  | 
(u + t) * w + (s + v) * x + u * (y + x) + v * (w + z)"  | 
|
| 48045 | 61  | 
by simp  | 
| 63652 | 62  | 
then show "(s * w + t * x) + (u * z + v * y) = (u * y + v * z) + (s * x + t * w)"  | 
| 48045 | 63  | 
by (simp add: algebra_simps)  | 
64  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
65  | 
|
| 48045 | 66  | 
instance  | 
| 63652 | 67  | 
by standard (transfer; clarsimp simp: algebra_simps)+  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
69  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
70  | 
|
| 63652 | 71  | 
abbreviation int :: "nat \<Rightarrow> int"  | 
72  | 
where "int \<equiv> of_nat"  | 
|
| 44709 | 73  | 
|
| 48045 | 74  | 
lemma int_def: "int n = Abs_Integ (n, 0)"  | 
| 63652 | 75  | 
by (induct n) (simp add: zero_int.abs_eq, simp add: one_int.abs_eq plus_int.abs_eq)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
76  | 
|
| 67399 | 77  | 
lemma int_transfer [transfer_rule]: "(rel_fun (=) pcr_int) (\<lambda>n. (n, 0)) int"  | 
| 63652 | 78  | 
by (simp add: rel_fun_def int.pcr_cr_eq cr_int_def int_def)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
79  | 
|
| 63652 | 80  | 
lemma int_diff_cases: obtains (diff) m n where "z = int m - int n"  | 
| 48045 | 81  | 
by transfer clarsimp  | 
82  | 
||
| 63652 | 83  | 
|
| 60758 | 84  | 
subsection \<open>Integers are totally ordered\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
85  | 
|
| 48045 | 86  | 
instantiation int :: linorder  | 
87  | 
begin  | 
|
88  | 
||
89  | 
lift_definition less_eq_int :: "int \<Rightarrow> int \<Rightarrow> bool"  | 
|
90  | 
is "\<lambda>(x, y) (u, v). x + v \<le> u + y"  | 
|
91  | 
by auto  | 
|
92  | 
||
93  | 
lift_definition less_int :: "int \<Rightarrow> int \<Rightarrow> bool"  | 
|
94  | 
is "\<lambda>(x, y) (u, v). x + v < u + y"  | 
|
95  | 
by auto  | 
|
96  | 
||
97  | 
instance  | 
|
| 61169 | 98  | 
by standard (transfer, force)+  | 
| 48045 | 99  | 
|
100  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
102  | 
instantiation int :: distrib_lattice  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
103  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
104  | 
|
| 63652 | 105  | 
definition "(inf :: int \<Rightarrow> int \<Rightarrow> int) = min"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
106  | 
|
| 63652 | 107  | 
definition "(sup :: int \<Rightarrow> int \<Rightarrow> int) = max"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
108  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
109  | 
instance  | 
| 63652 | 110  | 
by standard (auto simp add: inf_int_def sup_int_def max_min_distrib2)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
111  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
112  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
113  | 
|
| 60758 | 114  | 
subsection \<open>Ordering properties of arithmetic operations\<close>  | 
| 48045 | 115  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34055 
diff
changeset
 | 
116  | 
instance int :: ordered_cancel_ab_semigroup_add  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
117  | 
proof  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
118  | 
fix i j k :: int  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
119  | 
show "i \<le> j \<Longrightarrow> k + i \<le> k + j"  | 
| 48045 | 120  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
121  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
122  | 
|
| 63652 | 123  | 
text \<open>Strict Monotonicity of Multiplication.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
124  | 
|
| 63652 | 125  | 
text \<open>Strict, in 1st argument; proof is by induction on \<open>k > 0\<close>.\<close>  | 
126  | 
lemma zmult_zless_mono2_lemma: "i < j \<Longrightarrow> 0 < k \<Longrightarrow> int k * i < int k * j"  | 
|
127  | 
for i j :: int  | 
|
128  | 
proof (induct k)  | 
|
129  | 
case 0  | 
|
130  | 
then show ?case by simp  | 
|
131  | 
next  | 
|
132  | 
case (Suc k)  | 
|
133  | 
then show ?case  | 
|
134  | 
by (cases "k = 0") (simp_all add: distrib_right add_strict_mono)  | 
|
135  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
136  | 
|
| 63652 | 137  | 
lemma zero_le_imp_eq_int: "0 \<le> k \<Longrightarrow> \<exists>n. k = int n"  | 
138  | 
for k :: int  | 
|
139  | 
apply transfer  | 
|
140  | 
apply clarsimp  | 
|
141  | 
apply (rule_tac x="a - b" in exI)  | 
|
142  | 
apply simp  | 
|
143  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
144  | 
|
| 63652 | 145  | 
lemma zero_less_imp_eq_int: "0 < k \<Longrightarrow> \<exists>n>0. k = int n"  | 
146  | 
for k :: int  | 
|
147  | 
apply transfer  | 
|
148  | 
apply clarsimp  | 
|
149  | 
apply (rule_tac x="a - b" in exI)  | 
|
150  | 
apply simp  | 
|
151  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
152  | 
|
| 63652 | 153  | 
lemma zmult_zless_mono2: "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"  | 
154  | 
for i j k :: int  | 
|
155  | 
by (drule zero_less_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
156  | 
|
| 63652 | 157  | 
|
158  | 
text \<open>The integers form an ordered integral domain.\<close>  | 
|
159  | 
||
| 48045 | 160  | 
instantiation int :: linordered_idom  | 
161  | 
begin  | 
|
162  | 
||
| 63652 | 163  | 
definition zabs_def: "\<bar>i::int\<bar> = (if i < 0 then - i else i)"  | 
| 48045 | 164  | 
|
| 63652 | 165  | 
definition zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"  | 
| 48045 | 166  | 
|
| 63652 | 167  | 
instance  | 
168  | 
proof  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
169  | 
fix i j k :: int  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
170  | 
show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
171  | 
by (rule zmult_zless_mono2)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
172  | 
show "\<bar>i\<bar> = (if i < 0 then -i else i)"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
173  | 
by (simp only: zabs_def)  | 
| 61076 | 174  | 
show "sgn (i::int) = (if i=0 then 0 else if 0<i then 1 else - 1)"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
175  | 
by (simp only: zsgn_def)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
176  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
177  | 
|
| 48045 | 178  | 
end  | 
179  | 
||
| 63652 | 180  | 
lemma zless_imp_add1_zle: "w < z \<Longrightarrow> w + 1 \<le> z"  | 
181  | 
for w z :: int  | 
|
| 48045 | 182  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
183  | 
|
| 63652 | 184  | 
lemma zless_iff_Suc_zadd: "w < z \<longleftrightarrow> (\<exists>n. z = w + int (Suc n))"  | 
185  | 
for w z :: int  | 
|
186  | 
apply transfer  | 
|
187  | 
apply auto  | 
|
188  | 
apply (rename_tac a b c d)  | 
|
189  | 
apply (rule_tac x="c+b - Suc(a+d)" in exI)  | 
|
190  | 
apply arith  | 
|
191  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
192  | 
|
| 63652 | 193  | 
lemma zabs_less_one_iff [simp]: "\<bar>z\<bar> < 1 \<longleftrightarrow> z = 0" (is "?lhs \<longleftrightarrow> ?rhs")  | 
194  | 
for z :: int  | 
|
| 62347 | 195  | 
proof  | 
| 63652 | 196  | 
assume ?rhs  | 
197  | 
then show ?lhs by simp  | 
|
| 62347 | 198  | 
next  | 
| 63652 | 199  | 
assume ?lhs  | 
200  | 
with zless_imp_add1_zle [of "\<bar>z\<bar>" 1] have "\<bar>z\<bar> + 1 \<le> 1" by simp  | 
|
201  | 
then have "\<bar>z\<bar> \<le> 0" by simp  | 
|
202  | 
then show ?rhs by simp  | 
|
| 62347 | 203  | 
qed  | 
204  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
205  | 
|
| 61799 | 206  | 
subsection \<open>Embedding of the Integers into any \<open>ring_1\<close>: \<open>of_int\<close>\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
208  | 
context ring_1  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
209  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
210  | 
|
| 63652 | 211  | 
lift_definition of_int :: "int \<Rightarrow> 'a"  | 
212  | 
is "\<lambda>(i, j). of_nat i - of_nat j"  | 
|
| 48045 | 213  | 
by (clarsimp simp add: diff_eq_eq eq_diff_eq diff_add_eq  | 
| 63652 | 214  | 
of_nat_add [symmetric] simp del: of_nat_add)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
216  | 
lemma of_int_0 [simp]: "of_int 0 = 0"  | 
| 
48066
 
c6783c9b87bf
transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
 
huffman 
parents: 
48045 
diff
changeset
 | 
217  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
218  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
219  | 
lemma of_int_1 [simp]: "of_int 1 = 1"  | 
| 
48066
 
c6783c9b87bf
transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
 
huffman 
parents: 
48045 
diff
changeset
 | 
220  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
221  | 
|
| 63652 | 222  | 
lemma of_int_add [simp]: "of_int (w + z) = of_int w + of_int z"  | 
| 
48066
 
c6783c9b87bf
transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
 
huffman 
parents: 
48045 
diff
changeset
 | 
223  | 
by transfer (clarsimp simp add: algebra_simps)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
224  | 
|
| 63652 | 225  | 
lemma of_int_minus [simp]: "of_int (- z) = - (of_int z)"  | 
| 
48066
 
c6783c9b87bf
transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
 
huffman 
parents: 
48045 
diff
changeset
 | 
226  | 
by (transfer fixing: uminus) clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
227  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
228  | 
lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
54223 
diff
changeset
 | 
229  | 
using of_int_add [of w "- z"] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
230  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
231  | 
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"  | 
| 63652 | 232  | 
by (transfer fixing: times) (clarsimp simp add: algebra_simps)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
233  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61524 
diff
changeset
 | 
234  | 
lemma mult_of_int_commute: "of_int x * y = y * of_int x"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61524 
diff
changeset
 | 
235  | 
by (transfer fixing: times) (auto simp: algebra_simps mult_of_nat_commute)  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61524 
diff
changeset
 | 
236  | 
|
| 63652 | 237  | 
text \<open>Collapse nested embeddings.\<close>  | 
| 44709 | 238  | 
lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n"  | 
| 63652 | 239  | 
by (induct n) auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
240  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
241  | 
lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
242  | 
by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric])  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
243  | 
|
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
244  | 
lemma of_int_neg_numeral [code_post]: "of_int (- numeral k) = - numeral k"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
245  | 
by simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
246  | 
|
| 63652 | 247  | 
lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n"  | 
| 31015 | 248  | 
by (induct n) simp_all  | 
249  | 
||
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
250  | 
lemma of_int_of_bool [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
251  | 
"of_int (of_bool P) = of_bool P"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
252  | 
by auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
253  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
254  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
255  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
256  | 
context ring_char_0  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
257  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
258  | 
|
| 63652 | 259  | 
lemma of_int_eq_iff [simp]: "of_int w = of_int z \<longleftrightarrow> w = z"  | 
260  | 
by transfer (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
261  | 
|
| 63652 | 262  | 
text \<open>Special cases where either operand is zero.\<close>  | 
263  | 
lemma of_int_eq_0_iff [simp]: "of_int z = 0 \<longleftrightarrow> z = 0"  | 
|
| 36424 | 264  | 
using of_int_eq_iff [of z 0] by simp  | 
265  | 
||
| 63652 | 266  | 
lemma of_int_0_eq_iff [simp]: "0 = of_int z \<longleftrightarrow> z = 0"  | 
| 36424 | 267  | 
using of_int_eq_iff [of 0 z] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
268  | 
|
| 63652 | 269  | 
lemma of_int_eq_1_iff [iff]: "of_int z = 1 \<longleftrightarrow> z = 1"  | 
| 61234 | 270  | 
using of_int_eq_iff [of z 1] by simp  | 
271  | 
||
| 
66912
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
272  | 
lemma numeral_power_eq_of_int_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
273  | 
"numeral x ^ n = of_int y \<longleftrightarrow> numeral x ^ n = y"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
274  | 
using of_int_eq_iff[of "numeral x ^ n" y, unfolded of_int_numeral of_int_power] .  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
275  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
276  | 
lemma of_int_eq_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
277  | 
"of_int y = numeral x ^ n \<longleftrightarrow> y = numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
278  | 
using numeral_power_eq_of_int_cancel_iff [of x n y] by (metis (mono_tags))  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
279  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
280  | 
lemma neg_numeral_power_eq_of_int_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
281  | 
"(- numeral x) ^ n = of_int y \<longleftrightarrow> (- numeral x) ^ n = y"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
282  | 
using of_int_eq_iff[of "(- numeral x) ^ n" y]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
283  | 
by simp  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
284  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
285  | 
lemma of_int_eq_neg_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
286  | 
"of_int y = (- numeral x) ^ n \<longleftrightarrow> y = (- numeral x) ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
287  | 
using neg_numeral_power_eq_of_int_cancel_iff[of x n y] by (metis (mono_tags))  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
288  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
289  | 
lemma of_int_eq_of_int_power_cancel_iff[simp]: "(of_int b) ^ w = of_int x \<longleftrightarrow> b ^ w = x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
290  | 
by (metis of_int_power of_int_eq_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
291  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
292  | 
lemma of_int_power_eq_of_int_cancel_iff[simp]: "of_int x = (of_int b) ^ w \<longleftrightarrow> x = b ^ w"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
293  | 
by (metis of_int_eq_of_int_power_cancel_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
294  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
295  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
296  | 
|
| 36424 | 297  | 
context linordered_idom  | 
298  | 
begin  | 
|
299  | 
||
| 63652 | 300  | 
text \<open>Every \<open>linordered_idom\<close> has characteristic zero.\<close>  | 
| 36424 | 301  | 
subclass ring_char_0 ..  | 
302  | 
||
| 63652 | 303  | 
lemma of_int_le_iff [simp]: "of_int w \<le> of_int z \<longleftrightarrow> w \<le> z"  | 
304  | 
by (transfer fixing: less_eq)  | 
|
305  | 
(clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add)  | 
|
| 36424 | 306  | 
|
| 63652 | 307  | 
lemma of_int_less_iff [simp]: "of_int w < of_int z \<longleftrightarrow> w < z"  | 
| 36424 | 308  | 
by (simp add: less_le order_less_le)  | 
309  | 
||
| 63652 | 310  | 
lemma of_int_0_le_iff [simp]: "0 \<le> of_int z \<longleftrightarrow> 0 \<le> z"  | 
| 36424 | 311  | 
using of_int_le_iff [of 0 z] by simp  | 
312  | 
||
| 63652 | 313  | 
lemma of_int_le_0_iff [simp]: "of_int z \<le> 0 \<longleftrightarrow> z \<le> 0"  | 
| 36424 | 314  | 
using of_int_le_iff [of z 0] by simp  | 
315  | 
||
| 63652 | 316  | 
lemma of_int_0_less_iff [simp]: "0 < of_int z \<longleftrightarrow> 0 < z"  | 
| 36424 | 317  | 
using of_int_less_iff [of 0 z] by simp  | 
318  | 
||
| 63652 | 319  | 
lemma of_int_less_0_iff [simp]: "of_int z < 0 \<longleftrightarrow> z < 0"  | 
| 36424 | 320  | 
using of_int_less_iff [of z 0] by simp  | 
321  | 
||
| 63652 | 322  | 
lemma of_int_1_le_iff [simp]: "1 \<le> of_int z \<longleftrightarrow> 1 \<le> z"  | 
| 61234 | 323  | 
using of_int_le_iff [of 1 z] by simp  | 
324  | 
||
| 63652 | 325  | 
lemma of_int_le_1_iff [simp]: "of_int z \<le> 1 \<longleftrightarrow> z \<le> 1"  | 
| 61234 | 326  | 
using of_int_le_iff [of z 1] by simp  | 
327  | 
||
| 63652 | 328  | 
lemma of_int_1_less_iff [simp]: "1 < of_int z \<longleftrightarrow> 1 < z"  | 
| 61234 | 329  | 
using of_int_less_iff [of 1 z] by simp  | 
330  | 
||
| 63652 | 331  | 
lemma of_int_less_1_iff [simp]: "of_int z < 1 \<longleftrightarrow> z < 1"  | 
| 61234 | 332  | 
using of_int_less_iff [of z 1] by simp  | 
333  | 
||
| 
62128
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
334  | 
lemma of_int_pos: "z > 0 \<Longrightarrow> of_int z > 0"  | 
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
335  | 
by simp  | 
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
336  | 
|
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
337  | 
lemma of_int_nonneg: "z \<ge> 0 \<Longrightarrow> of_int z \<ge> 0"  | 
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
338  | 
by simp  | 
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
339  | 
|
| 63652 | 340  | 
lemma of_int_abs [simp]: "of_int \<bar>x\<bar> = \<bar>of_int x\<bar>"  | 
| 62347 | 341  | 
by (auto simp add: abs_if)  | 
342  | 
||
343  | 
lemma of_int_lessD:  | 
|
344  | 
assumes "\<bar>of_int n\<bar> < x"  | 
|
345  | 
shows "n = 0 \<or> x > 1"  | 
|
346  | 
proof (cases "n = 0")  | 
|
| 63652 | 347  | 
case True  | 
348  | 
then show ?thesis by simp  | 
|
| 62347 | 349  | 
next  | 
350  | 
case False  | 
|
351  | 
then have "\<bar>n\<bar> \<noteq> 0" by simp  | 
|
352  | 
then have "\<bar>n\<bar> > 0" by simp  | 
|
353  | 
then have "\<bar>n\<bar> \<ge> 1"  | 
|
354  | 
using zless_imp_add1_zle [of 0 "\<bar>n\<bar>"] by simp  | 
|
355  | 
then have "\<bar>of_int n\<bar> \<ge> 1"  | 
|
356  | 
unfolding of_int_1_le_iff [of "\<bar>n\<bar>", symmetric] by simp  | 
|
357  | 
then have "1 < x" using assms by (rule le_less_trans)  | 
|
358  | 
then show ?thesis ..  | 
|
359  | 
qed  | 
|
360  | 
||
361  | 
lemma of_int_leD:  | 
|
362  | 
assumes "\<bar>of_int n\<bar> \<le> x"  | 
|
363  | 
shows "n = 0 \<or> 1 \<le> x"  | 
|
364  | 
proof (cases "n = 0")  | 
|
| 63652 | 365  | 
case True  | 
366  | 
then show ?thesis by simp  | 
|
| 62347 | 367  | 
next  | 
368  | 
case False  | 
|
369  | 
then have "\<bar>n\<bar> \<noteq> 0" by simp  | 
|
370  | 
then have "\<bar>n\<bar> > 0" by simp  | 
|
371  | 
then have "\<bar>n\<bar> \<ge> 1"  | 
|
372  | 
using zless_imp_add1_zle [of 0 "\<bar>n\<bar>"] by simp  | 
|
373  | 
then have "\<bar>of_int n\<bar> \<ge> 1"  | 
|
374  | 
unfolding of_int_1_le_iff [of "\<bar>n\<bar>", symmetric] by simp  | 
|
375  | 
then have "1 \<le> x" using assms by (rule order_trans)  | 
|
376  | 
then show ?thesis ..  | 
|
377  | 
qed  | 
|
378  | 
||
| 
66912
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
379  | 
lemma numeral_power_le_of_int_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
380  | 
"numeral x ^ n \<le> of_int a \<longleftrightarrow> numeral x ^ n \<le> a"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
381  | 
by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_le_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
382  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
383  | 
lemma of_int_le_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
384  | 
"of_int a \<le> numeral x ^ n \<longleftrightarrow> a \<le> numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
385  | 
by (metis (mono_tags) local.numeral_power_eq_of_int_cancel_iff of_int_le_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
386  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
387  | 
lemma numeral_power_less_of_int_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
388  | 
"numeral x ^ n < of_int a \<longleftrightarrow> numeral x ^ n < a"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
389  | 
by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
390  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
391  | 
lemma of_int_less_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
392  | 
"of_int a < numeral x ^ n \<longleftrightarrow> a < numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
393  | 
by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
394  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
395  | 
lemma neg_numeral_power_le_of_int_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
396  | 
"(- numeral x) ^ n \<le> of_int a \<longleftrightarrow> (- numeral x) ^ n \<le> a"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
397  | 
by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
398  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
399  | 
lemma of_int_le_neg_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
400  | 
"of_int a \<le> (- numeral x) ^ n \<longleftrightarrow> a \<le> (- numeral x) ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
401  | 
by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
402  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
403  | 
lemma neg_numeral_power_less_of_int_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
404  | 
"(- numeral x) ^ n < of_int a \<longleftrightarrow> (- numeral x) ^ n < a"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
405  | 
using of_int_less_iff[of "(- numeral x) ^ n" a]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
406  | 
by simp  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
407  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
408  | 
lemma of_int_less_neg_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
409  | 
"of_int a < (- numeral x) ^ n \<longleftrightarrow> a < (- numeral x::int) ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
410  | 
using of_int_less_iff[of a "(- numeral x) ^ n"]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
411  | 
by simp  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
412  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
413  | 
lemma of_int_le_of_int_power_cancel_iff[simp]: "(of_int b) ^ w \<le> of_int x \<longleftrightarrow> b ^ w \<le> x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
414  | 
by (metis (mono_tags) of_int_le_iff of_int_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
415  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
416  | 
lemma of_int_power_le_of_int_cancel_iff[simp]: "of_int x \<le> (of_int b) ^ w\<longleftrightarrow> x \<le> b ^ w"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
417  | 
by (metis (mono_tags) of_int_le_iff of_int_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
418  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
419  | 
lemma of_int_less_of_int_power_cancel_iff[simp]: "(of_int b) ^ w < of_int x \<longleftrightarrow> b ^ w < x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
420  | 
by (metis (mono_tags) of_int_less_iff of_int_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
421  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
422  | 
lemma of_int_power_less_of_int_cancel_iff[simp]: "of_int x < (of_int b) ^ w\<longleftrightarrow> x < b ^ w"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
423  | 
by (metis (mono_tags) of_int_less_iff of_int_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
424  | 
|
| 
67969
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
425  | 
lemma of_int_max: "of_int (max x y) = max (of_int x) (of_int y)"  | 
| 
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
426  | 
by (auto simp: max_def)  | 
| 
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
427  | 
|
| 
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
428  | 
lemma of_int_min: "of_int (min x y) = min (of_int x) (of_int y)"  | 
| 
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
429  | 
by (auto simp: min_def)  | 
| 
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
430  | 
|
| 36424 | 431  | 
end  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
432  | 
|
| 61234 | 433  | 
text \<open>Comparisons involving @{term of_int}.\<close>
 | 
434  | 
||
| 63652 | 435  | 
lemma of_int_eq_numeral_iff [iff]: "of_int z = (numeral n :: 'a::ring_char_0) \<longleftrightarrow> z = numeral n"  | 
| 61234 | 436  | 
using of_int_eq_iff by fastforce  | 
437  | 
||
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
438  | 
lemma of_int_le_numeral_iff [simp]:  | 
| 63652 | 439  | 
"of_int z \<le> (numeral n :: 'a::linordered_idom) \<longleftrightarrow> z \<le> numeral n"  | 
| 61234 | 440  | 
using of_int_le_iff [of z "numeral n"] by simp  | 
441  | 
||
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
442  | 
lemma of_int_numeral_le_iff [simp]:  | 
| 63652 | 443  | 
"(numeral n :: 'a::linordered_idom) \<le> of_int z \<longleftrightarrow> numeral n \<le> z"  | 
| 61234 | 444  | 
using of_int_le_iff [of "numeral n"] by simp  | 
445  | 
||
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
446  | 
lemma of_int_less_numeral_iff [simp]:  | 
| 63652 | 447  | 
"of_int z < (numeral n :: 'a::linordered_idom) \<longleftrightarrow> z < numeral n"  | 
| 61234 | 448  | 
using of_int_less_iff [of z "numeral n"] by simp  | 
449  | 
||
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
450  | 
lemma of_int_numeral_less_iff [simp]:  | 
| 63652 | 451  | 
"(numeral n :: 'a::linordered_idom) < of_int z \<longleftrightarrow> numeral n < z"  | 
| 61234 | 452  | 
using of_int_less_iff [of "numeral n" z] by simp  | 
453  | 
||
| 63652 | 454  | 
lemma of_nat_less_of_int_iff: "(of_nat n::'a::linordered_idom) < of_int x \<longleftrightarrow> int n < x"  | 
| 
56889
 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 
hoelzl 
parents: 
56525 
diff
changeset
 | 
455  | 
by (metis of_int_of_nat_eq of_int_less_iff)  | 
| 
 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 
hoelzl 
parents: 
56525 
diff
changeset
 | 
456  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
457  | 
lemma of_int_eq_id [simp]: "of_int = id"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
458  | 
proof  | 
| 63652 | 459  | 
show "of_int z = id z" for z  | 
460  | 
by (cases z rule: int_diff_cases) simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
461  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
462  | 
|
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
463  | 
instance int :: no_top  | 
| 61169 | 464  | 
apply standard  | 
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
465  | 
apply (rule_tac x="x + 1" in exI)  | 
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
466  | 
apply simp  | 
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
467  | 
done  | 
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
468  | 
|
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
469  | 
instance int :: no_bot  | 
| 61169 | 470  | 
apply standard  | 
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
471  | 
apply (rule_tac x="x - 1" in exI)  | 
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
472  | 
apply simp  | 
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
473  | 
done  | 
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
474  | 
|
| 63652 | 475  | 
|
| 61799 | 476  | 
subsection \<open>Magnitude of an Integer, as a Natural Number: \<open>nat\<close>\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
477  | 
|
| 48045 | 478  | 
lift_definition nat :: "int \<Rightarrow> nat" is "\<lambda>(x, y). x - y"  | 
479  | 
by auto  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
480  | 
|
| 44709 | 481  | 
lemma nat_int [simp]: "nat (int n) = n"  | 
| 48045 | 482  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
483  | 
|
| 44709 | 484  | 
lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"  | 
| 48045 | 485  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
486  | 
|
| 63652 | 487  | 
lemma nat_0_le: "0 \<le> z \<Longrightarrow> int (nat z) = z"  | 
488  | 
by simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
489  | 
|
| 63652 | 490  | 
lemma nat_le_0 [simp]: "z \<le> 0 \<Longrightarrow> nat z = 0"  | 
| 48045 | 491  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
492  | 
|
| 63652 | 493  | 
lemma nat_le_eq_zle: "0 < w \<or> 0 \<le> z \<Longrightarrow> nat w \<le> nat z \<longleftrightarrow> w \<le> z"  | 
| 48045 | 494  | 
by transfer (clarsimp, arith)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
495  | 
|
| 63652 | 496  | 
text \<open>An alternative condition is @{term "0 \<le> w"}.\<close>
 | 
497  | 
lemma nat_mono_iff: "0 < z \<Longrightarrow> nat w < nat z \<longleftrightarrow> w < z"  | 
|
498  | 
by (simp add: nat_le_eq_zle linorder_not_le [symmetric])  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
499  | 
|
| 63652 | 500  | 
lemma nat_less_eq_zless: "0 \<le> w \<Longrightarrow> nat w < nat z \<longleftrightarrow> w < z"  | 
501  | 
by (simp add: nat_le_eq_zle linorder_not_le [symmetric])  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
502  | 
|
| 63652 | 503  | 
lemma zless_nat_conj [simp]: "nat w < nat z \<longleftrightarrow> 0 < z \<and> w < z"  | 
| 48045 | 504  | 
by transfer (clarsimp, arith)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
505  | 
|
| 
64714
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
506  | 
lemma nonneg_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
507  | 
assumes "0 \<le> k"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
508  | 
obtains n where "k = int n"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
509  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
510  | 
from assms have "k = int (nat k)"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
511  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
512  | 
then show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
513  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
514  | 
qed  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
515  | 
|
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
516  | 
lemma pos_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
517  | 
assumes "0 < k"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
518  | 
obtains n where "k = int n" and "n > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
519  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
520  | 
from assms have "0 \<le> k"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
521  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
522  | 
then obtain n where "k = int n"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
523  | 
by (rule nonneg_int_cases)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
524  | 
moreover have "n > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
525  | 
using \<open>k = int n\<close> assms by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
526  | 
ultimately show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
527  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
528  | 
qed  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
529  | 
|
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
530  | 
lemma nonpos_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
531  | 
assumes "k \<le> 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
532  | 
obtains n where "k = - int n"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
533  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
534  | 
from assms have "- k \<ge> 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
535  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
536  | 
then obtain n where "- k = int n"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
537  | 
by (rule nonneg_int_cases)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
538  | 
then have "k = - int n"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
539  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
540  | 
then show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
541  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
542  | 
qed  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
543  | 
|
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
544  | 
lemma neg_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
545  | 
assumes "k < 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
546  | 
obtains n where "k = - int n" and "n > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
547  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
548  | 
from assms have "- k > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
549  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
550  | 
then obtain n where "- k = int n" and "- k > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
551  | 
by (blast elim: pos_int_cases)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
552  | 
then have "k = - int n" and "n > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
553  | 
by simp_all  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
554  | 
then show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
555  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
556  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
557  | 
|
| 63652 | 558  | 
lemma nat_eq_iff: "nat w = m \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"  | 
| 48045 | 559  | 
by transfer (clarsimp simp add: le_imp_diff_is_add)  | 
| 60162 | 560  | 
|
| 63652 | 561  | 
lemma nat_eq_iff2: "m = nat w \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"  | 
| 54223 | 562  | 
using nat_eq_iff [of w m] by auto  | 
563  | 
||
| 63652 | 564  | 
lemma nat_0 [simp]: "nat 0 = 0"  | 
| 54223 | 565  | 
by (simp add: nat_eq_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
566  | 
|
| 63652 | 567  | 
lemma nat_1 [simp]: "nat 1 = Suc 0"  | 
| 54223 | 568  | 
by (simp add: nat_eq_iff)  | 
569  | 
||
| 63652 | 570  | 
lemma nat_numeral [simp]: "nat (numeral k) = numeral k"  | 
| 54223 | 571  | 
by (simp add: nat_eq_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
572  | 
|
| 63652 | 573  | 
lemma nat_neg_numeral [simp]: "nat (- numeral k) = 0"  | 
| 54223 | 574  | 
by simp  | 
575  | 
||
576  | 
lemma nat_2: "nat 2 = Suc (Suc 0)"  | 
|
577  | 
by simp  | 
|
| 60162 | 578  | 
|
| 63652 | 579  | 
lemma nat_less_iff: "0 \<le> w \<Longrightarrow> nat w < m \<longleftrightarrow> w < of_nat m"  | 
| 48045 | 580  | 
by transfer (clarsimp, arith)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
581  | 
|
| 44709 | 582  | 
lemma nat_le_iff: "nat x \<le> n \<longleftrightarrow> x \<le> int n"  | 
| 48045 | 583  | 
by transfer (clarsimp simp add: le_diff_conv)  | 
| 44707 | 584  | 
|
585  | 
lemma nat_mono: "x \<le> y \<Longrightarrow> nat x \<le> nat y"  | 
|
| 48045 | 586  | 
by transfer auto  | 
| 44707 | 587  | 
|
| 63652 | 588  | 
lemma nat_0_iff[simp]: "nat i = 0 \<longleftrightarrow> i \<le> 0"  | 
589  | 
for i :: int  | 
|
| 48045 | 590  | 
by transfer clarsimp  | 
| 29700 | 591  | 
|
| 63652 | 592  | 
lemma int_eq_iff: "of_nat m = z \<longleftrightarrow> m = nat z \<and> 0 \<le> z"  | 
593  | 
by (auto simp add: nat_eq_iff2)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
594  | 
|
| 63652 | 595  | 
lemma zero_less_nat_eq [simp]: "0 < nat z \<longleftrightarrow> 0 < z"  | 
596  | 
using zless_nat_conj [of 0] by auto  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
597  | 
|
| 63652 | 598  | 
lemma nat_add_distrib: "0 \<le> z \<Longrightarrow> 0 \<le> z' \<Longrightarrow> nat (z + z') = nat z + nat z'"  | 
| 48045 | 599  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
600  | 
|
| 63652 | 601  | 
lemma nat_diff_distrib': "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> nat (x - y) = nat x - nat y"  | 
| 54223 | 602  | 
by transfer clarsimp  | 
| 60162 | 603  | 
|
| 63652 | 604  | 
lemma nat_diff_distrib: "0 \<le> z' \<Longrightarrow> z' \<le> z \<Longrightarrow> nat (z - z') = nat z - nat z'"  | 
| 54223 | 605  | 
by (rule nat_diff_distrib') auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
606  | 
|
| 44709 | 607  | 
lemma nat_zminus_int [simp]: "nat (- int n) = 0"  | 
| 48045 | 608  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
609  | 
|
| 63652 | 610  | 
lemma le_nat_iff: "k \<ge> 0 \<Longrightarrow> n \<le> nat k \<longleftrightarrow> int n \<le> k"  | 
| 53065 | 611  | 
by transfer auto  | 
| 60162 | 612  | 
|
| 63652 | 613  | 
lemma zless_nat_eq_int_zless: "m < nat z \<longleftrightarrow> int m < z"  | 
| 48045 | 614  | 
by transfer (clarsimp simp add: less_diff_conv)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
615  | 
|
| 63652 | 616  | 
lemma (in ring_1) of_nat_nat [simp]: "0 \<le> z \<Longrightarrow> of_nat (nat z) = of_int z"  | 
| 
48066
 
c6783c9b87bf
transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
 
huffman 
parents: 
48045 
diff
changeset
 | 
617  | 
by transfer (clarsimp simp add: of_nat_diff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
618  | 
|
| 63652 | 619  | 
lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"  | 
| 54249 | 620  | 
by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)  | 
621  | 
||
| 66886 | 622  | 
lemma nat_abs_triangle_ineq:  | 
623  | 
"nat \<bar>k + l\<bar> \<le> nat \<bar>k\<bar> + nat \<bar>l\<bar>"  | 
|
624  | 
by (simp add: nat_add_distrib [symmetric] nat_le_eq_zle abs_triangle_ineq)  | 
|
625  | 
||
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
626  | 
lemma nat_of_bool [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
627  | 
"nat (of_bool P) = of_bool P"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
628  | 
by auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
629  | 
|
| 66836 | 630  | 
lemma split_nat [arith_split]: "P (nat i) \<longleftrightarrow> ((\<forall>n. i = int n \<longrightarrow> P n) \<and> (i < 0 \<longrightarrow> P 0))"  | 
631  | 
(is "?P = (?L \<and> ?R)")  | 
|
632  | 
for i :: int  | 
|
633  | 
proof (cases "i < 0")  | 
|
634  | 
case True  | 
|
635  | 
then show ?thesis  | 
|
636  | 
by auto  | 
|
637  | 
next  | 
|
638  | 
case False  | 
|
639  | 
have "?P = ?L"  | 
|
640  | 
proof  | 
|
641  | 
assume ?P  | 
|
642  | 
then show ?L using False by auto  | 
|
643  | 
next  | 
|
644  | 
assume ?L  | 
|
645  | 
moreover from False have "int (nat i) = i"  | 
|
646  | 
by (simp add: not_less)  | 
|
647  | 
ultimately show ?P  | 
|
648  | 
by simp  | 
|
649  | 
qed  | 
|
650  | 
with False show ?thesis by simp  | 
|
651  | 
qed  | 
|
652  | 
||
653  | 
lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))"  | 
|
654  | 
by (auto split: split_nat)  | 
|
655  | 
||
656  | 
lemma ex_nat: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> P (nat x))"  | 
|
657  | 
proof  | 
|
658  | 
assume "\<exists>x. P x"  | 
|
659  | 
then obtain x where "P x" ..  | 
|
660  | 
then have "int x \<ge> 0 \<and> P (nat (int x))" by simp  | 
|
661  | 
then show "\<exists>x\<ge>0. P (nat x)" ..  | 
|
662  | 
next  | 
|
663  | 
assume "\<exists>x\<ge>0. P (nat x)"  | 
|
664  | 
then show "\<exists>x. P x" by auto  | 
|
665  | 
qed  | 
|
666  | 
||
| 54249 | 667  | 
|
| 60758 | 668  | 
text \<open>For termination proofs:\<close>  | 
| 63652 | 669  | 
lemma measure_function_int[measure_function]: "is_measure (nat \<circ> abs)" ..  | 
| 29779 | 670  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
671  | 
|
| 63652 | 672  | 
subsection \<open>Lemmas about the Function @{term of_nat} and Orderings\<close>
 | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
673  | 
|
| 61076 | 674  | 
lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)"  | 
| 63652 | 675  | 
by (simp add: order_less_le del: of_nat_Suc)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
676  | 
|
| 44709 | 677  | 
lemma negative_zless [iff]: "- (int (Suc n)) < int m"  | 
| 63652 | 678  | 
by (rule negative_zless_0 [THEN order_less_le_trans], simp)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
679  | 
|
| 44709 | 680  | 
lemma negative_zle_0: "- int n \<le> 0"  | 
| 63652 | 681  | 
by (simp add: minus_le_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
682  | 
|
| 44709 | 683  | 
lemma negative_zle [iff]: "- int n \<le> int m"  | 
| 63652 | 684  | 
by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
685  | 
|
| 63652 | 686  | 
lemma not_zle_0_negative [simp]: "\<not> 0 \<le> - int (Suc n)"  | 
687  | 
by (subst le_minus_iff) (simp del: of_nat_Suc)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
688  | 
|
| 63652 | 689  | 
lemma int_zle_neg: "int n \<le> - int m \<longleftrightarrow> n = 0 \<and> m = 0"  | 
| 48045 | 690  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
691  | 
|
| 63652 | 692  | 
lemma not_int_zless_negative [simp]: "\<not> int n < - int m"  | 
693  | 
by (simp add: linorder_not_less)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
694  | 
|
| 63652 | 695  | 
lemma negative_eq_positive [simp]: "- int n = of_nat m \<longleftrightarrow> n = 0 \<and> m = 0"  | 
696  | 
by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
697  | 
|
| 63652 | 698  | 
lemma zle_iff_zadd: "w \<le> z \<longleftrightarrow> (\<exists>n. z = w + int n)"  | 
699  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 62348 | 700  | 
proof  | 
| 63652 | 701  | 
assume ?rhs  | 
702  | 
then show ?lhs by auto  | 
|
| 62348 | 703  | 
next  | 
| 63652 | 704  | 
assume ?lhs  | 
| 62348 | 705  | 
then have "0 \<le> z - w" by simp  | 
706  | 
then obtain n where "z - w = int n"  | 
|
707  | 
using zero_le_imp_eq_int [of "z - w"] by blast  | 
|
| 63652 | 708  | 
then have "z = w + int n" by simp  | 
709  | 
then show ?rhs ..  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
710  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
711  | 
|
| 44709 | 712  | 
lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z"  | 
| 63652 | 713  | 
by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
714  | 
|
| 63652 | 715  | 
text \<open>  | 
716  | 
This version is proved for all ordered rings, not just integers!  | 
|
717  | 
It is proved here because attribute \<open>arith_split\<close> is not available  | 
|
718  | 
in theory \<open>Rings\<close>.  | 
|
719  | 
But is it really better than just rewriting with \<open>abs_if\<close>?  | 
|
720  | 
\<close>  | 
|
721  | 
lemma abs_split [arith_split, no_atp]: "P \<bar>a\<bar> \<longleftrightarrow> (0 \<le> a \<longrightarrow> P a) \<and> (a < 0 \<longrightarrow> P (- a))"  | 
|
722  | 
for a :: "'a::linordered_idom"  | 
|
723  | 
by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
724  | 
|
| 44709 | 725  | 
lemma negD: "x < 0 \<Longrightarrow> \<exists>n. x = - (int (Suc n))"  | 
| 63652 | 726  | 
apply transfer  | 
727  | 
apply clarsimp  | 
|
728  | 
apply (rule_tac x="b - Suc a" in exI)  | 
|
729  | 
apply arith  | 
|
730  | 
done  | 
|
731  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
732  | 
|
| 60758 | 733  | 
subsection \<open>Cases and induction\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
734  | 
|
| 63652 | 735  | 
text \<open>  | 
736  | 
Now we replace the case analysis rule by a more conventional one:  | 
|
737  | 
whether an integer is negative or not.  | 
|
738  | 
\<close>  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
739  | 
|
| 63652 | 740  | 
text \<open>This version is symmetric in the two subgoals.\<close>  | 
741  | 
lemma int_cases2 [case_names nonneg nonpos, cases type: int]:  | 
|
742  | 
"(\<And>n. z = int n \<Longrightarrow> P) \<Longrightarrow> (\<And>n. z = - (int n) \<Longrightarrow> P) \<Longrightarrow> P"  | 
|
743  | 
by (cases "z < 0") (auto simp add: linorder_not_less dest!: negD nat_0_le [THEN sym])  | 
|
| 
59613
 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
59582 
diff
changeset
 | 
744  | 
|
| 63652 | 745  | 
text \<open>This is the default, with a negative case.\<close>  | 
746  | 
lemma int_cases [case_names nonneg neg, cases type: int]:  | 
|
747  | 
"(\<And>n. z = int n \<Longrightarrow> P) \<Longrightarrow> (\<And>n. z = - (int (Suc n)) \<Longrightarrow> P) \<Longrightarrow> P"  | 
|
748  | 
apply (cases "z < 0")  | 
|
749  | 
apply (blast dest!: negD)  | 
|
750  | 
apply (simp add: linorder_not_less del: of_nat_Suc)  | 
|
751  | 
apply auto  | 
|
752  | 
apply (blast dest: nat_0_le [THEN sym])  | 
|
753  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
754  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
755  | 
lemma int_cases3 [case_names zero pos neg]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
756  | 
fixes k :: int  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
757  | 
assumes "k = 0 \<Longrightarrow> P" and "\<And>n. k = int n \<Longrightarrow> n > 0 \<Longrightarrow> P"  | 
| 61204 | 758  | 
and "\<And>n. k = - int n \<Longrightarrow> n > 0 \<Longrightarrow> P"  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
759  | 
shows "P"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
760  | 
proof (cases k "0::int" rule: linorder_cases)  | 
| 63652 | 761  | 
case equal  | 
762  | 
with assms(1) show P by simp  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
763  | 
next  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
764  | 
case greater  | 
| 63539 | 765  | 
then have *: "nat k > 0" by simp  | 
766  | 
moreover from * have "k = int (nat k)" by auto  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
767  | 
ultimately show P using assms(2) by blast  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
768  | 
next  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
769  | 
case less  | 
| 63539 | 770  | 
then have *: "nat (- k) > 0" by simp  | 
771  | 
moreover from * have "k = - int (nat (- k))" by auto  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
772  | 
ultimately show P using assms(3) by blast  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
773  | 
qed  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
774  | 
|
| 63652 | 775  | 
lemma int_of_nat_induct [case_names nonneg neg, induct type: int]:  | 
776  | 
"(\<And>n. P (int n)) \<Longrightarrow> (\<And>n. P (- (int (Suc n)))) \<Longrightarrow> P z"  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
777  | 
by (cases z) auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
778  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
779  | 
lemma sgn_mult_dvd_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
780  | 
"sgn r * l dvd k \<longleftrightarrow> l dvd k \<and> (r = 0 \<longrightarrow> k = 0)" for k l r :: int  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
781  | 
by (cases r rule: int_cases3) auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
782  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
783  | 
lemma mult_sgn_dvd_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
784  | 
"l * sgn r dvd k \<longleftrightarrow> l dvd k \<and> (r = 0 \<longrightarrow> k = 0)" for k l r :: int  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
785  | 
using sgn_mult_dvd_iff [of r l k] by (simp add: ac_simps)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
786  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
787  | 
lemma dvd_sgn_mult_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
788  | 
"l dvd sgn r * k \<longleftrightarrow> l dvd k \<or> r = 0" for k l r :: int  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
789  | 
by (cases r rule: int_cases3) simp_all  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
790  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
791  | 
lemma dvd_mult_sgn_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
792  | 
"l dvd k * sgn r \<longleftrightarrow> l dvd k \<or> r = 0" for k l r :: int  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
793  | 
using dvd_sgn_mult_iff [of l r k] by (simp add: ac_simps)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
794  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
795  | 
lemma int_sgnE:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
796  | 
fixes k :: int  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
797  | 
obtains n and l where "k = sgn l * int n"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
798  | 
proof -  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
799  | 
have "k = sgn k * int (nat \<bar>k\<bar>)"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
800  | 
by (simp add: sgn_mult_abs)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
801  | 
then show ?thesis ..  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
802  | 
qed  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
803  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
804  | 
|
| 60758 | 805  | 
subsubsection \<open>Binary comparisons\<close>  | 
| 28958 | 806  | 
|
| 60758 | 807  | 
text \<open>Preliminaries\<close>  | 
| 28958 | 808  | 
|
| 60162 | 809  | 
lemma le_imp_0_less:  | 
| 63652 | 810  | 
fixes z :: int  | 
| 28958 | 811  | 
assumes le: "0 \<le> z"  | 
| 63652 | 812  | 
shows "0 < 1 + z"  | 
| 28958 | 813  | 
proof -  | 
814  | 
have "0 \<le> z" by fact  | 
|
| 63652 | 815  | 
also have "\<dots> < z + 1" by (rule less_add_one)  | 
816  | 
also have "\<dots> = 1 + z" by (simp add: ac_simps)  | 
|
| 28958 | 817  | 
finally show "0 < 1 + z" .  | 
818  | 
qed  | 
|
819  | 
||
| 63652 | 820  | 
lemma odd_less_0_iff: "1 + z + z < 0 \<longleftrightarrow> z < 0"  | 
821  | 
for z :: int  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
822  | 
proof (cases z)  | 
| 28958 | 823  | 
case (nonneg n)  | 
| 63652 | 824  | 
then show ?thesis  | 
825  | 
by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le])  | 
|
| 28958 | 826  | 
next  | 
827  | 
case (neg n)  | 
|
| 63652 | 828  | 
then show ?thesis  | 
829  | 
by (simp del: of_nat_Suc of_nat_add of_nat_1  | 
|
830  | 
add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric])  | 
|
| 28958 | 831  | 
qed  | 
832  | 
||
| 63652 | 833  | 
|
| 60758 | 834  | 
subsubsection \<open>Comparisons, for Ordered Rings\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
835  | 
|
| 63652 | 836  | 
lemma odd_nonzero: "1 + z + z \<noteq> 0"  | 
837  | 
for z :: int  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
838  | 
proof (cases z)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
839  | 
case (nonneg n)  | 
| 63652 | 840  | 
have le: "0 \<le> z + z"  | 
841  | 
by (simp add: nonneg add_increasing)  | 
|
842  | 
then show ?thesis  | 
|
| 67116 | 843  | 
using le_imp_0_less [OF le] by (auto simp: ac_simps)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
844  | 
next  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
845  | 
case (neg n)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
846  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
847  | 
proof  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
848  | 
assume eq: "1 + z + z = 0"  | 
| 63652 | 849  | 
have "0 < 1 + (int n + int n)"  | 
| 60162 | 850  | 
by (simp add: le_imp_0_less add_increasing)  | 
| 63652 | 851  | 
also have "\<dots> = - (1 + z + z)"  | 
| 60162 | 852  | 
by (simp add: neg add.assoc [symmetric])  | 
| 63652 | 853  | 
also have "\<dots> = 0" by (simp add: eq)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
854  | 
finally have "0<0" ..  | 
| 63652 | 855  | 
then show False by blast  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
856  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
857  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
858  | 
|
| 
30652
 
752329615264
distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
 
haftmann 
parents: 
30496 
diff
changeset
 | 
859  | 
|
| 60758 | 860  | 
subsection \<open>The Set of Integers\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
861  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
862  | 
context ring_1  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
863  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
864  | 
|
| 61070 | 865  | 
definition Ints :: "'a set"  ("\<int>")
 | 
866  | 
where "\<int> = range of_int"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
867  | 
|
| 35634 | 868  | 
lemma Ints_of_int [simp]: "of_int z \<in> \<int>"  | 
869  | 
by (simp add: Ints_def)  | 
|
870  | 
||
871  | 
lemma Ints_of_nat [simp]: "of_nat n \<in> \<int>"  | 
|
| 45533 | 872  | 
using Ints_of_int [of "of_nat n"] by simp  | 
| 35634 | 873  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
874  | 
lemma Ints_0 [simp]: "0 \<in> \<int>"  | 
| 45533 | 875  | 
using Ints_of_int [of "0"] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
876  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
877  | 
lemma Ints_1 [simp]: "1 \<in> \<int>"  | 
| 45533 | 878  | 
using Ints_of_int [of "1"] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
879  | 
|
| 
61552
 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 
eberlm 
parents: 
61531 
diff
changeset
 | 
880  | 
lemma Ints_numeral [simp]: "numeral n \<in> \<int>"  | 
| 
 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 
eberlm 
parents: 
61531 
diff
changeset
 | 
881  | 
by (subst of_nat_numeral [symmetric], rule Ints_of_nat)  | 
| 
 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 
eberlm 
parents: 
61531 
diff
changeset
 | 
882  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
883  | 
lemma Ints_add [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a + b \<in> \<int>"  | 
| 63652 | 884  | 
apply (auto simp add: Ints_def)  | 
885  | 
apply (rule range_eqI)  | 
|
886  | 
apply (rule of_int_add [symmetric])  | 
|
887  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
888  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
889  | 
lemma Ints_minus [simp]: "a \<in> \<int> \<Longrightarrow> -a \<in> \<int>"  | 
| 63652 | 890  | 
apply (auto simp add: Ints_def)  | 
891  | 
apply (rule range_eqI)  | 
|
892  | 
apply (rule of_int_minus [symmetric])  | 
|
893  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
894  | 
|
| 68721 | 895  | 
lemma minus_in_Ints_iff: "-x \<in> \<int> \<longleftrightarrow> x \<in> \<int>"  | 
896  | 
using Ints_minus[of x] Ints_minus[of "-x"] by auto  | 
|
897  | 
||
| 35634 | 898  | 
lemma Ints_diff [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a - b \<in> \<int>"  | 
| 63652 | 899  | 
apply (auto simp add: Ints_def)  | 
900  | 
apply (rule range_eqI)  | 
|
901  | 
apply (rule of_int_diff [symmetric])  | 
|
902  | 
done  | 
|
| 35634 | 903  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
904  | 
lemma Ints_mult [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a * b \<in> \<int>"  | 
| 63652 | 905  | 
apply (auto simp add: Ints_def)  | 
906  | 
apply (rule range_eqI)  | 
|
907  | 
apply (rule of_int_mult [symmetric])  | 
|
908  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
909  | 
|
| 35634 | 910  | 
lemma Ints_power [simp]: "a \<in> \<int> \<Longrightarrow> a ^ n \<in> \<int>"  | 
| 63652 | 911  | 
by (induct n) simp_all  | 
| 35634 | 912  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
913  | 
lemma Ints_cases [cases set: Ints]:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
914  | 
assumes "q \<in> \<int>"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
915  | 
obtains (of_int) z where "q = of_int z"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
916  | 
unfolding Ints_def  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
917  | 
proof -  | 
| 60758 | 918  | 
from \<open>q \<in> \<int>\<close> have "q \<in> range of_int" unfolding Ints_def .  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
919  | 
then obtain z where "q = of_int z" ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
920  | 
then show thesis ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
921  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
922  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
923  | 
lemma Ints_induct [case_names of_int, induct set: Ints]:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
924  | 
"q \<in> \<int> \<Longrightarrow> (\<And>z. P (of_int z)) \<Longrightarrow> P q"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
925  | 
by (rule Ints_cases) auto  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
926  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
927  | 
lemma Nats_subset_Ints: "\<nat> \<subseteq> \<int>"  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
928  | 
unfolding Nats_def Ints_def  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
929  | 
by (rule subsetI, elim imageE, hypsubst, subst of_int_of_nat_eq[symmetric], rule imageI) simp_all  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
930  | 
|
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
931  | 
lemma Nats_altdef1: "\<nat> = {of_int n |n. n \<ge> 0}"
 | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
932  | 
proof (intro subsetI equalityI)  | 
| 63652 | 933  | 
fix x :: 'a  | 
934  | 
  assume "x \<in> {of_int n |n. n \<ge> 0}"
 | 
|
935  | 
then obtain n where "x = of_int n" "n \<ge> 0"  | 
|
936  | 
by (auto elim!: Ints_cases)  | 
|
937  | 
then have "x = of_nat (nat n)"  | 
|
938  | 
by (subst of_nat_nat) simp_all  | 
|
939  | 
then show "x \<in> \<nat>"  | 
|
940  | 
by simp  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
941  | 
next  | 
| 63652 | 942  | 
fix x :: 'a  | 
943  | 
assume "x \<in> \<nat>"  | 
|
944  | 
then obtain n where "x = of_nat n"  | 
|
945  | 
by (auto elim!: Nats_cases)  | 
|
946  | 
then have "x = of_int (int n)" by simp  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
947  | 
also have "int n \<ge> 0" by simp  | 
| 63652 | 948  | 
  then have "of_int (int n) \<in> {of_int n |n. n \<ge> 0}" by blast
 | 
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
949  | 
  finally show "x \<in> {of_int n |n. n \<ge> 0}" .
 | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
950  | 
qed  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
951  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
952  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
953  | 
|
| 
64758
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
954  | 
lemma (in linordered_idom) Ints_abs [simp]:  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
955  | 
shows "a \<in> \<int> \<Longrightarrow> abs a \<in> \<int>"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
956  | 
by (auto simp: abs_if)  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
957  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
958  | 
lemma (in linordered_idom) Nats_altdef2: "\<nat> = {n \<in> \<int>. n \<ge> 0}"
 | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
959  | 
proof (intro subsetI equalityI)  | 
| 63652 | 960  | 
fix x :: 'a  | 
961  | 
  assume "x \<in> {n \<in> \<int>. n \<ge> 0}"
 | 
|
962  | 
then obtain n where "x = of_int n" "n \<ge> 0"  | 
|
963  | 
by (auto elim!: Ints_cases)  | 
|
964  | 
then have "x = of_nat (nat n)"  | 
|
965  | 
by (subst of_nat_nat) simp_all  | 
|
966  | 
then show "x \<in> \<nat>"  | 
|
967  | 
by simp  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
968  | 
qed (auto elim!: Nats_cases)  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
969  | 
|
| 64849 | 970  | 
lemma (in idom_divide) of_int_divide_in_Ints:  | 
971  | 
"of_int a div of_int b \<in> \<int>" if "b dvd a"  | 
|
972  | 
proof -  | 
|
973  | 
from that obtain c where "a = b * c" ..  | 
|
974  | 
then show ?thesis  | 
|
975  | 
by (cases "of_int b = 0") simp_all  | 
|
976  | 
qed  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
977  | 
|
| 60758 | 978  | 
text \<open>The premise involving @{term Ints} prevents @{term "a = 1/2"}.\<close>
 | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
979  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
980  | 
lemma Ints_double_eq_0_iff:  | 
| 63652 | 981  | 
fixes a :: "'a::ring_char_0"  | 
| 61070 | 982  | 
assumes in_Ints: "a \<in> \<int>"  | 
| 63652 | 983  | 
shows "a + a = 0 \<longleftrightarrow> a = 0"  | 
984  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
985  | 
proof -  | 
| 63652 | 986  | 
from in_Ints have "a \<in> range of_int"  | 
987  | 
unfolding Ints_def [symmetric] .  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
988  | 
then obtain z where a: "a = of_int z" ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
989  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
990  | 
proof  | 
| 63652 | 991  | 
assume ?rhs  | 
992  | 
then show ?lhs by simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
993  | 
next  | 
| 63652 | 994  | 
assume ?lhs  | 
995  | 
with a have "of_int (z + z) = (of_int 0 :: 'a)" by simp  | 
|
996  | 
then have "z + z = 0" by (simp only: of_int_eq_iff)  | 
|
| 67116 | 997  | 
then have "z = 0" by (simp only: double_zero)  | 
| 63652 | 998  | 
with a show ?rhs by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
999  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1000  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1001  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1002  | 
lemma Ints_odd_nonzero:  | 
| 63652 | 1003  | 
fixes a :: "'a::ring_char_0"  | 
| 61070 | 1004  | 
assumes in_Ints: "a \<in> \<int>"  | 
| 63652 | 1005  | 
shows "1 + a + a \<noteq> 0"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1006  | 
proof -  | 
| 63652 | 1007  | 
from in_Ints have "a \<in> range of_int"  | 
1008  | 
unfolding Ints_def [symmetric] .  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1009  | 
then obtain z where a: "a = of_int z" ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1010  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1011  | 
proof  | 
| 63652 | 1012  | 
assume "1 + a + a = 0"  | 
1013  | 
with a have "of_int (1 + z + z) = (of_int 0 :: 'a)" by simp  | 
|
1014  | 
then have "1 + z + z = 0" by (simp only: of_int_eq_iff)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1015  | 
with odd_nonzero show False by blast  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1016  | 
qed  | 
| 60162 | 1017  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1018  | 
|
| 61070 | 1019  | 
lemma Nats_numeral [simp]: "numeral w \<in> \<nat>"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1020  | 
using of_nat_in_Nats [of "numeral w"] by simp  | 
| 35634 | 1021  | 
|
| 60162 | 1022  | 
lemma Ints_odd_less_0:  | 
| 63652 | 1023  | 
fixes a :: "'a::linordered_idom"  | 
| 61070 | 1024  | 
assumes in_Ints: "a \<in> \<int>"  | 
| 63652 | 1025  | 
shows "1 + a + a < 0 \<longleftrightarrow> a < 0"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1026  | 
proof -  | 
| 63652 | 1027  | 
from in_Ints have "a \<in> range of_int"  | 
1028  | 
unfolding Ints_def [symmetric] .  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1029  | 
then obtain z where a: "a = of_int z" ..  | 
| 63652 | 1030  | 
with a have "1 + a + a < 0 \<longleftrightarrow> of_int (1 + z + z) < (of_int 0 :: 'a)"  | 
1031  | 
by simp  | 
|
1032  | 
also have "\<dots> \<longleftrightarrow> z < 0"  | 
|
1033  | 
by (simp only: of_int_less_iff odd_less_0_iff)  | 
|
1034  | 
also have "\<dots> \<longleftrightarrow> a < 0"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1035  | 
by (simp add: a)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1036  | 
finally show ?thesis .  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1037  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1038  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1039  | 
|
| 64272 | 1040  | 
subsection \<open>@{term sum} and @{term prod}\<close>
 | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1041  | 
|
| 69182 | 1042  | 
context semiring_1  | 
1043  | 
begin  | 
|
1044  | 
||
1045  | 
lemma of_nat_sum [simp]:  | 
|
1046  | 
"of_nat (sum f A) = (\<Sum>x\<in>A. of_nat (f x))"  | 
|
1047  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1048  | 
||
1049  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1050  | 
|
| 69182 | 1051  | 
context ring_1  | 
1052  | 
begin  | 
|
1053  | 
||
1054  | 
lemma of_int_sum [simp]:  | 
|
1055  | 
"of_int (sum f A) = (\<Sum>x\<in>A. of_int (f x))"  | 
|
1056  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1057  | 
||
1058  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1059  | 
|
| 69182 | 1060  | 
context comm_semiring_1  | 
1061  | 
begin  | 
|
1062  | 
||
1063  | 
lemma of_nat_prod [simp]:  | 
|
1064  | 
"of_nat (prod f A) = (\<Prod>x\<in>A. of_nat (f x))"  | 
|
1065  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1066  | 
||
1067  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1068  | 
|
| 69182 | 1069  | 
context comm_ring_1  | 
1070  | 
begin  | 
|
1071  | 
||
1072  | 
lemma of_int_prod [simp]:  | 
|
1073  | 
"of_int (prod f A) = (\<Prod>x\<in>A. of_int (f x))"  | 
|
1074  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1075  | 
||
1076  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1077  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1078  | 
|
| 60758 | 1079  | 
subsection \<open>Setting up simplification procedures\<close>  | 
| 30802 | 1080  | 
|
| 54249 | 1081  | 
lemmas of_int_simps =  | 
1082  | 
of_int_0 of_int_1 of_int_add of_int_mult  | 
|
1083  | 
||
| 48891 | 1084  | 
ML_file "Tools/int_arith.ML"  | 
| 60758 | 1085  | 
declaration \<open>K Int_Arith.setup\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1086  | 
|
| 63652 | 1087  | 
simproc_setup fast_arith  | 
1088  | 
  ("(m::'a::linordered_idom) < n" |
 | 
|
1089  | 
"(m::'a::linordered_idom) \<le> n" |  | 
|
1090  | 
"(m::'a::linordered_idom) = n") =  | 
|
| 61144 | 1091  | 
\<open>K Lin_Arith.simproc\<close>  | 
| 43595 | 1092  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1093  | 
|
| 60758 | 1094  | 
subsection\<open>More Inequality Reasoning\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1095  | 
|
| 63652 | 1096  | 
lemma zless_add1_eq: "w < z + 1 \<longleftrightarrow> w < z \<or> w = z"  | 
1097  | 
for w z :: int  | 
|
1098  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1099  | 
|
| 63652 | 1100  | 
lemma add1_zle_eq: "w + 1 \<le> z \<longleftrightarrow> w < z"  | 
1101  | 
for w z :: int  | 
|
1102  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1103  | 
|
| 63652 | 1104  | 
lemma zle_diff1_eq [simp]: "w \<le> z - 1 \<longleftrightarrow> w < z"  | 
1105  | 
for w z :: int  | 
|
1106  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1107  | 
|
| 63652 | 1108  | 
lemma zle_add1_eq_le [simp]: "w < z + 1 \<longleftrightarrow> w \<le> z"  | 
1109  | 
for w z :: int  | 
|
1110  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1111  | 
|
| 63652 | 1112  | 
lemma int_one_le_iff_zero_less: "1 \<le> z \<longleftrightarrow> 0 < z"  | 
1113  | 
for z :: int  | 
|
1114  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1115  | 
|
| 
64758
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1116  | 
lemma Ints_nonzero_abs_ge1:  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1117  | 
fixes x:: "'a :: linordered_idom"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1118  | 
assumes "x \<in> Ints" "x \<noteq> 0"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1119  | 
shows "1 \<le> abs x"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1120  | 
proof (rule Ints_cases [OF \<open>x \<in> Ints\<close>])  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1121  | 
fix z::int  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1122  | 
assume "x = of_int z"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1123  | 
with \<open>x \<noteq> 0\<close>  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1124  | 
show "1 \<le> \<bar>x\<bar>"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1125  | 
apply (auto simp add: abs_if)  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1126  | 
by (metis diff_0 of_int_1 of_int_le_iff of_int_minus zle_diff1_eq)  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1127  | 
qed  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1128  | 
|
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1129  | 
lemma Ints_nonzero_abs_less1:  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1130  | 
fixes x:: "'a :: linordered_idom"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1131  | 
shows "\<lbrakk>x \<in> Ints; abs x < 1\<rbrakk> \<Longrightarrow> x = 0"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1132  | 
using Ints_nonzero_abs_ge1 [of x] by auto  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1133  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1134  | 
|
| 63652 | 1135  | 
subsection \<open>The functions @{term nat} and @{term int}\<close>
 | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1136  | 
|
| 63652 | 1137  | 
text \<open>Simplify the term @{term "w + - z"}.\<close>
 | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1138  | 
|
| 63652 | 1139  | 
lemma one_less_nat_eq [simp]: "Suc 0 < nat z \<longleftrightarrow> 1 < z"  | 
| 60162 | 1140  | 
using zless_nat_conj [of 1 z] by auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1141  | 
|
| 67116 | 1142  | 
lemma int_eq_iff_numeral [simp]:  | 
1143  | 
"int m = numeral v \<longleftrightarrow> m = numeral v"  | 
|
1144  | 
by (simp add: int_eq_iff)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1145  | 
|
| 67116 | 1146  | 
lemma nat_abs_int_diff:  | 
1147  | 
"nat \<bar>int a - int b\<bar> = (if a \<le> b then b - a else a - b)"  | 
|
| 59000 | 1148  | 
by auto  | 
1149  | 
||
1150  | 
lemma nat_int_add: "nat (int a + int b) = a + b"  | 
|
1151  | 
by auto  | 
|
1152  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1153  | 
context ring_1  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1154  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1155  | 
|
| 
33056
 
791a4655cae3
renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
 
blanchet 
parents: 
32437 
diff
changeset
 | 
1156  | 
lemma of_int_of_nat [nitpick_simp]:  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1157  | 
"of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1158  | 
proof (cases "k < 0")  | 
| 63652 | 1159  | 
case True  | 
1160  | 
then have "0 \<le> - k" by simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1161  | 
then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1162  | 
with True show ?thesis by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1163  | 
next  | 
| 63652 | 1164  | 
case False  | 
1165  | 
then show ?thesis by (simp add: not_less)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1166  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1167  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1168  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1169  | 
|
| 64014 | 1170  | 
lemma transfer_rule_of_int:  | 
1171  | 
fixes R :: "'a::ring_1 \<Rightarrow> 'b::ring_1 \<Rightarrow> bool"  | 
|
1172  | 
assumes [transfer_rule]: "R 0 0" "R 1 1"  | 
|
1173  | 
"rel_fun R (rel_fun R R) plus plus"  | 
|
1174  | 
"rel_fun R R uminus uminus"  | 
|
1175  | 
shows "rel_fun HOL.eq R of_int of_int"  | 
|
1176  | 
proof -  | 
|
1177  | 
note transfer_rule_of_nat [transfer_rule]  | 
|
1178  | 
have [transfer_rule]: "rel_fun HOL.eq R of_nat of_nat"  | 
|
1179  | 
by transfer_prover  | 
|
1180  | 
show ?thesis  | 
|
1181  | 
by (unfold of_int_of_nat [abs_def]) transfer_prover  | 
|
1182  | 
qed  | 
|
1183  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1184  | 
lemma nat_mult_distrib:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1185  | 
fixes z z' :: int  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1186  | 
assumes "0 \<le> z"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1187  | 
shows "nat (z * z') = nat z * nat z'"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1188  | 
proof (cases "0 \<le> z'")  | 
| 63652 | 1189  | 
case False  | 
1190  | 
with assms have "z * z' \<le> 0"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1191  | 
by (simp add: not_le mult_le_0_iff)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1192  | 
then have "nat (z * z') = 0" by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1193  | 
moreover from False have "nat z' = 0" by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1194  | 
ultimately show ?thesis by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1195  | 
next  | 
| 63652 | 1196  | 
case True  | 
1197  | 
with assms have ge_0: "z * z' \<ge> 0" by (simp add: zero_le_mult_iff)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1198  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1199  | 
by (rule injD [of "of_nat :: nat \<Rightarrow> int", OF inj_of_nat])  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1200  | 
(simp only: of_nat_mult of_nat_nat [OF True]  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1201  | 
of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1202  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1203  | 
|
| 63652 | 1204  | 
lemma nat_mult_distrib_neg: "z \<le> 0 \<Longrightarrow> nat (z * z') = nat (- z) * nat (- z')"  | 
1205  | 
for z z' :: int  | 
|
1206  | 
apply (rule trans)  | 
|
1207  | 
apply (rule_tac [2] nat_mult_distrib)  | 
|
1208  | 
apply auto  | 
|
1209  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1210  | 
|
| 61944 | 1211  | 
lemma nat_abs_mult_distrib: "nat \<bar>w * z\<bar> = nat \<bar>w\<bar> * nat \<bar>z\<bar>"  | 
| 63652 | 1212  | 
by (cases "z = 0 \<or> w = 0")  | 
1213  | 
(auto simp add: abs_if nat_mult_distrib [symmetric]  | 
|
1214  | 
nat_mult_distrib_neg [symmetric] mult_less_0_iff)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1215  | 
|
| 63652 | 1216  | 
lemma int_in_range_abs [simp]: "int n \<in> range abs"  | 
| 60570 | 1217  | 
proof (rule range_eqI)  | 
| 63652 | 1218  | 
show "int n = \<bar>int n\<bar>" by simp  | 
| 60570 | 1219  | 
qed  | 
1220  | 
||
| 63652 | 1221  | 
lemma range_abs_Nats [simp]: "range abs = (\<nat> :: int set)"  | 
| 60570 | 1222  | 
proof -  | 
1223  | 
have "\<bar>k\<bar> \<in> \<nat>" for k :: int  | 
|
1224  | 
by (cases k) simp_all  | 
|
1225  | 
moreover have "k \<in> range abs" if "k \<in> \<nat>" for k :: int  | 
|
1226  | 
using that by induct simp  | 
|
1227  | 
ultimately show ?thesis by blast  | 
|
| 61204 | 1228  | 
qed  | 
| 60570 | 1229  | 
|
| 63652 | 1230  | 
lemma Suc_nat_eq_nat_zadd1: "0 \<le> z \<Longrightarrow> Suc (nat z) = nat (1 + z)"  | 
1231  | 
for z :: int  | 
|
1232  | 
by (rule sym) (simp add: nat_eq_iff)  | 
|
| 
47207
 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 
huffman 
parents: 
47192 
diff
changeset
 | 
1233  | 
|
| 
 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 
huffman 
parents: 
47192 
diff
changeset
 | 
1234  | 
lemma diff_nat_eq_if:  | 
| 63652 | 1235  | 
"nat z - nat z' =  | 
1236  | 
(if z' < 0 then nat z  | 
|
1237  | 
else  | 
|
1238  | 
let d = z - z'  | 
|
1239  | 
in if d < 0 then 0 else nat d)"  | 
|
1240  | 
by (simp add: Let_def nat_diff_distrib [symmetric])  | 
|
| 
47207
 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 
huffman 
parents: 
47192 
diff
changeset
 | 
1241  | 
|
| 63652 | 1242  | 
lemma nat_numeral_diff_1 [simp]: "numeral v - (1::nat) = nat (numeral v - 1)"  | 
| 
47207
 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 
huffman 
parents: 
47192 
diff
changeset
 | 
1243  | 
using diff_nat_numeral [of v Num.One] by simp  | 
| 
 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 
huffman 
parents: 
47192 
diff
changeset
 | 
1244  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1245  | 
|
| 63652 | 1246  | 
subsection \<open>Induction principles for int\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1247  | 
|
| 63652 | 1248  | 
text \<open>Well-founded segments of the integers.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1249  | 
|
| 63652 | 1250  | 
definition int_ge_less_than :: "int \<Rightarrow> (int \<times> int) set"  | 
1251  | 
  where "int_ge_less_than d = {(z', z). d \<le> z' \<and> z' < z}"
 | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1252  | 
|
| 63652 | 1253  | 
lemma wf_int_ge_less_than: "wf (int_ge_less_than d)"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1254  | 
proof -  | 
| 63652 | 1255  | 
have "int_ge_less_than d \<subseteq> measure (\<lambda>z. nat (z - d))"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1256  | 
by (auto simp add: int_ge_less_than_def)  | 
| 63652 | 1257  | 
then show ?thesis  | 
| 60162 | 1258  | 
by (rule wf_subset [OF wf_measure])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1259  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1260  | 
|
| 63652 | 1261  | 
text \<open>  | 
1262  | 
This variant looks odd, but is typical of the relations suggested  | 
|
1263  | 
by RankFinder.\<close>  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1264  | 
|
| 63652 | 1265  | 
definition int_ge_less_than2 :: "int \<Rightarrow> (int \<times> int) set"  | 
1266  | 
  where "int_ge_less_than2 d = {(z',z). d \<le> z \<and> z' < z}"
 | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1267  | 
|
| 63652 | 1268  | 
lemma wf_int_ge_less_than2: "wf (int_ge_less_than2 d)"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1269  | 
proof -  | 
| 63652 | 1270  | 
have "int_ge_less_than2 d \<subseteq> measure (\<lambda>z. nat (1 + z - d))"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1271  | 
by (auto simp add: int_ge_less_than2_def)  | 
| 63652 | 1272  | 
then show ?thesis  | 
| 60162 | 1273  | 
by (rule wf_subset [OF wf_measure])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1274  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1275  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1276  | 
(* `set:int': dummy construction *)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1277  | 
theorem int_ge_induct [case_names base step, induct set: int]:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1278  | 
fixes i :: int  | 
| 63652 | 1279  | 
assumes ge: "k \<le> i"  | 
1280  | 
and base: "P k"  | 
|
1281  | 
and step: "\<And>i. k \<le> i \<Longrightarrow> P i \<Longrightarrow> P (i + 1)"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1282  | 
shows "P i"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1283  | 
proof -  | 
| 63652 | 1284  | 
have "\<And>i::int. n = nat (i - k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i" for n  | 
1285  | 
proof (induct n)  | 
|
1286  | 
case 0  | 
|
1287  | 
then have "i = k" by arith  | 
|
1288  | 
with base show "P i" by simp  | 
|
1289  | 
next  | 
|
1290  | 
case (Suc n)  | 
|
1291  | 
then have "n = nat ((i - 1) - k)" by arith  | 
|
1292  | 
moreover have k: "k \<le> i - 1" using Suc.prems by arith  | 
|
1293  | 
ultimately have "P (i - 1)" by (rule Suc.hyps)  | 
|
1294  | 
from step [OF k this] show ?case by simp  | 
|
1295  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1296  | 
with ge show ?thesis by fast  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1297  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1298  | 
|
| 25928 | 1299  | 
(* `set:int': dummy construction *)  | 
1300  | 
theorem int_gr_induct [case_names base step, induct set: int]:  | 
|
| 63652 | 1301  | 
fixes i k :: int  | 
1302  | 
assumes gr: "k < i"  | 
|
1303  | 
and base: "P (k + 1)"  | 
|
1304  | 
and step: "\<And>i. k < i \<Longrightarrow> P i \<Longrightarrow> P (i + 1)"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1305  | 
shows "P i"  | 
| 63652 | 1306  | 
apply (rule int_ge_induct[of "k + 1"])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1307  | 
using gr apply arith  | 
| 63652 | 1308  | 
apply (rule base)  | 
1309  | 
apply (rule step)  | 
|
1310  | 
apply simp_all  | 
|
1311  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1312  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1313  | 
theorem int_le_induct [consumes 1, case_names base step]:  | 
| 63652 | 1314  | 
fixes i k :: int  | 
1315  | 
assumes le: "i \<le> k"  | 
|
1316  | 
and base: "P k"  | 
|
1317  | 
and step: "\<And>i. i \<le> k \<Longrightarrow> P i \<Longrightarrow> P (i - 1)"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1318  | 
shows "P i"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1319  | 
proof -  | 
| 63652 | 1320  | 
have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i" for n  | 
1321  | 
proof (induct n)  | 
|
1322  | 
case 0  | 
|
1323  | 
then have "i = k" by arith  | 
|
1324  | 
with base show "P i" by simp  | 
|
1325  | 
next  | 
|
1326  | 
case (Suc n)  | 
|
1327  | 
then have "n = nat (k - (i + 1))" by arith  | 
|
1328  | 
moreover have k: "i + 1 \<le> k" using Suc.prems by arith  | 
|
1329  | 
ultimately have "P (i + 1)" by (rule Suc.hyps)  | 
|
1330  | 
from step[OF k this] show ?case by simp  | 
|
1331  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1332  | 
with le show ?thesis by fast  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1333  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1334  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1335  | 
theorem int_less_induct [consumes 1, case_names base step]:  | 
| 63652 | 1336  | 
fixes i k :: int  | 
1337  | 
assumes less: "i < k"  | 
|
1338  | 
and base: "P (k - 1)"  | 
|
1339  | 
and step: "\<And>i. i < k \<Longrightarrow> P i \<Longrightarrow> P (i - 1)"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1340  | 
shows "P i"  | 
| 63652 | 1341  | 
apply (rule int_le_induct[of _ "k - 1"])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1342  | 
using less apply arith  | 
| 63652 | 1343  | 
apply (rule base)  | 
1344  | 
apply (rule step)  | 
|
1345  | 
apply simp_all  | 
|
1346  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1347  | 
|
| 
36811
 
4ab4aa5bee1c
renamed former Int.int_induct to Int.int_of_nat_induct, former Presburger.int_induct to Int.int_induct: is more conservative and more natural than the intermediate solution
 
haftmann 
parents: 
36801 
diff
changeset
 | 
1348  | 
theorem int_induct [case_names base step1 step2]:  | 
| 
36801
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1349  | 
fixes k :: int  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1350  | 
assumes base: "P k"  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1351  | 
and step1: "\<And>i. k \<le> i \<Longrightarrow> P i \<Longrightarrow> P (i + 1)"  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1352  | 
and step2: "\<And>i. k \<ge> i \<Longrightarrow> P i \<Longrightarrow> P (i - 1)"  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1353  | 
shows "P i"  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1354  | 
proof -  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1355  | 
have "i \<le> k \<or> i \<ge> k" by arith  | 
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1356  | 
then show ?thesis  | 
| 
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1357  | 
proof  | 
| 
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1358  | 
assume "i \<ge> k"  | 
| 63652 | 1359  | 
then show ?thesis  | 
1360  | 
using base by (rule int_ge_induct) (fact step1)  | 
|
| 
36801
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1361  | 
next  | 
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1362  | 
assume "i \<le> k"  | 
| 63652 | 1363  | 
then show ?thesis  | 
1364  | 
using base by (rule int_le_induct) (fact step2)  | 
|
| 
36801
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1365  | 
qed  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1366  | 
qed  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1367  | 
|
| 63652 | 1368  | 
|
1369  | 
subsection \<open>Intermediate value theorems\<close>  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1370  | 
|
| 67116 | 1371  | 
lemma nat_intermed_int_val:  | 
1372  | 
"\<exists>i. m \<le> i \<and> i \<le> n \<and> f i = k"  | 
|
1373  | 
if "\<forall>i. m \<le> i \<and> i < n \<longrightarrow> \<bar>f (Suc i) - f i\<bar> \<le> 1"  | 
|
1374  | 
"m \<le> n" "f m \<le> k" "k \<le> f n"  | 
|
1375  | 
for m n :: nat and k :: int  | 
|
1376  | 
proof -  | 
|
1377  | 
have "(\<forall>i<n. \<bar>f (Suc i) - f i\<bar> \<le> 1) \<Longrightarrow> f 0 \<le> k \<Longrightarrow> k \<le> f n  | 
|
1378  | 
\<Longrightarrow> (\<exists>i \<le> n. f i = k)"  | 
|
1379  | 
for n :: nat and f  | 
|
1380  | 
apply (induct n)  | 
|
1381  | 
apply auto  | 
|
1382  | 
apply (erule_tac x = n in allE)  | 
|
1383  | 
apply (case_tac "k = f (Suc n)")  | 
|
1384  | 
apply (auto simp add: abs_if split: if_split_asm intro: le_SucI)  | 
|
1385  | 
done  | 
|
1386  | 
from this [of "n - m" "f \<circ> plus m"] that show ?thesis  | 
|
1387  | 
apply auto  | 
|
1388  | 
apply (rule_tac x = "m + i" in exI)  | 
|
1389  | 
apply auto  | 
|
1390  | 
done  | 
|
1391  | 
qed  | 
|
1392  | 
||
1393  | 
lemma nat0_intermed_int_val:  | 
|
1394  | 
"\<exists>i\<le>n. f i = k"  | 
|
1395  | 
if "\<forall>i<n. \<bar>f (i + 1) - f i\<bar> \<le> 1" "f 0 \<le> k" "k \<le> f n"  | 
|
| 63652 | 1396  | 
for n :: nat and k :: int  | 
| 67116 | 1397  | 
using nat_intermed_int_val [of 0 n f k] that by auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1398  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1399  | 
|
| 63652 | 1400  | 
subsection \<open>Products and 1, by T. M. Rasmussen\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1401  | 
|
| 34055 | 1402  | 
lemma abs_zmult_eq_1:  | 
| 63652 | 1403  | 
fixes m n :: int  | 
| 34055 | 1404  | 
assumes mn: "\<bar>m * n\<bar> = 1"  | 
| 63652 | 1405  | 
shows "\<bar>m\<bar> = 1"  | 
| 34055 | 1406  | 
proof -  | 
| 63652 | 1407  | 
from mn have 0: "m \<noteq> 0" "n \<noteq> 0" by auto  | 
1408  | 
have "\<not> 2 \<le> \<bar>m\<bar>"  | 
|
| 34055 | 1409  | 
proof  | 
1410  | 
assume "2 \<le> \<bar>m\<bar>"  | 
|
| 63652 | 1411  | 
then have "2 * \<bar>n\<bar> \<le> \<bar>m\<bar> * \<bar>n\<bar>" by (simp add: mult_mono 0)  | 
1412  | 
also have "\<dots> = \<bar>m * n\<bar>" by (simp add: abs_mult)  | 
|
1413  | 
also from mn have "\<dots> = 1" by simp  | 
|
1414  | 
finally have "2 * \<bar>n\<bar> \<le> 1" .  | 
|
1415  | 
with 0 show "False" by arith  | 
|
| 34055 | 1416  | 
qed  | 
| 63652 | 1417  | 
with 0 show ?thesis by auto  | 
| 34055 | 1418  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1419  | 
|
| 63652 | 1420  | 
lemma pos_zmult_eq_1_iff_lemma: "m * n = 1 \<Longrightarrow> m = 1 \<or> m = - 1"  | 
1421  | 
for m n :: int  | 
|
1422  | 
using abs_zmult_eq_1 [of m n] by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1423  | 
|
| 
35815
 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 
boehmes 
parents: 
35634 
diff
changeset
 | 
1424  | 
lemma pos_zmult_eq_1_iff:  | 
| 63652 | 1425  | 
fixes m n :: int  | 
1426  | 
assumes "0 < m"  | 
|
1427  | 
shows "m * n = 1 \<longleftrightarrow> m = 1 \<and> n = 1"  | 
|
| 
35815
 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 
boehmes 
parents: 
35634 
diff
changeset
 | 
1428  | 
proof -  | 
| 63652 | 1429  | 
from assms have "m * n = 1 \<Longrightarrow> m = 1"  | 
1430  | 
by (auto dest: pos_zmult_eq_1_iff_lemma)  | 
|
1431  | 
then show ?thesis  | 
|
1432  | 
by (auto dest: pos_zmult_eq_1_iff_lemma)  | 
|
| 
35815
 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 
boehmes 
parents: 
35634 
diff
changeset
 | 
1433  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1434  | 
|
| 63652 | 1435  | 
lemma zmult_eq_1_iff: "m * n = 1 \<longleftrightarrow> (m = 1 \<and> n = 1) \<or> (m = - 1 \<and> n = - 1)"  | 
1436  | 
for m n :: int  | 
|
1437  | 
apply (rule iffI)  | 
|
1438  | 
apply (frule pos_zmult_eq_1_iff_lemma)  | 
|
1439  | 
apply (simp add: mult.commute [of m])  | 
|
1440  | 
apply (frule pos_zmult_eq_1_iff_lemma)  | 
|
1441  | 
apply auto  | 
|
1442  | 
done  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1443  | 
|
| 
33296
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1444  | 
lemma infinite_UNIV_int: "\<not> finite (UNIV::int set)"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1445  | 
proof  | 
| 
33296
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1446  | 
assume "finite (UNIV::int set)"  | 
| 61076 | 1447  | 
moreover have "inj (\<lambda>i::int. 2 * i)"  | 
| 
33296
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1448  | 
by (rule injI) simp  | 
| 61076 | 1449  | 
ultimately have "surj (\<lambda>i::int. 2 * i)"  | 
| 
33296
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1450  | 
by (rule finite_UNIV_inj_surj)  | 
| 
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1451  | 
then obtain i :: int where "1 = 2 * i" by (rule surjE)  | 
| 
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1452  | 
then show False by (simp add: pos_zmult_eq_1_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1453  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1454  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1455  | 
|
| 60758 | 1456  | 
subsection \<open>The divides relation\<close>  | 
| 33320 | 1457  | 
|
| 63652 | 1458  | 
lemma zdvd_antisym_nonneg: "0 \<le> m \<Longrightarrow> 0 \<le> n \<Longrightarrow> m dvd n \<Longrightarrow> n dvd m \<Longrightarrow> m = n"  | 
1459  | 
for m n :: int  | 
|
1460  | 
by (auto simp add: dvd_def mult.assoc zero_le_mult_iff zmult_eq_1_iff)  | 
|
| 33320 | 1461  | 
|
| 63652 | 1462  | 
lemma zdvd_antisym_abs:  | 
1463  | 
fixes a b :: int  | 
|
1464  | 
assumes "a dvd b" and "b dvd a"  | 
|
| 33320 | 1465  | 
shows "\<bar>a\<bar> = \<bar>b\<bar>"  | 
| 63652 | 1466  | 
proof (cases "a = 0")  | 
1467  | 
case True  | 
|
1468  | 
with assms show ?thesis by simp  | 
|
| 33657 | 1469  | 
next  | 
| 63652 | 1470  | 
case False  | 
1471  | 
from \<open>a dvd b\<close> obtain k where k: "b = a * k"  | 
|
1472  | 
unfolding dvd_def by blast  | 
|
1473  | 
from \<open>b dvd a\<close> obtain k' where k': "a = b * k'"  | 
|
1474  | 
unfolding dvd_def by blast  | 
|
1475  | 
from k k' have "a = a * k * k'" by simp  | 
|
1476  | 
with mult_cancel_left1[where c="a" and b="k*k'"] have kk': "k * k' = 1"  | 
|
1477  | 
using \<open>a \<noteq> 0\<close> by (simp add: mult.assoc)  | 
|
1478  | 
then have "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1"  | 
|
1479  | 
by (simp add: zmult_eq_1_iff)  | 
|
1480  | 
with k k' show ?thesis by auto  | 
|
| 33320 | 1481  | 
qed  | 
1482  | 
||
| 63652 | 1483  | 
lemma zdvd_zdiffD: "k dvd m - n \<Longrightarrow> k dvd n \<Longrightarrow> k dvd m"  | 
1484  | 
for k m n :: int  | 
|
| 60162 | 1485  | 
using dvd_add_right_iff [of k "- n" m] by simp  | 
| 33320 | 1486  | 
|
| 63652 | 1487  | 
lemma zdvd_reduce: "k dvd n + k * m \<longleftrightarrow> k dvd n"  | 
1488  | 
for k m n :: int  | 
|
| 
58649
 
a62065b5e1e2
generalized and consolidated some theorems concerning divisibility
 
haftmann 
parents: 
58512 
diff
changeset
 | 
1489  | 
using dvd_add_times_triv_right_iff [of k n m] by (simp add: ac_simps)  | 
| 33320 | 1490  | 
|
1491  | 
lemma dvd_imp_le_int:  | 
|
1492  | 
fixes d i :: int  | 
|
1493  | 
assumes "i \<noteq> 0" and "d dvd i"  | 
|
1494  | 
shows "\<bar>d\<bar> \<le> \<bar>i\<bar>"  | 
|
1495  | 
proof -  | 
|
| 60758 | 1496  | 
from \<open>d dvd i\<close> obtain k where "i = d * k" ..  | 
1497  | 
with \<open>i \<noteq> 0\<close> have "k \<noteq> 0" by auto  | 
|
| 33320 | 1498  | 
then have "1 \<le> \<bar>k\<bar>" and "0 \<le> \<bar>d\<bar>" by auto  | 
1499  | 
then have "\<bar>d\<bar> * 1 \<le> \<bar>d\<bar> * \<bar>k\<bar>" by (rule mult_left_mono)  | 
|
| 60758 | 1500  | 
with \<open>i = d * k\<close> show ?thesis by (simp add: abs_mult)  | 
| 33320 | 1501  | 
qed  | 
1502  | 
||
1503  | 
lemma zdvd_not_zless:  | 
|
1504  | 
fixes m n :: int  | 
|
1505  | 
assumes "0 < m" and "m < n"  | 
|
1506  | 
shows "\<not> n dvd m"  | 
|
1507  | 
proof  | 
|
1508  | 
from assms have "0 < n" by auto  | 
|
1509  | 
assume "n dvd m" then obtain k where k: "m = n * k" ..  | 
|
| 60758 | 1510  | 
with \<open>0 < m\<close> have "0 < n * k" by auto  | 
1511  | 
with \<open>0 < n\<close> have "0 < k" by (simp add: zero_less_mult_iff)  | 
|
1512  | 
with k \<open>0 < n\<close> \<open>m < n\<close> have "n * k < n * 1" by simp  | 
|
1513  | 
with \<open>0 < n\<close> \<open>0 < k\<close> show False unfolding mult_less_cancel_left by auto  | 
|
| 33320 | 1514  | 
qed  | 
1515  | 
||
| 63652 | 1516  | 
lemma zdvd_mult_cancel:  | 
1517  | 
fixes k m n :: int  | 
|
1518  | 
assumes d: "k * m dvd k * n"  | 
|
1519  | 
and "k \<noteq> 0"  | 
|
| 33320 | 1520  | 
shows "m dvd n"  | 
| 63652 | 1521  | 
proof -  | 
1522  | 
from d obtain h where h: "k * n = k * m * h"  | 
|
1523  | 
unfolding dvd_def by blast  | 
|
1524  | 
have "n = m * h"  | 
|
1525  | 
proof (rule ccontr)  | 
|
1526  | 
assume "\<not> ?thesis"  | 
|
1527  | 
with \<open>k \<noteq> 0\<close> have "k * n \<noteq> k * (m * h)" by simp  | 
|
1528  | 
with h show False  | 
|
1529  | 
by (simp add: mult.assoc)  | 
|
1530  | 
qed  | 
|
1531  | 
then show ?thesis by simp  | 
|
| 33320 | 1532  | 
qed  | 
1533  | 
||
| 67118 | 1534  | 
lemma int_dvd_int_iff [simp]:  | 
1535  | 
"int m dvd int n \<longleftrightarrow> m dvd n"  | 
|
| 33320 | 1536  | 
proof -  | 
| 67118 | 1537  | 
have "m dvd n" if "int n = int m * k" for k  | 
| 63652 | 1538  | 
proof (cases k)  | 
| 67118 | 1539  | 
case (nonneg q)  | 
1540  | 
with that have "n = m * q"  | 
|
| 63652 | 1541  | 
by (simp del: of_nat_mult add: of_nat_mult [symmetric])  | 
1542  | 
then show ?thesis ..  | 
|
1543  | 
next  | 
|
| 67118 | 1544  | 
case (neg q)  | 
1545  | 
with that have "int n = int m * (- int (Suc q))"  | 
|
| 63652 | 1546  | 
by simp  | 
| 67118 | 1547  | 
also have "\<dots> = - (int m * int (Suc q))"  | 
| 63652 | 1548  | 
by (simp only: mult_minus_right)  | 
| 67118 | 1549  | 
also have "\<dots> = - int (m * Suc q)"  | 
| 63652 | 1550  | 
by (simp only: of_nat_mult [symmetric])  | 
| 67118 | 1551  | 
finally have "- int (m * Suc q) = int n" ..  | 
| 63652 | 1552  | 
then show ?thesis  | 
1553  | 
by (simp only: negative_eq_positive) auto  | 
|
| 33320 | 1554  | 
qed  | 
| 67118 | 1555  | 
then show ?thesis by (auto simp add: dvd_def)  | 
| 33320 | 1556  | 
qed  | 
1557  | 
||
| 67118 | 1558  | 
lemma dvd_nat_abs_iff [simp]:  | 
1559  | 
"n dvd nat \<bar>k\<bar> \<longleftrightarrow> int n dvd k"  | 
|
1560  | 
proof -  | 
|
1561  | 
have "n dvd nat \<bar>k\<bar> \<longleftrightarrow> int n dvd int (nat \<bar>k\<bar>)"  | 
|
1562  | 
by (simp only: int_dvd_int_iff)  | 
|
1563  | 
then show ?thesis  | 
|
1564  | 
by simp  | 
|
1565  | 
qed  | 
|
1566  | 
||
1567  | 
lemma nat_abs_dvd_iff [simp]:  | 
|
1568  | 
"nat \<bar>k\<bar> dvd n \<longleftrightarrow> k dvd int n"  | 
|
1569  | 
proof -  | 
|
1570  | 
have "nat \<bar>k\<bar> dvd n \<longleftrightarrow> int (nat \<bar>k\<bar>) dvd int n"  | 
|
1571  | 
by (simp only: int_dvd_int_iff)  | 
|
1572  | 
then show ?thesis  | 
|
1573  | 
by simp  | 
|
1574  | 
qed  | 
|
1575  | 
||
1576  | 
lemma zdvd1_eq [simp]: "x dvd 1 \<longleftrightarrow> \<bar>x\<bar> = 1" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 63652 | 1577  | 
for x :: int  | 
| 33320 | 1578  | 
proof  | 
| 63652 | 1579  | 
assume ?lhs  | 
| 67118 | 1580  | 
then have "nat \<bar>x\<bar> dvd nat \<bar>1\<bar>"  | 
1581  | 
by (simp only: nat_abs_dvd_iff) simp  | 
|
1582  | 
then have "nat \<bar>x\<bar> = 1"  | 
|
1583  | 
by simp  | 
|
1584  | 
then show ?rhs  | 
|
1585  | 
by (cases "x < 0") simp_all  | 
|
| 33320 | 1586  | 
next  | 
| 63652 | 1587  | 
assume ?rhs  | 
| 67118 | 1588  | 
then have "x = 1 \<or> x = - 1"  | 
1589  | 
by auto  | 
|
1590  | 
then show ?lhs  | 
|
1591  | 
by (auto intro: dvdI)  | 
|
| 33320 | 1592  | 
qed  | 
1593  | 
||
| 60162 | 1594  | 
lemma zdvd_mult_cancel1:  | 
| 63652 | 1595  | 
fixes m :: int  | 
1596  | 
assumes mp: "m \<noteq> 0"  | 
|
1597  | 
shows "m * n dvd m \<longleftrightarrow> \<bar>n\<bar> = 1"  | 
|
1598  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 33320 | 1599  | 
proof  | 
| 63652 | 1600  | 
assume ?rhs  | 
1601  | 
then show ?lhs  | 
|
1602  | 
by (cases "n > 0") (auto simp add: minus_equation_iff)  | 
|
| 33320 | 1603  | 
next  | 
| 63652 | 1604  | 
assume ?lhs  | 
1605  | 
then have "m * n dvd m * 1" by simp  | 
|
1606  | 
from zdvd_mult_cancel[OF this mp] show ?rhs  | 
|
1607  | 
by (simp only: zdvd1_eq)  | 
|
| 33320 | 1608  | 
qed  | 
1609  | 
||
| 63652 | 1610  | 
lemma nat_dvd_iff: "nat z dvd m \<longleftrightarrow> (if 0 \<le> z then z dvd int m else m = 0)"  | 
| 67118 | 1611  | 
using nat_abs_dvd_iff [of z m] by (cases "z \<ge> 0") auto  | 
| 33320 | 1612  | 
|
| 63652 | 1613  | 
lemma eq_nat_nat_iff: "0 \<le> z \<Longrightarrow> 0 \<le> z' \<Longrightarrow> nat z = nat z' \<longleftrightarrow> z = z'"  | 
| 67116 | 1614  | 
by (auto elim: nonneg_int_cases)  | 
| 33341 | 1615  | 
|
| 63652 | 1616  | 
lemma nat_power_eq: "0 \<le> z \<Longrightarrow> nat (z ^ n) = nat z ^ n"  | 
| 33341 | 1617  | 
by (induct n) (simp_all add: nat_mult_distrib)  | 
1618  | 
||
| 
66912
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1619  | 
lemma numeral_power_eq_nat_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1620  | 
"numeral x ^ n = nat y \<longleftrightarrow> numeral x ^ n = y"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1621  | 
using nat_eq_iff2 by auto  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1622  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1623  | 
lemma nat_eq_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1624  | 
"nat y = numeral x ^ n \<longleftrightarrow> y = numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1625  | 
using numeral_power_eq_nat_cancel_iff[of x n y]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1626  | 
by (metis (mono_tags))  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1627  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1628  | 
lemma numeral_power_le_nat_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1629  | 
"numeral x ^ n \<le> nat a \<longleftrightarrow> numeral x ^ n \<le> a"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1630  | 
using nat_le_eq_zle[of "numeral x ^ n" a]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1631  | 
by (auto simp: nat_power_eq)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1632  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1633  | 
lemma nat_le_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1634  | 
"nat a \<le> numeral x ^ n \<longleftrightarrow> a \<le> numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1635  | 
by (simp add: nat_le_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1636  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1637  | 
lemma numeral_power_less_nat_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1638  | 
"numeral x ^ n < nat a \<longleftrightarrow> numeral x ^ n < a"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1639  | 
using nat_less_eq_zless[of "numeral x ^ n" a]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1640  | 
by (auto simp: nat_power_eq)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1641  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1642  | 
lemma nat_less_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1643  | 
"nat a < numeral x ^ n \<longleftrightarrow> a < numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1644  | 
using nat_less_eq_zless[of a "numeral x ^ n"]  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1645  | 
by (cases "a < 0") (auto simp: nat_power_eq less_le_trans[where y=0])  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66886 
diff
changeset
 | 
1646  | 
|
| 63652 | 1647  | 
lemma zdvd_imp_le: "z dvd n \<Longrightarrow> 0 < n \<Longrightarrow> z \<le> n"  | 
1648  | 
for n z :: int  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1649  | 
apply (cases n)  | 
| 67118 | 1650  | 
apply auto  | 
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1651  | 
apply (cases z)  | 
| 63652 | 1652  | 
apply (auto simp add: dvd_imp_le)  | 
| 33320 | 1653  | 
done  | 
1654  | 
||
| 36749 | 1655  | 
lemma zdvd_period:  | 
1656  | 
fixes a d :: int  | 
|
1657  | 
assumes "a dvd d"  | 
|
1658  | 
shows "a dvd (x + t) \<longleftrightarrow> a dvd ((x + c * d) + t)"  | 
|
| 63652 | 1659  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
| 36749 | 1660  | 
proof -  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1661  | 
from assms have "a dvd (x + t) \<longleftrightarrow> a dvd ((x + t) + c * d)"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1662  | 
by (simp add: dvd_add_left_iff)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1663  | 
then show ?thesis  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1664  | 
by (simp add: ac_simps)  | 
| 36749 | 1665  | 
qed  | 
1666  | 
||
| 33320 | 1667  | 
|
| 60758 | 1668  | 
subsection \<open>Finiteness of intervals\<close>  | 
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1669  | 
|
| 63652 | 1670  | 
lemma finite_interval_int1 [iff]: "finite {i :: int. a \<le> i \<and> i \<le> b}"
 | 
1671  | 
proof (cases "a \<le> b")  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1672  | 
case True  | 
| 63652 | 1673  | 
then show ?thesis  | 
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1674  | 
proof (induct b rule: int_ge_induct)  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1675  | 
case base  | 
| 63652 | 1676  | 
    have "{i. a \<le> i \<and> i \<le> a} = {a}" by auto
 | 
1677  | 
then show ?case by simp  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1678  | 
next  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1679  | 
case (step b)  | 
| 63652 | 1680  | 
    then have "{i. a \<le> i \<and> i \<le> b + 1} = {i. a \<le> i \<and> i \<le> b} \<union> {b + 1}" by auto
 | 
1681  | 
with step show ?case by simp  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1682  | 
qed  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1683  | 
next  | 
| 63652 | 1684  | 
case False  | 
1685  | 
then show ?thesis  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1686  | 
by (metis (lifting, no_types) Collect_empty_eq finite.emptyI order_trans)  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1687  | 
qed  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1688  | 
|
| 63652 | 1689  | 
lemma finite_interval_int2 [iff]: "finite {i :: int. a \<le> i \<and> i < b}"
 | 
1690  | 
by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1691  | 
|
| 63652 | 1692  | 
lemma finite_interval_int3 [iff]: "finite {i :: int. a < i \<and> i \<le> b}"
 | 
1693  | 
by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1694  | 
|
| 63652 | 1695  | 
lemma finite_interval_int4 [iff]: "finite {i :: int. a < i \<and> i < b}"
 | 
1696  | 
by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1697  | 
|
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
1698  | 
|
| 60758 | 1699  | 
subsection \<open>Configuration of the code generator\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1700  | 
|
| 60758 | 1701  | 
text \<open>Constructors\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1702  | 
|
| 63652 | 1703  | 
definition Pos :: "num \<Rightarrow> int"  | 
1704  | 
where [simp, code_abbrev]: "Pos = numeral"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1705  | 
|
| 63652 | 1706  | 
definition Neg :: "num \<Rightarrow> int"  | 
1707  | 
where [simp, code_abbrev]: "Neg n = - (Pos n)"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1708  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1709  | 
code_datatype "0::int" Pos Neg  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1710  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1711  | 
|
| 63652 | 1712  | 
text \<open>Auxiliary operations.\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1713  | 
|
| 63652 | 1714  | 
definition dup :: "int \<Rightarrow> int"  | 
1715  | 
where [simp]: "dup k = k + k"  | 
|
| 26507 | 1716  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1717  | 
lemma dup_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1718  | 
"dup 0 = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1719  | 
"dup (Pos n) = Pos (Num.Bit0 n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1720  | 
"dup (Neg n) = Neg (Num.Bit0 n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1721  | 
by (simp_all add: numeral_Bit0)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1722  | 
|
| 63652 | 1723  | 
definition sub :: "num \<Rightarrow> num \<Rightarrow> int"  | 
1724  | 
where [simp]: "sub m n = numeral m - numeral n"  | 
|
| 26507 | 1725  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1726  | 
lemma sub_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1727  | 
"sub Num.One Num.One = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1728  | 
"sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1729  | 
"sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1730  | 
"sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1731  | 
"sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1732  | 
"sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1733  | 
"sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1734  | 
"sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1735  | 
"sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"  | 
| 
66035
 
de6cd60b1226
replace non-arithmetic terms by fresh variables before replaying linear-arithmetic proofs: avoid failed proof replays due to an overambitious simpset which may cause proof replay to diverge from the pre-computed proof trace
 
boehmes 
parents: 
64996 
diff
changeset
 | 
1736  | 
by (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM)  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1737  | 
|
| 63652 | 1738  | 
text \<open>Implementations.\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1739  | 
|
| 64996 | 1740  | 
lemma one_int_code [code]: "1 = Pos Num.One"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1741  | 
by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1742  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1743  | 
lemma plus_int_code [code]:  | 
| 63652 | 1744  | 
"k + 0 = k"  | 
1745  | 
"0 + l = l"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1746  | 
"Pos m + Pos n = Pos (m + n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1747  | 
"Pos m + Neg n = sub m n"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1748  | 
"Neg m + Pos n = sub n m"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1749  | 
"Neg m + Neg n = Neg (m + n)"  | 
| 63652 | 1750  | 
for k l :: int  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1751  | 
by simp_all  | 
| 26507 | 1752  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1753  | 
lemma uminus_int_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1754  | 
"uminus 0 = (0::int)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1755  | 
"uminus (Pos m) = Neg m"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1756  | 
"uminus (Neg m) = Pos m"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1757  | 
by simp_all  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1758  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1759  | 
lemma minus_int_code [code]:  | 
| 63652 | 1760  | 
"k - 0 = k"  | 
1761  | 
"0 - l = uminus l"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1762  | 
"Pos m - Pos n = sub m n"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1763  | 
"Pos m - Neg n = Pos (m + n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1764  | 
"Neg m - Pos n = Neg (m + n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1765  | 
"Neg m - Neg n = sub n m"  | 
| 63652 | 1766  | 
for k l :: int  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1767  | 
by simp_all  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1768  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1769  | 
lemma times_int_code [code]:  | 
| 63652 | 1770  | 
"k * 0 = 0"  | 
1771  | 
"0 * l = 0"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1772  | 
"Pos m * Pos n = Pos (m * n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1773  | 
"Pos m * Neg n = Neg (m * n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1774  | 
"Neg m * Pos n = Neg (m * n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1775  | 
"Neg m * Neg n = Pos (m * n)"  | 
| 63652 | 1776  | 
for k l :: int  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1777  | 
by simp_all  | 
| 26507 | 1778  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37887 
diff
changeset
 | 
1779  | 
instantiation int :: equal  | 
| 26507 | 1780  | 
begin  | 
1781  | 
||
| 63652 | 1782  | 
definition "HOL.equal k l \<longleftrightarrow> k = (l::int)"  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37887 
diff
changeset
 | 
1783  | 
|
| 61169 | 1784  | 
instance  | 
1785  | 
by standard (rule equal_int_def)  | 
|
| 26507 | 1786  | 
|
1787  | 
end  | 
|
1788  | 
||
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1789  | 
lemma equal_int_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1790  | 
"HOL.equal 0 (0::int) \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1791  | 
"HOL.equal 0 (Pos l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1792  | 
"HOL.equal 0 (Neg l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1793  | 
"HOL.equal (Pos k) 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1794  | 
"HOL.equal (Pos k) (Pos l) \<longleftrightarrow> HOL.equal k l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1795  | 
"HOL.equal (Pos k) (Neg l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1796  | 
"HOL.equal (Neg k) 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1797  | 
"HOL.equal (Neg k) (Pos l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1798  | 
"HOL.equal (Neg k) (Neg l) \<longleftrightarrow> HOL.equal k l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1799  | 
by (auto simp add: equal)  | 
| 26507 | 1800  | 
|
| 63652 | 1801  | 
lemma equal_int_refl [code nbe]: "HOL.equal k k \<longleftrightarrow> True"  | 
1802  | 
for k :: int  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1803  | 
by (fact equal_refl)  | 
| 26507 | 1804  | 
|
| 28562 | 1805  | 
lemma less_eq_int_code [code]:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1806  | 
"0 \<le> (0::int) \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1807  | 
"0 \<le> Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1808  | 
"0 \<le> Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1809  | 
"Pos k \<le> 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1810  | 
"Pos k \<le> Pos l \<longleftrightarrow> k \<le> l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1811  | 
"Pos k \<le> Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1812  | 
"Neg k \<le> 0 \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1813  | 
"Neg k \<le> Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1814  | 
"Neg k \<le> Neg l \<longleftrightarrow> l \<le> k"  | 
| 28958 | 1815  | 
by simp_all  | 
| 26507 | 1816  | 
|
| 28562 | 1817  | 
lemma less_int_code [code]:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1818  | 
"0 < (0::int) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1819  | 
"0 < Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1820  | 
"0 < Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1821  | 
"Pos k < 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1822  | 
"Pos k < Pos l \<longleftrightarrow> k < l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1823  | 
"Pos k < Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1824  | 
"Neg k < 0 \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1825  | 
"Neg k < Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1826  | 
"Neg k < Neg l \<longleftrightarrow> l < k"  | 
| 28958 | 1827  | 
by simp_all  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1828  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1829  | 
lemma nat_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1830  | 
"nat (Int.Neg k) = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1831  | 
"nat 0 = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1832  | 
"nat (Int.Pos k) = nat_of_num k"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
1833  | 
by (simp_all add: nat_of_num_numeral)  | 
| 25928 | 1834  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1835  | 
lemma (in ring_1) of_int_code [code]:  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
1836  | 
"of_int (Int.Neg k) = - numeral k"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1837  | 
"of_int 0 = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1838  | 
"of_int (Int.Pos k) = numeral k"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1839  | 
by simp_all  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1840  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1841  | 
|
| 63652 | 1842  | 
text \<open>Serializer setup.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1843  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51994 
diff
changeset
 | 
1844  | 
code_identifier  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51994 
diff
changeset
 | 
1845  | 
code_module Int \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1846  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1847  | 
quickcheck_params [default_type = int]  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1848  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1849  | 
hide_const (open) Pos Neg sub dup  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1850  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1851  | 
|
| 61799 | 1852  | 
text \<open>De-register \<open>int\<close> as a quotient type:\<close>  | 
| 48045 | 1853  | 
|
| 
53652
 
18fbca265e2e
use lifting_forget for deregistering numeric types as a quotient type
 
kuncar 
parents: 
53065 
diff
changeset
 | 
1854  | 
lifting_update int.lifting  | 
| 
 
18fbca265e2e
use lifting_forget for deregistering numeric types as a quotient type
 
kuncar 
parents: 
53065 
diff
changeset
 | 
1855  | 
lifting_forget int.lifting  | 
| 48045 | 1856  | 
|
| 67116 | 1857  | 
|
1858  | 
subsection \<open>Duplicates\<close>  | 
|
1859  | 
||
1860  | 
lemmas int_sum = of_nat_sum [where 'a=int]  | 
|
1861  | 
lemmas int_prod = of_nat_prod [where 'a=int]  | 
|
1862  | 
lemmas zle_int = of_nat_le_iff [where 'a=int]  | 
|
1863  | 
lemmas int_int_eq = of_nat_eq_iff [where 'a=int]  | 
|
1864  | 
lemmas nonneg_eq_int = nonneg_int_cases  | 
|
1865  | 
lemmas double_eq_0_iff = double_zero  | 
|
1866  | 
||
1867  | 
lemmas int_distrib =  | 
|
1868  | 
distrib_right [of z1 z2 w]  | 
|
1869  | 
distrib_left [of w z1 z2]  | 
|
1870  | 
left_diff_distrib [of z1 z2 w]  | 
|
1871  | 
right_diff_distrib [of w z1 z2]  | 
|
1872  | 
for z1 z2 w :: int  | 
|
1873  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1874  | 
end  | 
| 67116 | 1875  |