| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 08 Nov 2023 11:08:03 +0100 | |
| changeset 78928 | 6c2c60b852e0 | 
| parent 78698 | 1b9388e6eb75 | 
| child 78935 | 5e788ff7a489 | 
| 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  | 
| 74979 | 9  | 
imports Quotient Groups_Big 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)"  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
56  | 
proof (unfold intrel_def, clarify)  | 
| 48045 | 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  | 
|
| 70927 | 77  | 
lemma int_transfer [transfer_rule]:  | 
78  | 
includes lifting_syntax  | 
|
79  | 
shows "rel_fun (=) pcr_int (\<lambda>n. (n, 0)) int"  | 
|
| 63652 | 80  | 
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
 | 
81  | 
|
| 63652 | 82  | 
lemma int_diff_cases: obtains (diff) m n where "z = int m - int n"  | 
| 48045 | 83  | 
by transfer clarsimp  | 
84  | 
||
| 63652 | 85  | 
|
| 60758 | 86  | 
subsection \<open>Integers are totally ordered\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
87  | 
|
| 48045 | 88  | 
instantiation int :: linorder  | 
89  | 
begin  | 
|
90  | 
||
91  | 
lift_definition less_eq_int :: "int \<Rightarrow> int \<Rightarrow> bool"  | 
|
92  | 
is "\<lambda>(x, y) (u, v). x + v \<le> u + y"  | 
|
93  | 
by auto  | 
|
94  | 
||
95  | 
lift_definition less_int :: "int \<Rightarrow> int \<Rightarrow> bool"  | 
|
96  | 
is "\<lambda>(x, y) (u, v). x + v < u + y"  | 
|
97  | 
by auto  | 
|
98  | 
||
99  | 
instance  | 
|
| 61169 | 100  | 
by standard (transfer, force)+  | 
| 48045 | 101  | 
|
102  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
104  | 
instantiation int :: distrib_lattice  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
105  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
106  | 
|
| 63652 | 107  | 
definition "(inf :: int \<Rightarrow> int \<Rightarrow> int) = min"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
108  | 
|
| 63652 | 109  | 
definition "(sup :: int \<Rightarrow> int \<Rightarrow> int) = max"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
111  | 
instance  | 
| 63652 | 112  | 
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
 | 
113  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
114  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
115  | 
|
| 60758 | 116  | 
subsection \<open>Ordering properties of arithmetic operations\<close>  | 
| 48045 | 117  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34055 
diff
changeset
 | 
118  | 
instance int :: ordered_cancel_ab_semigroup_add  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
119  | 
proof  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
120  | 
fix i j k :: int  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
121  | 
show "i \<le> j \<Longrightarrow> k + i \<le> k + j"  | 
| 48045 | 122  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
123  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
124  | 
|
| 63652 | 125  | 
text \<open>Strict Monotonicity of Multiplication.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
126  | 
|
| 63652 | 127  | 
text \<open>Strict, in 1st argument; proof is by induction on \<open>k > 0\<close>.\<close>  | 
128  | 
lemma zmult_zless_mono2_lemma: "i < j \<Longrightarrow> 0 < k \<Longrightarrow> int k * i < int k * j"  | 
|
129  | 
for i j :: int  | 
|
130  | 
proof (induct k)  | 
|
131  | 
case 0  | 
|
132  | 
then show ?case by simp  | 
|
133  | 
next  | 
|
134  | 
case (Suc k)  | 
|
135  | 
then show ?case  | 
|
136  | 
by (cases "k = 0") (simp_all add: distrib_right add_strict_mono)  | 
|
137  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
138  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
139  | 
lemma zero_le_imp_eq_int:  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
140  | 
assumes "k \<ge> (0::int)" shows "\<exists>n. k = int n"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
141  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
142  | 
have "b \<le> a \<Longrightarrow> \<exists>n::nat. a = n + b" for a b  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
143  | 
using exI[of _ "a - b"] by simp  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
144  | 
with assms show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
145  | 
by transfer auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
146  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
147  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
148  | 
lemma zero_less_imp_eq_int:  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
149  | 
assumes "k > (0::int)" shows "\<exists>n>0. k = int n"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
150  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
151  | 
have "b < a \<Longrightarrow> \<exists>n::nat. n>0 \<and> a = n + b" for a b  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
152  | 
using exI[of _ "a - b"] by simp  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
153  | 
with assms show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
154  | 
by transfer auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
155  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
156  | 
|
| 63652 | 157  | 
lemma zmult_zless_mono2: "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"  | 
158  | 
for i j k :: int  | 
|
159  | 
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
 | 
160  | 
|
| 63652 | 161  | 
|
162  | 
text \<open>The integers form an ordered integral domain.\<close>  | 
|
163  | 
||
| 48045 | 164  | 
instantiation int :: linordered_idom  | 
165  | 
begin  | 
|
166  | 
||
| 63652 | 167  | 
definition zabs_def: "\<bar>i::int\<bar> = (if i < 0 then - i else i)"  | 
| 48045 | 168  | 
|
| 63652 | 169  | 
definition zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"  | 
| 48045 | 170  | 
|
| 63652 | 171  | 
instance  | 
172  | 
proof  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
173  | 
fix i j k :: int  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
174  | 
show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
175  | 
by (rule zmult_zless_mono2)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
176  | 
show "\<bar>i\<bar> = (if i < 0 then -i else i)"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
177  | 
by (simp only: zabs_def)  | 
| 61076 | 178  | 
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
 | 
179  | 
by (simp only: zsgn_def)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
180  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
181  | 
|
| 48045 | 182  | 
end  | 
183  | 
||
| 63652 | 184  | 
lemma zless_imp_add1_zle: "w < z \<Longrightarrow> w + 1 \<le> z"  | 
185  | 
for w z :: int  | 
|
| 48045 | 186  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
187  | 
|
| 63652 | 188  | 
lemma zless_iff_Suc_zadd: "w < z \<longleftrightarrow> (\<exists>n. z = w + int (Suc n))"  | 
189  | 
for w z :: int  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
190  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
191  | 
have "\<And>a b c d. a + d < c + b \<Longrightarrow> \<exists>n. c + b = Suc (a + n + d)"  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
192  | 
proof -  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
193  | 
fix a b c d :: nat  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
194  | 
assume "a + d < c + b"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
195  | 
then have "c + b = Suc (a + (c + b - Suc (a + d)) + d) "  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
196  | 
by arith  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
197  | 
then show "\<exists>n. c + b = Suc (a + n + d)"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
198  | 
by (rule exI)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
199  | 
qed  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
200  | 
then show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
201  | 
by transfer auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
202  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
203  | 
|
| 63652 | 204  | 
lemma zabs_less_one_iff [simp]: "\<bar>z\<bar> < 1 \<longleftrightarrow> z = 0" (is "?lhs \<longleftrightarrow> ?rhs")  | 
205  | 
for z :: int  | 
|
| 62347 | 206  | 
proof  | 
| 63652 | 207  | 
assume ?rhs  | 
208  | 
then show ?lhs by simp  | 
|
| 62347 | 209  | 
next  | 
| 63652 | 210  | 
assume ?lhs  | 
211  | 
with zless_imp_add1_zle [of "\<bar>z\<bar>" 1] have "\<bar>z\<bar> + 1 \<le> 1" by simp  | 
|
212  | 
then have "\<bar>z\<bar> \<le> 0" by simp  | 
|
213  | 
then show ?rhs by simp  | 
|
| 62347 | 214  | 
qed  | 
215  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
216  | 
|
| 61799 | 217  | 
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
 | 
218  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
219  | 
context ring_1  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
220  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
221  | 
|
| 63652 | 222  | 
lift_definition of_int :: "int \<Rightarrow> 'a"  | 
223  | 
is "\<lambda>(i, j). of_nat i - of_nat j"  | 
|
| 48045 | 224  | 
by (clarsimp simp add: diff_eq_eq eq_diff_eq diff_add_eq  | 
| 63652 | 225  | 
of_nat_add [symmetric] simp del: of_nat_add)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
226  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
227  | 
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
 | 
228  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
229  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
230  | 
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
 | 
231  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
232  | 
|
| 63652 | 233  | 
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
 | 
234  | 
by transfer (clarsimp simp add: algebra_simps)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
235  | 
|
| 63652 | 236  | 
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
 | 
237  | 
by (transfer fixing: uminus) clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
239  | 
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
 | 
240  | 
using of_int_add [of w "- z"] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
241  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
242  | 
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"  | 
| 63652 | 243  | 
by (transfer fixing: times) (clarsimp simp add: algebra_simps)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
244  | 
|
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61524 
diff
changeset
 | 
245  | 
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
 | 
246  | 
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
 | 
247  | 
|
| 63652 | 248  | 
text \<open>Collapse nested embeddings.\<close>  | 
| 44709 | 249  | 
lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n"  | 
| 63652 | 250  | 
by (induct n) auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
251  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
252  | 
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
 | 
253  | 
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
 | 
254  | 
|
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
255  | 
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
 | 
256  | 
by simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
257  | 
|
| 63652 | 258  | 
lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n"  | 
| 31015 | 259  | 
by (induct n) simp_all  | 
260  | 
||
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
261  | 
lemma of_int_of_bool [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
262  | 
"of_int (of_bool P) = of_bool P"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
263  | 
by auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
264  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
265  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
266  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
267  | 
context ring_char_0  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
268  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
269  | 
|
| 63652 | 270  | 
lemma of_int_eq_iff [simp]: "of_int w = of_int z \<longleftrightarrow> w = z"  | 
271  | 
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
 | 
272  | 
|
| 63652 | 273  | 
text \<open>Special cases where either operand is zero.\<close>  | 
274  | 
lemma of_int_eq_0_iff [simp]: "of_int z = 0 \<longleftrightarrow> z = 0"  | 
|
| 36424 | 275  | 
using of_int_eq_iff [of z 0] by simp  | 
276  | 
||
| 63652 | 277  | 
lemma of_int_0_eq_iff [simp]: "0 = of_int z \<longleftrightarrow> z = 0"  | 
| 36424 | 278  | 
using of_int_eq_iff [of 0 z] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
279  | 
|
| 63652 | 280  | 
lemma of_int_eq_1_iff [iff]: "of_int z = 1 \<longleftrightarrow> z = 1"  | 
| 61234 | 281  | 
using of_int_eq_iff [of z 1] by simp  | 
282  | 
||
| 
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
 | 
283  | 
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
 | 
284  | 
"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
 | 
285  | 
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
 | 
286  | 
|
| 
 
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  | 
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
 | 
288  | 
"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
 | 
289  | 
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
 | 
290  | 
|
| 
 
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  | 
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
 | 
292  | 
"(- 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
 | 
293  | 
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
 | 
294  | 
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
 | 
295  | 
|
| 
 
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
 | 
296  | 
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
 | 
297  | 
"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
 | 
298  | 
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
 | 
299  | 
|
| 
 
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
 | 
300  | 
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
 | 
301  | 
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
 | 
302  | 
|
| 
 
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
 | 
303  | 
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
 | 
304  | 
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
 | 
305  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
306  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
307  | 
|
| 36424 | 308  | 
context linordered_idom  | 
309  | 
begin  | 
|
310  | 
||
| 63652 | 311  | 
text \<open>Every \<open>linordered_idom\<close> has characteristic zero.\<close>  | 
| 36424 | 312  | 
subclass ring_char_0 ..  | 
313  | 
||
| 63652 | 314  | 
lemma of_int_le_iff [simp]: "of_int w \<le> of_int z \<longleftrightarrow> w \<le> z"  | 
315  | 
by (transfer fixing: less_eq)  | 
|
316  | 
(clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add)  | 
|
| 36424 | 317  | 
|
| 63652 | 318  | 
lemma of_int_less_iff [simp]: "of_int w < of_int z \<longleftrightarrow> w < z"  | 
| 36424 | 319  | 
by (simp add: less_le order_less_le)  | 
320  | 
||
| 63652 | 321  | 
lemma of_int_0_le_iff [simp]: "0 \<le> of_int z \<longleftrightarrow> 0 \<le> z"  | 
| 36424 | 322  | 
using of_int_le_iff [of 0 z] by simp  | 
323  | 
||
| 63652 | 324  | 
lemma of_int_le_0_iff [simp]: "of_int z \<le> 0 \<longleftrightarrow> z \<le> 0"  | 
| 36424 | 325  | 
using of_int_le_iff [of z 0] by simp  | 
326  | 
||
| 63652 | 327  | 
lemma of_int_0_less_iff [simp]: "0 < of_int z \<longleftrightarrow> 0 < z"  | 
| 36424 | 328  | 
using of_int_less_iff [of 0 z] by simp  | 
329  | 
||
| 63652 | 330  | 
lemma of_int_less_0_iff [simp]: "of_int z < 0 \<longleftrightarrow> z < 0"  | 
| 36424 | 331  | 
using of_int_less_iff [of z 0] by simp  | 
332  | 
||
| 63652 | 333  | 
lemma of_int_1_le_iff [simp]: "1 \<le> of_int z \<longleftrightarrow> 1 \<le> z"  | 
| 61234 | 334  | 
using of_int_le_iff [of 1 z] by simp  | 
335  | 
||
| 63652 | 336  | 
lemma of_int_le_1_iff [simp]: "of_int z \<le> 1 \<longleftrightarrow> z \<le> 1"  | 
| 61234 | 337  | 
using of_int_le_iff [of z 1] by simp  | 
338  | 
||
| 63652 | 339  | 
lemma of_int_1_less_iff [simp]: "1 < of_int z \<longleftrightarrow> 1 < z"  | 
| 61234 | 340  | 
using of_int_less_iff [of 1 z] by simp  | 
341  | 
||
| 63652 | 342  | 
lemma of_int_less_1_iff [simp]: "of_int z < 1 \<longleftrightarrow> z < 1"  | 
| 61234 | 343  | 
using of_int_less_iff [of z 1] by simp  | 
344  | 
||
| 
62128
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
345  | 
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
 | 
346  | 
by simp  | 
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
347  | 
|
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
348  | 
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
 | 
349  | 
by simp  | 
| 
 
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
 
eberlm 
parents: 
61944 
diff
changeset
 | 
350  | 
|
| 63652 | 351  | 
lemma of_int_abs [simp]: "of_int \<bar>x\<bar> = \<bar>of_int x\<bar>"  | 
| 62347 | 352  | 
by (auto simp add: abs_if)  | 
353  | 
||
354  | 
lemma of_int_lessD:  | 
|
355  | 
assumes "\<bar>of_int n\<bar> < x"  | 
|
356  | 
shows "n = 0 \<or> x > 1"  | 
|
357  | 
proof (cases "n = 0")  | 
|
| 63652 | 358  | 
case True  | 
359  | 
then show ?thesis by simp  | 
|
| 62347 | 360  | 
next  | 
361  | 
case False  | 
|
362  | 
then have "\<bar>n\<bar> \<noteq> 0" by simp  | 
|
363  | 
then have "\<bar>n\<bar> > 0" by simp  | 
|
364  | 
then have "\<bar>n\<bar> \<ge> 1"  | 
|
365  | 
using zless_imp_add1_zle [of 0 "\<bar>n\<bar>"] by simp  | 
|
366  | 
then have "\<bar>of_int n\<bar> \<ge> 1"  | 
|
367  | 
unfolding of_int_1_le_iff [of "\<bar>n\<bar>", symmetric] by simp  | 
|
368  | 
then have "1 < x" using assms by (rule le_less_trans)  | 
|
369  | 
then show ?thesis ..  | 
|
370  | 
qed  | 
|
371  | 
||
372  | 
lemma of_int_leD:  | 
|
373  | 
assumes "\<bar>of_int n\<bar> \<le> x"  | 
|
374  | 
shows "n = 0 \<or> 1 \<le> x"  | 
|
375  | 
proof (cases "n = 0")  | 
|
| 63652 | 376  | 
case True  | 
377  | 
then show ?thesis by simp  | 
|
| 62347 | 378  | 
next  | 
379  | 
case False  | 
|
380  | 
then have "\<bar>n\<bar> \<noteq> 0" by simp  | 
|
381  | 
then have "\<bar>n\<bar> > 0" by simp  | 
|
382  | 
then have "\<bar>n\<bar> \<ge> 1"  | 
|
383  | 
using zless_imp_add1_zle [of 0 "\<bar>n\<bar>"] by simp  | 
|
384  | 
then have "\<bar>of_int n\<bar> \<ge> 1"  | 
|
385  | 
unfolding of_int_1_le_iff [of "\<bar>n\<bar>", symmetric] by simp  | 
|
386  | 
then have "1 \<le> x" using assms by (rule order_trans)  | 
|
387  | 
then show ?thesis ..  | 
|
388  | 
qed  | 
|
389  | 
||
| 
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
 | 
390  | 
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
 | 
391  | 
"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
 | 
392  | 
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
 | 
393  | 
|
| 
 
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  | 
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
 | 
395  | 
"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
 | 
396  | 
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
 | 
397  | 
|
| 
 
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  | 
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
 | 
399  | 
"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
 | 
400  | 
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
 | 
401  | 
|
| 
 
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  | 
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
 | 
403  | 
"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
 | 
404  | 
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
 | 
405  | 
|
| 
 
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  | 
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
 | 
407  | 
"(- 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
 | 
408  | 
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
 | 
409  | 
|
| 
 
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  | 
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
 | 
411  | 
"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
 | 
412  | 
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
 | 
413  | 
|
| 
 
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  | 
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
 | 
415  | 
"(- 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
 | 
416  | 
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
 | 
417  | 
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
 | 
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_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
 | 
420  | 
"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
 | 
421  | 
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
 | 
422  | 
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
 | 
423  | 
|
| 
 
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  | 
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
 | 
425  | 
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
 | 
426  | 
|
| 
 
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
 | 
427  | 
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
 | 
428  | 
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
 | 
429  | 
|
| 
 
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
 | 
430  | 
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
 | 
431  | 
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
 | 
432  | 
|
| 
 
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
 | 
433  | 
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
 | 
434  | 
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
 | 
435  | 
|
| 
67969
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
436  | 
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
 | 
437  | 
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
 | 
438  | 
|
| 
 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
439  | 
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
 | 
440  | 
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
 | 
441  | 
|
| 36424 | 442  | 
end  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
443  | 
|
| 
69791
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
444  | 
context division_ring  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
445  | 
begin  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
446  | 
|
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
447  | 
lemmas mult_inverse_of_int_commute =  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
448  | 
mult_commute_imp_mult_inverse_commute[OF mult_of_int_commute]  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
449  | 
|
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
450  | 
end  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
451  | 
|
| 69593 | 452  | 
text \<open>Comparisons involving \<^term>\<open>of_int\<close>.\<close>  | 
| 61234 | 453  | 
|
| 63652 | 454  | 
lemma of_int_eq_numeral_iff [iff]: "of_int z = (numeral n :: 'a::ring_char_0) \<longleftrightarrow> z = numeral n"  | 
| 61234 | 455  | 
using of_int_eq_iff by fastforce  | 
456  | 
||
| 
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
 | 
457  | 
lemma of_int_le_numeral_iff [simp]:  | 
| 63652 | 458  | 
"of_int z \<le> (numeral n :: 'a::linordered_idom) \<longleftrightarrow> z \<le> numeral n"  | 
| 61234 | 459  | 
using of_int_le_iff [of z "numeral n"] by simp  | 
460  | 
||
| 
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
 | 
461  | 
lemma of_int_numeral_le_iff [simp]:  | 
| 63652 | 462  | 
"(numeral n :: 'a::linordered_idom) \<le> of_int z \<longleftrightarrow> numeral n \<le> z"  | 
| 61234 | 463  | 
using of_int_le_iff [of "numeral n"] by simp  | 
464  | 
||
| 
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
 | 
465  | 
lemma of_int_less_numeral_iff [simp]:  | 
| 63652 | 466  | 
"of_int z < (numeral n :: 'a::linordered_idom) \<longleftrightarrow> z < numeral n"  | 
| 61234 | 467  | 
using of_int_less_iff [of z "numeral n"] by simp  | 
468  | 
||
| 
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
 | 
469  | 
lemma of_int_numeral_less_iff [simp]:  | 
| 63652 | 470  | 
"(numeral n :: 'a::linordered_idom) < of_int z \<longleftrightarrow> numeral n < z"  | 
| 61234 | 471  | 
using of_int_less_iff [of "numeral n" z] by simp  | 
472  | 
||
| 63652 | 473  | 
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
 | 
474  | 
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
 | 
475  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
476  | 
lemma of_int_eq_id [simp]: "of_int = id"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
477  | 
proof  | 
| 63652 | 478  | 
show "of_int z = id z" for z  | 
479  | 
by (cases z rule: int_diff_cases) simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
480  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
481  | 
|
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
482  | 
instance int :: no_top  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
483  | 
proof  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
484  | 
fix x::int  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
485  | 
have "x < x + 1"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
486  | 
by simp  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
487  | 
then show "\<exists>y. x < y"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
488  | 
by (rule exI)  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
489  | 
qed  | 
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
490  | 
|
| 
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
491  | 
instance int :: no_bot  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
492  | 
proof  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
493  | 
fix x::int  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
494  | 
have "x - 1< x"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
495  | 
by simp  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
496  | 
then show "\<exists>y. y < x"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
497  | 
by (rule exI)  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
498  | 
qed  | 
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
51185 
diff
changeset
 | 
499  | 
|
| 63652 | 500  | 
|
| 61799 | 501  | 
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
 | 
502  | 
|
| 48045 | 503  | 
lift_definition nat :: "int \<Rightarrow> nat" is "\<lambda>(x, y). x - y"  | 
504  | 
by auto  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
505  | 
|
| 44709 | 506  | 
lemma nat_int [simp]: "nat (int n) = n"  | 
| 48045 | 507  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
508  | 
|
| 44709 | 509  | 
lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"  | 
| 48045 | 510  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
511  | 
|
| 63652 | 512  | 
lemma nat_0_le: "0 \<le> z \<Longrightarrow> int (nat z) = z"  | 
513  | 
by simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
514  | 
|
| 63652 | 515  | 
lemma nat_le_0 [simp]: "z \<le> 0 \<Longrightarrow> nat z = 0"  | 
| 48045 | 516  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
517  | 
|
| 63652 | 518  | 
lemma nat_le_eq_zle: "0 < w \<or> 0 \<le> z \<Longrightarrow> nat w \<le> nat z \<longleftrightarrow> w \<le> z"  | 
| 48045 | 519  | 
by transfer (clarsimp, arith)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
520  | 
|
| 69593 | 521  | 
text \<open>An alternative condition is \<^term>\<open>0 \<le> w\<close>.\<close>  | 
| 63652 | 522  | 
lemma nat_mono_iff: "0 < z \<Longrightarrow> nat w < nat z \<longleftrightarrow> w < z"  | 
523  | 
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
 | 
524  | 
|
| 63652 | 525  | 
lemma nat_less_eq_zless: "0 \<le> w \<Longrightarrow> nat w < nat z \<longleftrightarrow> w < z"  | 
526  | 
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
 | 
527  | 
|
| 63652 | 528  | 
lemma zless_nat_conj [simp]: "nat w < nat z \<longleftrightarrow> 0 < z \<and> w < z"  | 
| 48045 | 529  | 
by transfer (clarsimp, arith)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
530  | 
|
| 
64714
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
531  | 
lemma nonneg_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
532  | 
assumes "0 \<le> k"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
533  | 
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
 | 
534  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
535  | 
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
 | 
536  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
537  | 
then show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
538  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
539  | 
qed  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
540  | 
|
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
541  | 
lemma pos_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
542  | 
assumes "0 < k"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
543  | 
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
 | 
544  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
545  | 
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
 | 
546  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
547  | 
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
 | 
548  | 
by (rule nonneg_int_cases)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
549  | 
moreover have "n > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
550  | 
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
 | 
551  | 
ultimately show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
552  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
553  | 
qed  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
554  | 
|
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
555  | 
lemma nonpos_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
556  | 
assumes "k \<le> 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
557  | 
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
 | 
558  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
559  | 
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
 | 
560  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
561  | 
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
 | 
562  | 
by (rule nonneg_int_cases)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
563  | 
then have "k = - int n"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
564  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
565  | 
then show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
566  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
567  | 
qed  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
568  | 
|
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
569  | 
lemma neg_int_cases:  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
570  | 
assumes "k < 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
571  | 
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
 | 
572  | 
proof -  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
573  | 
from assms have "- k > 0"  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
574  | 
by simp  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
575  | 
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
 | 
576  | 
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
 | 
577  | 
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
 | 
578  | 
by simp_all  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
579  | 
then show thesis  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
580  | 
by (rule that)  | 
| 
 
53bab28983f1
complete set of cases rules for integers known to be (non-)positive/negative;
 
haftmann 
parents: 
64272 
diff
changeset
 | 
581  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
582  | 
|
| 63652 | 583  | 
lemma nat_eq_iff: "nat w = m \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"  | 
| 48045 | 584  | 
by transfer (clarsimp simp add: le_imp_diff_is_add)  | 
| 60162 | 585  | 
|
| 63652 | 586  | 
lemma nat_eq_iff2: "m = nat w \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"  | 
| 54223 | 587  | 
using nat_eq_iff [of w m] by auto  | 
588  | 
||
| 63652 | 589  | 
lemma nat_0 [simp]: "nat 0 = 0"  | 
| 54223 | 590  | 
by (simp add: nat_eq_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
591  | 
|
| 63652 | 592  | 
lemma nat_1 [simp]: "nat 1 = Suc 0"  | 
| 54223 | 593  | 
by (simp add: nat_eq_iff)  | 
594  | 
||
| 63652 | 595  | 
lemma nat_numeral [simp]: "nat (numeral k) = numeral k"  | 
| 54223 | 596  | 
by (simp add: nat_eq_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
597  | 
|
| 63652 | 598  | 
lemma nat_neg_numeral [simp]: "nat (- numeral k) = 0"  | 
| 54223 | 599  | 
by simp  | 
600  | 
||
601  | 
lemma nat_2: "nat 2 = Suc (Suc 0)"  | 
|
602  | 
by simp  | 
|
| 60162 | 603  | 
|
| 63652 | 604  | 
lemma nat_less_iff: "0 \<le> w \<Longrightarrow> nat w < m \<longleftrightarrow> w < of_nat m"  | 
| 48045 | 605  | 
by transfer (clarsimp, arith)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
606  | 
|
| 44709 | 607  | 
lemma nat_le_iff: "nat x \<le> n \<longleftrightarrow> x \<le> int n"  | 
| 48045 | 608  | 
by transfer (clarsimp simp add: le_diff_conv)  | 
| 44707 | 609  | 
|
610  | 
lemma nat_mono: "x \<le> y \<Longrightarrow> nat x \<le> nat y"  | 
|
| 48045 | 611  | 
by transfer auto  | 
| 44707 | 612  | 
|
| 63652 | 613  | 
lemma nat_0_iff[simp]: "nat i = 0 \<longleftrightarrow> i \<le> 0"  | 
614  | 
for i :: int  | 
|
| 48045 | 615  | 
by transfer clarsimp  | 
| 29700 | 616  | 
|
| 63652 | 617  | 
lemma int_eq_iff: "of_nat m = z \<longleftrightarrow> m = nat z \<and> 0 \<le> z"  | 
618  | 
by (auto simp add: nat_eq_iff2)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
619  | 
|
| 63652 | 620  | 
lemma zero_less_nat_eq [simp]: "0 < nat z \<longleftrightarrow> 0 < z"  | 
621  | 
using zless_nat_conj [of 0] by auto  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
622  | 
|
| 63652 | 623  | 
lemma nat_add_distrib: "0 \<le> z \<Longrightarrow> 0 \<le> z' \<Longrightarrow> nat (z + z') = nat z + nat z'"  | 
| 48045 | 624  | 
by transfer clarsimp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
625  | 
|
| 63652 | 626  | 
lemma nat_diff_distrib': "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> nat (x - y) = nat x - nat y"  | 
| 54223 | 627  | 
by transfer clarsimp  | 
| 60162 | 628  | 
|
| 63652 | 629  | 
lemma nat_diff_distrib: "0 \<le> z' \<Longrightarrow> z' \<le> z \<Longrightarrow> nat (z - z') = nat z - nat z'"  | 
| 54223 | 630  | 
by (rule nat_diff_distrib') auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
631  | 
|
| 44709 | 632  | 
lemma nat_zminus_int [simp]: "nat (- int n) = 0"  | 
| 48045 | 633  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
634  | 
|
| 63652 | 635  | 
lemma le_nat_iff: "k \<ge> 0 \<Longrightarrow> n \<le> nat k \<longleftrightarrow> int n \<le> k"  | 
| 53065 | 636  | 
by transfer auto  | 
| 60162 | 637  | 
|
| 63652 | 638  | 
lemma zless_nat_eq_int_zless: "m < nat z \<longleftrightarrow> int m < z"  | 
| 48045 | 639  | 
by transfer (clarsimp simp add: less_diff_conv)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
640  | 
|
| 63652 | 641  | 
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
 | 
642  | 
by transfer (clarsimp simp add: of_nat_diff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
643  | 
|
| 63652 | 644  | 
lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"  | 
| 54249 | 645  | 
by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)  | 
646  | 
||
| 66886 | 647  | 
lemma nat_abs_triangle_ineq:  | 
648  | 
"nat \<bar>k + l\<bar> \<le> nat \<bar>k\<bar> + nat \<bar>l\<bar>"  | 
|
649  | 
by (simp add: nat_add_distrib [symmetric] nat_le_eq_zle abs_triangle_ineq)  | 
|
650  | 
||
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
651  | 
lemma nat_of_bool [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
652  | 
"nat (of_bool P) = of_bool P"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
653  | 
by auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
654  | 
|
| 75878 | 655  | 
lemma split_nat [linarith_split]: "P (nat i) \<longleftrightarrow> ((\<forall>n. i = int n \<longrightarrow> P n) \<and> (i < 0 \<longrightarrow> P 0))"  | 
| 66836 | 656  | 
(is "?P = (?L \<and> ?R)")  | 
657  | 
for i :: int  | 
|
658  | 
proof (cases "i < 0")  | 
|
659  | 
case True  | 
|
660  | 
then show ?thesis  | 
|
661  | 
by auto  | 
|
662  | 
next  | 
|
663  | 
case False  | 
|
664  | 
have "?P = ?L"  | 
|
665  | 
proof  | 
|
666  | 
assume ?P  | 
|
667  | 
then show ?L using False by auto  | 
|
668  | 
next  | 
|
669  | 
assume ?L  | 
|
670  | 
moreover from False have "int (nat i) = i"  | 
|
671  | 
by (simp add: not_less)  | 
|
672  | 
ultimately show ?P  | 
|
673  | 
by simp  | 
|
674  | 
qed  | 
|
675  | 
with False show ?thesis by simp  | 
|
676  | 
qed  | 
|
677  | 
||
678  | 
lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))"  | 
|
679  | 
by (auto split: split_nat)  | 
|
680  | 
||
681  | 
lemma ex_nat: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> P (nat x))"  | 
|
682  | 
proof  | 
|
683  | 
assume "\<exists>x. P x"  | 
|
684  | 
then obtain x where "P x" ..  | 
|
685  | 
then have "int x \<ge> 0 \<and> P (nat (int x))" by simp  | 
|
686  | 
then show "\<exists>x\<ge>0. P (nat x)" ..  | 
|
687  | 
next  | 
|
688  | 
assume "\<exists>x\<ge>0. P (nat x)"  | 
|
689  | 
then show "\<exists>x. P x" by auto  | 
|
690  | 
qed  | 
|
691  | 
||
| 54249 | 692  | 
|
| 60758 | 693  | 
text \<open>For termination proofs:\<close>  | 
| 63652 | 694  | 
lemma measure_function_int[measure_function]: "is_measure (nat \<circ> abs)" ..  | 
| 29779 | 695  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
696  | 
|
| 69593 | 697  | 
subsection \<open>Lemmas about the Function \<^term>\<open>of_nat\<close> and Orderings\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
698  | 
|
| 61076 | 699  | 
lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)"  | 
| 63652 | 700  | 
by (simp add: order_less_le del: of_nat_Suc)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
701  | 
|
| 44709 | 702  | 
lemma negative_zless [iff]: "- (int (Suc n)) < int m"  | 
| 63652 | 703  | 
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
 | 
704  | 
|
| 44709 | 705  | 
lemma negative_zle_0: "- int n \<le> 0"  | 
| 63652 | 706  | 
by (simp add: minus_le_iff)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
707  | 
|
| 44709 | 708  | 
lemma negative_zle [iff]: "- int n \<le> int m"  | 
| 63652 | 709  | 
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
 | 
710  | 
|
| 63652 | 711  | 
lemma not_zle_0_negative [simp]: "\<not> 0 \<le> - int (Suc n)"  | 
712  | 
by (subst le_minus_iff) (simp del: of_nat_Suc)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
713  | 
|
| 63652 | 714  | 
lemma int_zle_neg: "int n \<le> - int m \<longleftrightarrow> n = 0 \<and> m = 0"  | 
| 48045 | 715  | 
by transfer simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
716  | 
|
| 63652 | 717  | 
lemma not_int_zless_negative [simp]: "\<not> int n < - int m"  | 
718  | 
by (simp add: linorder_not_less)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
719  | 
|
| 63652 | 720  | 
lemma negative_eq_positive [simp]: "- int n = of_nat m \<longleftrightarrow> n = 0 \<and> m = 0"  | 
721  | 
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
 | 
722  | 
|
| 63652 | 723  | 
lemma zle_iff_zadd: "w \<le> z \<longleftrightarrow> (\<exists>n. z = w + int n)"  | 
724  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 62348 | 725  | 
proof  | 
| 63652 | 726  | 
assume ?rhs  | 
727  | 
then show ?lhs by auto  | 
|
| 62348 | 728  | 
next  | 
| 63652 | 729  | 
assume ?lhs  | 
| 62348 | 730  | 
then have "0 \<le> z - w" by simp  | 
731  | 
then obtain n where "z - w = int n"  | 
|
732  | 
using zero_le_imp_eq_int [of "z - w"] by blast  | 
|
| 63652 | 733  | 
then have "z = w + int n" by simp  | 
734  | 
then show ?rhs ..  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
735  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
736  | 
|
| 44709 | 737  | 
lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z"  | 
| 63652 | 738  | 
by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
739  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
740  | 
lemma negD:  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
741  | 
assumes "x < 0" shows "\<exists>n. x = - (int (Suc n))"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
742  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
743  | 
have "\<And>a b. a < b \<Longrightarrow> \<exists>n. Suc (a + n) = b"  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
744  | 
proof -  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
745  | 
fix a b:: nat  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
746  | 
assume "a < b"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
747  | 
then have "Suc (a + (b - Suc a)) = b"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
748  | 
by arith  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
749  | 
then show "\<exists>n. Suc (a + n) = b"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
750  | 
by (rule exI)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
751  | 
qed  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
752  | 
with assms show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
753  | 
by transfer auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
754  | 
qed  | 
| 63652 | 755  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
756  | 
|
| 60758 | 757  | 
subsection \<open>Cases and induction\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
758  | 
|
| 63652 | 759  | 
text \<open>  | 
760  | 
Now we replace the case analysis rule by a more conventional one:  | 
|
761  | 
whether an integer is negative or not.  | 
|
762  | 
\<close>  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
763  | 
|
| 63652 | 764  | 
text \<open>This version is symmetric in the two subgoals.\<close>  | 
765  | 
lemma int_cases2 [case_names nonneg nonpos, cases type: int]:  | 
|
766  | 
"(\<And>n. z = int n \<Longrightarrow> P) \<Longrightarrow> (\<And>n. z = - (int n) \<Longrightarrow> P) \<Longrightarrow> P"  | 
|
767  | 
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
 | 
768  | 
|
| 63652 | 769  | 
text \<open>This is the default, with a negative case.\<close>  | 
770  | 
lemma int_cases [case_names nonneg neg, cases type: int]:  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
771  | 
assumes pos: "\<And>n. z = int n \<Longrightarrow> P" and neg: "\<And>n. z = - (int (Suc n)) \<Longrightarrow> P"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
772  | 
shows P  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
773  | 
proof (cases "z < 0")  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
774  | 
case True  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
775  | 
with neg show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
776  | 
by (blast dest!: negD)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
777  | 
next  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
778  | 
case False  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
779  | 
with pos show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
780  | 
by (force simp add: linorder_not_less dest: nat_0_le [THEN sym])  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
781  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
782  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
783  | 
lemma int_cases3 [case_names zero pos neg]:  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
784  | 
fixes k :: int  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
785  | 
assumes "k = 0 \<Longrightarrow> P" and "\<And>n. k = int n \<Longrightarrow> n > 0 \<Longrightarrow> P"  | 
| 61204 | 786  | 
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
 | 
787  | 
shows "P"  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
788  | 
proof (cases k "0::int" rule: linorder_cases)  | 
| 63652 | 789  | 
case equal  | 
790  | 
with assms(1) show P by simp  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
791  | 
next  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
792  | 
case greater  | 
| 63539 | 793  | 
then have *: "nat k > 0" by simp  | 
794  | 
moreover from * have "k = int (nat k)" by auto  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
795  | 
ultimately show P using assms(2) by blast  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
796  | 
next  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
797  | 
case less  | 
| 63539 | 798  | 
then have *: "nat (- k) > 0" by simp  | 
799  | 
moreover from * have "k = - int (nat (- k))" by auto  | 
|
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
800  | 
ultimately show P using assms(3) by blast  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
801  | 
qed  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
802  | 
|
| 63652 | 803  | 
lemma int_of_nat_induct [case_names nonneg neg, induct type: int]:  | 
804  | 
"(\<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
 | 
805  | 
by (cases z) auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
806  | 
|
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
807  | 
lemma sgn_mult_dvd_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
808  | 
"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
 | 
809  | 
by (cases r rule: int_cases3) auto  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
810  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
811  | 
lemma mult_sgn_dvd_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
812  | 
"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
 | 
813  | 
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
 | 
814  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
815  | 
lemma dvd_sgn_mult_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
816  | 
"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
 | 
817  | 
by (cases r rule: int_cases3) simp_all  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
818  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
819  | 
lemma dvd_mult_sgn_iff [simp]:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
820  | 
"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
 | 
821  | 
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
 | 
822  | 
|
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
823  | 
lemma int_sgnE:  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
824  | 
fixes k :: int  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
825  | 
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
 | 
826  | 
proof -  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
827  | 
have "k = sgn k * int (nat \<bar>k\<bar>)"  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
828  | 
by (simp add: sgn_mult_abs)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
829  | 
then show ?thesis ..  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
830  | 
qed  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
831  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
832  | 
|
| 60758 | 833  | 
subsubsection \<open>Binary comparisons\<close>  | 
| 28958 | 834  | 
|
| 60758 | 835  | 
text \<open>Preliminaries\<close>  | 
| 28958 | 836  | 
|
| 60162 | 837  | 
lemma le_imp_0_less:  | 
| 63652 | 838  | 
fixes z :: int  | 
| 28958 | 839  | 
assumes le: "0 \<le> z"  | 
| 63652 | 840  | 
shows "0 < 1 + z"  | 
| 28958 | 841  | 
proof -  | 
842  | 
have "0 \<le> z" by fact  | 
|
| 63652 | 843  | 
also have "\<dots> < z + 1" by (rule less_add_one)  | 
844  | 
also have "\<dots> = 1 + z" by (simp add: ac_simps)  | 
|
| 28958 | 845  | 
finally show "0 < 1 + z" .  | 
846  | 
qed  | 
|
847  | 
||
| 63652 | 848  | 
lemma odd_less_0_iff: "1 + z + z < 0 \<longleftrightarrow> z < 0"  | 
849  | 
for z :: int  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
850  | 
proof (cases z)  | 
| 28958 | 851  | 
case (nonneg n)  | 
| 63652 | 852  | 
then show ?thesis  | 
853  | 
by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le])  | 
|
| 28958 | 854  | 
next  | 
855  | 
case (neg n)  | 
|
| 63652 | 856  | 
then show ?thesis  | 
857  | 
by (simp del: of_nat_Suc of_nat_add of_nat_1  | 
|
858  | 
add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric])  | 
|
| 28958 | 859  | 
qed  | 
860  | 
||
| 63652 | 861  | 
|
| 60758 | 862  | 
subsubsection \<open>Comparisons, for Ordered Rings\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
863  | 
|
| 63652 | 864  | 
lemma odd_nonzero: "1 + z + z \<noteq> 0"  | 
865  | 
for z :: int  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
866  | 
proof (cases z)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
867  | 
case (nonneg n)  | 
| 63652 | 868  | 
have le: "0 \<le> z + z"  | 
869  | 
by (simp add: nonneg add_increasing)  | 
|
870  | 
then show ?thesis  | 
|
| 67116 | 871  | 
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
 | 
872  | 
next  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
873  | 
case (neg n)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
874  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
875  | 
proof  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
876  | 
assume eq: "1 + z + z = 0"  | 
| 63652 | 877  | 
have "0 < 1 + (int n + int n)"  | 
| 60162 | 878  | 
by (simp add: le_imp_0_less add_increasing)  | 
| 63652 | 879  | 
also have "\<dots> = - (1 + z + z)"  | 
| 60162 | 880  | 
by (simp add: neg add.assoc [symmetric])  | 
| 63652 | 881  | 
also have "\<dots> = 0" by (simp add: eq)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
882  | 
finally have "0<0" ..  | 
| 63652 | 883  | 
then show False by blast  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
884  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
885  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
886  | 
|
| 
30652
 
752329615264
distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
 
haftmann 
parents: 
30496 
diff
changeset
 | 
887  | 
|
| 60758 | 888  | 
subsection \<open>The Set of Integers\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
889  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
890  | 
context ring_1  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
891  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
892  | 
|
| 61070 | 893  | 
definition Ints :: "'a set"  ("\<int>")
 | 
894  | 
where "\<int> = range of_int"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
895  | 
|
| 35634 | 896  | 
lemma Ints_of_int [simp]: "of_int z \<in> \<int>"  | 
897  | 
by (simp add: Ints_def)  | 
|
898  | 
||
899  | 
lemma Ints_of_nat [simp]: "of_nat n \<in> \<int>"  | 
|
| 45533 | 900  | 
using Ints_of_int [of "of_nat n"] by simp  | 
| 35634 | 901  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
902  | 
lemma Ints_0 [simp]: "0 \<in> \<int>"  | 
| 45533 | 903  | 
using Ints_of_int [of "0"] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
904  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
905  | 
lemma Ints_1 [simp]: "1 \<in> \<int>"  | 
| 45533 | 906  | 
using Ints_of_int [of "1"] by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
907  | 
|
| 
61552
 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 
eberlm 
parents: 
61531 
diff
changeset
 | 
908  | 
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
 | 
909  | 
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
 | 
910  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
911  | 
lemma Ints_add [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a + b \<in> \<int>"  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
912  | 
by (force simp add: Ints_def simp flip: of_int_add intro: range_eqI)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
913  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
914  | 
lemma Ints_minus [simp]: "a \<in> \<int> \<Longrightarrow> -a \<in> \<int>"  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
915  | 
by (force simp add: Ints_def simp flip: of_int_minus intro: range_eqI)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
916  | 
|
| 68721 | 917  | 
lemma minus_in_Ints_iff: "-x \<in> \<int> \<longleftrightarrow> x \<in> \<int>"  | 
918  | 
using Ints_minus[of x] Ints_minus[of "-x"] by auto  | 
|
919  | 
||
| 35634 | 920  | 
lemma Ints_diff [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a - b \<in> \<int>"  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
921  | 
by (force simp add: Ints_def simp flip: of_int_diff intro: range_eqI)  | 
| 35634 | 922  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
923  | 
lemma Ints_mult [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a * b \<in> \<int>"  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
924  | 
by (force simp add: Ints_def simp flip: of_int_mult intro: range_eqI)  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
925  | 
|
| 35634 | 926  | 
lemma Ints_power [simp]: "a \<in> \<int> \<Longrightarrow> a ^ n \<in> \<int>"  | 
| 63652 | 927  | 
by (induct n) simp_all  | 
| 35634 | 928  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
929  | 
lemma Ints_cases [cases set: Ints]:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
930  | 
assumes "q \<in> \<int>"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
931  | 
obtains (of_int) z where "q = of_int z"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
932  | 
unfolding Ints_def  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
933  | 
proof -  | 
| 60758 | 934  | 
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
 | 
935  | 
then obtain z where "q = of_int z" ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
936  | 
then show thesis ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
937  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
938  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
939  | 
lemma Ints_induct [case_names of_int, induct set: Ints]:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
940  | 
"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
 | 
941  | 
by (rule Ints_cases) auto  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
942  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
943  | 
lemma Nats_subset_Ints: "\<nat> \<subseteq> \<int>"  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
944  | 
unfolding Nats_def Ints_def  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
945  | 
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
 | 
946  | 
|
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
947  | 
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
 | 
948  | 
proof (intro subsetI equalityI)  | 
| 63652 | 949  | 
fix x :: 'a  | 
950  | 
  assume "x \<in> {of_int n |n. n \<ge> 0}"
 | 
|
951  | 
then obtain n where "x = of_int n" "n \<ge> 0"  | 
|
952  | 
by (auto elim!: Ints_cases)  | 
|
953  | 
then have "x = of_nat (nat n)"  | 
|
954  | 
by (subst of_nat_nat) simp_all  | 
|
955  | 
then show "x \<in> \<nat>"  | 
|
956  | 
by simp  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
957  | 
next  | 
| 63652 | 958  | 
fix x :: 'a  | 
959  | 
assume "x \<in> \<nat>"  | 
|
960  | 
then obtain n where "x = of_nat n"  | 
|
961  | 
by (auto elim!: Nats_cases)  | 
|
962  | 
then have "x = of_int (int n)" by simp  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
963  | 
also have "int n \<ge> 0" by simp  | 
| 63652 | 964  | 
  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
 | 
965  | 
  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
 | 
966  | 
qed  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
967  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
968  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
969  | 
|
| 
73109
 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72512 
diff
changeset
 | 
970  | 
lemma Ints_sum [intro]: "(\<And>x. x \<in> A \<Longrightarrow> f x \<in> \<int>) \<Longrightarrow> sum f A \<in> \<int>"  | 
| 
 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72512 
diff
changeset
 | 
971  | 
by (induction A rule: infinite_finite_induct) auto  | 
| 
 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72512 
diff
changeset
 | 
972  | 
|
| 
 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72512 
diff
changeset
 | 
973  | 
lemma Ints_prod [intro]: "(\<And>x. x \<in> A \<Longrightarrow> f x \<in> \<int>) \<Longrightarrow> prod f A \<in> \<int>"  | 
| 
 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72512 
diff
changeset
 | 
974  | 
by (induction A rule: infinite_finite_induct) auto  | 
| 
 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72512 
diff
changeset
 | 
975  | 
|
| 
64758
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
976  | 
lemma (in linordered_idom) Ints_abs [simp]:  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
977  | 
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
 | 
978  | 
by (auto simp: abs_if)  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
979  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
980  | 
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
 | 
981  | 
proof (intro subsetI equalityI)  | 
| 63652 | 982  | 
fix x :: 'a  | 
983  | 
  assume "x \<in> {n \<in> \<int>. n \<ge> 0}"
 | 
|
984  | 
then obtain n where "x = of_int n" "n \<ge> 0"  | 
|
985  | 
by (auto elim!: Ints_cases)  | 
|
986  | 
then have "x = of_nat (nat n)"  | 
|
987  | 
by (subst of_nat_nat) simp_all  | 
|
988  | 
then show "x \<in> \<nat>"  | 
|
989  | 
by simp  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
990  | 
qed (auto elim!: Nats_cases)  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
991  | 
|
| 64849 | 992  | 
lemma (in idom_divide) of_int_divide_in_Ints:  | 
993  | 
"of_int a div of_int b \<in> \<int>" if "b dvd a"  | 
|
994  | 
proof -  | 
|
995  | 
from that obtain c where "a = b * c" ..  | 
|
996  | 
then show ?thesis  | 
|
997  | 
by (cases "of_int b = 0") simp_all  | 
|
998  | 
qed  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61234 
diff
changeset
 | 
999  | 
|
| 69593 | 1000  | 
text \<open>The premise involving \<^term>\<open>Ints\<close> prevents \<^term>\<open>a = 1/2\<close>.\<close>  | 
| 
25919
 
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_double_eq_0_iff:  | 
| 63652 | 1003  | 
fixes a :: "'a::ring_char_0"  | 
| 61070 | 1004  | 
assumes in_Ints: "a \<in> \<int>"  | 
| 63652 | 1005  | 
shows "a + a = 0 \<longleftrightarrow> a = 0"  | 
1006  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1007  | 
proof -  | 
| 63652 | 1008  | 
from in_Ints have "a \<in> range of_int"  | 
1009  | 
unfolding Ints_def [symmetric] .  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1010  | 
then obtain z where a: "a = of_int z" ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1011  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1012  | 
proof  | 
| 63652 | 1013  | 
assume ?rhs  | 
1014  | 
then show ?lhs by simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1015  | 
next  | 
| 63652 | 1016  | 
assume ?lhs  | 
1017  | 
with a have "of_int (z + z) = (of_int 0 :: 'a)" by simp  | 
|
1018  | 
then have "z + z = 0" by (simp only: of_int_eq_iff)  | 
|
| 67116 | 1019  | 
then have "z = 0" by (simp only: double_zero)  | 
| 63652 | 1020  | 
with a show ?rhs by simp  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1021  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1022  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1023  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1024  | 
lemma Ints_odd_nonzero:  | 
| 63652 | 1025  | 
fixes a :: "'a::ring_char_0"  | 
| 61070 | 1026  | 
assumes in_Ints: "a \<in> \<int>"  | 
| 63652 | 1027  | 
shows "1 + a + a \<noteq> 0"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1028  | 
proof -  | 
| 63652 | 1029  | 
from in_Ints have "a \<in> range of_int"  | 
1030  | 
unfolding Ints_def [symmetric] .  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1031  | 
then obtain z where a: "a = of_int z" ..  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1032  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1033  | 
proof  | 
| 63652 | 1034  | 
assume "1 + a + a = 0"  | 
1035  | 
with a have "of_int (1 + z + z) = (of_int 0 :: 'a)" by simp  | 
|
1036  | 
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
 | 
1037  | 
with odd_nonzero show False by blast  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1038  | 
qed  | 
| 60162 | 1039  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1040  | 
|
| 61070 | 1041  | 
lemma Nats_numeral [simp]: "numeral w \<in> \<nat>"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
1042  | 
using of_nat_in_Nats [of "numeral w"] by simp  | 
| 35634 | 1043  | 
|
| 60162 | 1044  | 
lemma Ints_odd_less_0:  | 
| 63652 | 1045  | 
fixes a :: "'a::linordered_idom"  | 
| 61070 | 1046  | 
assumes in_Ints: "a \<in> \<int>"  | 
| 63652 | 1047  | 
shows "1 + a + a < 0 \<longleftrightarrow> a < 0"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1048  | 
proof -  | 
| 63652 | 1049  | 
from in_Ints have "a \<in> range of_int"  | 
1050  | 
unfolding Ints_def [symmetric] .  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1051  | 
then obtain z where a: "a = of_int z" ..  | 
| 63652 | 1052  | 
with a have "1 + a + a < 0 \<longleftrightarrow> of_int (1 + z + z) < (of_int 0 :: 'a)"  | 
1053  | 
by simp  | 
|
1054  | 
also have "\<dots> \<longleftrightarrow> z < 0"  | 
|
1055  | 
by (simp only: of_int_less_iff odd_less_0_iff)  | 
|
1056  | 
also have "\<dots> \<longleftrightarrow> a < 0"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1057  | 
by (simp add: a)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1058  | 
finally show ?thesis .  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1059  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1060  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1061  | 
|
| 69593 | 1062  | 
subsection \<open>\<^term>\<open>sum\<close> and \<^term>\<open>prod\<close>\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1063  | 
|
| 69182 | 1064  | 
context semiring_1  | 
1065  | 
begin  | 
|
1066  | 
||
1067  | 
lemma of_nat_sum [simp]:  | 
|
1068  | 
"of_nat (sum f A) = (\<Sum>x\<in>A. of_nat (f x))"  | 
|
1069  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1070  | 
||
1071  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1072  | 
|
| 69182 | 1073  | 
context ring_1  | 
1074  | 
begin  | 
|
1075  | 
||
1076  | 
lemma of_int_sum [simp]:  | 
|
1077  | 
"of_int (sum f A) = (\<Sum>x\<in>A. of_int (f x))"  | 
|
1078  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1079  | 
||
1080  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1081  | 
|
| 69182 | 1082  | 
context comm_semiring_1  | 
1083  | 
begin  | 
|
1084  | 
||
1085  | 
lemma of_nat_prod [simp]:  | 
|
1086  | 
"of_nat (prod f A) = (\<Prod>x\<in>A. of_nat (f x))"  | 
|
1087  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1088  | 
||
1089  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1090  | 
|
| 69182 | 1091  | 
context comm_ring_1  | 
1092  | 
begin  | 
|
1093  | 
||
1094  | 
lemma of_int_prod [simp]:  | 
|
1095  | 
"of_int (prod f A) = (\<Prod>x\<in>A. of_int (f x))"  | 
|
1096  | 
by (induction A rule: infinite_finite_induct) auto  | 
|
1097  | 
||
1098  | 
end  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1099  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1100  | 
|
| 60758 | 1101  | 
subsection \<open>Setting up simplification procedures\<close>  | 
| 30802 | 1102  | 
|
| 
70356
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1103  | 
ML_file \<open>Tools/int_arith.ML\<close>  | 
| 54249 | 1104  | 
|
| 
70356
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1105  | 
declaration \<open>K (  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1106  | 
Lin_Arith.add_discrete_type \<^type_name>\<open>Int.int\<close>  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1107  | 
  #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle}
 | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1108  | 
  #> Lin_Arith.add_inj_thms @{thms of_nat_le_iff [THEN iffD2] of_nat_eq_iff [THEN iffD2]}
 | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1109  | 
#> Lin_Arith.add_inj_const (\<^const_name>\<open>of_nat\<close>, \<^typ>\<open>nat \<Rightarrow> int\<close>)  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1110  | 
#> Lin_Arith.add_simps  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1111  | 
      @{thms of_int_0 of_int_1 of_int_add of_int_mult of_int_numeral of_int_neg_numeral nat_0 nat_1 diff_nat_numeral nat_numeral
 | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1112  | 
neg_less_iff_less  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1113  | 
True_implies_equals  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1114  | 
distrib_left [where a = "numeral v" for v]  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1115  | 
distrib_left [where a = "- numeral v" for v]  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1116  | 
div_by_1 div_0  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1117  | 
times_divide_eq_right times_divide_eq_left  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1118  | 
minus_divide_left [THEN sym] minus_divide_right [THEN sym]  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1119  | 
add_divide_distrib diff_divide_distrib  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1120  | 
of_int_minus of_int_diff  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1121  | 
of_int_of_nat_eq}  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1122  | 
#> Lin_Arith.add_simprocs [Int_Arith.zero_one_idom_simproc]  | 
| 
 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 
haftmann 
parents: 
70354 
diff
changeset
 | 
1123  | 
)\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1124  | 
|
| 63652 | 1125  | 
simproc_setup fast_arith  | 
1126  | 
  ("(m::'a::linordered_idom) < n" |
 | 
|
1127  | 
"(m::'a::linordered_idom) \<le> n" |  | 
|
1128  | 
"(m::'a::linordered_idom) = n") =  | 
|
| 61144 | 1129  | 
\<open>K Lin_Arith.simproc\<close>  | 
| 43595 | 1130  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1131  | 
|
| 60758 | 1132  | 
subsection\<open>More Inequality Reasoning\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1133  | 
|
| 63652 | 1134  | 
lemma zless_add1_eq: "w < z + 1 \<longleftrightarrow> w < z \<or> w = z"  | 
1135  | 
for w z :: int  | 
|
1136  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1137  | 
|
| 63652 | 1138  | 
lemma add1_zle_eq: "w + 1 \<le> z \<longleftrightarrow> w < z"  | 
1139  | 
for w z :: int  | 
|
1140  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1141  | 
|
| 63652 | 1142  | 
lemma zle_diff1_eq [simp]: "w \<le> z - 1 \<longleftrightarrow> w < z"  | 
1143  | 
for w z :: int  | 
|
1144  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1145  | 
|
| 63652 | 1146  | 
lemma zle_add1_eq_le [simp]: "w < z + 1 \<longleftrightarrow> w \<le> z"  | 
1147  | 
for w z :: int  | 
|
1148  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1149  | 
|
| 63652 | 1150  | 
lemma int_one_le_iff_zero_less: "1 \<le> z \<longleftrightarrow> 0 < z"  | 
1151  | 
for z :: int  | 
|
1152  | 
by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1153  | 
|
| 
64758
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1154  | 
lemma Ints_nonzero_abs_ge1:  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1155  | 
fixes x:: "'a :: linordered_idom"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1156  | 
assumes "x \<in> Ints" "x \<noteq> 0"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1157  | 
shows "1 \<le> abs x"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1158  | 
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
 | 
1159  | 
fix z::int  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1160  | 
assume "x = of_int z"  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
1161  | 
with \<open>x \<noteq> 0\<close>  | 
| 
64758
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1162  | 
show "1 \<le> \<bar>x\<bar>"  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
1163  | 
apply (auto simp: abs_if)  | 
| 
64758
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1164  | 
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
 | 
1165  | 
qed  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1166  | 
|
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1167  | 
lemma Ints_nonzero_abs_less1:  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1168  | 
fixes x:: "'a :: linordered_idom"  | 
| 
 
3b33d2fc5fc0
A few new lemmas and needed adaptations
 
paulson <lp15@cam.ac.uk> 
parents: 
64714 
diff
changeset
 | 
1169  | 
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
 | 
1170  | 
using Ints_nonzero_abs_ge1 [of x] by auto  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70356 
diff
changeset
 | 
1171  | 
|
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70356 
diff
changeset
 | 
1172  | 
lemma Ints_eq_abs_less1:  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70356 
diff
changeset
 | 
1173  | 
fixes x:: "'a :: linordered_idom"  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70356 
diff
changeset
 | 
1174  | 
shows "\<lbrakk>x \<in> Ints; y \<in> Ints\<rbrakk> \<Longrightarrow> x = y \<longleftrightarrow> abs (x-y) < 1"  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70356 
diff
changeset
 | 
1175  | 
using eq_iff_diff_eq_0 by (fastforce intro: Ints_nonzero_abs_less1)  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70356 
diff
changeset
 | 
1176  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1177  | 
|
| 69593 | 1178  | 
subsection \<open>The functions \<^term>\<open>nat\<close> and \<^term>\<open>int\<close>\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1179  | 
|
| 69593 | 1180  | 
text \<open>Simplify the term \<^term>\<open>w + - z\<close>.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1181  | 
|
| 63652 | 1182  | 
lemma one_less_nat_eq [simp]: "Suc 0 < nat z \<longleftrightarrow> 1 < z"  | 
| 60162 | 1183  | 
using zless_nat_conj [of 1 z] by auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1184  | 
|
| 67116 | 1185  | 
lemma int_eq_iff_numeral [simp]:  | 
1186  | 
"int m = numeral v \<longleftrightarrow> m = numeral v"  | 
|
1187  | 
by (simp add: int_eq_iff)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1188  | 
|
| 67116 | 1189  | 
lemma nat_abs_int_diff:  | 
1190  | 
"nat \<bar>int a - int b\<bar> = (if a \<le> b then b - a else a - b)"  | 
|
| 59000 | 1191  | 
by auto  | 
1192  | 
||
1193  | 
lemma nat_int_add: "nat (int a + int b) = a + b"  | 
|
1194  | 
by auto  | 
|
1195  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1196  | 
context ring_1  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1197  | 
begin  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1198  | 
|
| 
33056
 
791a4655cae3
renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
 
blanchet 
parents: 
32437 
diff
changeset
 | 
1199  | 
lemma of_int_of_nat [nitpick_simp]:  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1200  | 
"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
 | 
1201  | 
proof (cases "k < 0")  | 
| 63652 | 1202  | 
case True  | 
1203  | 
then have "0 \<le> - k" by simp  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1204  | 
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
 | 
1205  | 
with True show ?thesis by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1206  | 
next  | 
| 63652 | 1207  | 
case False  | 
1208  | 
then show ?thesis by (simp add: not_less)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1209  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1210  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1211  | 
end  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1212  | 
|
| 64014 | 1213  | 
lemma transfer_rule_of_int:  | 
| 70927 | 1214  | 
includes lifting_syntax  | 
| 64014 | 1215  | 
fixes R :: "'a::ring_1 \<Rightarrow> 'b::ring_1 \<Rightarrow> bool"  | 
1216  | 
assumes [transfer_rule]: "R 0 0" "R 1 1"  | 
|
| 70927 | 1217  | 
"(R ===> R ===> R) (+) (+)"  | 
1218  | 
"(R ===> R) uminus uminus"  | 
|
1219  | 
shows "((=) ===> R) of_int of_int"  | 
|
| 64014 | 1220  | 
proof -  | 
| 70927 | 1221  | 
note assms  | 
| 64014 | 1222  | 
note transfer_rule_of_nat [transfer_rule]  | 
| 70927 | 1223  | 
have [transfer_rule]: "((=) ===> R) of_nat of_nat"  | 
| 64014 | 1224  | 
by transfer_prover  | 
1225  | 
show ?thesis  | 
|
1226  | 
by (unfold of_int_of_nat [abs_def]) transfer_prover  | 
|
1227  | 
qed  | 
|
1228  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1229  | 
lemma nat_mult_distrib:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1230  | 
fixes z z' :: int  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1231  | 
assumes "0 \<le> z"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1232  | 
shows "nat (z * z') = nat z * nat z'"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1233  | 
proof (cases "0 \<le> z'")  | 
| 63652 | 1234  | 
case False  | 
1235  | 
with assms have "z * z' \<le> 0"  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1236  | 
by (simp add: not_le mult_le_0_iff)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1237  | 
then have "nat (z * z') = 0" by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1238  | 
moreover from False have "nat z' = 0" by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1239  | 
ultimately show ?thesis by simp  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1240  | 
next  | 
| 63652 | 1241  | 
case True  | 
1242  | 
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
 | 
1243  | 
show ?thesis  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1244  | 
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
 | 
1245  | 
(simp only: of_nat_mult of_nat_nat [OF True]  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1246  | 
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
 | 
1247  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1248  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1249  | 
lemma nat_mult_distrib_neg:  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1250  | 
assumes "z \<le> (0::int)" shows "nat (z * z') = nat (- z) * nat (- z')" (is "?L = ?R")  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1251  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1252  | 
have "?L = nat (- z * - z')"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1253  | 
using assms by auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1254  | 
also have "... = ?R"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1255  | 
by (rule nat_mult_distrib) (use assms in auto)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1256  | 
finally show ?thesis .  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1257  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1258  | 
|
| 61944 | 1259  | 
lemma nat_abs_mult_distrib: "nat \<bar>w * z\<bar> = nat \<bar>w\<bar> * nat \<bar>z\<bar>"  | 
| 63652 | 1260  | 
by (cases "z = 0 \<or> w = 0")  | 
1261  | 
(auto simp add: abs_if nat_mult_distrib [symmetric]  | 
|
1262  | 
nat_mult_distrib_neg [symmetric] mult_less_0_iff)  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1263  | 
|
| 63652 | 1264  | 
lemma int_in_range_abs [simp]: "int n \<in> range abs"  | 
| 60570 | 1265  | 
proof (rule range_eqI)  | 
| 63652 | 1266  | 
show "int n = \<bar>int n\<bar>" by simp  | 
| 60570 | 1267  | 
qed  | 
1268  | 
||
| 63652 | 1269  | 
lemma range_abs_Nats [simp]: "range abs = (\<nat> :: int set)"  | 
| 60570 | 1270  | 
proof -  | 
1271  | 
have "\<bar>k\<bar> \<in> \<nat>" for k :: int  | 
|
1272  | 
by (cases k) simp_all  | 
|
1273  | 
moreover have "k \<in> range abs" if "k \<in> \<nat>" for k :: int  | 
|
1274  | 
using that by induct simp  | 
|
1275  | 
ultimately show ?thesis by blast  | 
|
| 61204 | 1276  | 
qed  | 
| 60570 | 1277  | 
|
| 63652 | 1278  | 
lemma Suc_nat_eq_nat_zadd1: "0 \<le> z \<Longrightarrow> Suc (nat z) = nat (1 + z)"  | 
1279  | 
for z :: int  | 
|
1280  | 
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
 | 
1281  | 
|
| 
 
9368aa814518
move lemmas from Nat_Numeral to Int.thy and Num.thy
 
huffman 
parents: 
47192 
diff
changeset
 | 
1282  | 
lemma diff_nat_eq_if:  | 
| 63652 | 1283  | 
"nat z - nat z' =  | 
1284  | 
(if z' < 0 then nat z  | 
|
1285  | 
else  | 
|
1286  | 
let d = z - z'  | 
|
1287  | 
in if d < 0 then 0 else nat d)"  | 
|
1288  | 
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
 | 
1289  | 
|
| 63652 | 1290  | 
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
 | 
1291  | 
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
 | 
1292  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1293  | 
|
| 63652 | 1294  | 
subsection \<open>Induction principles for int\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1295  | 
|
| 63652 | 1296  | 
text \<open>Well-founded segments of the integers.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1297  | 
|
| 63652 | 1298  | 
definition int_ge_less_than :: "int \<Rightarrow> (int \<times> int) set"  | 
1299  | 
  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
 | 
1300  | 
|
| 63652 | 1301  | 
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
 | 
1302  | 
proof -  | 
| 63652 | 1303  | 
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
 | 
1304  | 
by (auto simp add: int_ge_less_than_def)  | 
| 63652 | 1305  | 
then show ?thesis  | 
| 60162 | 1306  | 
by (rule wf_subset [OF wf_measure])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1307  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1308  | 
|
| 63652 | 1309  | 
text \<open>  | 
1310  | 
This variant looks odd, but is typical of the relations suggested  | 
|
1311  | 
by RankFinder.\<close>  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1312  | 
|
| 63652 | 1313  | 
definition int_ge_less_than2 :: "int \<Rightarrow> (int \<times> int) set"  | 
1314  | 
  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
 | 
1315  | 
|
| 63652 | 1316  | 
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
 | 
1317  | 
proof -  | 
| 63652 | 1318  | 
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
 | 
1319  | 
by (auto simp add: int_ge_less_than2_def)  | 
| 63652 | 1320  | 
then show ?thesis  | 
| 60162 | 1321  | 
by (rule wf_subset [OF wf_measure])  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1322  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1323  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1324  | 
(* `set:int': dummy construction *)  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1325  | 
theorem int_ge_induct [case_names base step, induct set: int]:  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1326  | 
fixes i :: int  | 
| 63652 | 1327  | 
assumes ge: "k \<le> i"  | 
1328  | 
and base: "P k"  | 
|
1329  | 
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
 | 
1330  | 
shows "P i"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1331  | 
proof -  | 
| 63652 | 1332  | 
have "\<And>i::int. n = nat (i - k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i" for n  | 
1333  | 
proof (induct n)  | 
|
1334  | 
case 0  | 
|
1335  | 
then have "i = k" by arith  | 
|
1336  | 
with base show "P i" by simp  | 
|
1337  | 
next  | 
|
1338  | 
case (Suc n)  | 
|
1339  | 
then have "n = nat ((i - 1) - k)" by arith  | 
|
1340  | 
moreover have k: "k \<le> i - 1" using Suc.prems by arith  | 
|
1341  | 
ultimately have "P (i - 1)" by (rule Suc.hyps)  | 
|
1342  | 
from step [OF k this] show ?case by simp  | 
|
1343  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1344  | 
with ge show ?thesis by fast  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1345  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1346  | 
|
| 25928 | 1347  | 
(* `set:int': dummy construction *)  | 
1348  | 
theorem int_gr_induct [case_names base step, induct set: int]:  | 
|
| 63652 | 1349  | 
fixes i k :: int  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1350  | 
assumes "k < i" "P (k + 1)" "\<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
 | 
1351  | 
shows "P i"  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1352  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1353  | 
have "k+1 \<le> i"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1354  | 
using assms by auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1355  | 
then show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1356  | 
by (induction i rule: int_ge_induct) (auto simp: assms)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1357  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1358  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1359  | 
theorem int_le_induct [consumes 1, case_names base step]:  | 
| 63652 | 1360  | 
fixes i k :: int  | 
1361  | 
assumes le: "i \<le> k"  | 
|
1362  | 
and base: "P k"  | 
|
1363  | 
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
 | 
1364  | 
shows "P i"  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1365  | 
proof -  | 
| 63652 | 1366  | 
have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i" for n  | 
1367  | 
proof (induct n)  | 
|
1368  | 
case 0  | 
|
1369  | 
then have "i = k" by arith  | 
|
1370  | 
with base show "P i" by simp  | 
|
1371  | 
next  | 
|
1372  | 
case (Suc n)  | 
|
1373  | 
then have "n = nat (k - (i + 1))" by arith  | 
|
1374  | 
moreover have k: "i + 1 \<le> k" using Suc.prems by arith  | 
|
1375  | 
ultimately have "P (i + 1)" by (rule Suc.hyps)  | 
|
1376  | 
from step[OF k this] show ?case by simp  | 
|
1377  | 
qed  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1378  | 
with le show ?thesis by fast  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1379  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1380  | 
|
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1381  | 
theorem int_less_induct [consumes 1, case_names base step]:  | 
| 63652 | 1382  | 
fixes i k :: int  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1383  | 
assumes "i < k" "P (k - 1)" "\<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
 | 
1384  | 
shows "P i"  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1385  | 
proof -  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1386  | 
have "i \<le> k-1"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1387  | 
using assms by auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1388  | 
then show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1389  | 
by (induction i rule: int_le_induct) (auto simp: assms)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1390  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1391  | 
|
| 
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
 | 
1392  | 
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
 | 
1393  | 
fixes k :: int  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1394  | 
assumes base: "P k"  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1395  | 
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
 | 
1396  | 
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
 | 
1397  | 
shows "P i"  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1398  | 
proof -  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1399  | 
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
 | 
1400  | 
then show ?thesis  | 
| 
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1401  | 
proof  | 
| 
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1402  | 
assume "i \<ge> k"  | 
| 63652 | 1403  | 
then show ?thesis  | 
1404  | 
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
 | 
1405  | 
next  | 
| 
42676
 
8724f20bf69c
proper case_names for int_cases, int_of_nat_induct;
 
wenzelm 
parents: 
42411 
diff
changeset
 | 
1406  | 
assume "i \<le> k"  | 
| 63652 | 1407  | 
then show ?thesis  | 
1408  | 
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
 | 
1409  | 
qed  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1410  | 
qed  | 
| 
 
3560de0fe851
moved int induction lemma to theory Int as int_bidirectional_induct
 
haftmann 
parents: 
36749 
diff
changeset
 | 
1411  | 
|
| 63652 | 1412  | 
|
1413  | 
subsection \<open>Intermediate value theorems\<close>  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1414  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1415  | 
lemma nat_ivt_aux:  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1416  | 
"\<lbrakk>\<forall>i<n. \<bar>f (Suc i) - f i\<bar> \<le> 1; f 0 \<le> k; k \<le> f n\<rbrakk> \<Longrightarrow> \<exists>i \<le> n. f i = k"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1417  | 
for m n :: nat and k :: int  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1418  | 
proof (induct n)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1419  | 
case (Suc n)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1420  | 
show ?case  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1421  | 
proof (cases "k = f (Suc n)")  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1422  | 
case False  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1423  | 
with Suc have "k \<le> f n"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1424  | 
by auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1425  | 
with Suc show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1426  | 
by (auto simp add: abs_if split: if_split_asm intro: le_SucI)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1427  | 
qed (use Suc in auto)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1428  | 
qed auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1429  | 
|
| 67116 | 1430  | 
lemma nat_intermed_int_val:  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1431  | 
fixes m n :: nat and k :: int  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1432  | 
assumes "\<forall>i. m \<le> i \<and> i < n \<longrightarrow> \<bar>f (Suc i) - f i\<bar> \<le> 1" "m \<le> n" "f m \<le> k" "k \<le> f n"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1433  | 
shows "\<exists>i. m \<le> i \<and> i \<le> n \<and> f i = k"  | 
| 67116 | 1434  | 
proof -  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1435  | 
obtain i where "i \<le> n - m" "k = f (m + i)"  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1436  | 
using nat_ivt_aux [of "n - m" "f \<circ> plus m" k] assms by auto  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1437  | 
with assms show ?thesis  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74979 
diff
changeset
 | 
1438  | 
using exI[of _ "m + i"] by auto  | 
| 67116 | 1439  | 
qed  | 
1440  | 
||
1441  | 
lemma nat0_intermed_int_val:  | 
|
1442  | 
"\<exists>i\<le>n. f i = k"  | 
|
1443  | 
if "\<forall>i<n. \<bar>f (i + 1) - f i\<bar> \<le> 1" "f 0 \<le> k" "k \<le> f n"  | 
|
| 63652 | 1444  | 
for n :: nat and k :: int  | 
| 67116 | 1445  | 
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
 | 
1446  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1447  | 
|
| 63652 | 1448  | 
subsection \<open>Products and 1, by T. M. Rasmussen\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1449  | 
|
| 34055 | 1450  | 
lemma abs_zmult_eq_1:  | 
| 63652 | 1451  | 
fixes m n :: int  | 
| 34055 | 1452  | 
assumes mn: "\<bar>m * n\<bar> = 1"  | 
| 63652 | 1453  | 
shows "\<bar>m\<bar> = 1"  | 
| 34055 | 1454  | 
proof -  | 
| 63652 | 1455  | 
from mn have 0: "m \<noteq> 0" "n \<noteq> 0" by auto  | 
1456  | 
have "\<not> 2 \<le> \<bar>m\<bar>"  | 
|
| 34055 | 1457  | 
proof  | 
1458  | 
assume "2 \<le> \<bar>m\<bar>"  | 
|
| 63652 | 1459  | 
then have "2 * \<bar>n\<bar> \<le> \<bar>m\<bar> * \<bar>n\<bar>" by (simp add: mult_mono 0)  | 
1460  | 
also have "\<dots> = \<bar>m * n\<bar>" by (simp add: abs_mult)  | 
|
1461  | 
also from mn have "\<dots> = 1" by simp  | 
|
1462  | 
finally have "2 * \<bar>n\<bar> \<le> 1" .  | 
|
1463  | 
with 0 show "False" by arith  | 
|
| 34055 | 1464  | 
qed  | 
| 63652 | 1465  | 
with 0 show ?thesis by auto  | 
| 34055 | 1466  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1467  | 
|
| 63652 | 1468  | 
lemma pos_zmult_eq_1_iff_lemma: "m * n = 1 \<Longrightarrow> m = 1 \<or> m = - 1"  | 
1469  | 
for m n :: int  | 
|
1470  | 
using abs_zmult_eq_1 [of m n] by arith  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1471  | 
|
| 
35815
 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 
boehmes 
parents: 
35634 
diff
changeset
 | 
1472  | 
lemma pos_zmult_eq_1_iff:  | 
| 63652 | 1473  | 
fixes m n :: int  | 
1474  | 
assumes "0 < m"  | 
|
1475  | 
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
 | 
1476  | 
proof -  | 
| 63652 | 1477  | 
from assms have "m * n = 1 \<Longrightarrow> m = 1"  | 
1478  | 
by (auto dest: pos_zmult_eq_1_iff_lemma)  | 
|
1479  | 
then show ?thesis  | 
|
1480  | 
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
 | 
1481  | 
qed  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1482  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1483  | 
lemma zmult_eq_1_iff: "m * n = 1 \<longleftrightarrow> (m = 1 \<and> n = 1) \<or> (m = - 1 \<and> n = - 1)" (is "?L = ?R")  | 
| 63652 | 1484  | 
for m n :: int  | 
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1485  | 
proof  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1486  | 
assume L: ?L show ?R  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1487  | 
using pos_zmult_eq_1_iff_lemma [OF L] L by force  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1488  | 
qed auto  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1489  | 
|
| 78685 | 1490  | 
lemma zmult_eq_neg1_iff: "a * b = (-1 :: int) \<longleftrightarrow> a = 1 \<and> b = -1 \<or> a = -1 \<and> b = 1"  | 
1491  | 
using zmult_eq_1_iff[of a "-b"] by auto  | 
|
1492  | 
||
| 
69700
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69605 
diff
changeset
 | 
1493  | 
lemma infinite_UNIV_int [simp]: "\<not> finite (UNIV::int set)"  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1494  | 
proof  | 
| 
33296
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1495  | 
assume "finite (UNIV::int set)"  | 
| 61076 | 1496  | 
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
 | 
1497  | 
by (rule injI) simp  | 
| 61076 | 1498  | 
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
 | 
1499  | 
by (rule finite_UNIV_inj_surj)  | 
| 
 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 
haftmann 
parents: 
33056 
diff
changeset
 | 
1500  | 
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
 | 
1501  | 
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
 | 
1502  | 
qed  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1503  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
1504  | 
|
| 60758 | 1505  | 
subsection \<open>The divides relation\<close>  | 
| 33320 | 1506  | 
|
| 63652 | 1507  | 
lemma zdvd_antisym_nonneg: "0 \<le> m \<Longrightarrow> 0 \<le> n \<Longrightarrow> m dvd n \<Longrightarrow> n dvd m \<Longrightarrow> m = n"  | 
1508  | 
for m n :: int  | 
|
1509  | 
by (auto simp add: dvd_def mult.assoc zero_le_mult_iff zmult_eq_1_iff)  | 
|
| 33320 | 1510  | 
|
| 63652 | 1511  | 
lemma zdvd_antisym_abs:  | 
1512  | 
fixes a b :: int  | 
|
1513  | 
assumes "a dvd b" and "b dvd a"  | 
|
| 33320 | 1514  | 
shows "\<bar>a\<bar> = \<bar>b\<bar>"  | 
| 63652 | 1515  | 
proof (cases "a = 0")  | 
1516  | 
case True  | 
|
1517  | 
with assms show ?thesis by simp  | 
|
| 33657 | 1518  | 
next  | 
| 63652 | 1519  | 
case False  | 
1520  | 
from \<open>a dvd b\<close> obtain k where k: "b = a * k"  | 
|
1521  | 
unfolding dvd_def by blast  | 
|
1522  | 
from \<open>b dvd a\<close> obtain k' where k': "a = b * k'"  | 
|
1523  | 
unfolding dvd_def by blast  | 
|
1524  | 
from k k' have "a = a * k * k'" by simp  | 
|
1525  | 
with mult_cancel_left1[where c="a" and b="k*k'"] have kk': "k * k' = 1"  | 
|
1526  | 
using \<open>a \<noteq> 0\<close> by (simp add: mult.assoc)  | 
|
1527  | 
then have "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1"  | 
|
1528  | 
by (simp add: zmult_eq_1_iff)  | 
|
1529  | 
with k k' show ?thesis by auto  | 
|
| 33320 | 1530  | 
qed  | 
1531  | 
||
| 63652 | 1532  | 
lemma zdvd_zdiffD: "k dvd m - n \<Longrightarrow> k dvd n \<Longrightarrow> k dvd m"  | 
1533  | 
for k m n :: int  | 
|
| 60162 | 1534  | 
using dvd_add_right_iff [of k "- n" m] by simp  | 
| 33320 | 1535  | 
|
| 63652 | 1536  | 
lemma zdvd_reduce: "k dvd n + k * m \<longleftrightarrow> k dvd n"  | 
1537  | 
for k m n :: int  | 
|
| 
58649
 
a62065b5e1e2
generalized and consolidated some theorems concerning divisibility
 
haftmann 
parents: 
58512 
diff
changeset
 | 
1538  | 
using dvd_add_times_triv_right_iff [of k n m] by (simp add: ac_simps)  | 
| 33320 | 1539  | 
|
1540  | 
lemma dvd_imp_le_int:  | 
|
1541  | 
fixes d i :: int  | 
|
1542  | 
assumes "i \<noteq> 0" and "d dvd i"  | 
|
1543  | 
shows "\<bar>d\<bar> \<le> \<bar>i\<bar>"  | 
|
1544  | 
proof -  | 
|
| 60758 | 1545  | 
from \<open>d dvd i\<close> obtain k where "i = d * k" ..  | 
1546  | 
with \<open>i \<noteq> 0\<close> have "k \<noteq> 0" by auto  | 
|
| 33320 | 1547  | 
then have "1 \<le> \<bar>k\<bar>" and "0 \<le> \<bar>d\<bar>" by auto  | 
1548  | 
then have "\<bar>d\<bar> * 1 \<le> \<bar>d\<bar> * \<bar>k\<bar>" by (rule mult_left_mono)  | 
|
| 60758 | 1549  | 
with \<open>i = d * k\<close> show ?thesis by (simp add: abs_mult)  | 
| 33320 | 1550  | 
qed  | 
1551  | 
||
1552  | 
lemma zdvd_not_zless:  | 
|
1553  | 
fixes m n :: int  | 
|
1554  | 
assumes "0 < m" and "m < n"  | 
|
1555  | 
shows "\<not> n dvd m"  | 
|
1556  | 
proof  | 
|
1557  | 
from assms have "0 < n" by auto  | 
|
1558  | 
assume "n dvd m" then obtain k where k: "m = n * k" ..  | 
|
| 60758 | 1559  | 
with \<open>0 < m\<close> have "0 < n * k" by auto  | 
1560  | 
with \<open>0 < n\<close> have "0 < k" by (simp add: zero_less_mult_iff)  | 
|
1561  | 
with k \<open>0 < n\<close> \<open>m < n\<close> have "n * k < n * 1" by simp  | 
|
1562  | 
with \<open>0 < n\<close> \<open>0 < k\<close> show False unfolding mult_less_cancel_left by auto  | 
|
| 33320 | 1563  | 
qed  | 
1564  | 
||
| 63652 | 1565  | 
lemma zdvd_mult_cancel:  | 
1566  | 
fixes k m n :: int  | 
|
1567  | 
assumes d: "k * m dvd k * n"  | 
|
1568  | 
and "k \<noteq> 0"  | 
|
| 33320 | 1569  | 
shows "m dvd n"  | 
| 63652 | 1570  | 
proof -  | 
1571  | 
from d obtain h where h: "k * n = k * m * h"  | 
|
1572  | 
unfolding dvd_def by blast  | 
|
1573  | 
have "n = m * h"  | 
|
1574  | 
proof (rule ccontr)  | 
|
1575  | 
assume "\<not> ?thesis"  | 
|
1576  | 
with \<open>k \<noteq> 0\<close> have "k * n \<noteq> k * (m * h)" by simp  | 
|
1577  | 
with h show False  | 
|
1578  | 
by (simp add: mult.assoc)  | 
|
1579  | 
qed  | 
|
1580  | 
then show ?thesis by simp  | 
|
| 33320 | 1581  | 
qed  | 
1582  | 
||
| 67118 | 1583  | 
lemma int_dvd_int_iff [simp]:  | 
1584  | 
"int m dvd int n \<longleftrightarrow> m dvd n"  | 
|
| 33320 | 1585  | 
proof -  | 
| 67118 | 1586  | 
have "m dvd n" if "int n = int m * k" for k  | 
| 63652 | 1587  | 
proof (cases k)  | 
| 67118 | 1588  | 
case (nonneg q)  | 
1589  | 
with that have "n = m * q"  | 
|
| 63652 | 1590  | 
by (simp del: of_nat_mult add: of_nat_mult [symmetric])  | 
1591  | 
then show ?thesis ..  | 
|
1592  | 
next  | 
|
| 67118 | 1593  | 
case (neg q)  | 
1594  | 
with that have "int n = int m * (- int (Suc q))"  | 
|
| 63652 | 1595  | 
by simp  | 
| 67118 | 1596  | 
also have "\<dots> = - (int m * int (Suc q))"  | 
| 63652 | 1597  | 
by (simp only: mult_minus_right)  | 
| 67118 | 1598  | 
also have "\<dots> = - int (m * Suc q)"  | 
| 63652 | 1599  | 
by (simp only: of_nat_mult [symmetric])  | 
| 67118 | 1600  | 
finally have "- int (m * Suc q) = int n" ..  | 
| 63652 | 1601  | 
then show ?thesis  | 
1602  | 
by (simp only: negative_eq_positive) auto  | 
|
| 33320 | 1603  | 
qed  | 
| 67118 | 1604  | 
then show ?thesis by (auto simp add: dvd_def)  | 
| 33320 | 1605  | 
qed  | 
1606  | 
||
| 67118 | 1607  | 
lemma dvd_nat_abs_iff [simp]:  | 
1608  | 
"n dvd nat \<bar>k\<bar> \<longleftrightarrow> int n dvd k"  | 
|
1609  | 
proof -  | 
|
1610  | 
have "n dvd nat \<bar>k\<bar> \<longleftrightarrow> int n dvd int (nat \<bar>k\<bar>)"  | 
|
1611  | 
by (simp only: int_dvd_int_iff)  | 
|
1612  | 
then show ?thesis  | 
|
1613  | 
by simp  | 
|
1614  | 
qed  | 
|
1615  | 
||
1616  | 
lemma nat_abs_dvd_iff [simp]:  | 
|
1617  | 
"nat \<bar>k\<bar> dvd n \<longleftrightarrow> k dvd int n"  | 
|
1618  | 
proof -  | 
|
1619  | 
have "nat \<bar>k\<bar> dvd n \<longleftrightarrow> int (nat \<bar>k\<bar>) dvd int n"  | 
|
1620  | 
by (simp only: int_dvd_int_iff)  | 
|
1621  | 
then show ?thesis  | 
|
1622  | 
by simp  | 
|
1623  | 
qed  | 
|
1624  | 
||
1625  | 
lemma zdvd1_eq [simp]: "x dvd 1 \<longleftrightarrow> \<bar>x\<bar> = 1" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 63652 | 1626  | 
for x :: int  | 
| 33320 | 1627  | 
proof  | 
| 63652 | 1628  | 
assume ?lhs  | 
| 67118 | 1629  | 
then have "nat \<bar>x\<bar> dvd nat \<bar>1\<bar>"  | 
1630  | 
by (simp only: nat_abs_dvd_iff) simp  | 
|
1631  | 
then have "nat \<bar>x\<bar> = 1"  | 
|
1632  | 
by simp  | 
|
1633  | 
then show ?rhs  | 
|
1634  | 
by (cases "x < 0") simp_all  | 
|
| 33320 | 1635  | 
next  | 
| 63652 | 1636  | 
assume ?rhs  | 
| 67118 | 1637  | 
then have "x = 1 \<or> x = - 1"  | 
1638  | 
by auto  | 
|
1639  | 
then show ?lhs  | 
|
1640  | 
by (auto intro: dvdI)  | 
|
| 33320 | 1641  | 
qed  | 
1642  | 
||
| 60162 | 1643  | 
lemma zdvd_mult_cancel1:  | 
| 63652 | 1644  | 
fixes m :: int  | 
1645  | 
assumes mp: "m \<noteq> 0"  | 
|
1646  | 
shows "m * n dvd m \<longleftrightarrow> \<bar>n\<bar> = 1"  | 
|
1647  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 33320 | 1648  | 
proof  | 
| 63652 | 1649  | 
assume ?rhs  | 
1650  | 
then show ?lhs  | 
|
1651  | 
by (cases "n > 0") (auto simp add: minus_equation_iff)  | 
|
| 33320 | 1652  | 
next  | 
| 63652 | 1653  | 
assume ?lhs  | 
1654  | 
then have "m * n dvd m * 1" by simp  | 
|
1655  | 
from zdvd_mult_cancel[OF this mp] show ?rhs  | 
|
1656  | 
by (simp only: zdvd1_eq)  | 
|
| 33320 | 1657  | 
qed  | 
1658  | 
||
| 63652 | 1659  | 
lemma nat_dvd_iff: "nat z dvd m \<longleftrightarrow> (if 0 \<le> z then z dvd int m else m = 0)"  | 
| 67118 | 1660  | 
using nat_abs_dvd_iff [of z m] by (cases "z \<ge> 0") auto  | 
| 33320 | 1661  | 
|
| 63652 | 1662  | 
lemma eq_nat_nat_iff: "0 \<le> z \<Longrightarrow> 0 \<le> z' \<Longrightarrow> nat z = nat z' \<longleftrightarrow> z = z'"  | 
| 67116 | 1663  | 
by (auto elim: nonneg_int_cases)  | 
| 33341 | 1664  | 
|
| 63652 | 1665  | 
lemma nat_power_eq: "0 \<le> z \<Longrightarrow> nat (z ^ n) = nat z ^ n"  | 
| 33341 | 1666  | 
by (induct n) (simp_all add: nat_mult_distrib)  | 
1667  | 
||
| 
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
 | 
1668  | 
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
 | 
1669  | 
"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
 | 
1670  | 
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
 | 
1671  | 
|
| 
 
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
 | 
1672  | 
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
 | 
1673  | 
"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
 | 
1674  | 
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
 | 
1675  | 
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
 | 
1676  | 
|
| 
 
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
 | 
1677  | 
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
 | 
1678  | 
"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
 | 
1679  | 
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
 | 
1680  | 
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
 | 
1681  | 
|
| 
 
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
 | 
1682  | 
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
 | 
1683  | 
"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
 | 
1684  | 
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
 | 
1685  | 
|
| 
 
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
 | 
1686  | 
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
 | 
1687  | 
"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
 | 
1688  | 
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
 | 
1689  | 
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
 | 
1690  | 
|
| 
 
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
 | 
1691  | 
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
 | 
1692  | 
"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
 | 
1693  | 
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
 | 
1694  | 
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
 | 
1695  | 
|
| 
71616
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1696  | 
lemma zdvd_imp_le: "z \<le> n" if "z dvd n" "0 < n" for n z :: int  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1697  | 
proof (cases n)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1698  | 
case (nonneg n)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1699  | 
show ?thesis  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1700  | 
by (cases z) (use nonneg dvd_imp_le that in auto)  | 
| 
 
a9de39608b1a
more tidying up of old apply-proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
70927 
diff
changeset
 | 
1701  | 
qed (use that in auto)  | 
| 33320 | 1702  | 
|
| 36749 | 1703  | 
lemma zdvd_period:  | 
1704  | 
fixes a d :: int  | 
|
1705  | 
assumes "a dvd d"  | 
|
1706  | 
shows "a dvd (x + t) \<longleftrightarrow> a dvd ((x + c * d) + t)"  | 
|
| 63652 | 1707  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
| 36749 | 1708  | 
proof -  | 
| 
66816
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1709  | 
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
 | 
1710  | 
by (simp add: dvd_add_left_iff)  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1711  | 
then show ?thesis  | 
| 
 
212a3334e7da
more fundamental definition of div and mod on int
 
haftmann 
parents: 
66035 
diff
changeset
 | 
1712  | 
by (simp add: ac_simps)  | 
| 36749 | 1713  | 
qed  | 
1714  | 
||
| 33320 | 1715  | 
|
| 
71837
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1716  | 
subsection \<open>Powers with integer exponents\<close>  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1717  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1718  | 
text \<open>  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1719  | 
The following allows writing powers with an integer exponent. While the type signature  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1720  | 
is very generic, most theorems will assume that the underlying type is a division ring or  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1721  | 
a field.  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1722  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1723  | 
The notation `powi' is inspired by the `powr' notation for real/complex exponentiation.  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1724  | 
\<close>  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1725  | 
definition power_int :: "'a :: {inverse, power} \<Rightarrow> int \<Rightarrow> 'a" (infixr "powi" 80) where
 | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1726  | 
"power_int x n = (if n \<ge> 0 then x ^ nat n else inverse x ^ (nat (-n)))"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1727  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1728  | 
lemma power_int_0_right [simp]: "power_int x 0 = 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1729  | 
and power_int_1_right [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1730  | 
        "power_int (y :: 'a :: {power, inverse, monoid_mult}) 1 = y"
 | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1731  | 
and power_int_minus1_right [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1732  | 
        "power_int (y :: 'a :: {power, inverse, monoid_mult}) (-1) = inverse y"
 | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1733  | 
by (simp_all add: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1734  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1735  | 
lemma power_int_of_nat [simp]: "power_int x (int n) = x ^ n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1736  | 
by (simp add: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1737  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1738  | 
lemma power_int_numeral [simp]: "power_int x (numeral n) = x ^ numeral n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1739  | 
by (simp add: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1740  | 
|
| 78698 | 1741  | 
lemma powi_numeral_reduce: "x powi numeral n = x * x powi int (pred_numeral n)"  | 
1742  | 
by (simp add: numeral_eq_Suc)  | 
|
1743  | 
||
1744  | 
lemma powi_minus_numeral_reduce: "x powi - (numeral n) = inverse x * x powi - int(pred_numeral n)"  | 
|
1745  | 
by (simp add: numeral_eq_Suc power_int_def)  | 
|
1746  | 
||
| 
71837
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1747  | 
lemma int_cases4 [case_names nonneg neg]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1748  | 
fixes m :: int  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1749  | 
obtains n where "m = int n" | n where "n > 0" "m = -int n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1750  | 
proof (cases "m \<ge> 0")  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1751  | 
case True  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1752  | 
thus ?thesis using that(1)[of "nat m"] by auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1753  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1754  | 
case False  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1755  | 
thus ?thesis using that(2)[of "nat (-m)"] by auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1756  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1757  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1758  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1759  | 
context  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1760  | 
  assumes "SORT_CONSTRAINT('a::division_ring)"
 | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1761  | 
begin  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1762  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1763  | 
lemma power_int_minus: "power_int (x::'a) (-n) = inverse (power_int x n)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1764  | 
by (auto simp: power_int_def power_inverse)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1765  | 
|
| 77351 | 1766  | 
lemma power_int_minus_divide: "power_int (x::'a) (-n) = 1 / (power_int x n)"  | 
1767  | 
by (simp add: divide_inverse power_int_minus)  | 
|
1768  | 
||
| 
71837
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1769  | 
lemma power_int_eq_0_iff [simp]: "power_int (x::'a) n = 0 \<longleftrightarrow> x = 0 \<and> n \<noteq> 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1770  | 
by (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1771  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1772  | 
lemma power_int_0_left_If: "power_int (0 :: 'a) m = (if m = 0 then 1 else 0)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1773  | 
by (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1774  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1775  | 
lemma power_int_0_left [simp]: "m \<noteq> 0 \<Longrightarrow> power_int (0 :: 'a) m = 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1776  | 
by (simp add: power_int_0_left_If)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1777  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1778  | 
lemma power_int_1_left [simp]: "power_int 1 n = (1 :: 'a :: division_ring)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1779  | 
by (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1780  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1781  | 
lemma power_diff_conv_inverse: "x \<noteq> 0 \<Longrightarrow> m \<le> n \<Longrightarrow> (x :: 'a) ^ (n - m) = x ^ n * inverse x ^ m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1782  | 
by (simp add: field_simps flip: power_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1783  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1784  | 
lemma power_mult_inverse_distrib: "x ^ m * inverse (x :: 'a) = inverse x * x ^ m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1785  | 
proof (cases "x = 0")  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1786  | 
case [simp]: False  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1787  | 
show ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1788  | 
proof (cases m)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1789  | 
case (Suc m')  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1790  | 
have "x ^ Suc m' * inverse x = x ^ m'"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1791  | 
by (subst power_Suc2) (auto simp: mult.assoc)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1792  | 
also have "\<dots> = inverse x * x ^ Suc m'"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1793  | 
by (subst power_Suc) (auto simp: mult.assoc [symmetric])  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1794  | 
finally show ?thesis using Suc by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1795  | 
qed auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1796  | 
qed auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1797  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1798  | 
lemma power_mult_power_inverse_commute:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1799  | 
"x ^ m * inverse (x :: 'a) ^ n = inverse x ^ n * x ^ m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1800  | 
proof (induction n)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1801  | 
case (Suc n)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1802  | 
have "x ^ m * inverse x ^ Suc n = (x ^ m * inverse x ^ n) * inverse x"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1803  | 
by (simp only: power_Suc2 mult.assoc)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1804  | 
also have "x ^ m * inverse x ^ n = inverse x ^ n * x ^ m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1805  | 
by (rule Suc)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1806  | 
also have "\<dots> * inverse x = (inverse x ^ n * inverse x) * x ^ m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1807  | 
by (simp add: mult.assoc power_mult_inverse_distrib)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1808  | 
also have "\<dots> = inverse x ^ (Suc n) * x ^ m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1809  | 
by (simp only: power_Suc2)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1810  | 
finally show ?case .  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1811  | 
qed auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1812  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1813  | 
lemma power_int_add:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1814  | 
assumes "x \<noteq> 0 \<or> m + n \<noteq> 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1815  | 
shows "power_int (x::'a) (m + n) = power_int x m * power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1816  | 
proof (cases "x = 0")  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1817  | 
case True  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1818  | 
thus ?thesis using assms by (auto simp: power_int_0_left_If)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1819  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1820  | 
case [simp]: False  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1821  | 
show ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1822  | 
proof (cases m n rule: int_cases4[case_product int_cases4])  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1823  | 
case (nonneg_nonneg a b)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1824  | 
thus ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1825  | 
by (auto simp: power_int_def nat_add_distrib power_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1826  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1827  | 
case (nonneg_neg a b)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1828  | 
thus ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1829  | 
by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1830  | 
power_mult_power_inverse_commute)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1831  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1832  | 
case (neg_nonneg a b)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1833  | 
thus ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1834  | 
by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1835  | 
power_mult_power_inverse_commute)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1836  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1837  | 
case (neg_neg a b)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1838  | 
thus ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1839  | 
by (auto simp: power_int_def nat_add_distrib add.commute simp flip: power_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1840  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1841  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1842  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1843  | 
lemma power_int_add_1:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1844  | 
assumes "x \<noteq> 0 \<or> m \<noteq> -1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1845  | 
shows "power_int (x::'a) (m + 1) = power_int x m * x"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1846  | 
using assms by (subst power_int_add) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1847  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1848  | 
lemma power_int_add_1':  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1849  | 
assumes "x \<noteq> 0 \<or> m \<noteq> -1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1850  | 
shows "power_int (x::'a) (m + 1) = x * power_int x m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1851  | 
using assms by (subst add.commute, subst power_int_add) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1852  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1853  | 
lemma power_int_commutes: "power_int (x :: 'a) n * x = x * power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1854  | 
by (cases "x = 0") (auto simp flip: power_int_add_1 power_int_add_1')  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1855  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1856  | 
lemma power_int_inverse [field_simps, field_split_simps, divide_simps]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1857  | 
"power_int (inverse (x :: 'a)) n = inverse (power_int x n)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1858  | 
by (auto simp: power_int_def power_inverse)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1859  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1860  | 
lemma power_int_mult: "power_int (x :: 'a) (m * n) = power_int (power_int x m) n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1861  | 
by (auto simp: power_int_def zero_le_mult_iff simp flip: power_mult power_inverse nat_mult_distrib)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1862  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1863  | 
end  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1864  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1865  | 
context  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1866  | 
  assumes "SORT_CONSTRAINT('a::field)"
 | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1867  | 
begin  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1868  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1869  | 
lemma power_int_diff:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1870  | 
assumes "x \<noteq> 0 \<or> m \<noteq> n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1871  | 
shows "power_int (x::'a) (m - n) = power_int x m / power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1872  | 
using power_int_add[of x m "-n"] assms by (auto simp: field_simps power_int_minus)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1873  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1874  | 
lemma power_int_minus_mult: "x \<noteq> 0 \<or> n \<noteq> 0 \<Longrightarrow> power_int (x :: 'a) (n - 1) * x = power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1875  | 
by (auto simp flip: power_int_add_1)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1876  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1877  | 
lemma power_int_mult_distrib: "power_int (x * y :: 'a) m = power_int x m * power_int y m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1878  | 
by (auto simp: power_int_def power_mult_distrib)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1879  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1880  | 
lemmas power_int_mult_distrib_numeral1 = power_int_mult_distrib [where x = "numeral w" for w, simp]  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1881  | 
lemmas power_int_mult_distrib_numeral2 = power_int_mult_distrib [where y = "numeral w" for w, simp]  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1882  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1883  | 
lemma power_int_divide_distrib: "power_int (x / y :: 'a) m = power_int x m / power_int y m"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1884  | 
using power_int_mult_distrib[of x "inverse y" m] unfolding power_int_inverse  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1885  | 
by (simp add: field_simps)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1886  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1887  | 
end  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1888  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1889  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1890  | 
lemma power_int_add_numeral [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1891  | 
"power_int x (numeral m) * power_int x (numeral n) = power_int x (numeral (m + n))"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1892  | 
for x :: "'a :: division_ring"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1893  | 
by (simp add: power_int_add [symmetric])  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1894  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1895  | 
lemma power_int_add_numeral2 [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1896  | 
"power_int x (numeral m) * (power_int x (numeral n) * b) = power_int x (numeral (m + n)) * b"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1897  | 
for x :: "'a :: division_ring"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1898  | 
by (simp add: mult.assoc [symmetric])  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1899  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1900  | 
lemma power_int_mult_numeral [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1901  | 
"power_int (power_int x (numeral m)) (numeral n) = power_int x (numeral (m * n))"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1902  | 
for x :: "'a :: division_ring"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1903  | 
by (simp only: numeral_mult power_int_mult)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1904  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1905  | 
lemma power_int_not_zero: "(x :: 'a :: division_ring) \<noteq> 0 \<or> n = 0 \<Longrightarrow> power_int x n \<noteq> 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1906  | 
by (subst power_int_eq_0_iff) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1907  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1908  | 
lemma power_int_one_over [field_simps, field_split_simps, divide_simps]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1909  | 
"power_int (1 / x :: 'a :: division_ring) n = 1 / power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1910  | 
using power_int_inverse[of x] by (simp add: divide_inverse)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1911  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1912  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1913  | 
context  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1914  | 
  assumes "SORT_CONSTRAINT('a :: linordered_field)"
 | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1915  | 
begin  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1916  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1917  | 
lemma power_int_numeral_neg_numeral [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1918  | 
"power_int (numeral m) (-numeral n) = (inverse (numeral (Num.pow m n)) :: 'a)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1919  | 
by (simp add: power_int_minus)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1920  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1921  | 
lemma zero_less_power_int [simp]: "0 < (x :: 'a) \<Longrightarrow> 0 < power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1922  | 
by (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1923  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1924  | 
lemma zero_le_power_int [simp]: "0 \<le> (x :: 'a) \<Longrightarrow> 0 \<le> power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1925  | 
by (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1926  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1927  | 
lemma power_int_mono: "(x :: 'a) \<le> y \<Longrightarrow> n \<ge> 0 \<Longrightarrow> 0 \<le> x \<Longrightarrow> power_int x n \<le> power_int y n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1928  | 
by (cases n rule: int_cases4) (auto intro: power_mono)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1929  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1930  | 
lemma one_le_power_int [simp]: "1 \<le> (x :: 'a) \<Longrightarrow> n \<ge> 0 \<Longrightarrow> 1 \<le> power_int x n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1931  | 
using power_int_mono [of 1 x n] by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1932  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1933  | 
lemma power_int_le_one: "0 \<le> (x :: 'a) \<Longrightarrow> n \<ge> 0 \<Longrightarrow> x \<le> 1 \<Longrightarrow> power_int x n \<le> 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1934  | 
using power_int_mono [of x 1 n] by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1935  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1936  | 
lemma power_int_le_imp_le_exp:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1937  | 
assumes gt1: "1 < (x :: 'a :: linordered_field)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1938  | 
assumes "power_int x m \<le> power_int x n" "n \<ge> 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1939  | 
shows "m \<le> n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1940  | 
proof (cases "m < 0")  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1941  | 
case True  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1942  | 
with \<open>n \<ge> 0\<close> show ?thesis by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1943  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1944  | 
case False  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1945  | 
with assms have "x ^ nat m \<le> x ^ nat n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1946  | 
by (simp add: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1947  | 
from gt1 and this show ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1948  | 
using False \<open>n \<ge> 0\<close> by auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1949  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1950  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1951  | 
lemma power_int_le_imp_less_exp:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1952  | 
assumes gt1: "1 < (x :: 'a :: linordered_field)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1953  | 
assumes "power_int x m < power_int x n" "n \<ge> 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1954  | 
shows "m < n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1955  | 
proof (cases "m < 0")  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1956  | 
case True  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1957  | 
with \<open>n \<ge> 0\<close> show ?thesis by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1958  | 
next  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1959  | 
case False  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1960  | 
with assms have "x ^ nat m < x ^ nat n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1961  | 
by (simp add: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1962  | 
from gt1 and this show ?thesis  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1963  | 
using False \<open>n \<ge> 0\<close> by auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1964  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1965  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1966  | 
lemma power_int_strict_mono:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1967  | 
"(a :: 'a :: linordered_field) < b \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 < n \<Longrightarrow> power_int a n < power_int b n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1968  | 
by (auto simp: power_int_def intro!: power_strict_mono)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1969  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1970  | 
lemma power_int_mono_iff [simp]:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1971  | 
fixes a b :: "'a :: linordered_field"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1972  | 
shows "\<lbrakk>a \<ge> 0; b \<ge> 0; n > 0\<rbrakk> \<Longrightarrow> power_int a n \<le> power_int b n \<longleftrightarrow> a \<le> b"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1973  | 
by (auto simp: power_int_def intro!: power_strict_mono)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1974  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1975  | 
lemma power_int_strict_increasing:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1976  | 
fixes a :: "'a :: linordered_field"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1977  | 
assumes "n < N" "1 < a"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1978  | 
shows "power_int a N > power_int a n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1979  | 
proof -  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1980  | 
have *: "a ^ nat (N - n) > a ^ 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1981  | 
using assms by (intro power_strict_increasing) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1982  | 
have "power_int a N = power_int a n * power_int a (N - n)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1983  | 
using assms by (simp flip: power_int_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1984  | 
also have "\<dots> > power_int a n * 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1985  | 
using assms *  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1986  | 
by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1987  | 
finally show ?thesis by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1988  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1989  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1990  | 
lemma power_int_increasing:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1991  | 
fixes a :: "'a :: linordered_field"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1992  | 
assumes "n \<le> N" "a \<ge> 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1993  | 
shows "power_int a N \<ge> power_int a n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1994  | 
proof -  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1995  | 
have *: "a ^ nat (N - n) \<ge> a ^ 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1996  | 
using assms by (intro power_increasing) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1997  | 
have "power_int a N = power_int a n * power_int a (N - n)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1998  | 
using assms by (simp flip: power_int_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
1999  | 
also have "\<dots> \<ge> power_int a n * 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2000  | 
using assms * by (intro mult_left_mono) (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2001  | 
finally show ?thesis by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2002  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2003  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2004  | 
lemma power_int_strict_decreasing:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2005  | 
fixes a :: "'a :: linordered_field"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2006  | 
assumes "n < N" "0 < a" "a < 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2007  | 
shows "power_int a N < power_int a n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2008  | 
proof -  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2009  | 
have *: "a ^ nat (N - n) < a ^ 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2010  | 
using assms by (intro power_strict_decreasing) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2011  | 
have "power_int a N = power_int a n * power_int a (N - n)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2012  | 
using assms by (simp flip: power_int_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2013  | 
also have "\<dots> < power_int a n * 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2014  | 
using assms *  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2015  | 
by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2016  | 
finally show ?thesis by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2017  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2018  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2019  | 
lemma power_int_decreasing:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2020  | 
fixes a :: "'a :: linordered_field"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2021  | 
assumes "n \<le> N" "0 \<le> a" "a \<le> 1" "a \<noteq> 0 \<or> N \<noteq> 0 \<or> n = 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2022  | 
shows "power_int a N \<le> power_int a n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2023  | 
proof (cases "a = 0")  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2024  | 
case False  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2025  | 
have *: "a ^ nat (N - n) \<le> a ^ 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2026  | 
using assms by (intro power_decreasing) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2027  | 
have "power_int a N = power_int a n * power_int a (N - n)"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2028  | 
using assms False by (simp flip: power_int_add)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2029  | 
also have "\<dots> \<le> power_int a n * 1"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2030  | 
using assms * by (intro mult_left_mono) (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2031  | 
finally show ?thesis by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2032  | 
qed (use assms in \<open>auto simp: power_int_0_left_If\<close>)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2033  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2034  | 
lemma one_less_power_int: "1 < (a :: 'a) \<Longrightarrow> 0 < n \<Longrightarrow> 1 < power_int a n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2035  | 
using power_int_strict_increasing[of 0 n a] by simp  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2036  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2037  | 
lemma power_int_abs: "\<bar>power_int a n :: 'a\<bar> = power_int \<bar>a\<bar> n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2038  | 
by (auto simp: power_int_def power_abs)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2039  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2040  | 
lemma power_int_sgn [simp]: "sgn (power_int a n :: 'a) = power_int (sgn a) n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2041  | 
by (auto simp: power_int_def)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2042  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2043  | 
lemma abs_power_int_minus [simp]: "\<bar>power_int (- a) n :: 'a\<bar> = \<bar>power_int a n\<bar>"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2044  | 
by (simp add: power_int_abs)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2045  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2046  | 
lemma power_int_strict_antimono:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2047  | 
assumes "(a :: 'a :: linordered_field) < b" "0 < a" "n < 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2048  | 
shows "power_int a n > power_int b n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2049  | 
proof -  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2050  | 
have "inverse (power_int a (-n)) > inverse (power_int b (-n))"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2051  | 
using assms by (intro less_imp_inverse_less power_int_strict_mono zero_less_power_int) auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2052  | 
thus ?thesis by (simp add: power_int_minus)  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2053  | 
qed  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2054  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2055  | 
lemma power_int_antimono:  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2056  | 
assumes "(a :: 'a :: linordered_field) \<le> b" "0 < a" "n < 0"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2057  | 
shows "power_int a n \<ge> power_int b n"  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2058  | 
using power_int_strict_antimono[of a b n] assms by (cases "a = b") auto  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2059  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2060  | 
end  | 
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2061  | 
|
| 
 
dca11678c495
new constant power_int in HOL
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
71616 
diff
changeset
 | 
2062  | 
|
| 60758 | 2063  | 
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
 | 
2064  | 
|
| 63652 | 2065  | 
lemma finite_interval_int1 [iff]: "finite {i :: int. a \<le> i \<and> i \<le> b}"
 | 
2066  | 
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
 | 
2067  | 
case True  | 
| 63652 | 2068  | 
then show ?thesis  | 
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
2069  | 
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
 | 
2070  | 
case base  | 
| 63652 | 2071  | 
    have "{i. a \<le> i \<and> i \<le> a} = {a}" by auto
 | 
2072  | 
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
 | 
2073  | 
next  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
2074  | 
case (step b)  | 
| 63652 | 2075  | 
    then have "{i. a \<le> i \<and> i \<le> b + 1} = {i. a \<le> i \<and> i \<le> b} \<union> {b + 1}" by auto
 | 
2076  | 
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
 | 
2077  | 
qed  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
2078  | 
next  | 
| 63652 | 2079  | 
case False  | 
2080  | 
then show ?thesis  | 
|
| 
46756
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
2081  | 
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
 | 
2082  | 
qed  | 
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
2083  | 
|
| 63652 | 2084  | 
lemma finite_interval_int2 [iff]: "finite {i :: int. a \<le> i \<and> i < b}"
 | 
2085  | 
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
 | 
2086  | 
|
| 63652 | 2087  | 
lemma finite_interval_int3 [iff]: "finite {i :: int. a < i \<and> i \<le> b}"
 | 
2088  | 
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
 | 
2089  | 
|
| 63652 | 2090  | 
lemma finite_interval_int4 [iff]: "finite {i :: int. a < i \<and> i < b}"
 | 
2091  | 
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
 | 
2092  | 
|
| 
 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 
bulwahn 
parents: 
46027 
diff
changeset
 | 
2093  | 
|
| 60758 | 2094  | 
subsection \<open>Configuration of the code generator\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2095  | 
|
| 60758 | 2096  | 
text \<open>Constructors\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2097  | 
|
| 63652 | 2098  | 
definition Pos :: "num \<Rightarrow> int"  | 
2099  | 
where [simp, code_abbrev]: "Pos = numeral"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2100  | 
|
| 63652 | 2101  | 
definition Neg :: "num \<Rightarrow> int"  | 
2102  | 
where [simp, code_abbrev]: "Neg n = - (Pos n)"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2103  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2104  | 
code_datatype "0::int" Pos Neg  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2105  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2106  | 
|
| 63652 | 2107  | 
text \<open>Auxiliary operations.\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2108  | 
|
| 63652 | 2109  | 
definition dup :: "int \<Rightarrow> int"  | 
2110  | 
where [simp]: "dup k = k + k"  | 
|
| 26507 | 2111  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2112  | 
lemma dup_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2113  | 
"dup 0 = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2114  | 
"dup (Pos n) = Pos (Num.Bit0 n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2115  | 
"dup (Neg n) = Neg (Num.Bit0 n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2116  | 
by (simp_all add: numeral_Bit0)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2117  | 
|
| 63652 | 2118  | 
definition sub :: "num \<Rightarrow> num \<Rightarrow> int"  | 
2119  | 
where [simp]: "sub m n = numeral m - numeral n"  | 
|
| 26507 | 2120  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2121  | 
lemma sub_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2122  | 
"sub Num.One Num.One = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2123  | 
"sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2124  | 
"sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2125  | 
"sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2126  | 
"sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2127  | 
"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
 | 
2128  | 
"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
 | 
2129  | 
"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
 | 
2130  | 
"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
 | 
2131  | 
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
 | 
2132  | 
|
| 72512 | 2133  | 
lemma sub_BitM_One_eq:  | 
2134  | 
\<open>(Num.sub (Num.BitM n) num.One) = 2 * (Num.sub n Num.One :: int)\<close>  | 
|
2135  | 
by (cases n) simp_all  | 
|
2136  | 
||
| 63652 | 2137  | 
text \<open>Implementations.\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2138  | 
|
| 64996 | 2139  | 
lemma one_int_code [code]: "1 = Pos Num.One"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2140  | 
by simp  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2141  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2142  | 
lemma plus_int_code [code]:  | 
| 63652 | 2143  | 
"k + 0 = k"  | 
2144  | 
"0 + l = l"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2145  | 
"Pos m + Pos n = Pos (m + n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2146  | 
"Pos m + Neg n = sub m n"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2147  | 
"Neg m + Pos n = sub n m"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2148  | 
"Neg m + Neg n = Neg (m + n)"  | 
| 63652 | 2149  | 
for k l :: int  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2150  | 
by simp_all  | 
| 26507 | 2151  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2152  | 
lemma uminus_int_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2153  | 
"uminus 0 = (0::int)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2154  | 
"uminus (Pos m) = Neg m"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2155  | 
"uminus (Neg m) = Pos m"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2156  | 
by simp_all  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2157  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2158  | 
lemma minus_int_code [code]:  | 
| 63652 | 2159  | 
"k - 0 = k"  | 
2160  | 
"0 - l = uminus l"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2161  | 
"Pos m - Pos n = sub m n"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2162  | 
"Pos m - Neg n = Pos (m + n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2163  | 
"Neg m - Pos n = Neg (m + n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2164  | 
"Neg m - Neg n = sub n m"  | 
| 63652 | 2165  | 
for k l :: int  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2166  | 
by simp_all  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2167  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2168  | 
lemma times_int_code [code]:  | 
| 63652 | 2169  | 
"k * 0 = 0"  | 
2170  | 
"0 * l = 0"  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2171  | 
"Pos m * Pos n = Pos (m * n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2172  | 
"Pos m * Neg n = Neg (m * n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2173  | 
"Neg m * Pos n = Neg (m * n)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2174  | 
"Neg m * Neg n = Pos (m * n)"  | 
| 63652 | 2175  | 
for k l :: int  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2176  | 
by simp_all  | 
| 26507 | 2177  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37887 
diff
changeset
 | 
2178  | 
instantiation int :: equal  | 
| 26507 | 2179  | 
begin  | 
2180  | 
||
| 63652 | 2181  | 
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
 | 
2182  | 
|
| 61169 | 2183  | 
instance  | 
2184  | 
by standard (rule equal_int_def)  | 
|
| 26507 | 2185  | 
|
2186  | 
end  | 
|
2187  | 
||
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2188  | 
lemma equal_int_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2189  | 
"HOL.equal 0 (0::int) \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2190  | 
"HOL.equal 0 (Pos l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2191  | 
"HOL.equal 0 (Neg l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2192  | 
"HOL.equal (Pos k) 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2193  | 
"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
 | 
2194  | 
"HOL.equal (Pos k) (Neg l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2195  | 
"HOL.equal (Neg k) 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2196  | 
"HOL.equal (Neg k) (Pos l) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2197  | 
"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
 | 
2198  | 
by (auto simp add: equal)  | 
| 26507 | 2199  | 
|
| 63652 | 2200  | 
lemma equal_int_refl [code nbe]: "HOL.equal k k \<longleftrightarrow> True"  | 
2201  | 
for k :: int  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2202  | 
by (fact equal_refl)  | 
| 26507 | 2203  | 
|
| 28562 | 2204  | 
lemma less_eq_int_code [code]:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2205  | 
"0 \<le> (0::int) \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2206  | 
"0 \<le> Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2207  | 
"0 \<le> Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2208  | 
"Pos k \<le> 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2209  | 
"Pos k \<le> Pos l \<longleftrightarrow> k \<le> l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2210  | 
"Pos k \<le> Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2211  | 
"Neg k \<le> 0 \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2212  | 
"Neg k \<le> Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2213  | 
"Neg k \<le> Neg l \<longleftrightarrow> l \<le> k"  | 
| 28958 | 2214  | 
by simp_all  | 
| 26507 | 2215  | 
|
| 28562 | 2216  | 
lemma less_int_code [code]:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2217  | 
"0 < (0::int) \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2218  | 
"0 < Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2219  | 
"0 < Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2220  | 
"Pos k < 0 \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2221  | 
"Pos k < Pos l \<longleftrightarrow> k < l"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2222  | 
"Pos k < Neg l \<longleftrightarrow> False"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2223  | 
"Neg k < 0 \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2224  | 
"Neg k < Pos l \<longleftrightarrow> True"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2225  | 
"Neg k < Neg l \<longleftrightarrow> l < k"  | 
| 28958 | 2226  | 
by simp_all  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2227  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2228  | 
lemma nat_code [code]:  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2229  | 
"nat (Int.Neg k) = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2230  | 
"nat 0 = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2231  | 
"nat (Int.Pos k) = nat_of_num k"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
2232  | 
by (simp_all add: nat_of_num_numeral)  | 
| 25928 | 2233  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2234  | 
lemma (in ring_1) of_int_code [code]:  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54249 
diff
changeset
 | 
2235  | 
"of_int (Int.Neg k) = - numeral k"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2236  | 
"of_int 0 = 0"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2237  | 
"of_int (Int.Pos k) = numeral k"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2238  | 
by simp_all  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2239  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2240  | 
|
| 63652 | 2241  | 
text \<open>Serializer setup.\<close>  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2242  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51994 
diff
changeset
 | 
2243  | 
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
 | 
2244  | 
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
 | 
2245  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2246  | 
quickcheck_params [default_type = int]  | 
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2247  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46756 
diff
changeset
 | 
2248  | 
hide_const (open) Pos Neg sub dup  | 
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2249  | 
|
| 
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2250  | 
|
| 61799 | 2251  | 
text \<open>De-register \<open>int\<close> as a quotient type:\<close>  | 
| 48045 | 2252  | 
|
| 
53652
 
18fbca265e2e
use lifting_forget for deregistering numeric types as a quotient type
 
kuncar 
parents: 
53065 
diff
changeset
 | 
2253  | 
lifting_update int.lifting  | 
| 
 
18fbca265e2e
use lifting_forget for deregistering numeric types as a quotient type
 
kuncar 
parents: 
53065 
diff
changeset
 | 
2254  | 
lifting_forget int.lifting  | 
| 48045 | 2255  | 
|
| 67116 | 2256  | 
|
2257  | 
subsection \<open>Duplicates\<close>  | 
|
2258  | 
||
2259  | 
lemmas int_sum = of_nat_sum [where 'a=int]  | 
|
2260  | 
lemmas int_prod = of_nat_prod [where 'a=int]  | 
|
2261  | 
lemmas zle_int = of_nat_le_iff [where 'a=int]  | 
|
2262  | 
lemmas int_int_eq = of_nat_eq_iff [where 'a=int]  | 
|
2263  | 
lemmas nonneg_eq_int = nonneg_int_cases  | 
|
2264  | 
lemmas double_eq_0_iff = double_zero  | 
|
2265  | 
||
2266  | 
lemmas int_distrib =  | 
|
2267  | 
distrib_right [of z1 z2 w]  | 
|
2268  | 
distrib_left [of w z1 z2]  | 
|
2269  | 
left_diff_distrib [of z1 z2 w]  | 
|
2270  | 
right_diff_distrib [of w z1 z2]  | 
|
2271  | 
for z1 z2 w :: int  | 
|
2272  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents:  
diff
changeset
 | 
2273  | 
end  | 
| 67116 | 2274  |