| author | traytel | 
| Fri, 28 Feb 2020 21:23:11 +0100 | |
| changeset 71494 | cbe0b6b0bed8 | 
| parent 71398 | e0237f2eb49d | 
| child 73109 | 783406dd051e | 
| permissions | -rw-r--r-- | 
| 63489 | 1  | 
(* Title: HOL/GCD.thy  | 
2  | 
Author: Christophe Tabacznyj  | 
|
3  | 
Author: Lawrence C. Paulson  | 
|
4  | 
Author: Amine Chaieb  | 
|
5  | 
Author: Thomas M. Rasmussen  | 
|
6  | 
Author: Jeremy Avigad  | 
|
7  | 
Author: Tobias Nipkow  | 
|
| 31706 | 8  | 
|
| 32479 | 9  | 
This file deals with the functions gcd and lcm. Definitions and  | 
10  | 
lemmas are proved uniformly for the natural numbers and integers.  | 
|
| 31706 | 11  | 
|
12  | 
This file combines and revises a number of prior developments.  | 
|
13  | 
||
14  | 
The original theories "GCD" and "Primes" were by Christophe Tabacznyj  | 
|
| 58623 | 15  | 
and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
 | 
| 31706 | 16  | 
gcd, lcm, and prime for the natural numbers.  | 
17  | 
||
18  | 
The original theory "IntPrimes" was by Thomas M. Rasmussen, and  | 
|
19  | 
extended gcd, lcm, primes to the integers. Amine Chaieb provided  | 
|
20  | 
another extension of the notions to the integers, and added a number  | 
|
21  | 
of results to "Primes" and "GCD". IntPrimes also defined and developed  | 
|
22  | 
the congruence relations on the integers. The notion was extended to  | 
|
| 34915 | 23  | 
the natural numbers by Chaieb.  | 
| 31706 | 24  | 
|
| 
32036
 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 
avigad 
parents: 
31952 
diff
changeset
 | 
25  | 
Jeremy Avigad combined all of these, made everything uniform for the  | 
| 
 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 
avigad 
parents: 
31952 
diff
changeset
 | 
26  | 
natural numbers and the integers, and added a number of new theorems.  | 
| 
 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 
avigad 
parents: 
31952 
diff
changeset
 | 
27  | 
|
| 31798 | 28  | 
Tobias Nipkow cleaned up a lot.  | 
| 21256 | 29  | 
*)  | 
30  | 
||
| 60758 | 31  | 
section \<open>Greatest common divisor and least common multiple\<close>  | 
| 21256 | 32  | 
|
33  | 
theory GCD  | 
|
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
34  | 
imports Groups_List  | 
| 31706 | 35  | 
begin  | 
36  | 
||
| 64850 | 37  | 
subsection \<open>Abstract bounded quasi semilattices as common foundation\<close>  | 
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
38  | 
|
| 
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
39  | 
locale bounded_quasi_semilattice = abel_semigroup +  | 
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
40  | 
  fixes top :: 'a  ("\<^bold>\<top>") and bot :: 'a  ("\<^bold>\<bottom>")
 | 
| 64850 | 41  | 
and normalize :: "'a \<Rightarrow> 'a"  | 
42  | 
assumes idem_normalize [simp]: "a \<^bold>* a = normalize a"  | 
|
43  | 
and normalize_left_idem [simp]: "normalize a \<^bold>* b = a \<^bold>* b"  | 
|
44  | 
and normalize_idem [simp]: "normalize (a \<^bold>* b) = a \<^bold>* b"  | 
|
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
45  | 
and normalize_top [simp]: "normalize \<^bold>\<top> = \<^bold>\<top>"  | 
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
46  | 
and normalize_bottom [simp]: "normalize \<^bold>\<bottom> = \<^bold>\<bottom>"  | 
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
47  | 
and top_left_normalize [simp]: "\<^bold>\<top> \<^bold>* a = normalize a"  | 
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
48  | 
and bottom_left_bottom [simp]: "\<^bold>\<bottom> \<^bold>* a = \<^bold>\<bottom>"  | 
| 64850 | 49  | 
begin  | 
50  | 
||
51  | 
lemma left_idem [simp]:  | 
|
52  | 
"a \<^bold>* (a \<^bold>* b) = a \<^bold>* b"  | 
|
53  | 
using assoc [of a a b, symmetric] by simp  | 
|
54  | 
||
55  | 
lemma right_idem [simp]:  | 
|
56  | 
"(a \<^bold>* b) \<^bold>* b = a \<^bold>* b"  | 
|
57  | 
using left_idem [of b a] by (simp add: ac_simps)  | 
|
58  | 
||
59  | 
lemma comp_fun_idem: "comp_fun_idem f"  | 
|
60  | 
by standard (simp_all add: fun_eq_iff ac_simps)  | 
|
61  | 
||
62  | 
interpretation comp_fun_idem f  | 
|
63  | 
by (fact comp_fun_idem)  | 
|
64  | 
||
65  | 
lemma top_right_normalize [simp]:  | 
|
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
66  | 
"a \<^bold>* \<^bold>\<top> = normalize a"  | 
| 64850 | 67  | 
using top_left_normalize [of a] by (simp add: ac_simps)  | 
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
68  | 
|
| 64850 | 69  | 
lemma bottom_right_bottom [simp]:  | 
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
70  | 
"a \<^bold>* \<^bold>\<bottom> = \<^bold>\<bottom>"  | 
| 64850 | 71  | 
using bottom_left_bottom [of a] by (simp add: ac_simps)  | 
72  | 
||
73  | 
lemma normalize_right_idem [simp]:  | 
|
74  | 
"a \<^bold>* normalize b = a \<^bold>* b"  | 
|
75  | 
using normalize_left_idem [of b a] by (simp add: ac_simps)  | 
|
76  | 
||
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
77  | 
end  | 
| 64850 | 78  | 
|
79  | 
locale bounded_quasi_semilattice_set = bounded_quasi_semilattice  | 
|
80  | 
begin  | 
|
81  | 
||
82  | 
interpretation comp_fun_idem f  | 
|
83  | 
by (fact comp_fun_idem)  | 
|
84  | 
||
85  | 
definition F :: "'a set \<Rightarrow> 'a"  | 
|
86  | 
where  | 
|
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
87  | 
eq_fold: "F A = (if finite A then Finite_Set.fold f \<^bold>\<top> A else \<^bold>\<bottom>)"  | 
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
88  | 
|
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
89  | 
lemma infinite [simp]:  | 
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
90  | 
"infinite A \<Longrightarrow> F A = \<^bold>\<bottom>"  | 
| 
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
91  | 
by (simp add: eq_fold)  | 
| 64850 | 92  | 
|
93  | 
lemma set_eq_fold [code]:  | 
|
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
94  | 
"F (set xs) = fold f xs \<^bold>\<top>"  | 
| 64850 | 95  | 
by (simp add: eq_fold fold_set_fold)  | 
96  | 
||
97  | 
lemma empty [simp]:  | 
|
| 
65555
 
85ed070017b7
include GCD as integral part of computational algebra in session HOL
 
haftmann 
parents: 
65552 
diff
changeset
 | 
98  | 
  "F {} = \<^bold>\<top>"
 | 
| 64850 | 99  | 
by (simp add: eq_fold)  | 
100  | 
||
101  | 
lemma insert [simp]:  | 
|
102  | 
"F (insert a A) = a \<^bold>* F A"  | 
|
103  | 
by (cases "finite A") (simp_all add: eq_fold)  | 
|
104  | 
||
105  | 
lemma normalize [simp]:  | 
|
106  | 
"normalize (F A) = F A"  | 
|
107  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
|
108  | 
||
109  | 
lemma in_idem:  | 
|
110  | 
assumes "a \<in> A"  | 
|
111  | 
shows "a \<^bold>* F A = F A"  | 
|
112  | 
using assms by (induct A rule: infinite_finite_induct)  | 
|
| 68708 | 113  | 
(auto simp: left_commute [of a])  | 
| 64850 | 114  | 
|
115  | 
lemma union:  | 
|
116  | 
"F (A \<union> B) = F A \<^bold>* F B"  | 
|
117  | 
by (induct A rule: infinite_finite_induct)  | 
|
118  | 
(simp_all add: ac_simps)  | 
|
119  | 
||
120  | 
lemma remove:  | 
|
121  | 
assumes "a \<in> A"  | 
|
122  | 
  shows "F A = a \<^bold>* F (A - {a})"
 | 
|
123  | 
proof -  | 
|
124  | 
from assms obtain B where "A = insert a B" and "a \<notin> B"  | 
|
125  | 
by (blast dest: mk_disjoint_insert)  | 
|
126  | 
with assms show ?thesis by simp  | 
|
127  | 
qed  | 
|
128  | 
||
129  | 
lemma insert_remove:  | 
|
130  | 
  "F (insert a A) = a \<^bold>* F (A - {a})"
 | 
|
131  | 
by (cases "a \<in> A") (simp_all add: insert_absorb remove)  | 
|
132  | 
||
133  | 
lemma subset:  | 
|
134  | 
assumes "B \<subseteq> A"  | 
|
135  | 
shows "F B \<^bold>* F A = F A"  | 
|
136  | 
using assms by (simp add: union [symmetric] Un_absorb1)  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
137  | 
|
| 64850 | 138  | 
end  | 
| 63489 | 139  | 
|
| 62345 | 140  | 
subsection \<open>Abstract GCD and LCM\<close>  | 
| 31706 | 141  | 
|
| 31992 | 142  | 
class gcd = zero + one + dvd +  | 
| 41550 | 143  | 
fixes gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
144  | 
and lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
|
| 31706 | 145  | 
|
| 60686 | 146  | 
class Gcd = gcd +  | 
| 63025 | 147  | 
fixes Gcd :: "'a set \<Rightarrow> 'a"  | 
148  | 
and Lcm :: "'a set \<Rightarrow> 'a"  | 
|
| 62350 | 149  | 
|
150  | 
syntax  | 
|
| 63025 | 151  | 
  "_GCD1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3GCD _./ _)" [0, 10] 10)
 | 
152  | 
  "_GCD"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3GCD _\<in>_./ _)" [0, 0, 10] 10)
 | 
|
153  | 
  "_LCM1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3LCM _./ _)" [0, 10] 10)
 | 
|
154  | 
  "_LCM"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3LCM _\<in>_./ _)" [0, 0, 10] 10)
 | 
|
| 
68795
 
210b687cdbbe
dropped redundant syntax translation rules for big operators
 
haftmann 
parents: 
68708 
diff
changeset
 | 
155  | 
|
| 62350 | 156  | 
translations  | 
| 
68796
 
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
 
haftmann 
parents: 
68795 
diff
changeset
 | 
157  | 
"GCD x y. f" \<rightleftharpoons> "GCD x. GCD y. f"  | 
| 
 
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
 
haftmann 
parents: 
68795 
diff
changeset
 | 
158  | 
"GCD x. f" \<rightleftharpoons> "CONST Gcd (CONST range (\<lambda>x. f))"  | 
| 
 
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
 
haftmann 
parents: 
68795 
diff
changeset
 | 
159  | 
"GCD x\<in>A. f" \<rightleftharpoons> "CONST Gcd ((\<lambda>x. f) ` A)"  | 
| 
 
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
 
haftmann 
parents: 
68795 
diff
changeset
 | 
160  | 
"LCM x y. f" \<rightleftharpoons> "LCM x. LCM y. f"  | 
| 
 
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
 
haftmann 
parents: 
68795 
diff
changeset
 | 
161  | 
"LCM x. f" \<rightleftharpoons> "CONST Lcm (CONST range (\<lambda>x. f))"  | 
| 
 
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
 
haftmann 
parents: 
68795 
diff
changeset
 | 
162  | 
"LCM x\<in>A. f" \<rightleftharpoons> "CONST Lcm ((\<lambda>x. f) ` A)"  | 
| 60686 | 163  | 
|
164  | 
class semiring_gcd = normalization_semidom + gcd +  | 
|
| 59008 | 165  | 
assumes gcd_dvd1 [iff]: "gcd a b dvd a"  | 
| 59977 | 166  | 
and gcd_dvd2 [iff]: "gcd a b dvd b"  | 
167  | 
and gcd_greatest: "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> c dvd gcd a b"  | 
|
| 60686 | 168  | 
and normalize_gcd [simp]: "normalize (gcd a b) = gcd a b"  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
169  | 
and lcm_gcd: "lcm a b = normalize (a * b div gcd a b)"  | 
| 63489 | 170  | 
begin  | 
171  | 
||
172  | 
lemma gcd_greatest_iff [simp]: "a dvd gcd b c \<longleftrightarrow> a dvd b \<and> a dvd c"  | 
|
| 60686 | 173  | 
by (blast intro!: gcd_greatest intro: dvd_trans)  | 
174  | 
||
| 63489 | 175  | 
lemma gcd_dvdI1: "a dvd c \<Longrightarrow> gcd a b dvd c"  | 
| 60689 | 176  | 
by (rule dvd_trans) (rule gcd_dvd1)  | 
177  | 
||
| 63489 | 178  | 
lemma gcd_dvdI2: "b dvd c \<Longrightarrow> gcd a b dvd c"  | 
| 60689 | 179  | 
by (rule dvd_trans) (rule gcd_dvd2)  | 
180  | 
||
| 63489 | 181  | 
lemma dvd_gcdD1: "a dvd gcd b c \<Longrightarrow> a dvd b"  | 
| 62345 | 182  | 
using gcd_dvd1 [of b c] by (blast intro: dvd_trans)  | 
183  | 
||
| 63489 | 184  | 
lemma dvd_gcdD2: "a dvd gcd b c \<Longrightarrow> a dvd c"  | 
| 62345 | 185  | 
using gcd_dvd2 [of b c] by (blast intro: dvd_trans)  | 
186  | 
||
| 63489 | 187  | 
lemma gcd_0_left [simp]: "gcd 0 a = normalize a"  | 
| 
60688
 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 
haftmann 
parents: 
60687 
diff
changeset
 | 
188  | 
by (rule associated_eqI) simp_all  | 
| 60686 | 189  | 
|
| 63489 | 190  | 
lemma gcd_0_right [simp]: "gcd a 0 = normalize a"  | 
| 
60688
 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 
haftmann 
parents: 
60687 
diff
changeset
 | 
191  | 
by (rule associated_eqI) simp_all  | 
| 63489 | 192  | 
|
193  | 
lemma gcd_eq_0_iff [simp]: "gcd a b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"  | 
|
194  | 
(is "?P \<longleftrightarrow> ?Q")  | 
|
| 60686 | 195  | 
proof  | 
| 63489 | 196  | 
assume ?P  | 
197  | 
then have "0 dvd gcd a b"  | 
|
198  | 
by simp  | 
|
199  | 
then have "0 dvd a" and "0 dvd b"  | 
|
200  | 
by (blast intro: dvd_trans)+  | 
|
201  | 
then show ?Q  | 
|
202  | 
by simp  | 
|
| 60686 | 203  | 
next  | 
| 63489 | 204  | 
assume ?Q  | 
205  | 
then show ?P  | 
|
206  | 
by simp  | 
|
| 60686 | 207  | 
qed  | 
208  | 
||
| 63489 | 209  | 
lemma unit_factor_gcd: "unit_factor (gcd a b) = (if a = 0 \<and> b = 0 then 0 else 1)"  | 
| 60686 | 210  | 
proof (cases "gcd a b = 0")  | 
| 63489 | 211  | 
case True  | 
212  | 
then show ?thesis by simp  | 
|
| 60686 | 213  | 
next  | 
214  | 
case False  | 
|
215  | 
have "unit_factor (gcd a b) * normalize (gcd a b) = gcd a b"  | 
|
216  | 
by (rule unit_factor_mult_normalize)  | 
|
217  | 
then have "unit_factor (gcd a b) * gcd a b = gcd a b"  | 
|
218  | 
by simp  | 
|
219  | 
then have "unit_factor (gcd a b) * gcd a b div gcd a b = gcd a b div gcd a b"  | 
|
220  | 
by simp  | 
|
| 63489 | 221  | 
with False show ?thesis  | 
222  | 
by simp  | 
|
| 60686 | 223  | 
qed  | 
224  | 
||
| 67051 | 225  | 
lemma is_unit_gcd_iff [simp]:  | 
226  | 
"is_unit (gcd a b) \<longleftrightarrow> gcd a b = 1"  | 
|
| 68708 | 227  | 
by (cases "a = 0 \<and> b = 0") (auto simp: unit_factor_gcd dest: is_unit_unit_factor)  | 
| 60690 | 228  | 
|
| 61605 | 229  | 
sublocale gcd: abel_semigroup gcd  | 
| 60686 | 230  | 
proof  | 
231  | 
fix a b c  | 
|
232  | 
show "gcd a b = gcd b a"  | 
|
| 
60688
 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 
haftmann 
parents: 
60687 
diff
changeset
 | 
233  | 
by (rule associated_eqI) simp_all  | 
| 60686 | 234  | 
from gcd_dvd1 have "gcd (gcd a b) c dvd a"  | 
235  | 
by (rule dvd_trans) simp  | 
|
236  | 
moreover from gcd_dvd1 have "gcd (gcd a b) c dvd b"  | 
|
237  | 
by (rule dvd_trans) simp  | 
|
238  | 
ultimately have P1: "gcd (gcd a b) c dvd gcd a (gcd b c)"  | 
|
239  | 
by (auto intro!: gcd_greatest)  | 
|
240  | 
from gcd_dvd2 have "gcd a (gcd b c) dvd b"  | 
|
241  | 
by (rule dvd_trans) simp  | 
|
242  | 
moreover from gcd_dvd2 have "gcd a (gcd b c) dvd c"  | 
|
243  | 
by (rule dvd_trans) simp  | 
|
244  | 
ultimately have P2: "gcd a (gcd b c) dvd gcd (gcd a b) c"  | 
|
245  | 
by (auto intro!: gcd_greatest)  | 
|
| 
60688
 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 
haftmann 
parents: 
60687 
diff
changeset
 | 
246  | 
from P1 P2 show "gcd (gcd a b) c = gcd a (gcd b c)"  | 
| 
 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 
haftmann 
parents: 
60687 
diff
changeset
 | 
247  | 
by (rule associated_eqI) simp_all  | 
| 60686 | 248  | 
qed  | 
249  | 
||
| 64850 | 250  | 
sublocale gcd: bounded_quasi_semilattice gcd 0 1 normalize  | 
251  | 
proof  | 
|
252  | 
show "gcd a a = normalize a" for a  | 
|
253  | 
proof -  | 
|
254  | 
have "a dvd gcd a a"  | 
|
255  | 
by (rule gcd_greatest) simp_all  | 
|
256  | 
then show ?thesis  | 
|
257  | 
by (auto intro: associated_eqI)  | 
|
258  | 
qed  | 
|
259  | 
show "gcd (normalize a) b = gcd a b" for a b  | 
|
260  | 
using gcd_dvd1 [of "normalize a" b]  | 
|
| 
60688
 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 
haftmann 
parents: 
60687 
diff
changeset
 | 
261  | 
by (auto intro: associated_eqI)  | 
| 67051 | 262  | 
show "gcd 1 a = 1" for a  | 
| 64850 | 263  | 
by (rule associated_eqI) simp_all  | 
264  | 
qed simp_all  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
265  | 
|
| 64850 | 266  | 
lemma gcd_self: "gcd a a = normalize a"  | 
267  | 
by (fact gcd.idem_normalize)  | 
|
268  | 
||
269  | 
lemma gcd_left_idem: "gcd a (gcd a b) = gcd a b"  | 
|
270  | 
by (fact gcd.left_idem)  | 
|
271  | 
||
272  | 
lemma gcd_right_idem: "gcd (gcd a b) b = gcd a b"  | 
|
273  | 
by (fact gcd.right_idem)  | 
|
274  | 
||
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
275  | 
lemma gcdI:  | 
| 63489 | 276  | 
assumes "c dvd a" and "c dvd b"  | 
277  | 
and greatest: "\<And>d. d dvd a \<Longrightarrow> d dvd b \<Longrightarrow> d dvd c"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
278  | 
and "normalize c = c"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
279  | 
shows "c = gcd a b"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
280  | 
by (rule associated_eqI) (auto simp: assms intro: gcd_greatest)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
281  | 
|
| 63489 | 282  | 
lemma gcd_unique:  | 
283  | 
"d dvd a \<and> d dvd b \<and> normalize d = d \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
284  | 
by rule (auto intro: gcdI simp: gcd_greatest)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
285  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
286  | 
lemma gcd_dvd_prod: "gcd a b dvd k * b"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
287  | 
using mult_dvd_mono [of 1] by auto  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
288  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
289  | 
lemma gcd_proj2_if_dvd: "b dvd a \<Longrightarrow> gcd a b = normalize b"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
290  | 
by (rule gcdI [symmetric]) simp_all  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
291  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
292  | 
lemma gcd_proj1_if_dvd: "a dvd b \<Longrightarrow> gcd a b = normalize a"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
293  | 
by (rule gcdI [symmetric]) simp_all  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
294  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
295  | 
lemma gcd_proj1_iff: "gcd m n = normalize m \<longleftrightarrow> m dvd n"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
296  | 
proof  | 
| 63489 | 297  | 
assume *: "gcd m n = normalize m"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
298  | 
show "m dvd n"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
299  | 
proof (cases "m = 0")  | 
| 63489 | 300  | 
case True  | 
301  | 
with * show ?thesis by simp  | 
|
302  | 
next  | 
|
303  | 
case [simp]: False  | 
|
304  | 
from * have **: "m = gcd m n * unit_factor m"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
305  | 
by (simp add: unit_eq_div2)  | 
| 63489 | 306  | 
show ?thesis  | 
307  | 
by (subst **) (simp add: mult_unit_dvd_iff)  | 
|
308  | 
qed  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
309  | 
next  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
310  | 
assume "m dvd n"  | 
| 63489 | 311  | 
then show "gcd m n = normalize m"  | 
312  | 
by (rule gcd_proj1_if_dvd)  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
313  | 
qed  | 
| 63489 | 314  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
315  | 
lemma gcd_proj2_iff: "gcd m n = normalize n \<longleftrightarrow> n dvd m"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
316  | 
using gcd_proj1_iff [of n m] by (simp add: ac_simps)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
317  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
318  | 
lemma gcd_mult_left: "gcd (c * a) (c * b) = normalize (c * gcd a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
319  | 
proof (cases "c = 0")  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
320  | 
case True  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
321  | 
then show ?thesis by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
322  | 
next  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
323  | 
case False  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
324  | 
then have *: "c * gcd a b dvd gcd (c * a) (c * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
325  | 
by (auto intro: gcd_greatest)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
326  | 
moreover from False * have "gcd (c * a) (c * b) dvd c * gcd a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
327  | 
by (metis div_dvd_iff_mult dvd_mult_left gcd_dvd1 gcd_dvd2 gcd_greatest mult_commute)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
328  | 
ultimately have "normalize (gcd (c * a) (c * b)) = normalize (c * gcd a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
329  | 
by (auto intro: associated_eqI)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
330  | 
then show ?thesis  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
331  | 
by (simp add: normalize_mult)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
332  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
333  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
334  | 
lemma gcd_mult_right: "gcd (a * c) (b * c) = normalize (gcd b a * c)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
335  | 
using gcd_mult_left [of c a b] by (simp add: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
336  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
337  | 
lemma dvd_lcm1 [iff]: "a dvd lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
338  | 
by (metis div_mult_swap dvd_mult2 dvd_normalize_iff dvd_refl gcd_dvd2 lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
339  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
340  | 
lemma dvd_lcm2 [iff]: "b dvd lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
341  | 
by (metis dvd_div_mult dvd_mult dvd_normalize_iff dvd_refl gcd_dvd1 lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
342  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
343  | 
lemma dvd_lcmI1: "a dvd b \<Longrightarrow> a dvd lcm b c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
344  | 
by (rule dvd_trans) (assumption, blast)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
345  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
346  | 
lemma dvd_lcmI2: "a dvd c \<Longrightarrow> a dvd lcm b c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
347  | 
by (rule dvd_trans) (assumption, blast)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
348  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
349  | 
lemma lcm_dvdD1: "lcm a b dvd c \<Longrightarrow> a dvd c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
350  | 
using dvd_lcm1 [of a b] by (blast intro: dvd_trans)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
351  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
352  | 
lemma lcm_dvdD2: "lcm a b dvd c \<Longrightarrow> b dvd c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
353  | 
using dvd_lcm2 [of a b] by (blast intro: dvd_trans)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
354  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
355  | 
lemma lcm_least:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
356  | 
assumes "a dvd c" and "b dvd c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
357  | 
shows "lcm a b dvd c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
358  | 
proof (cases "c = 0")  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
359  | 
case True  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
360  | 
then show ?thesis by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
361  | 
next  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
362  | 
case False  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
363  | 
then have *: "is_unit (unit_factor c)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
364  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
365  | 
show ?thesis  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
366  | 
proof (cases "gcd a b = 0")  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
367  | 
case True  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
368  | 
with assms show ?thesis by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
369  | 
next  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
370  | 
case False  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
371  | 
have "a * b dvd normalize (c * gcd a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
372  | 
using assms by (subst gcd_mult_left [symmetric]) (auto intro!: gcd_greatest simp: mult_ac)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
373  | 
with False have "(a * b div gcd a b) dvd c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
374  | 
by (subst div_dvd_iff_mult) auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
375  | 
thus ?thesis by (simp add: lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
376  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
377  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
378  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
379  | 
lemma lcm_least_iff [simp]: "lcm a b dvd c \<longleftrightarrow> a dvd c \<and> b dvd c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
380  | 
by (blast intro!: lcm_least intro: dvd_trans)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
381  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
382  | 
lemma normalize_lcm [simp]: "normalize (lcm a b) = lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
383  | 
by (simp add: lcm_gcd dvd_normalize_div)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
384  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
385  | 
lemma lcm_0_left [simp]: "lcm 0 a = 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
386  | 
by (simp add: lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
387  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
388  | 
lemma lcm_0_right [simp]: "lcm a 0 = 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
389  | 
by (simp add: lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
390  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
391  | 
lemma lcm_eq_0_iff: "lcm a b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
392  | 
(is "?P \<longleftrightarrow> ?Q")  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
393  | 
proof  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
394  | 
assume ?P  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
395  | 
then have "0 dvd lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
396  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
397  | 
also have "lcm a b dvd (a * b)"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
398  | 
by simp  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
399  | 
finally show ?Q  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
400  | 
by auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
401  | 
next  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
402  | 
assume ?Q  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
403  | 
then show ?P  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
404  | 
by auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
405  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
406  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
407  | 
lemma zero_eq_lcm_iff: "0 = lcm a b \<longleftrightarrow> a = 0 \<or> b = 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
408  | 
using lcm_eq_0_iff[of a b] by auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
409  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
410  | 
lemma lcm_eq_1_iff [simp]: "lcm a b = 1 \<longleftrightarrow> is_unit a \<and> is_unit b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
411  | 
by (auto intro: associated_eqI)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
412  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
413  | 
lemma unit_factor_lcm: "unit_factor (lcm a b) = (if a = 0 \<or> b = 0 then 0 else 1)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
414  | 
using lcm_eq_0_iff[of a b] by (cases "lcm a b = 0") (auto simp: lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
415  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
416  | 
sublocale lcm: abel_semigroup lcm  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
417  | 
proof  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
418  | 
fix a b c  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
419  | 
show "lcm a b = lcm b a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
420  | 
by (simp add: lcm_gcd ac_simps normalize_mult dvd_normalize_div)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
421  | 
have "lcm (lcm a b) c dvd lcm a (lcm b c)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
422  | 
and "lcm a (lcm b c) dvd lcm (lcm a b) c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
423  | 
by (auto intro: lcm_least  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
424  | 
dvd_trans [of b "lcm b c" "lcm a (lcm b c)"]  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
425  | 
dvd_trans [of c "lcm b c" "lcm a (lcm b c)"]  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
426  | 
dvd_trans [of a "lcm a b" "lcm (lcm a b) c"]  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
427  | 
dvd_trans [of b "lcm a b" "lcm (lcm a b) c"])  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
428  | 
then show "lcm (lcm a b) c = lcm a (lcm b c)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
429  | 
by (rule associated_eqI) simp_all  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
430  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
431  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
432  | 
sublocale lcm: bounded_quasi_semilattice lcm 1 0 normalize  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
433  | 
proof  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
434  | 
show "lcm a a = normalize a" for a  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
435  | 
proof -  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
436  | 
have "lcm a a dvd a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
437  | 
by (rule lcm_least) simp_all  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
438  | 
then show ?thesis  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
439  | 
by (auto intro: associated_eqI)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
440  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
441  | 
show "lcm (normalize a) b = lcm a b" for a b  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
442  | 
using dvd_lcm1 [of "normalize a" b] unfolding normalize_dvd_iff  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
443  | 
by (auto intro: associated_eqI)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
444  | 
show "lcm 1 a = normalize a" for a  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
445  | 
by (rule associated_eqI) simp_all  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
446  | 
qed simp_all  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
447  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
448  | 
lemma lcm_self: "lcm a a = normalize a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
449  | 
by (fact lcm.idem_normalize)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
450  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
451  | 
lemma lcm_left_idem: "lcm a (lcm a b) = lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
452  | 
by (fact lcm.left_idem)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
453  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
454  | 
lemma lcm_right_idem: "lcm (lcm a b) b = lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
455  | 
by (fact lcm.right_idem)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
456  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
457  | 
lemma gcd_lcm:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
458  | 
assumes "a \<noteq> 0" and "b \<noteq> 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
459  | 
shows "gcd a b = normalize (a * b div lcm a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
460  | 
proof -  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
461  | 
from assms have [simp]: "a * b div gcd a b \<noteq> 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
462  | 
by (subst dvd_div_eq_0_iff) auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
463  | 
let ?u = "unit_factor (a * b div gcd a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
464  | 
have "gcd a b * normalize (a * b div gcd a b) =  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
465  | 
gcd a b * ((a * b div gcd a b) * (1 div ?u))"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
466  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
467  | 
also have "\<dots> = a * b div ?u"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
468  | 
by (subst mult.assoc [symmetric]) auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
469  | 
also have "\<dots> dvd a * b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
470  | 
by (subst div_unit_dvd_iff) auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
471  | 
finally have "gcd a b dvd ((a * b) div lcm a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
472  | 
by (intro dvd_mult_imp_div) (auto simp: lcm_gcd)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
473  | 
moreover have "a * b div lcm a b dvd a" and "a * b div lcm a b dvd b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
474  | 
using assms by (subst div_dvd_iff_mult; simp add: lcm_eq_0_iff mult.commute[of b "lcm a b"])+  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
475  | 
ultimately have "normalize (gcd a b) = normalize (a * b div lcm a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
476  | 
apply -  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
477  | 
apply (rule associated_eqI)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
478  | 
using assms  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
479  | 
apply (auto simp: div_dvd_iff_mult zero_eq_lcm_iff)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
480  | 
done  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
481  | 
thus ?thesis by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
482  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
483  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
484  | 
lemma lcm_1_left: "lcm 1 a = normalize a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
485  | 
by (fact lcm.top_left_normalize)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
486  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
487  | 
lemma lcm_1_right: "lcm a 1 = normalize a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
488  | 
by (fact lcm.top_right_normalize)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
489  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
490  | 
lemma lcm_mult_left: "lcm (c * a) (c * b) = normalize (c * lcm a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
491  | 
proof (cases "c = 0")  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
492  | 
case True  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
493  | 
then show ?thesis by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
494  | 
next  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
495  | 
case False  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
496  | 
then have *: "lcm (c * a) (c * b) dvd c * lcm a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
497  | 
by (auto intro: lcm_least)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
498  | 
moreover have "lcm a b dvd lcm (c * a) (c * b) div c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
499  | 
by (intro lcm_least) (auto intro!: dvd_mult_imp_div simp: mult_ac)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
500  | 
hence "c * lcm a b dvd lcm (c * a) (c * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
501  | 
using False by (subst (asm) dvd_div_iff_mult) (auto simp: mult_ac intro: dvd_lcmI1)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
502  | 
ultimately have "normalize (lcm (c * a) (c * b)) = normalize (c * lcm a b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
503  | 
by (auto intro: associated_eqI)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
504  | 
then show ?thesis  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
505  | 
by (simp add: normalize_mult)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
506  | 
qed  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
507  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
508  | 
lemma lcm_mult_right: "lcm (a * c) (b * c) = normalize (lcm b a * c)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
509  | 
using lcm_mult_left [of c a b] by (simp add: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
510  | 
|
| 63489 | 511  | 
lemma lcm_mult_unit1: "is_unit a \<Longrightarrow> lcm (b * a) c = lcm b c"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
512  | 
by (rule associated_eqI) (simp_all add: mult_unit_dvd_iff dvd_lcmI1)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
513  | 
|
| 63489 | 514  | 
lemma lcm_mult_unit2: "is_unit a \<Longrightarrow> lcm b (c * a) = lcm b c"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
515  | 
using lcm_mult_unit1 [of a c b] by (simp add: ac_simps)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
516  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
517  | 
lemma lcm_div_unit1:  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
518  | 
"is_unit a \<Longrightarrow> lcm (b div a) c = lcm b c"  | 
| 63489 | 519  | 
by (erule is_unitE [of _ b]) (simp add: lcm_mult_unit1)  | 
520  | 
||
521  | 
lemma lcm_div_unit2: "is_unit a \<Longrightarrow> lcm b (c div a) = lcm b c"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
522  | 
by (erule is_unitE [of _ c]) (simp add: lcm_mult_unit2)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
523  | 
|
| 64850 | 524  | 
lemma normalize_lcm_left: "lcm (normalize a) b = lcm a b"  | 
525  | 
by (fact lcm.normalize_left_idem)  | 
|
526  | 
||
527  | 
lemma normalize_lcm_right: "lcm a (normalize b) = lcm a b"  | 
|
528  | 
by (fact lcm.normalize_right_idem)  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
529  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
530  | 
lemma comp_fun_idem_gcd: "comp_fun_idem gcd"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
531  | 
by standard (simp_all add: fun_eq_iff ac_simps)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
532  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
533  | 
lemma comp_fun_idem_lcm: "comp_fun_idem lcm"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
534  | 
by standard (simp_all add: fun_eq_iff ac_simps)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
535  | 
|
| 63489 | 536  | 
lemma gcd_dvd_antisym: "gcd a b dvd gcd c d \<Longrightarrow> gcd c d dvd gcd a b \<Longrightarrow> gcd a b = gcd c d"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
537  | 
proof (rule gcdI)  | 
| 63489 | 538  | 
assume *: "gcd a b dvd gcd c d"  | 
539  | 
and **: "gcd c d dvd gcd a b"  | 
|
540  | 
have "gcd c d dvd c"  | 
|
541  | 
by simp  | 
|
542  | 
with * show "gcd a b dvd c"  | 
|
543  | 
by (rule dvd_trans)  | 
|
544  | 
have "gcd c d dvd d"  | 
|
545  | 
by simp  | 
|
546  | 
with * show "gcd a b dvd d"  | 
|
547  | 
by (rule dvd_trans)  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
548  | 
show "normalize (gcd a b) = gcd a b"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
549  | 
by simp  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
550  | 
fix l assume "l dvd c" and "l dvd d"  | 
| 63489 | 551  | 
then have "l dvd gcd c d"  | 
552  | 
by (rule gcd_greatest)  | 
|
553  | 
from this and ** show "l dvd gcd a b"  | 
|
554  | 
by (rule dvd_trans)  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
555  | 
qed  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
556  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
557  | 
declare unit_factor_lcm [simp]  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
558  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
559  | 
lemma lcmI:  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
560  | 
assumes "a dvd c" and "b dvd c" and "\<And>d. a dvd d \<Longrightarrow> b dvd d \<Longrightarrow> c dvd d"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
561  | 
and "normalize c = c"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
562  | 
shows "c = lcm a b"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
563  | 
by (rule associated_eqI) (auto simp: assms intro: lcm_least)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
564  | 
|
| 63489 | 565  | 
lemma gcd_dvd_lcm [simp]: "gcd a b dvd lcm a b"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
566  | 
using gcd_dvd2 by (rule dvd_lcmI2)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
567  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
568  | 
lemmas lcm_0 = lcm_0_right  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
569  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
570  | 
lemma lcm_unique:  | 
| 63489 | 571  | 
"a dvd d \<and> b dvd d \<and> normalize d = d \<and> (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
572  | 
by rule (auto intro: lcmI simp: lcm_least lcm_eq_0_iff)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
573  | 
|
| 68708 | 574  | 
lemma lcm_proj1_if_dvd:  | 
575  | 
assumes "b dvd a" shows "lcm a b = normalize a"  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
576  | 
proof -  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
577  | 
have "normalize (lcm a b) = normalize a"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
578  | 
by (rule associatedI) (use assms in auto)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
579  | 
thus ?thesis by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
580  | 
qed  | 
| 63489 | 581  | 
|
582  | 
lemma lcm_proj2_if_dvd: "a dvd b \<Longrightarrow> lcm a b = normalize b"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
583  | 
using lcm_proj1_if_dvd [of a b] by (simp add: ac_simps)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
584  | 
|
| 63489 | 585  | 
lemma lcm_proj1_iff: "lcm m n = normalize m \<longleftrightarrow> n dvd m"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
586  | 
proof  | 
| 63489 | 587  | 
assume *: "lcm m n = normalize m"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
588  | 
show "n dvd m"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
589  | 
proof (cases "m = 0")  | 
| 63489 | 590  | 
case True  | 
591  | 
then show ?thesis by simp  | 
|
592  | 
next  | 
|
593  | 
case [simp]: False  | 
|
594  | 
from * have **: "m = lcm m n * unit_factor m"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
595  | 
by (simp add: unit_eq_div2)  | 
| 63489 | 596  | 
show ?thesis by (subst **) simp  | 
597  | 
qed  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
598  | 
next  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
599  | 
assume "n dvd m"  | 
| 63489 | 600  | 
then show "lcm m n = normalize m"  | 
601  | 
by (rule lcm_proj1_if_dvd)  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
602  | 
qed  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
603  | 
|
| 63489 | 604  | 
lemma lcm_proj2_iff: "lcm m n = normalize n \<longleftrightarrow> m dvd n"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
605  | 
using lcm_proj1_iff [of n m] by (simp add: ac_simps)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
606  | 
|
| 
69785
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
607  | 
lemma gcd_mono: "a dvd c \<Longrightarrow> b dvd d \<Longrightarrow> gcd a b dvd gcd c d"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
608  | 
by (simp add: gcd_dvdI1 gcd_dvdI2)  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
609  | 
|
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
610  | 
lemma lcm_mono: "a dvd c \<Longrightarrow> b dvd d \<Longrightarrow> lcm a b dvd lcm c d"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
611  | 
by (simp add: dvd_lcmI1 dvd_lcmI2)  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
612  | 
|
| 63924 | 613  | 
lemma dvd_productE:  | 
| 67051 | 614  | 
assumes "p dvd a * b"  | 
| 63924 | 615  | 
obtains x y where "p = x * y" "x dvd a" "y dvd b"  | 
616  | 
proof (cases "a = 0")  | 
|
617  | 
case True  | 
|
618  | 
thus ?thesis by (intro that[of p 1]) simp_all  | 
|
619  | 
next  | 
|
620  | 
case False  | 
|
621  | 
define x y where "x = gcd a p" and "y = p div x"  | 
|
622  | 
have "p = x * y" by (simp add: x_def y_def)  | 
|
623  | 
moreover have "x dvd a" by (simp add: x_def)  | 
|
624  | 
moreover from assms have "p dvd gcd (b * a) (b * p)"  | 
|
625  | 
by (intro gcd_greatest) (simp_all add: mult.commute)  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
626  | 
hence "p dvd b * gcd a p" by (subst (asm) gcd_mult_left) auto  | 
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
627  | 
with False have "y dvd b"  | 
| 63924 | 628  | 
by (simp add: x_def y_def div_dvd_iff_mult assms)  | 
629  | 
ultimately show ?thesis by (rule that)  | 
|
630  | 
qed  | 
|
631  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
632  | 
lemma gcd_mult_unit1:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
633  | 
assumes "is_unit a" shows "gcd (b * a) c = gcd b c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
634  | 
proof -  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
635  | 
have "gcd (b * a) c dvd b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
636  | 
using assms dvd_mult_unit_iff by blast  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
637  | 
then show ?thesis  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
638  | 
by (rule gcdI) simp_all  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
639  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
640  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
641  | 
lemma gcd_mult_unit2: "is_unit a \<Longrightarrow> gcd b (c * a) = gcd b c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
642  | 
using gcd.commute gcd_mult_unit1 by auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
643  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
644  | 
lemma gcd_div_unit1: "is_unit a \<Longrightarrow> gcd (b div a) c = gcd b c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
645  | 
by (erule is_unitE [of _ b]) (simp add: gcd_mult_unit1)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
646  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
647  | 
lemma gcd_div_unit2: "is_unit a \<Longrightarrow> gcd b (c div a) = gcd b c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
648  | 
by (erule is_unitE [of _ c]) (simp add: gcd_mult_unit2)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
649  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
650  | 
lemma normalize_gcd_left: "gcd (normalize a) b = gcd a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
651  | 
by (fact gcd.normalize_left_idem)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
652  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
653  | 
lemma normalize_gcd_right: "gcd a (normalize b) = gcd a b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
654  | 
by (fact gcd.normalize_right_idem)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
655  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
656  | 
lemma gcd_add1 [simp]: "gcd (m + n) n = gcd m n"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
657  | 
by (rule gcdI [symmetric]) (simp_all add: dvd_add_left_iff)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
658  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
659  | 
lemma gcd_add2 [simp]: "gcd m (m + n) = gcd m n"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
660  | 
using gcd_add1 [of n m] by (simp add: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
661  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
662  | 
lemma gcd_add_mult: "gcd m (k * m + n) = gcd m n"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
663  | 
by (rule gcdI [symmetric]) (simp_all add: dvd_add_right_iff)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
664  | 
|
| 60686 | 665  | 
end  | 
666  | 
||
| 62345 | 667  | 
class ring_gcd = comm_ring_1 + semiring_gcd  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
668  | 
begin  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
669  | 
|
| 63489 | 670  | 
lemma gcd_neg1 [simp]: "gcd (-a) b = gcd a b"  | 
671  | 
by (rule sym, rule gcdI) (simp_all add: gcd_greatest)  | 
|
672  | 
||
673  | 
lemma gcd_neg2 [simp]: "gcd a (-b) = gcd a b"  | 
|
674  | 
by (rule sym, rule gcdI) (simp_all add: gcd_greatest)  | 
|
675  | 
||
676  | 
lemma gcd_neg_numeral_1 [simp]: "gcd (- numeral n) a = gcd (numeral n) a"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
677  | 
by (fact gcd_neg1)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
678  | 
|
| 63489 | 679  | 
lemma gcd_neg_numeral_2 [simp]: "gcd a (- numeral n) = gcd a (numeral n)"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
680  | 
by (fact gcd_neg2)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
681  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
682  | 
lemma gcd_diff1: "gcd (m - n) n = gcd m n"  | 
| 63489 | 683  | 
by (subst diff_conv_add_uminus, subst gcd_neg2[symmetric], subst gcd_add1, simp)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
684  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
685  | 
lemma gcd_diff2: "gcd (n - m) n = gcd m n"  | 
| 63489 | 686  | 
by (subst gcd_neg1[symmetric]) (simp only: minus_diff_eq gcd_diff1)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
687  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
688  | 
lemma lcm_neg1 [simp]: "lcm (-a) b = lcm a b"  | 
| 63489 | 689  | 
by (rule sym, rule lcmI) (simp_all add: lcm_least lcm_eq_0_iff)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
690  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
691  | 
lemma lcm_neg2 [simp]: "lcm a (-b) = lcm a b"  | 
| 63489 | 692  | 
by (rule sym, rule lcmI) (simp_all add: lcm_least lcm_eq_0_iff)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
693  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
694  | 
lemma lcm_neg_numeral_1 [simp]: "lcm (- numeral n) a = lcm (numeral n) a"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
695  | 
by (fact lcm_neg1)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
696  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
697  | 
lemma lcm_neg_numeral_2 [simp]: "lcm a (- numeral n) = lcm a (numeral n)"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
698  | 
by (fact lcm_neg2)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
699  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
700  | 
end  | 
| 62345 | 701  | 
|
| 60686 | 702  | 
class semiring_Gcd = semiring_gcd + Gcd +  | 
703  | 
assumes Gcd_dvd: "a \<in> A \<Longrightarrow> Gcd A dvd a"  | 
|
704  | 
and Gcd_greatest: "(\<And>b. b \<in> A \<Longrightarrow> a dvd b) \<Longrightarrow> a dvd Gcd A"  | 
|
705  | 
and normalize_Gcd [simp]: "normalize (Gcd A) = Gcd A"  | 
|
| 62345 | 706  | 
assumes dvd_Lcm: "a \<in> A \<Longrightarrow> a dvd Lcm A"  | 
707  | 
and Lcm_least: "(\<And>b. b \<in> A \<Longrightarrow> b dvd a) \<Longrightarrow> Lcm A dvd a"  | 
|
708  | 
and normalize_Lcm [simp]: "normalize (Lcm A) = Lcm A"  | 
|
| 60686 | 709  | 
begin  | 
710  | 
||
| 63489 | 711  | 
lemma Lcm_Gcd: "Lcm A = Gcd {b. \<forall>a\<in>A. a dvd b}"
 | 
| 62345 | 712  | 
by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least)  | 
713  | 
||
| 63489 | 714  | 
lemma Gcd_Lcm: "Gcd A = Lcm {b. \<forall>a\<in>A. b dvd a}"
 | 
| 62345 | 715  | 
by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least)  | 
716  | 
||
| 63489 | 717  | 
lemma Gcd_empty [simp]: "Gcd {} = 0"
 | 
| 60686 | 718  | 
by (rule dvd_0_left, rule Gcd_greatest) simp  | 
719  | 
||
| 63489 | 720  | 
lemma Lcm_empty [simp]: "Lcm {} = 1"
 | 
| 62345 | 721  | 
by (auto intro: associated_eqI Lcm_least)  | 
722  | 
||
| 63489 | 723  | 
lemma Gcd_insert [simp]: "Gcd (insert a A) = gcd a (Gcd A)"  | 
| 62345 | 724  | 
proof -  | 
725  | 
have "Gcd (insert a A) dvd gcd a (Gcd A)"  | 
|
726  | 
by (auto intro: Gcd_dvd Gcd_greatest)  | 
|
727  | 
moreover have "gcd a (Gcd A) dvd Gcd (insert a A)"  | 
|
728  | 
proof (rule Gcd_greatest)  | 
|
729  | 
fix b  | 
|
730  | 
assume "b \<in> insert a A"  | 
|
731  | 
then show "gcd a (Gcd A) dvd b"  | 
|
732  | 
proof  | 
|
| 63489 | 733  | 
assume "b = a"  | 
734  | 
then show ?thesis  | 
|
735  | 
by simp  | 
|
| 62345 | 736  | 
next  | 
737  | 
assume "b \<in> A"  | 
|
| 63489 | 738  | 
then have "Gcd A dvd b"  | 
739  | 
by (rule Gcd_dvd)  | 
|
740  | 
moreover have "gcd a (Gcd A) dvd Gcd A"  | 
|
741  | 
by simp  | 
|
742  | 
ultimately show ?thesis  | 
|
743  | 
by (blast intro: dvd_trans)  | 
|
| 62345 | 744  | 
qed  | 
745  | 
qed  | 
|
746  | 
ultimately show ?thesis  | 
|
747  | 
by (auto intro: associated_eqI)  | 
|
748  | 
qed  | 
|
749  | 
||
| 63489 | 750  | 
lemma Lcm_insert [simp]: "Lcm (insert a A) = lcm a (Lcm A)"  | 
| 62345 | 751  | 
proof (rule sym)  | 
752  | 
have "lcm a (Lcm A) dvd Lcm (insert a A)"  | 
|
753  | 
by (auto intro: dvd_Lcm Lcm_least)  | 
|
754  | 
moreover have "Lcm (insert a A) dvd lcm a (Lcm A)"  | 
|
755  | 
proof (rule Lcm_least)  | 
|
756  | 
fix b  | 
|
757  | 
assume "b \<in> insert a A"  | 
|
758  | 
then show "b dvd lcm a (Lcm A)"  | 
|
759  | 
proof  | 
|
| 63489 | 760  | 
assume "b = a"  | 
761  | 
then show ?thesis by simp  | 
|
| 62345 | 762  | 
next  | 
763  | 
assume "b \<in> A"  | 
|
| 63489 | 764  | 
then have "b dvd Lcm A"  | 
765  | 
by (rule dvd_Lcm)  | 
|
766  | 
moreover have "Lcm A dvd lcm a (Lcm A)"  | 
|
767  | 
by simp  | 
|
768  | 
ultimately show ?thesis  | 
|
769  | 
by (blast intro: dvd_trans)  | 
|
| 62345 | 770  | 
qed  | 
771  | 
qed  | 
|
772  | 
ultimately show "lcm a (Lcm A) = Lcm (insert a A)"  | 
|
773  | 
by (rule associated_eqI) (simp_all add: lcm_eq_0_iff)  | 
|
774  | 
qed  | 
|
775  | 
||
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
776  | 
lemma LcmI:  | 
| 63489 | 777  | 
assumes "\<And>a. a \<in> A \<Longrightarrow> a dvd b"  | 
778  | 
and "\<And>c. (\<And>a. a \<in> A \<Longrightarrow> a dvd c) \<Longrightarrow> b dvd c"  | 
|
779  | 
and "normalize b = b"  | 
|
780  | 
shows "b = Lcm A"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
781  | 
by (rule associated_eqI) (auto simp: assms dvd_Lcm intro: Lcm_least)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
782  | 
|
| 63489 | 783  | 
lemma Lcm_subset: "A \<subseteq> B \<Longrightarrow> Lcm A dvd Lcm B"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
784  | 
by (blast intro: Lcm_least dvd_Lcm)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
785  | 
|
| 63489 | 786  | 
lemma Lcm_Un: "Lcm (A \<union> B) = lcm (Lcm A) (Lcm B)"  | 
| 68708 | 787  | 
proof -  | 
788  | 
have "\<And>d. \<lbrakk>Lcm A dvd d; Lcm B dvd d\<rbrakk> \<Longrightarrow> Lcm (A \<union> B) dvd d"  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
789  | 
by (meson UnE Lcm_least dvd_Lcm dvd_trans)  | 
| 68708 | 790  | 
then show ?thesis  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
791  | 
by (meson Lcm_subset lcm_unique normalize_Lcm sup.cobounded1 sup.cobounded2)  | 
| 68708 | 792  | 
qed  | 
| 63489 | 793  | 
|
794  | 
lemma Gcd_0_iff [simp]: "Gcd A = 0 \<longleftrightarrow> A \<subseteq> {0}"
 | 
|
795  | 
(is "?P \<longleftrightarrow> ?Q")  | 
|
| 60686 | 796  | 
proof  | 
797  | 
assume ?P  | 
|
798  | 
show ?Q  | 
|
799  | 
proof  | 
|
800  | 
fix a  | 
|
801  | 
assume "a \<in> A"  | 
|
| 63489 | 802  | 
then have "Gcd A dvd a"  | 
803  | 
by (rule Gcd_dvd)  | 
|
804  | 
with \<open>?P\<close> have "a = 0"  | 
|
805  | 
by simp  | 
|
806  | 
    then show "a \<in> {0}"
 | 
|
807  | 
by simp  | 
|
| 60686 | 808  | 
qed  | 
809  | 
next  | 
|
810  | 
assume ?Q  | 
|
811  | 
have "0 dvd Gcd A"  | 
|
812  | 
proof (rule Gcd_greatest)  | 
|
813  | 
fix a  | 
|
814  | 
assume "a \<in> A"  | 
|
| 63489 | 815  | 
with \<open>?Q\<close> have "a = 0"  | 
816  | 
by auto  | 
|
817  | 
then show "0 dvd a"  | 
|
818  | 
by simp  | 
|
| 60686 | 819  | 
qed  | 
| 63489 | 820  | 
then show ?P  | 
821  | 
by simp  | 
|
| 60686 | 822  | 
qed  | 
823  | 
||
| 63489 | 824  | 
lemma Lcm_1_iff [simp]: "Lcm A = 1 \<longleftrightarrow> (\<forall>a\<in>A. is_unit a)"  | 
825  | 
(is "?P \<longleftrightarrow> ?Q")  | 
|
| 60686 | 826  | 
proof  | 
827  | 
assume ?P  | 
|
828  | 
show ?Q  | 
|
829  | 
proof  | 
|
830  | 
fix a  | 
|
831  | 
assume "a \<in> A"  | 
|
832  | 
then have "a dvd Lcm A"  | 
|
833  | 
by (rule dvd_Lcm)  | 
|
834  | 
with \<open>?P\<close> show "is_unit a"  | 
|
835  | 
by simp  | 
|
836  | 
qed  | 
|
837  | 
next  | 
|
838  | 
assume ?Q  | 
|
839  | 
then have "is_unit (Lcm A)"  | 
|
840  | 
by (blast intro: Lcm_least)  | 
|
841  | 
then have "normalize (Lcm A) = 1"  | 
|
842  | 
by (rule is_unit_normalize)  | 
|
843  | 
then show ?P  | 
|
844  | 
by simp  | 
|
845  | 
qed  | 
|
846  | 
||
| 63489 | 847  | 
lemma unit_factor_Lcm: "unit_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)"  | 
| 62345 | 848  | 
proof (cases "Lcm A = 0")  | 
| 63489 | 849  | 
case True  | 
850  | 
then show ?thesis  | 
|
851  | 
by simp  | 
|
| 62345 | 852  | 
next  | 
853  | 
case False  | 
|
854  | 
with unit_factor_normalize have "unit_factor (normalize (Lcm A)) = 1"  | 
|
855  | 
by blast  | 
|
856  | 
with False show ?thesis  | 
|
857  | 
by simp  | 
|
858  | 
qed  | 
|
859  | 
||
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
860  | 
lemma unit_factor_Gcd: "unit_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)"  | 
| 63489 | 861  | 
by (simp add: Gcd_Lcm unit_factor_Lcm)  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
862  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
863  | 
lemma GcdI:  | 
| 63489 | 864  | 
assumes "\<And>a. a \<in> A \<Longrightarrow> b dvd a"  | 
865  | 
and "\<And>c. (\<And>a. a \<in> A \<Longrightarrow> c dvd a) \<Longrightarrow> c dvd b"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
866  | 
and "normalize b = b"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
867  | 
shows "b = Gcd A"  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
868  | 
by (rule associated_eqI) (auto simp: assms Gcd_dvd intro: Gcd_greatest)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
869  | 
|
| 62345 | 870  | 
lemma Gcd_eq_1_I:  | 
871  | 
assumes "is_unit a" and "a \<in> A"  | 
|
872  | 
shows "Gcd A = 1"  | 
|
873  | 
proof -  | 
|
874  | 
from assms have "is_unit (Gcd A)"  | 
|
875  | 
by (blast intro: Gcd_dvd dvd_unit_imp_unit)  | 
|
876  | 
then have "normalize (Gcd A) = 1"  | 
|
877  | 
by (rule is_unit_normalize)  | 
|
878  | 
then show ?thesis  | 
|
879  | 
by simp  | 
|
880  | 
qed  | 
|
881  | 
||
| 60686 | 882  | 
lemma Lcm_eq_0_I:  | 
883  | 
assumes "0 \<in> A"  | 
|
884  | 
shows "Lcm A = 0"  | 
|
885  | 
proof -  | 
|
886  | 
from assms have "0 dvd Lcm A"  | 
|
887  | 
by (rule dvd_Lcm)  | 
|
888  | 
then show ?thesis  | 
|
889  | 
by simp  | 
|
890  | 
qed  | 
|
891  | 
||
| 63489 | 892  | 
lemma Gcd_UNIV [simp]: "Gcd UNIV = 1"  | 
| 62345 | 893  | 
using dvd_refl by (rule Gcd_eq_1_I) simp  | 
894  | 
||
| 63489 | 895  | 
lemma Lcm_UNIV [simp]: "Lcm UNIV = 0"  | 
| 61929 | 896  | 
by (rule Lcm_eq_0_I) simp  | 
| 60686 | 897  | 
|
| 61929 | 898  | 
lemma Lcm_0_iff:  | 
899  | 
assumes "finite A"  | 
|
900  | 
shows "Lcm A = 0 \<longleftrightarrow> 0 \<in> A"  | 
|
901  | 
proof (cases "A = {}")
 | 
|
| 63489 | 902  | 
case True  | 
903  | 
then show ?thesis by simp  | 
|
| 61929 | 904  | 
next  | 
| 63489 | 905  | 
case False  | 
906  | 
with assms show ?thesis  | 
|
| 68708 | 907  | 
by (induct A rule: finite_ne_induct) (auto simp: lcm_eq_0_iff)  | 
| 60686 | 908  | 
qed  | 
| 61929 | 909  | 
|
| 63489 | 910  | 
lemma Gcd_image_normalize [simp]: "Gcd (normalize ` A) = Gcd A"  | 
| 62345 | 911  | 
proof -  | 
912  | 
have "Gcd (normalize ` A) dvd a" if "a \<in> A" for a  | 
|
913  | 
proof -  | 
|
| 63489 | 914  | 
from that obtain B where "A = insert a B"  | 
915  | 
by blast  | 
|
| 62350 | 916  | 
moreover have "gcd (normalize a) (Gcd (normalize ` B)) dvd normalize a"  | 
| 62345 | 917  | 
by (rule gcd_dvd1)  | 
918  | 
ultimately show "Gcd (normalize ` A) dvd a"  | 
|
919  | 
by simp  | 
|
920  | 
qed  | 
|
921  | 
then have "Gcd (normalize ` A) dvd Gcd A" and "Gcd A dvd Gcd (normalize ` A)"  | 
|
922  | 
by (auto intro!: Gcd_greatest intro: Gcd_dvd)  | 
|
923  | 
then show ?thesis  | 
|
924  | 
by (auto intro: associated_eqI)  | 
|
925  | 
qed  | 
|
926  | 
||
| 62346 | 927  | 
lemma Gcd_eqI:  | 
928  | 
assumes "normalize a = a"  | 
|
929  | 
assumes "\<And>b. b \<in> A \<Longrightarrow> a dvd b"  | 
|
930  | 
and "\<And>c. (\<And>b. b \<in> A \<Longrightarrow> c dvd b) \<Longrightarrow> c dvd a"  | 
|
931  | 
shows "Gcd A = a"  | 
|
932  | 
using assms by (blast intro: associated_eqI Gcd_greatest Gcd_dvd normalize_Gcd)  | 
|
933  | 
||
| 63489 | 934  | 
lemma dvd_GcdD: "x dvd Gcd A \<Longrightarrow> y \<in> A \<Longrightarrow> x dvd y"  | 
935  | 
using Gcd_dvd dvd_trans by blast  | 
|
936  | 
||
937  | 
lemma dvd_Gcd_iff: "x dvd Gcd A \<longleftrightarrow> (\<forall>y\<in>A. x dvd y)"  | 
|
| 63359 | 938  | 
by (blast dest: dvd_GcdD intro: Gcd_greatest)  | 
939  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
940  | 
lemma Gcd_mult: "Gcd ((*) c ` A) = normalize (c * Gcd A)"  | 
| 63359 | 941  | 
proof (cases "c = 0")  | 
| 63489 | 942  | 
case True  | 
943  | 
then show ?thesis by auto  | 
|
944  | 
next  | 
|
| 63359 | 945  | 
case [simp]: False  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
69038 
diff
changeset
 | 
946  | 
have "Gcd ((*) c ` A) div c dvd Gcd A"  | 
| 63359 | 947  | 
by (intro Gcd_greatest, subst div_dvd_iff_mult)  | 
948  | 
(auto intro!: Gcd_greatest Gcd_dvd simp: mult.commute[of _ c])  | 
|
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
69038 
diff
changeset
 | 
949  | 
then have "Gcd ((*) c ` A) dvd c * Gcd A"  | 
| 63359 | 950  | 
by (subst (asm) div_dvd_iff_mult) (auto intro: Gcd_greatest simp: mult_ac)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
951  | 
moreover have "c * Gcd A dvd Gcd ((*) c ` A)"  | 
| 63359 | 952  | 
by (intro Gcd_greatest) (auto intro: mult_dvd_mono Gcd_dvd)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
953  | 
ultimately have "normalize (Gcd ((*) c ` A)) = normalize (c * Gcd A)"  | 
| 63359 | 954  | 
by (rule associatedI)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
955  | 
then show ?thesis by simp  | 
| 63489 | 956  | 
qed  | 
| 63359 | 957  | 
|
| 62346 | 958  | 
lemma Lcm_eqI:  | 
959  | 
assumes "normalize a = a"  | 
|
| 63489 | 960  | 
and "\<And>b. b \<in> A \<Longrightarrow> b dvd a"  | 
| 62346 | 961  | 
and "\<And>c. (\<And>b. b \<in> A \<Longrightarrow> b dvd c) \<Longrightarrow> a dvd c"  | 
962  | 
shows "Lcm A = a"  | 
|
963  | 
using assms by (blast intro: associated_eqI Lcm_least dvd_Lcm normalize_Lcm)  | 
|
964  | 
||
| 63489 | 965  | 
lemma Lcm_dvdD: "Lcm A dvd x \<Longrightarrow> y \<in> A \<Longrightarrow> y dvd x"  | 
966  | 
using dvd_Lcm dvd_trans by blast  | 
|
967  | 
||
968  | 
lemma Lcm_dvd_iff: "Lcm A dvd x \<longleftrightarrow> (\<forall>y\<in>A. y dvd x)"  | 
|
| 63359 | 969  | 
by (blast dest: Lcm_dvdD intro: Lcm_least)  | 
970  | 
||
| 63489 | 971  | 
lemma Lcm_mult:  | 
| 63359 | 972  | 
  assumes "A \<noteq> {}"
 | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
973  | 
shows "Lcm ((*) c ` A) = normalize (c * Lcm A)"  | 
| 63359 | 974  | 
proof (cases "c = 0")  | 
975  | 
case True  | 
|
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
69038 
diff
changeset
 | 
976  | 
  with assms have "(*) c ` A = {0}"
 | 
| 63489 | 977  | 
by auto  | 
978  | 
with True show ?thesis by auto  | 
|
| 63359 | 979  | 
next  | 
980  | 
case [simp]: False  | 
|
| 63489 | 981  | 
from assms obtain x where x: "x \<in> A"  | 
982  | 
by blast  | 
|
983  | 
have "c dvd c * x"  | 
|
984  | 
by simp  | 
|
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
69038 
diff
changeset
 | 
985  | 
also from x have "c * x dvd Lcm ((*) c ` A)"  | 
| 63489 | 986  | 
by (intro dvd_Lcm) auto  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
69038 
diff
changeset
 | 
987  | 
finally have dvd: "c dvd Lcm ((*) c ` A)" .  | 
| 69768 | 988  | 
moreover have "Lcm A dvd Lcm ((*) c ` A) div c"  | 
| 63359 | 989  | 
by (intro Lcm_least dvd_mult_imp_div)  | 
| 63489 | 990  | 
(auto intro!: Lcm_least dvd_Lcm simp: mult.commute[of _ c])  | 
| 69768 | 991  | 
ultimately have "c * Lcm A dvd Lcm ((*) c ` A)"  | 
992  | 
by auto  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
993  | 
moreover have "Lcm ((*) c ` A) dvd c * Lcm A"  | 
| 63359 | 994  | 
by (intro Lcm_least) (auto intro: mult_dvd_mono dvd_Lcm)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
995  | 
ultimately have "normalize (c * Lcm A) = normalize (Lcm ((*) c ` A))"  | 
| 63359 | 996  | 
by (rule associatedI)  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
997  | 
then show ?thesis by simp  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
998  | 
qed  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
999  | 
|
| 63489 | 1000  | 
lemma Lcm_no_units: "Lcm A = Lcm (A - {a. is_unit a})"
 | 
1001  | 
proof -  | 
|
1002  | 
  have "(A - {a. is_unit a}) \<union> {a\<in>A. is_unit a} = A"
 | 
|
1003  | 
by blast  | 
|
1004  | 
  then have "Lcm A = lcm (Lcm (A - {a. is_unit a})) (Lcm {a\<in>A. is_unit a})"
 | 
|
1005  | 
by (simp add: Lcm_Un [symmetric])  | 
|
1006  | 
  also have "Lcm {a\<in>A. is_unit a} = 1"
 | 
|
1007  | 
by simp  | 
|
1008  | 
finally show ?thesis  | 
|
1009  | 
by simp  | 
|
1010  | 
qed  | 
|
1011  | 
||
1012  | 
lemma Lcm_0_iff': "Lcm A = 0 \<longleftrightarrow> (\<nexists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l))"  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1013  | 
by (metis Lcm_least dvd_0_left dvd_Lcm)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1014  | 
|
| 63489 | 1015  | 
lemma Lcm_no_multiple: "(\<forall>m. m \<noteq> 0 \<longrightarrow> (\<exists>a\<in>A. \<not> a dvd m)) \<Longrightarrow> Lcm A = 0"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1016  | 
by (auto simp: Lcm_0_iff')  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1017  | 
|
| 63489 | 1018  | 
lemma Lcm_singleton [simp]: "Lcm {a} = normalize a"
 | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1019  | 
by simp  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1020  | 
|
| 63489 | 1021  | 
lemma Lcm_2 [simp]: "Lcm {a, b} = lcm a b"
 | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1022  | 
by simp  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1023  | 
|
| 63489 | 1024  | 
lemma Gcd_1: "1 \<in> A \<Longrightarrow> Gcd A = 1"  | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1025  | 
by (auto intro!: Gcd_eq_1_I)  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1026  | 
|
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1027  | 
lemma Gcd_singleton [simp]: "Gcd {a} = normalize a"
 | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1028  | 
by simp  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1029  | 
|
| 63489 | 1030  | 
lemma Gcd_2 [simp]: "Gcd {a, b} = gcd a b"
 | 
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1031  | 
by simp  | 
| 
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
1032  | 
|
| 
69785
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1033  | 
lemma Gcd_mono:  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1034  | 
assumes "\<And>x. x \<in> A \<Longrightarrow> f x dvd g x"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1035  | 
shows "(GCD x\<in>A. f x) dvd (GCD x\<in>A. g x)"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1036  | 
proof (intro Gcd_greatest, safe)  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1037  | 
fix x assume "x \<in> A"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1038  | 
hence "(GCD x\<in>A. f x) dvd f x"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1039  | 
by (intro Gcd_dvd) auto  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1040  | 
also have "f x dvd g x"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1041  | 
using \<open>x \<in> A\<close> assms by blast  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1042  | 
finally show "(GCD x\<in>A. f x) dvd \<dots>" .  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1043  | 
qed  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1044  | 
|
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1045  | 
lemma Lcm_mono:  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1046  | 
assumes "\<And>x. x \<in> A \<Longrightarrow> f x dvd g x"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1047  | 
shows "(LCM x\<in>A. f x) dvd (LCM x\<in>A. g x)"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1048  | 
proof (intro Lcm_least, safe)  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1049  | 
fix x assume "x \<in> A"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1050  | 
hence "f x dvd g x" by (rule assms)  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1051  | 
also have "g x dvd (LCM x\<in>A. g x)"  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1052  | 
using \<open>x \<in> A\<close> by (intro dvd_Lcm) auto  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1053  | 
finally show "f x dvd \<dots>" .  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1054  | 
qed  | 
| 
 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69768 
diff
changeset
 | 
1055  | 
|
| 62350 | 1056  | 
end  | 
| 62345 | 1057  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
1058  | 
|
| 64850 | 1059  | 
subsection \<open>An aside: GCD and LCM on finite sets for incomplete gcd rings\<close>  | 
1060  | 
||
1061  | 
context semiring_gcd  | 
|
1062  | 
begin  | 
|
1063  | 
||
1064  | 
sublocale Gcd_fin: bounded_quasi_semilattice_set gcd 0 1 normalize  | 
|
1065  | 
defines  | 
|
| 69038 | 1066  | 
  Gcd_fin ("Gcd\<^sub>f\<^sub>i\<^sub>n") = "Gcd_fin.F :: 'a set \<Rightarrow> 'a" ..
 | 
| 64850 | 1067  | 
|
1068  | 
abbreviation gcd_list :: "'a list \<Rightarrow> 'a"  | 
|
1069  | 
where "gcd_list xs \<equiv> Gcd\<^sub>f\<^sub>i\<^sub>n (set xs)"  | 
|
1070  | 
||
1071  | 
sublocale Lcm_fin: bounded_quasi_semilattice_set lcm 1 0 normalize  | 
|
1072  | 
defines  | 
|
| 69038 | 1073  | 
  Lcm_fin ("Lcm\<^sub>f\<^sub>i\<^sub>n") = Lcm_fin.F ..
 | 
| 64850 | 1074  | 
|
1075  | 
abbreviation lcm_list :: "'a list \<Rightarrow> 'a"  | 
|
1076  | 
where "lcm_list xs \<equiv> Lcm\<^sub>f\<^sub>i\<^sub>n (set xs)"  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
1077  | 
|
| 64850 | 1078  | 
lemma Gcd_fin_dvd:  | 
1079  | 
"a \<in> A \<Longrightarrow> Gcd\<^sub>f\<^sub>i\<^sub>n A dvd a"  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
1080  | 
by (induct A rule: infinite_finite_induct)  | 
| 64850 | 1081  | 
(auto intro: dvd_trans)  | 
1082  | 
||
1083  | 
lemma dvd_Lcm_fin:  | 
|
1084  | 
"a \<in> A \<Longrightarrow> a dvd Lcm\<^sub>f\<^sub>i\<^sub>n A"  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
1085  | 
by (induct A rule: infinite_finite_induct)  | 
| 64850 | 1086  | 
(auto intro: dvd_trans)  | 
1087  | 
||
1088  | 
lemma Gcd_fin_greatest:  | 
|
1089  | 
"a dvd Gcd\<^sub>f\<^sub>i\<^sub>n A" if "finite A" and "\<And>b. b \<in> A \<Longrightarrow> a dvd b"  | 
|
1090  | 
using that by (induct A) simp_all  | 
|
1091  | 
||
1092  | 
lemma Lcm_fin_least:  | 
|
1093  | 
"Lcm\<^sub>f\<^sub>i\<^sub>n A dvd a" if "finite A" and "\<And>b. b \<in> A \<Longrightarrow> b dvd a"  | 
|
1094  | 
using that by (induct A) simp_all  | 
|
1095  | 
||
1096  | 
lemma gcd_list_greatest:  | 
|
1097  | 
"a dvd gcd_list bs" if "\<And>b. b \<in> set bs \<Longrightarrow> a dvd b"  | 
|
1098  | 
by (rule Gcd_fin_greatest) (simp_all add: that)  | 
|
1099  | 
||
1100  | 
lemma lcm_list_least:  | 
|
1101  | 
"lcm_list bs dvd a" if "\<And>b. b \<in> set bs \<Longrightarrow> b dvd a"  | 
|
1102  | 
by (rule Lcm_fin_least) (simp_all add: that)  | 
|
1103  | 
||
1104  | 
lemma dvd_Gcd_fin_iff:  | 
|
1105  | 
"b dvd Gcd\<^sub>f\<^sub>i\<^sub>n A \<longleftrightarrow> (\<forall>a\<in>A. b dvd a)" if "finite A"  | 
|
1106  | 
using that by (auto intro: Gcd_fin_greatest Gcd_fin_dvd dvd_trans [of b "Gcd\<^sub>f\<^sub>i\<^sub>n A"])  | 
|
1107  | 
||
1108  | 
lemma dvd_gcd_list_iff:  | 
|
1109  | 
"b dvd gcd_list xs \<longleftrightarrow> (\<forall>a\<in>set xs. b dvd a)"  | 
|
1110  | 
by (simp add: dvd_Gcd_fin_iff)  | 
|
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
64850 
diff
changeset
 | 
1111  | 
|
| 64850 | 1112  | 
lemma Lcm_fin_dvd_iff:  | 
1113  | 
"Lcm\<^sub>f\<^sub>i\<^sub>n A dvd b \<longleftrightarrow> (\<forall>a\<in>A. a dvd b)" if "finite A"  | 
|
1114  | 
using that by (auto intro: Lcm_fin_least dvd_Lcm_fin dvd_trans [of _ "Lcm\<^sub>f\<^sub>i\<^sub>n A" b])  | 
|
1115  | 
||
1116  | 
lemma lcm_list_dvd_iff:  | 
|
1117  | 
"lcm_list xs dvd b \<longleftrightarrow> (\<forall>a\<in>set xs. a dvd b)"  | 
|
1118  | 
by (simp add: Lcm_fin_dvd_iff)  | 
|
1119  | 
||
1120  | 
lemma Gcd_fin_mult:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1121  | 
"Gcd\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize (b * Gcd\<^sub>f\<^sub>i\<^sub>n A)" if "finite A"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1122  | 
using that by induction (auto simp: gcd_mult_left)  | 
| 64850 | 1123  | 
|
1124  | 
lemma Lcm_fin_mult:  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1125  | 
  "Lcm\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize (b * Lcm\<^sub>f\<^sub>i\<^sub>n A)" if "A \<noteq> {}"
 | 
| 64850 | 1126  | 
proof (cases "b = 0")  | 
1127  | 
case True  | 
|
1128  | 
  moreover from that have "times 0 ` A = {0}"
 | 
|
1129  | 
by auto  | 
|
1130  | 
ultimately show ?thesis  | 
|
1131  | 
by simp  | 
|
1132  | 
next  | 
|
1133  | 
case False  | 
|
1134  | 
show ?thesis proof (cases "finite A")  | 
|
1135  | 
case False  | 
|
| 66936 | 1136  | 
moreover have "inj_on (times b) A"  | 
1137  | 
using \<open>b \<noteq> 0\<close> by (rule inj_on_mult)  | 
|
| 64850 | 1138  | 
ultimately have "infinite (times b ` A)"  | 
1139  | 
by (simp add: finite_image_iff)  | 
|
1140  | 
with False show ?thesis  | 
|
1141  | 
by simp  | 
|
1142  | 
next  | 
|
1143  | 
case True  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1144  | 
then show ?thesis using that  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1145  | 
by (induct A rule: finite_ne_induct) (auto simp: lcm_mult_left)  | 
| 64850 | 1146  | 
qed  | 
1147  | 
qed  | 
|
1148  | 
||
| 65811 | 1149  | 
lemma unit_factor_Gcd_fin:  | 
1150  | 
"unit_factor (Gcd\<^sub>f\<^sub>i\<^sub>n A) = of_bool (Gcd\<^sub>f\<^sub>i\<^sub>n A \<noteq> 0)"  | 
|
1151  | 
by (rule normalize_idem_imp_unit_factor_eq) simp  | 
|
1152  | 
||
1153  | 
lemma unit_factor_Lcm_fin:  | 
|
1154  | 
"unit_factor (Lcm\<^sub>f\<^sub>i\<^sub>n A) = of_bool (Lcm\<^sub>f\<^sub>i\<^sub>n A \<noteq> 0)"  | 
|
1155  | 
by (rule normalize_idem_imp_unit_factor_eq) simp  | 
|
1156  | 
||
1157  | 
lemma is_unit_Gcd_fin_iff [simp]:  | 
|
1158  | 
"is_unit (Gcd\<^sub>f\<^sub>i\<^sub>n A) \<longleftrightarrow> Gcd\<^sub>f\<^sub>i\<^sub>n A = 1"  | 
|
1159  | 
by (rule normalize_idem_imp_is_unit_iff) simp  | 
|
1160  | 
||
1161  | 
lemma is_unit_Lcm_fin_iff [simp]:  | 
|
1162  | 
"is_unit (Lcm\<^sub>f\<^sub>i\<^sub>n A) \<longleftrightarrow> Lcm\<^sub>f\<^sub>i\<^sub>n A = 1"  | 
|
1163  | 
by (rule normalize_idem_imp_is_unit_iff) simp  | 
|
1164  | 
||
1165  | 
lemma Gcd_fin_0_iff:  | 
|
1166  | 
  "Gcd\<^sub>f\<^sub>i\<^sub>n A = 0 \<longleftrightarrow> A \<subseteq> {0} \<and> finite A"
 | 
|
1167  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
|
1168  | 
||
1169  | 
lemma Lcm_fin_0_iff:  | 
|
1170  | 
"Lcm\<^sub>f\<^sub>i\<^sub>n A = 0 \<longleftrightarrow> 0 \<in> A" if "finite A"  | 
|
| 68708 | 1171  | 
using that by (induct A) (auto simp: lcm_eq_0_iff)  | 
| 65811 | 1172  | 
|
1173  | 
lemma Lcm_fin_1_iff:  | 
|
1174  | 
"Lcm\<^sub>f\<^sub>i\<^sub>n A = 1 \<longleftrightarrow> (\<forall>a\<in>A. is_unit a) \<and> finite A"  | 
|
1175  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
|
1176  | 
||
| 64850 | 1177  | 
end  | 
1178  | 
||
1179  | 
context semiring_Gcd  | 
|
1180  | 
begin  | 
|
1181  | 
||
1182  | 
lemma Gcd_fin_eq_Gcd [simp]:  | 
|
1183  | 
"Gcd\<^sub>f\<^sub>i\<^sub>n A = Gcd A" if "finite A" for A :: "'a set"  | 
|
1184  | 
using that by induct simp_all  | 
|
1185  | 
||
1186  | 
lemma Gcd_set_eq_fold [code_unfold]:  | 
|
1187  | 
"Gcd (set xs) = fold gcd xs 0"  | 
|
1188  | 
by (simp add: Gcd_fin.set_eq_fold [symmetric])  | 
|
1189  | 
||
1190  | 
lemma Lcm_fin_eq_Lcm [simp]:  | 
|
1191  | 
"Lcm\<^sub>f\<^sub>i\<^sub>n A = Lcm A" if "finite A" for A :: "'a set"  | 
|
1192  | 
using that by induct simp_all  | 
|
1193  | 
||
1194  | 
lemma Lcm_set_eq_fold [code_unfold]:  | 
|
1195  | 
"Lcm (set xs) = fold lcm xs 1"  | 
|
1196  | 
by (simp add: Lcm_fin.set_eq_fold [symmetric])  | 
|
1197  | 
||
1198  | 
end  | 
|
| 63489 | 1199  | 
|
| 67051 | 1200  | 
|
1201  | 
subsection \<open>Coprimality\<close>  | 
|
1202  | 
||
1203  | 
context semiring_gcd  | 
|
1204  | 
begin  | 
|
1205  | 
||
1206  | 
lemma coprime_imp_gcd_eq_1 [simp]:  | 
|
1207  | 
"gcd a b = 1" if "coprime a b"  | 
|
1208  | 
proof -  | 
|
1209  | 
define t r s where "t = gcd a b" and "r = a div t" and "s = b div t"  | 
|
1210  | 
then have "a = t * r" and "b = t * s"  | 
|
1211  | 
by simp_all  | 
|
1212  | 
with that have "coprime (t * r) (t * s)"  | 
|
1213  | 
by simp  | 
|
1214  | 
then show ?thesis  | 
|
1215  | 
by (simp add: t_def)  | 
|
1216  | 
qed  | 
|
1217  | 
||
| 
68270
 
2bc921b2159b
treat gcd_eq_1_imp_coprime analogously to mod_0_imp_dvd
 
haftmann 
parents: 
67399 
diff
changeset
 | 
1218  | 
lemma gcd_eq_1_imp_coprime [dest!]:  | 
| 67051 | 1219  | 
"coprime a b" if "gcd a b = 1"  | 
1220  | 
proof (rule coprimeI)  | 
|
1221  | 
fix c  | 
|
1222  | 
assume "c dvd a" and "c dvd b"  | 
|
1223  | 
then have "c dvd gcd a b"  | 
|
1224  | 
by (rule gcd_greatest)  | 
|
1225  | 
with that show "is_unit c"  | 
|
1226  | 
by simp  | 
|
1227  | 
qed  | 
|
1228  | 
||
1229  | 
lemma coprime_iff_gcd_eq_1 [presburger, code]:  | 
|
1230  | 
"coprime a b \<longleftrightarrow> gcd a b = 1"  | 
|
1231  | 
by rule (simp_all add: gcd_eq_1_imp_coprime)  | 
|
1232  | 
||
1233  | 
lemma is_unit_gcd [simp]:  | 
|
1234  | 
"is_unit (gcd a b) \<longleftrightarrow> coprime a b"  | 
|
1235  | 
by (simp add: coprime_iff_gcd_eq_1)  | 
|
1236  | 
||
1237  | 
lemma coprime_add_one_left [simp]: "coprime (a + 1) a"  | 
|
1238  | 
by (simp add: gcd_eq_1_imp_coprime ac_simps)  | 
|
1239  | 
||
1240  | 
lemma coprime_add_one_right [simp]: "coprime a (a + 1)"  | 
|
1241  | 
using coprime_add_one_left [of a] by (simp add: ac_simps)  | 
|
1242  | 
||
1243  | 
lemma coprime_mult_left_iff [simp]:  | 
|
1244  | 
"coprime (a * b) c \<longleftrightarrow> coprime a c \<and> coprime b c"  | 
|
1245  | 
proof  | 
|
1246  | 
assume "coprime (a * b) c"  | 
|
1247  | 
with coprime_common_divisor [of "a * b" c]  | 
|
1248  | 
have *: "is_unit d" if "d dvd a * b" and "d dvd c" for d  | 
|
1249  | 
using that by blast  | 
|
1250  | 
have "coprime a c"  | 
|
1251  | 
by (rule coprimeI, rule *) simp_all  | 
|
1252  | 
moreover have "coprime b c"  | 
|
1253  | 
by (rule coprimeI, rule *) simp_all  | 
|
1254  | 
ultimately show "coprime a c \<and> coprime b c" ..  | 
|
1255  | 
next  | 
|
1256  | 
assume "coprime a c \<and> coprime b c"  | 
|
1257  | 
then have "coprime a c" "coprime b c"  | 
|
1258  | 
by simp_all  | 
|
1259  | 
show "coprime (a * b) c"  | 
|
1260  | 
proof (rule coprimeI)  | 
|
1261  | 
fix d  | 
|
1262  | 
assume "d dvd a * b"  | 
|
1263  | 
then obtain r s where d: "d = r * s" "r dvd a" "s dvd b"  | 
|
1264  | 
by (rule dvd_productE)  | 
|
1265  | 
assume "d dvd c"  | 
|
1266  | 
with d have "r * s dvd c"  | 
|
1267  | 
by simp  | 
|
1268  | 
then have "r dvd c" "s dvd c"  | 
|
1269  | 
by (auto intro: dvd_mult_left dvd_mult_right)  | 
|
1270  | 
from \<open>coprime a c\<close> \<open>r dvd a\<close> \<open>r dvd c\<close>  | 
|
1271  | 
have "is_unit r"  | 
|
1272  | 
by (rule coprime_common_divisor)  | 
|
1273  | 
moreover from \<open>coprime b c\<close> \<open>s dvd b\<close> \<open>s dvd c\<close>  | 
|
1274  | 
have "is_unit s"  | 
|
1275  | 
by (rule coprime_common_divisor)  | 
|
1276  | 
ultimately show "is_unit d"  | 
|
1277  | 
by (simp add: d is_unit_mult_iff)  | 
|
1278  | 
qed  | 
|
1279  | 
qed  | 
|
1280  | 
||
1281  | 
lemma coprime_mult_right_iff [simp]:  | 
|
1282  | 
"coprime c (a * b) \<longleftrightarrow> coprime c a \<and> coprime c b"  | 
|
1283  | 
using coprime_mult_left_iff [of a b c] by (simp add: ac_simps)  | 
|
1284  | 
||
1285  | 
lemma coprime_power_left_iff [simp]:  | 
|
1286  | 
"coprime (a ^ n) b \<longleftrightarrow> coprime a b \<or> n = 0"  | 
|
1287  | 
proof (cases "n = 0")  | 
|
1288  | 
case True  | 
|
1289  | 
then show ?thesis  | 
|
1290  | 
by simp  | 
|
1291  | 
next  | 
|
1292  | 
case False  | 
|
1293  | 
then have "n > 0"  | 
|
1294  | 
by simp  | 
|
1295  | 
then show ?thesis  | 
|
1296  | 
by (induction n rule: nat_induct_non_zero) simp_all  | 
|
1297  | 
qed  | 
|
1298  | 
||
1299  | 
lemma coprime_power_right_iff [simp]:  | 
|
1300  | 
"coprime a (b ^ n) \<longleftrightarrow> coprime a b \<or> n = 0"  | 
|
1301  | 
using coprime_power_left_iff [of b n a] by (simp add: ac_simps)  | 
|
1302  | 
||
1303  | 
lemma prod_coprime_left:  | 
|
1304  | 
"coprime (\<Prod>i\<in>A. f i) a" if "\<And>i. i \<in> A \<Longrightarrow> coprime (f i) a"  | 
|
1305  | 
using that by (induct A rule: infinite_finite_induct) simp_all  | 
|
1306  | 
||
1307  | 
lemma prod_coprime_right:  | 
|
1308  | 
"coprime a (\<Prod>i\<in>A. f i)" if "\<And>i. i \<in> A \<Longrightarrow> coprime a (f i)"  | 
|
1309  | 
using that prod_coprime_left [of A f a] by (simp add: ac_simps)  | 
|
1310  | 
||
1311  | 
lemma prod_list_coprime_left:  | 
|
1312  | 
"coprime (prod_list xs) a" if "\<And>x. x \<in> set xs \<Longrightarrow> coprime x a"  | 
|
1313  | 
using that by (induct xs) simp_all  | 
|
1314  | 
||
1315  | 
lemma prod_list_coprime_right:  | 
|
1316  | 
"coprime a (prod_list xs)" if "\<And>x. x \<in> set xs \<Longrightarrow> coprime a x"  | 
|
1317  | 
using that prod_list_coprime_left [of xs a] by (simp add: ac_simps)  | 
|
1318  | 
||
1319  | 
lemma coprime_dvd_mult_left_iff:  | 
|
1320  | 
"a dvd b * c \<longleftrightarrow> a dvd b" if "coprime a c"  | 
|
1321  | 
proof  | 
|
1322  | 
assume "a dvd b"  | 
|
1323  | 
then show "a dvd b * c"  | 
|
1324  | 
by simp  | 
|
1325  | 
next  | 
|
1326  | 
assume "a dvd b * c"  | 
|
1327  | 
show "a dvd b"  | 
|
1328  | 
proof (cases "b = 0")  | 
|
1329  | 
case True  | 
|
1330  | 
then show ?thesis  | 
|
1331  | 
by simp  | 
|
1332  | 
next  | 
|
1333  | 
case False  | 
|
1334  | 
then have unit: "is_unit (unit_factor b)"  | 
|
1335  | 
by simp  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1336  | 
from \<open>coprime a c\<close>  | 
| 67051 | 1337  | 
have "gcd (b * a) (b * c) * unit_factor b = b"  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1338  | 
by (subst gcd_mult_left) (simp add: ac_simps)  | 
| 67051 | 1339  | 
moreover from \<open>a dvd b * c\<close>  | 
1340  | 
have "a dvd gcd (b * a) (b * c) * unit_factor b"  | 
|
1341  | 
by (simp add: dvd_mult_unit_iff unit)  | 
|
1342  | 
ultimately show ?thesis  | 
|
1343  | 
by simp  | 
|
1344  | 
qed  | 
|
1345  | 
qed  | 
|
1346  | 
||
1347  | 
lemma coprime_dvd_mult_right_iff:  | 
|
1348  | 
"a dvd c * b \<longleftrightarrow> a dvd b" if "coprime a c"  | 
|
1349  | 
using that coprime_dvd_mult_left_iff [of a c b] by (simp add: ac_simps)  | 
|
1350  | 
||
1351  | 
lemma divides_mult:  | 
|
1352  | 
"a * b dvd c" if "a dvd c" and "b dvd c" and "coprime a b"  | 
|
1353  | 
proof -  | 
|
1354  | 
from \<open>b dvd c\<close> obtain b' where "c = b * b'" ..  | 
|
1355  | 
with \<open>a dvd c\<close> have "a dvd b' * b"  | 
|
1356  | 
by (simp add: ac_simps)  | 
|
1357  | 
with \<open>coprime a b\<close> have "a dvd b'"  | 
|
1358  | 
by (simp add: coprime_dvd_mult_left_iff)  | 
|
1359  | 
then obtain a' where "b' = a * a'" ..  | 
|
1360  | 
with \<open>c = b * b'\<close> have "c = (a * b) * a'"  | 
|
1361  | 
by (simp add: ac_simps)  | 
|
1362  | 
then show ?thesis ..  | 
|
1363  | 
qed  | 
|
1364  | 
||
1365  | 
lemma div_gcd_coprime:  | 
|
1366  | 
assumes "a \<noteq> 0 \<or> b \<noteq> 0"  | 
|
1367  | 
shows "coprime (a div gcd a b) (b div gcd a b)"  | 
|
1368  | 
proof -  | 
|
1369  | 
let ?g = "gcd a b"  | 
|
1370  | 
let ?a' = "a div ?g"  | 
|
1371  | 
let ?b' = "b div ?g"  | 
|
1372  | 
let ?g' = "gcd ?a' ?b'"  | 
|
1373  | 
have dvdg: "?g dvd a" "?g dvd b"  | 
|
1374  | 
by simp_all  | 
|
1375  | 
have dvdg': "?g' dvd ?a'" "?g' dvd ?b'"  | 
|
1376  | 
by simp_all  | 
|
1377  | 
from dvdg dvdg' obtain ka kb ka' kb' where  | 
|
1378  | 
kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"  | 
|
1379  | 
unfolding dvd_def by blast  | 
|
1380  | 
from this [symmetric] have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'"  | 
|
1381  | 
by (simp_all add: mult.assoc mult.left_commute [of "gcd a b"])  | 
|
1382  | 
then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"  | 
|
| 68708 | 1383  | 
by (auto simp: dvd_mult_div_cancel [OF dvdg(1)] dvd_mult_div_cancel [OF dvdg(2)] dvd_def)  | 
| 67051 | 1384  | 
have "?g \<noteq> 0"  | 
1385  | 
using assms by simp  | 
|
1386  | 
moreover from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .  | 
|
1387  | 
ultimately show ?thesis  | 
|
1388  | 
using dvd_times_left_cancel_iff [of "gcd a b" _ 1]  | 
|
1389  | 
by simp (simp only: coprime_iff_gcd_eq_1)  | 
|
1390  | 
qed  | 
|
1391  | 
||
1392  | 
lemma gcd_coprime:  | 
|
1393  | 
assumes c: "gcd a b \<noteq> 0"  | 
|
1394  | 
and a: "a = a' * gcd a b"  | 
|
1395  | 
and b: "b = b' * gcd a b"  | 
|
1396  | 
shows "coprime a' b'"  | 
|
1397  | 
proof -  | 
|
1398  | 
from c have "a \<noteq> 0 \<or> b \<noteq> 0"  | 
|
1399  | 
by simp  | 
|
1400  | 
with div_gcd_coprime have "coprime (a div gcd a b) (b div gcd a b)" .  | 
|
1401  | 
also from assms have "a div gcd a b = a'"  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1402  | 
using dvd_div_eq_mult gcd_dvd1 by blast  | 
| 67051 | 1403  | 
also from assms have "b div gcd a b = b'"  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1404  | 
using dvd_div_eq_mult gcd_dvd1 by blast  | 
| 67051 | 1405  | 
finally show ?thesis .  | 
1406  | 
qed  | 
|
1407  | 
||
1408  | 
lemma gcd_coprime_exists:  | 
|
1409  | 
assumes "gcd a b \<noteq> 0"  | 
|
1410  | 
shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"  | 
|
| 68708 | 1411  | 
proof -  | 
1412  | 
have "coprime (a div gcd a b) (b div gcd a b)"  | 
|
1413  | 
using assms div_gcd_coprime by auto  | 
|
1414  | 
then show ?thesis  | 
|
1415  | 
by force  | 
|
1416  | 
qed  | 
|
| 67051 | 1417  | 
|
1418  | 
lemma pow_divides_pow_iff [simp]:  | 
|
1419  | 
"a ^ n dvd b ^ n \<longleftrightarrow> a dvd b" if "n > 0"  | 
|
1420  | 
proof (cases "gcd a b = 0")  | 
|
1421  | 
case True  | 
|
1422  | 
then show ?thesis  | 
|
1423  | 
by simp  | 
|
1424  | 
next  | 
|
1425  | 
case False  | 
|
1426  | 
show ?thesis  | 
|
1427  | 
proof  | 
|
1428  | 
let ?d = "gcd a b"  | 
|
1429  | 
from \<open>n > 0\<close> obtain m where m: "n = Suc m"  | 
|
1430  | 
by (cases n) simp_all  | 
|
1431  | 
from False have zn: "?d ^ n \<noteq> 0"  | 
|
1432  | 
by (rule power_not_zero)  | 
|
1433  | 
from gcd_coprime_exists [OF False]  | 
|
1434  | 
obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "coprime a' b'"  | 
|
1435  | 
by blast  | 
|
1436  | 
assume "a ^ n dvd b ^ n"  | 
|
1437  | 
then have "(a' * ?d) ^ n dvd (b' * ?d) ^ n"  | 
|
1438  | 
by (simp add: ab'(1,2)[symmetric])  | 
|
1439  | 
then have "?d^n * a'^n dvd ?d^n * b'^n"  | 
|
1440  | 
by (simp only: power_mult_distrib ac_simps)  | 
|
1441  | 
with zn have "a' ^ n dvd b' ^ n"  | 
|
1442  | 
by simp  | 
|
1443  | 
then have "a' dvd b' ^ n"  | 
|
1444  | 
using dvd_trans[of a' "a'^n" "b'^n"] by (simp add: m)  | 
|
1445  | 
then have "a' dvd b' ^ m * b'"  | 
|
1446  | 
by (simp add: m ac_simps)  | 
|
1447  | 
moreover have "coprime a' (b' ^ n)"  | 
|
1448  | 
using \<open>coprime a' b'\<close> by simp  | 
|
1449  | 
then have "a' dvd b'"  | 
|
1450  | 
using \<open>a' dvd b' ^ n\<close> coprime_dvd_mult_left_iff dvd_mult by blast  | 
|
1451  | 
then have "a' * ?d dvd b' * ?d"  | 
|
1452  | 
by (rule mult_dvd_mono) simp  | 
|
1453  | 
with ab'(1,2) show "a dvd b"  | 
|
1454  | 
by simp  | 
|
1455  | 
next  | 
|
1456  | 
assume "a dvd b"  | 
|
1457  | 
with \<open>n > 0\<close> show "a ^ n dvd b ^ n"  | 
|
1458  | 
by (induction rule: nat_induct_non_zero)  | 
|
1459  | 
(simp_all add: mult_dvd_mono)  | 
|
1460  | 
qed  | 
|
1461  | 
qed  | 
|
1462  | 
||
1463  | 
lemma coprime_crossproduct:  | 
|
1464  | 
fixes a b c d :: 'a  | 
|
1465  | 
assumes "coprime a d" and "coprime b c"  | 
|
1466  | 
shows "normalize a * normalize c = normalize b * normalize d \<longleftrightarrow>  | 
|
1467  | 
normalize a = normalize b \<and> normalize c = normalize d"  | 
|
1468  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
1469  | 
proof  | 
|
1470  | 
assume ?rhs  | 
|
1471  | 
then show ?lhs by simp  | 
|
1472  | 
next  | 
|
1473  | 
assume ?lhs  | 
|
1474  | 
from \<open>?lhs\<close> have "normalize a dvd normalize b * normalize d"  | 
|
1475  | 
by (auto intro: dvdI dest: sym)  | 
|
1476  | 
with \<open>coprime a d\<close> have "a dvd b"  | 
|
1477  | 
by (simp add: coprime_dvd_mult_left_iff normalize_mult [symmetric])  | 
|
1478  | 
from \<open>?lhs\<close> have "normalize b dvd normalize a * normalize c"  | 
|
1479  | 
by (auto intro: dvdI dest: sym)  | 
|
1480  | 
with \<open>coprime b c\<close> have "b dvd a"  | 
|
1481  | 
by (simp add: coprime_dvd_mult_left_iff normalize_mult [symmetric])  | 
|
1482  | 
from \<open>?lhs\<close> have "normalize c dvd normalize d * normalize b"  | 
|
1483  | 
by (auto intro: dvdI dest: sym simp add: mult.commute)  | 
|
1484  | 
with \<open>coprime b c\<close> have "c dvd d"  | 
|
1485  | 
by (simp add: coprime_dvd_mult_left_iff coprime_commute normalize_mult [symmetric])  | 
|
1486  | 
from \<open>?lhs\<close> have "normalize d dvd normalize c * normalize a"  | 
|
1487  | 
by (auto intro: dvdI dest: sym simp add: mult.commute)  | 
|
1488  | 
with \<open>coprime a d\<close> have "d dvd c"  | 
|
1489  | 
by (simp add: coprime_dvd_mult_left_iff coprime_commute normalize_mult [symmetric])  | 
|
1490  | 
from \<open>a dvd b\<close> \<open>b dvd a\<close> have "normalize a = normalize b"  | 
|
1491  | 
by (rule associatedI)  | 
|
1492  | 
moreover from \<open>c dvd d\<close> \<open>d dvd c\<close> have "normalize c = normalize d"  | 
|
1493  | 
by (rule associatedI)  | 
|
1494  | 
ultimately show ?rhs ..  | 
|
1495  | 
qed  | 
|
1496  | 
||
1497  | 
lemma gcd_mult_left_left_cancel:  | 
|
1498  | 
"gcd (c * a) b = gcd a b" if "coprime b c"  | 
|
1499  | 
proof -  | 
|
1500  | 
have "coprime (gcd b (a * c)) c"  | 
|
1501  | 
by (rule coprimeI) (auto intro: that coprime_common_divisor)  | 
|
1502  | 
then have "gcd b (a * c) dvd a"  | 
|
1503  | 
using coprime_dvd_mult_left_iff [of "gcd b (a * c)" c a]  | 
|
1504  | 
by simp  | 
|
1505  | 
then show ?thesis  | 
|
1506  | 
by (auto intro: associated_eqI simp add: ac_simps)  | 
|
1507  | 
qed  | 
|
1508  | 
||
1509  | 
lemma gcd_mult_left_right_cancel:  | 
|
1510  | 
"gcd (a * c) b = gcd a b" if "coprime b c"  | 
|
1511  | 
using that gcd_mult_left_left_cancel [of b c a]  | 
|
1512  | 
by (simp add: ac_simps)  | 
|
1513  | 
||
1514  | 
lemma gcd_mult_right_left_cancel:  | 
|
1515  | 
"gcd a (c * b) = gcd a b" if "coprime a c"  | 
|
1516  | 
using that gcd_mult_left_left_cancel [of a c b]  | 
|
1517  | 
by (simp add: ac_simps)  | 
|
1518  | 
||
1519  | 
lemma gcd_mult_right_right_cancel:  | 
|
1520  | 
"gcd a (b * c) = gcd a b" if "coprime a c"  | 
|
1521  | 
using that gcd_mult_right_left_cancel [of a c b]  | 
|
1522  | 
by (simp add: ac_simps)  | 
|
1523  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1524  | 
lemma gcd_exp_weak:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1525  | 
"gcd (a ^ n) (b ^ n) = normalize (gcd a b ^ n)"  | 
| 67051 | 1526  | 
proof (cases "a = 0 \<and> b = 0 \<or> n = 0")  | 
1527  | 
case True  | 
|
1528  | 
then show ?thesis  | 
|
1529  | 
by (cases n) simp_all  | 
|
1530  | 
next  | 
|
1531  | 
case False  | 
|
1532  | 
then have "coprime (a div gcd a b) (b div gcd a b)" and "n > 0"  | 
|
1533  | 
by (auto intro: div_gcd_coprime)  | 
|
1534  | 
then have "coprime ((a div gcd a b) ^ n) ((b div gcd a b) ^ n)"  | 
|
1535  | 
by simp  | 
|
1536  | 
then have "1 = gcd ((a div gcd a b) ^ n) ((b div gcd a b) ^ n)"  | 
|
1537  | 
by simp  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1538  | 
then have "normalize (gcd a b ^ n) = normalize (gcd a b ^ n * \<dots>)"  | 
| 67051 | 1539  | 
by simp  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1540  | 
also have "\<dots> = gcd (gcd a b ^ n * (a div gcd a b) ^ n) (gcd a b ^ n * (b div gcd a b) ^ n)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1541  | 
by (rule gcd_mult_left [symmetric])  | 
| 67051 | 1542  | 
also have "(gcd a b) ^ n * (a div gcd a b) ^ n = a ^ n"  | 
1543  | 
by (simp add: ac_simps div_power dvd_power_same)  | 
|
1544  | 
also have "(gcd a b) ^ n * (b div gcd a b) ^ n = b ^ n"  | 
|
1545  | 
by (simp add: ac_simps div_power dvd_power_same)  | 
|
1546  | 
finally show ?thesis by simp  | 
|
1547  | 
qed  | 
|
1548  | 
||
1549  | 
lemma division_decomp:  | 
|
1550  | 
assumes "a dvd b * c"  | 
|
1551  | 
shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"  | 
|
1552  | 
proof (cases "gcd a b = 0")  | 
|
1553  | 
case True  | 
|
1554  | 
then have "a = 0 \<and> b = 0"  | 
|
1555  | 
by simp  | 
|
1556  | 
then have "a = 0 * c \<and> 0 dvd b \<and> c dvd c"  | 
|
1557  | 
by simp  | 
|
1558  | 
then show ?thesis by blast  | 
|
1559  | 
next  | 
|
1560  | 
case False  | 
|
1561  | 
let ?d = "gcd a b"  | 
|
1562  | 
from gcd_coprime_exists [OF False]  | 
|
1563  | 
obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "coprime a' b'"  | 
|
1564  | 
by blast  | 
|
1565  | 
from ab'(1) have "a' dvd a" ..  | 
|
1566  | 
with assms have "a' dvd b * c"  | 
|
1567  | 
using dvd_trans [of a' a "b * c"] by simp  | 
|
1568  | 
from assms ab'(1,2) have "a' * ?d dvd (b' * ?d) * c"  | 
|
1569  | 
by simp  | 
|
1570  | 
then have "?d * a' dvd ?d * (b' * c)"  | 
|
1571  | 
by (simp add: mult_ac)  | 
|
1572  | 
with \<open>?d \<noteq> 0\<close> have "a' dvd b' * c"  | 
|
1573  | 
by simp  | 
|
1574  | 
then have "a' dvd c"  | 
|
1575  | 
using \<open>coprime a' b'\<close> by (simp add: coprime_dvd_mult_right_iff)  | 
|
1576  | 
with ab'(1) have "a = ?d * a' \<and> ?d dvd b \<and> a' dvd c"  | 
|
1577  | 
by (simp add: ac_simps)  | 
|
1578  | 
then show ?thesis by blast  | 
|
1579  | 
qed  | 
|
1580  | 
||
1581  | 
lemma lcm_coprime: "coprime a b \<Longrightarrow> lcm a b = normalize (a * b)"  | 
|
1582  | 
by (subst lcm_gcd) simp  | 
|
1583  | 
||
1584  | 
end  | 
|
1585  | 
||
1586  | 
context ring_gcd  | 
|
1587  | 
begin  | 
|
1588  | 
||
1589  | 
lemma coprime_minus_left_iff [simp]:  | 
|
1590  | 
"coprime (- a) b \<longleftrightarrow> coprime a b"  | 
|
1591  | 
by (rule; rule coprimeI) (auto intro: coprime_common_divisor)  | 
|
1592  | 
||
1593  | 
lemma coprime_minus_right_iff [simp]:  | 
|
1594  | 
"coprime a (- b) \<longleftrightarrow> coprime a b"  | 
|
1595  | 
using coprime_minus_left_iff [of b a] by (simp add: ac_simps)  | 
|
1596  | 
||
1597  | 
lemma coprime_diff_one_left [simp]: "coprime (a - 1) a"  | 
|
1598  | 
using coprime_add_one_right [of "a - 1"] by simp  | 
|
1599  | 
||
1600  | 
lemma coprime_doff_one_right [simp]: "coprime a (a - 1)"  | 
|
1601  | 
using coprime_diff_one_left [of a] by (simp add: ac_simps)  | 
|
1602  | 
||
1603  | 
end  | 
|
1604  | 
||
1605  | 
context semiring_Gcd  | 
|
1606  | 
begin  | 
|
1607  | 
||
1608  | 
lemma Lcm_coprime:  | 
|
1609  | 
assumes "finite A"  | 
|
1610  | 
    and "A \<noteq> {}"
 | 
|
1611  | 
and "\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> coprime a b"  | 
|
1612  | 
shows "Lcm A = normalize (\<Prod>A)"  | 
|
1613  | 
using assms  | 
|
1614  | 
proof (induct rule: finite_ne_induct)  | 
|
1615  | 
case singleton  | 
|
1616  | 
then show ?case by simp  | 
|
1617  | 
next  | 
|
1618  | 
case (insert a A)  | 
|
1619  | 
have "Lcm (insert a A) = lcm a (Lcm A)"  | 
|
1620  | 
by simp  | 
|
1621  | 
also from insert have "Lcm A = normalize (\<Prod>A)"  | 
|
1622  | 
by blast  | 
|
1623  | 
also have "lcm a \<dots> = lcm a (\<Prod>A)"  | 
|
1624  | 
by (cases "\<Prod>A = 0") (simp_all add: lcm_div_unit2)  | 
|
1625  | 
also from insert have "coprime a (\<Prod>A)"  | 
|
1626  | 
by (subst coprime_commute, intro prod_coprime_left) auto  | 
|
1627  | 
with insert have "lcm a (\<Prod>A) = normalize (\<Prod>(insert a A))"  | 
|
1628  | 
by (simp add: lcm_coprime)  | 
|
1629  | 
finally show ?case .  | 
|
1630  | 
qed  | 
|
1631  | 
||
1632  | 
lemma Lcm_coprime':  | 
|
1633  | 
"card A \<noteq> 0 \<Longrightarrow>  | 
|
1634  | 
(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> coprime a b) \<Longrightarrow>  | 
|
1635  | 
Lcm A = normalize (\<Prod>A)"  | 
|
1636  | 
by (rule Lcm_coprime) (simp_all add: card_eq_0_iff)  | 
|
1637  | 
||
1638  | 
end  | 
|
1639  | 
||
1640  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1641  | 
subsection \<open>GCD and LCM for multiplicative normalisation functions\<close>  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1642  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1643  | 
class semiring_gcd_mult_normalize = semiring_gcd + normalization_semidom_multiplicative  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1644  | 
begin  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1645  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1646  | 
lemma mult_gcd_left: "c * gcd a b = unit_factor c * gcd (c * a) (c * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1647  | 
by (simp add: gcd_mult_left normalize_mult mult.assoc [symmetric])  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1648  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1649  | 
lemma mult_gcd_right: "gcd a b * c = gcd (a * c) (b * c) * unit_factor c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1650  | 
using mult_gcd_left [of c a b] by (simp add: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1651  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1652  | 
lemma gcd_mult_distrib': "normalize c * gcd a b = gcd (c * a) (c * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1653  | 
by (subst gcd_mult_left) (simp_all add: normalize_mult)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1654  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1655  | 
lemma gcd_mult_distrib: "k * gcd a b = gcd (k * a) (k * b) * unit_factor k"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1656  | 
proof-  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1657  | 
have "normalize k * gcd a b = gcd (k * a) (k * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1658  | 
by (simp add: gcd_mult_distrib')  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1659  | 
then have "normalize k * gcd a b * unit_factor k = gcd (k * a) (k * b) * unit_factor k"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1660  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1661  | 
then have "normalize k * unit_factor k * gcd a b = gcd (k * a) (k * b) * unit_factor k"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1662  | 
by (simp only: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1663  | 
then show ?thesis  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1664  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1665  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1666  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1667  | 
lemma gcd_mult_lcm [simp]: "gcd a b * lcm a b = normalize a * normalize b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1668  | 
by (simp add: lcm_gcd normalize_mult dvd_normalize_div)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1669  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1670  | 
lemma lcm_mult_gcd [simp]: "lcm a b * gcd a b = normalize a * normalize b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1671  | 
using gcd_mult_lcm [of a b] by (simp add: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1672  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1673  | 
lemma mult_lcm_left: "c * lcm a b = unit_factor c * lcm (c * a) (c * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1674  | 
by (simp add: lcm_mult_left mult.assoc [symmetric] normalize_mult)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1675  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1676  | 
lemma mult_lcm_right: "lcm a b * c = lcm (a * c) (b * c) * unit_factor c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1677  | 
using mult_lcm_left [of c a b] by (simp add: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1678  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1679  | 
lemma lcm_gcd_prod: "lcm a b * gcd a b = normalize (a * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1680  | 
by (simp add: lcm_gcd dvd_normalize_div normalize_mult)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1681  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1682  | 
lemma lcm_mult_distrib': "normalize c * lcm a b = lcm (c * a) (c * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1683  | 
by (subst lcm_mult_left) (simp add: normalize_mult)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1684  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1685  | 
lemma lcm_mult_distrib: "k * lcm a b = lcm (k * a) (k * b) * unit_factor k"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1686  | 
proof-  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1687  | 
have "normalize k * lcm a b = lcm (k * a) (k * b)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1688  | 
by (simp add: lcm_mult_distrib')  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1689  | 
then have "normalize k * lcm a b * unit_factor k = lcm (k * a) (k * b) * unit_factor k"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1690  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1691  | 
then have "normalize k * unit_factor k * lcm a b = lcm (k * a) (k * b) * unit_factor k"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1692  | 
by (simp only: ac_simps)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1693  | 
then show ?thesis  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1694  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1695  | 
qed  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1696  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1697  | 
lemma coprime_crossproduct':  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1698  | 
fixes a b c d  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1699  | 
assumes "b \<noteq> 0"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1700  | 
assumes unit_factors: "unit_factor b = unit_factor d"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1701  | 
assumes coprime: "coprime a b" "coprime c d"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1702  | 
shows "a * d = b * c \<longleftrightarrow> a = c \<and> b = d"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1703  | 
proof safe  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1704  | 
assume eq: "a * d = b * c"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1705  | 
hence "normalize a * normalize d = normalize c * normalize b"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1706  | 
by (simp only: normalize_mult [symmetric] mult_ac)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1707  | 
with coprime have "normalize b = normalize d"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1708  | 
by (subst (asm) coprime_crossproduct) simp_all  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1709  | 
from this and unit_factors show "b = d"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1710  | 
by (rule normalize_unit_factor_eqI)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1711  | 
from eq have "a * d = c * d" by (simp only: \<open>b = d\<close> mult_ac)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1712  | 
with \<open>b \<noteq> 0\<close> \<open>b = d\<close> show "a = c" by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1713  | 
qed (simp_all add: mult_ac)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1714  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1715  | 
lemma gcd_exp [simp]:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1716  | 
"gcd (a ^ n) (b ^ n) = gcd a b ^ n"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1717  | 
using gcd_exp_weak[of a n b] by (simp add: normalize_power)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1718  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1719  | 
end  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1720  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1721  | 
|
| 69593 | 1722  | 
subsection \<open>GCD and LCM on \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close>\<close>  | 
| 59008 | 1723  | 
|
| 31706 | 1724  | 
instantiation nat :: gcd  | 
1725  | 
begin  | 
|
| 21256 | 1726  | 
|
| 62345 | 1727  | 
fun gcd_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 63489 | 1728  | 
where "gcd_nat x y = (if y = 0 then x else gcd y (x mod y))"  | 
| 31706 | 1729  | 
|
| 62345 | 1730  | 
definition lcm_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 63489 | 1731  | 
where "lcm_nat x y = x * y div (gcd x y)"  | 
1732  | 
||
1733  | 
instance ..  | 
|
| 31706 | 1734  | 
|
1735  | 
end  | 
|
1736  | 
||
1737  | 
instantiation int :: gcd  | 
|
1738  | 
begin  | 
|
| 21256 | 1739  | 
|
| 62345 | 1740  | 
definition gcd_int :: "int \<Rightarrow> int \<Rightarrow> int"  | 
1741  | 
where "gcd_int x y = int (gcd (nat \<bar>x\<bar>) (nat \<bar>y\<bar>))"  | 
|
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1742  | 
|
| 62345 | 1743  | 
definition lcm_int :: "int \<Rightarrow> int \<Rightarrow> int"  | 
1744  | 
where "lcm_int x y = int (lcm (nat \<bar>x\<bar>) (nat \<bar>y\<bar>))"  | 
|
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1745  | 
|
| 61944 | 1746  | 
instance ..  | 
| 31706 | 1747  | 
|
1748  | 
end  | 
|
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1749  | 
|
| 67118 | 1750  | 
lemma gcd_int_int_eq [simp]:  | 
1751  | 
"gcd (int m) (int n) = int (gcd m n)"  | 
|
1752  | 
by (simp add: gcd_int_def)  | 
|
1753  | 
||
1754  | 
lemma gcd_nat_abs_left_eq [simp]:  | 
|
1755  | 
"gcd (nat \<bar>k\<bar>) n = nat (gcd k (int n))"  | 
|
1756  | 
by (simp add: gcd_int_def)  | 
|
1757  | 
||
1758  | 
lemma gcd_nat_abs_right_eq [simp]:  | 
|
1759  | 
"gcd n (nat \<bar>k\<bar>) = nat (gcd (int n) k)"  | 
|
1760  | 
by (simp add: gcd_int_def)  | 
|
1761  | 
||
1762  | 
lemma abs_gcd_int [simp]:  | 
|
1763  | 
"\<bar>gcd x y\<bar> = gcd x y"  | 
|
1764  | 
for x y :: int  | 
|
1765  | 
by (simp only: gcd_int_def)  | 
|
1766  | 
||
1767  | 
lemma gcd_abs1_int [simp]:  | 
|
1768  | 
"gcd \<bar>x\<bar> y = gcd x y"  | 
|
1769  | 
for x y :: int  | 
|
1770  | 
by (simp only: gcd_int_def) simp  | 
|
1771  | 
||
1772  | 
lemma gcd_abs2_int [simp]:  | 
|
1773  | 
"gcd x \<bar>y\<bar> = gcd x y"  | 
|
1774  | 
for x y :: int  | 
|
1775  | 
by (simp only: gcd_int_def) simp  | 
|
1776  | 
||
1777  | 
lemma lcm_int_int_eq [simp]:  | 
|
1778  | 
"lcm (int m) (int n) = int (lcm m n)"  | 
|
1779  | 
by (simp add: lcm_int_def)  | 
|
1780  | 
||
1781  | 
lemma lcm_nat_abs_left_eq [simp]:  | 
|
1782  | 
"lcm (nat \<bar>k\<bar>) n = nat (lcm k (int n))"  | 
|
1783  | 
by (simp add: lcm_int_def)  | 
|
1784  | 
||
1785  | 
lemma lcm_nat_abs_right_eq [simp]:  | 
|
1786  | 
"lcm n (nat \<bar>k\<bar>) = nat (lcm (int n) k)"  | 
|
1787  | 
by (simp add: lcm_int_def)  | 
|
1788  | 
||
1789  | 
lemma lcm_abs1_int [simp]:  | 
|
1790  | 
"lcm \<bar>x\<bar> y = lcm x y"  | 
|
1791  | 
for x y :: int  | 
|
1792  | 
by (simp only: lcm_int_def) simp  | 
|
1793  | 
||
1794  | 
lemma lcm_abs2_int [simp]:  | 
|
1795  | 
"lcm x \<bar>y\<bar> = lcm x y"  | 
|
1796  | 
for x y :: int  | 
|
1797  | 
by (simp only: lcm_int_def) simp  | 
|
1798  | 
||
1799  | 
lemma abs_lcm_int [simp]: "\<bar>lcm i j\<bar> = lcm i j"  | 
|
1800  | 
for i j :: int  | 
|
1801  | 
by (simp only: lcm_int_def)  | 
|
1802  | 
||
| 68708 | 1803  | 
lemma gcd_nat_induct [case_names base step]:  | 
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1804  | 
fixes m n :: nat  | 
| 
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1805  | 
assumes "\<And>m. P m 0"  | 
| 
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1806  | 
and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"  | 
| 
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1807  | 
shows "P m n"  | 
| 68708 | 1808  | 
proof (induction m n rule: gcd_nat.induct)  | 
1809  | 
case (1 x y)  | 
|
1810  | 
then show ?case  | 
|
1811  | 
using assms neq0_conv by blast  | 
|
1812  | 
qed  | 
|
| 63489 | 1813  | 
|
1814  | 
lemma gcd_neg1_int [simp]: "gcd (- x) y = gcd x y"  | 
|
1815  | 
for x y :: int  | 
|
| 67118 | 1816  | 
by (simp only: gcd_int_def) simp  | 
| 31706 | 1817  | 
|
| 63489 | 1818  | 
lemma gcd_neg2_int [simp]: "gcd x (- y) = gcd x y"  | 
1819  | 
for x y :: int  | 
|
| 67118 | 1820  | 
by (simp only: gcd_int_def) simp  | 
| 31706 | 1821  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31814 
diff
changeset
 | 
1822  | 
lemma gcd_cases_int:  | 
| 63489 | 1823  | 
fixes x y :: int  | 
1824  | 
assumes "x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> P (gcd x y)"  | 
|
1825  | 
and "x \<ge> 0 \<Longrightarrow> y \<le> 0 \<Longrightarrow> P (gcd x (- y))"  | 
|
1826  | 
and "x \<le> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> P (gcd (- x) y)"  | 
|
1827  | 
and "x \<le> 0 \<Longrightarrow> y \<le> 0 \<Longrightarrow> P (gcd (- x) (- y))"  | 
|
| 31706 | 1828  | 
shows "P (gcd x y)"  | 
| 63489 | 1829  | 
using assms by auto arith  | 
| 21256 | 1830  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31814 
diff
changeset
 | 
1831  | 
lemma gcd_ge_0_int [simp]: "gcd (x::int) y >= 0"  | 
| 63489 | 1832  | 
for x y :: int  | 
| 31706 | 1833  | 
by (simp add: gcd_int_def)  | 
1834  | 
||
| 63489 | 1835  | 
lemma lcm_neg1_int: "lcm (- x) y = lcm x y"  | 
1836  | 
for x y :: int  | 
|
| 67118 | 1837  | 
by (simp only: lcm_int_def) simp  | 
| 31706 | 1838  | 
|
| 63489 | 1839  | 
lemma lcm_neg2_int: "lcm x (- y) = lcm x y"  | 
1840  | 
for x y :: int  | 
|
| 67118 | 1841  | 
by (simp only: lcm_int_def) simp  | 
| 31814 | 1842  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31814 
diff
changeset
 | 
1843  | 
lemma lcm_cases_int:  | 
| 63489 | 1844  | 
fixes x y :: int  | 
1845  | 
assumes "x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> P (lcm x y)"  | 
|
1846  | 
and "x \<ge> 0 \<Longrightarrow> y \<le> 0 \<Longrightarrow> P (lcm x (- y))"  | 
|
1847  | 
and "x \<le> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> P (lcm (- x) y)"  | 
|
1848  | 
and "x \<le> 0 \<Longrightarrow> y \<le> 0 \<Longrightarrow> P (lcm (- x) (- y))"  | 
|
| 31706 | 1849  | 
shows "P (lcm x y)"  | 
| 68708 | 1850  | 
using assms by (auto simp: lcm_neg1_int lcm_neg2_int) arith  | 
| 31706 | 1851  | 
|
| 63489 | 1852  | 
lemma lcm_ge_0_int [simp]: "lcm x y \<ge> 0"  | 
1853  | 
for x y :: int  | 
|
| 67118 | 1854  | 
by (simp only: lcm_int_def)  | 
| 31706 | 1855  | 
|
| 63489 | 1856  | 
lemma gcd_0_nat: "gcd x 0 = x"  | 
1857  | 
for x :: nat  | 
|
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1858  | 
by simp  | 
| 
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1859  | 
|
| 63489 | 1860  | 
lemma gcd_0_int [simp]: "gcd x 0 = \<bar>x\<bar>"  | 
1861  | 
for x :: int  | 
|
1862  | 
by (auto simp: gcd_int_def)  | 
|
1863  | 
||
1864  | 
lemma gcd_0_left_nat: "gcd 0 x = x"  | 
|
1865  | 
for x :: nat  | 
|
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1866  | 
by simp  | 
| 
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
1867  | 
|
| 63489 | 1868  | 
lemma gcd_0_left_int [simp]: "gcd 0 x = \<bar>x\<bar>"  | 
1869  | 
for x :: int  | 
|
| 67118 | 1870  | 
by (auto simp: gcd_int_def)  | 
| 63489 | 1871  | 
|
1872  | 
lemma gcd_red_nat: "gcd x y = gcd y (x mod y)"  | 
|
1873  | 
for x y :: nat  | 
|
1874  | 
by (cases "y = 0") auto  | 
|
1875  | 
||
1876  | 
||
1877  | 
text \<open>Weaker, but useful for the simplifier.\<close>  | 
|
1878  | 
||
1879  | 
lemma gcd_non_0_nat: "y \<noteq> 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"  | 
|
1880  | 
for x y :: nat  | 
|
| 21263 | 1881  | 
by simp  | 
| 21256 | 1882  | 
|
| 63489 | 1883  | 
lemma gcd_1_nat [simp]: "gcd m 1 = 1"  | 
1884  | 
for m :: nat  | 
|
| 60690 | 1885  | 
by simp  | 
| 31706 | 1886  | 
|
| 63489 | 1887  | 
lemma gcd_Suc_0 [simp]: "gcd m (Suc 0) = Suc 0"  | 
1888  | 
for m :: nat  | 
|
1889  | 
by simp  | 
|
1890  | 
||
1891  | 
lemma gcd_1_int [simp]: "gcd m 1 = 1"  | 
|
1892  | 
for m :: int  | 
|
| 31706 | 1893  | 
by (simp add: gcd_int_def)  | 
| 
30082
 
43c5b7bfc791
make more proofs work whether or not One_nat_def is a simp rule
 
huffman 
parents: 
30042 
diff
changeset
 | 
1894  | 
|
| 63489 | 1895  | 
lemma gcd_idem_nat: "gcd x x = x"  | 
1896  | 
for x :: nat  | 
|
1897  | 
by simp  | 
|
1898  | 
||
1899  | 
lemma gcd_idem_int: "gcd x x = \<bar>x\<bar>"  | 
|
1900  | 
for x :: int  | 
|
| 68708 | 1901  | 
by (auto simp: gcd_int_def)  | 
| 31706 | 1902  | 
|
1903  | 
declare gcd_nat.simps [simp del]  | 
|
| 21256 | 1904  | 
|
| 60758 | 1905  | 
text \<open>  | 
| 69593 | 1906  | 
\<^medskip> \<^term>\<open>gcd m n\<close> divides \<open>m\<close> and \<open>n\<close>.  | 
| 63489 | 1907  | 
The conjunctions don't seem provable separately.  | 
| 60758 | 1908  | 
\<close>  | 
| 21256 | 1909  | 
|
| 59008 | 1910  | 
instance nat :: semiring_gcd  | 
1911  | 
proof  | 
|
1912  | 
fix m n :: nat  | 
|
1913  | 
show "gcd m n dvd m" and "gcd m n dvd n"  | 
|
1914  | 
proof (induct m n rule: gcd_nat_induct)  | 
|
| 68708 | 1915  | 
case (step m n)  | 
| 59008 | 1916  | 
then have "gcd n (m mod n) dvd m"  | 
| 68708 | 1917  | 
by (metis dvd_mod_imp_dvd)  | 
1918  | 
with step show "gcd m n dvd m"  | 
|
| 59008 | 1919  | 
by (simp add: gcd_non_0_nat)  | 
1920  | 
qed (simp_all add: gcd_0_nat gcd_non_0_nat)  | 
|
1921  | 
next  | 
|
1922  | 
fix m n k :: nat  | 
|
1923  | 
assume "k dvd m" and "k dvd n"  | 
|
1924  | 
then show "k dvd gcd m n"  | 
|
1925  | 
by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat dvd_mod gcd_0_nat)  | 
|
| 60686 | 1926  | 
qed (simp_all add: lcm_nat_def)  | 
| 
59667
 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59545 
diff
changeset
 | 
1927  | 
|
| 59008 | 1928  | 
instance int :: ring_gcd  | 
| 67118 | 1929  | 
proof  | 
1930  | 
fix k l r :: int  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1931  | 
show [simp]: "gcd k l dvd k" "gcd k l dvd l"  | 
| 67118 | 1932  | 
using gcd_dvd1 [of "nat \<bar>k\<bar>" "nat \<bar>l\<bar>"]  | 
1933  | 
gcd_dvd2 [of "nat \<bar>k\<bar>" "nat \<bar>l\<bar>"]  | 
|
1934  | 
by simp_all  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1935  | 
show "lcm k l = normalize (k * l div gcd k l)"  | 
| 67118 | 1936  | 
using lcm_gcd [of "nat \<bar>k\<bar>" "nat \<bar>l\<bar>"]  | 
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1937  | 
by (simp add: nat_eq_iff of_nat_div abs_mult abs_div)  | 
| 67118 | 1938  | 
assume "r dvd k" "r dvd l"  | 
1939  | 
then show "r dvd gcd k l"  | 
|
1940  | 
using gcd_greatest [of "nat \<bar>r\<bar>" "nat \<bar>k\<bar>" "nat \<bar>l\<bar>"]  | 
|
1941  | 
by simp  | 
|
1942  | 
qed simp  | 
|
| 63489 | 1943  | 
|
1944  | 
lemma gcd_le1_nat [simp]: "a \<noteq> 0 \<Longrightarrow> gcd a b \<le> a"  | 
|
1945  | 
for a b :: nat  | 
|
1946  | 
by (rule dvd_imp_le) auto  | 
|
1947  | 
||
1948  | 
lemma gcd_le2_nat [simp]: "b \<noteq> 0 \<Longrightarrow> gcd a b \<le> b"  | 
|
1949  | 
for a b :: nat  | 
|
1950  | 
by (rule dvd_imp_le) auto  | 
|
1951  | 
||
1952  | 
lemma gcd_le1_int [simp]: "a > 0 \<Longrightarrow> gcd a b \<le> a"  | 
|
1953  | 
for a b :: int  | 
|
1954  | 
by (rule zdvd_imp_le) auto  | 
|
1955  | 
||
1956  | 
lemma gcd_le2_int [simp]: "b > 0 \<Longrightarrow> gcd a b \<le> b"  | 
|
1957  | 
for a b :: int  | 
|
1958  | 
by (rule zdvd_imp_le) auto  | 
|
1959  | 
||
1960  | 
lemma gcd_pos_nat [simp]: "gcd m n > 0 \<longleftrightarrow> m \<noteq> 0 \<or> n \<noteq> 0"  | 
|
1961  | 
for m n :: nat  | 
|
1962  | 
using gcd_eq_0_iff [of m n] by arith  | 
|
1963  | 
||
1964  | 
lemma gcd_pos_int [simp]: "gcd m n > 0 \<longleftrightarrow> m \<noteq> 0 \<or> n \<noteq> 0"  | 
|
1965  | 
for m n :: int  | 
|
1966  | 
using gcd_eq_0_iff [of m n] gcd_ge_0_int [of m n] by arith  | 
|
1967  | 
||
1968  | 
lemma gcd_unique_nat: "d dvd a \<and> d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"  | 
|
1969  | 
for d a :: nat  | 
|
| 68708 | 1970  | 
using gcd_unique by fastforce  | 
| 63489 | 1971  | 
|
1972  | 
lemma gcd_unique_int:  | 
|
1973  | 
"d \<ge> 0 \<and> d dvd a \<and> d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"  | 
|
1974  | 
for d a :: int  | 
|
| 68708 | 1975  | 
using zdvd_antisym_nonneg by auto  | 
| 
30082
 
43c5b7bfc791
make more proofs work whether or not One_nat_def is a simp rule
 
huffman 
parents: 
30042 
diff
changeset
 | 
1976  | 
|
| 61913 | 1977  | 
interpretation gcd_nat:  | 
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
1978  | 
semilattice_neutr_order gcd "0::nat" Rings.dvd "\<lambda>m n. m dvd n \<and> m \<noteq> n"  | 
| 68708 | 1979  | 
by standard (auto simp: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans)  | 
| 31798 | 1980  | 
|
| 63489 | 1981  | 
lemma gcd_proj1_if_dvd_int [simp]: "x dvd y \<Longrightarrow> gcd x y = \<bar>x\<bar>"  | 
1982  | 
for x y :: int  | 
|
| 67118 | 1983  | 
by (metis abs_dvd_iff gcd_0_left_int gcd_unique_int)  | 
| 31798 | 1984  | 
|
| 63489 | 1985  | 
lemma gcd_proj2_if_dvd_int [simp]: "y dvd x \<Longrightarrow> gcd x y = \<bar>y\<bar>"  | 
1986  | 
for x y :: int  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
1987  | 
by (metis gcd_proj1_if_dvd_int gcd.commute)  | 
| 31798 | 1988  | 
|
| 63489 | 1989  | 
|
1990  | 
text \<open>\<^medskip> Multiplication laws.\<close>  | 
|
1991  | 
||
1992  | 
lemma gcd_mult_distrib_nat: "k * gcd m n = gcd (k * m) (k * n)"  | 
|
1993  | 
for k m n :: nat  | 
|
1994  | 
  \<comment> \<open>@{cite \<open>page 27\<close> davenport92}\<close>
 | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1995  | 
by (simp add: gcd_mult_left)  | 
| 63489 | 1996  | 
|
1997  | 
lemma gcd_mult_distrib_int: "\<bar>k\<bar> * gcd m n = gcd (k * m) (k * n)"  | 
|
1998  | 
for k m n :: int  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
1999  | 
by (simp add: gcd_mult_left abs_mult)  | 
| 21256 | 2000  | 
|
| 63489 | 2001  | 
text \<open>\medskip Addition laws.\<close>  | 
2002  | 
||
2003  | 
(* TODO: add the other variations? *)  | 
|
2004  | 
||
2005  | 
lemma gcd_diff1_nat: "m \<ge> n \<Longrightarrow> gcd (m - n) n = gcd m n"  | 
|
2006  | 
for m n :: nat  | 
|
| 
62429
 
25271ff79171
Tuned Euclidean Rings/GCD rings
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62353 
diff
changeset
 | 
2007  | 
by (subst gcd_add1 [symmetric]) auto  | 
| 31706 | 2008  | 
|
| 63489 | 2009  | 
lemma gcd_diff2_nat: "n \<ge> m \<Longrightarrow> gcd (n - m) n = gcd m n"  | 
2010  | 
for m n :: nat  | 
|
| 68708 | 2011  | 
by (metis gcd.commute gcd_add2 gcd_diff1_nat le_add_diff_inverse2)  | 
2012  | 
||
2013  | 
lemma gcd_non_0_int:  | 
|
2014  | 
fixes x y :: int  | 
|
2015  | 
assumes "y > 0" shows "gcd x y = gcd y (x mod y)"  | 
|
2016  | 
proof (cases "x mod y = 0")  | 
|
2017  | 
case False  | 
|
2018  | 
then have neg: "x mod y = y - (- x) mod y"  | 
|
2019  | 
by (simp add: zmod_zminus1_eq_if)  | 
|
2020  | 
have xy: "0 \<le> x mod y"  | 
|
2021  | 
by (simp add: assms)  | 
|
2022  | 
show ?thesis  | 
|
2023  | 
proof (cases "x < 0")  | 
|
2024  | 
case True  | 
|
2025  | 
have "nat (- x mod y) \<le> nat y"  | 
|
2026  | 
by (simp add: assms dual_order.order_iff_strict)  | 
|
2027  | 
moreover have "gcd (nat (- x)) (nat y) = gcd (nat (- x mod y)) (nat y)"  | 
|
2028  | 
using True assms gcd_non_0_nat nat_mod_distrib by auto  | 
|
2029  | 
ultimately have "gcd (nat (- x)) (nat y) = gcd (nat y) (nat (x mod y))"  | 
|
2030  | 
using assms  | 
|
2031  | 
by (simp add: neg nat_diff_distrib') (metis gcd.commute gcd_diff2_nat)  | 
|
2032  | 
with assms \<open>0 \<le> x mod y\<close> show ?thesis  | 
|
2033  | 
by (simp add: True dual_order.order_iff_strict gcd_int_def)  | 
|
2034  | 
next  | 
|
2035  | 
case False  | 
|
2036  | 
with assms xy have "gcd (nat x) (nat y) = gcd (nat y) (nat x mod nat y)"  | 
|
2037  | 
using gcd_red_nat by blast  | 
|
2038  | 
with False assms show ?thesis  | 
|
2039  | 
by (simp add: gcd_int_def nat_mod_distrib)  | 
|
2040  | 
qed  | 
|
2041  | 
qed (use assms in auto)  | 
|
| 21256 | 2042  | 
|
| 63489 | 2043  | 
lemma gcd_red_int: "gcd x y = gcd y (x mod y)"  | 
2044  | 
for x y :: int  | 
|
| 68708 | 2045  | 
proof (cases y "0::int" rule: linorder_cases)  | 
2046  | 
case less  | 
|
2047  | 
with gcd_non_0_int [of "- y" "- x"] show ?thesis  | 
|
2048  | 
by auto  | 
|
2049  | 
next  | 
|
2050  | 
case greater  | 
|
2051  | 
with gcd_non_0_int [of y x] show ?thesis  | 
|
2052  | 
by auto  | 
|
2053  | 
qed auto  | 
|
| 63489 | 2054  | 
|
2055  | 
(* TODO: differences, and all variations of addition rules  | 
|
| 31706 | 2056  | 
as simplification rules for nat and int *)  | 
2057  | 
||
| 63489 | 2058  | 
(* TODO: add the three variations of these, and for ints? *)  | 
2059  | 
||
2060  | 
lemma finite_divisors_nat [simp]: (* FIXME move *)  | 
|
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2061  | 
fixes m :: nat  | 
| 63489 | 2062  | 
assumes "m > 0"  | 
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2063  | 
  shows "finite {d. d dvd m}"
 | 
| 31734 | 2064  | 
proof-  | 
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2065  | 
  from assms have "{d. d dvd m} \<subseteq> {d. d \<le> m}"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2066  | 
by (auto dest: dvd_imp_le)  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2067  | 
then show ?thesis  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2068  | 
using finite_Collect_le_nat by (rule finite_subset)  | 
| 31734 | 2069  | 
qed  | 
2070  | 
||
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2071  | 
lemma finite_divisors_int [simp]:  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2072  | 
fixes i :: int  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2073  | 
assumes "i \<noteq> 0"  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2074  | 
  shows "finite {d. d dvd i}"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2075  | 
proof -  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2076  | 
  have "{d. \<bar>d\<bar> \<le> \<bar>i\<bar>} = {- \<bar>i\<bar>..\<bar>i\<bar>}"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2077  | 
by (auto simp: abs_if)  | 
| 63489 | 2078  | 
  then have "finite {d. \<bar>d\<bar> \<le> \<bar>i\<bar>}"
 | 
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2079  | 
by simp  | 
| 63489 | 2080  | 
from finite_subset [OF _ this] show ?thesis  | 
2081  | 
using assms by (simp add: dvd_imp_le_int subset_iff)  | 
|
| 31734 | 2082  | 
qed  | 
2083  | 
||
| 63489 | 2084  | 
lemma Max_divisors_self_nat [simp]: "n \<noteq> 0 \<Longrightarrow> Max {d::nat. d dvd n} = n"
 | 
| 68708 | 2085  | 
by (fastforce intro: antisym Max_le_iff[THEN iffD2] simp: dvd_imp_le)  | 
2086  | 
||
2087  | 
lemma Max_divisors_self_int [simp]:  | 
|
2088  | 
  assumes "n \<noteq> 0" shows "Max {d::int. d dvd n} = \<bar>n\<bar>"
 | 
|
2089  | 
proof (rule antisym)  | 
|
2090  | 
  show "Max {d. d dvd n} \<le> \<bar>n\<bar>"
 | 
|
2091  | 
using assms by (auto intro: abs_le_D1 dvd_imp_le_int intro!: Max_le_iff [THEN iffD2])  | 
|
2092  | 
qed (simp add: assms)  | 
|
2093  | 
||
2094  | 
lemma gcd_is_Max_divisors_nat:  | 
|
2095  | 
fixes m n :: nat  | 
|
2096  | 
  assumes "n > 0" shows "gcd m n = Max {d. d dvd m \<and> d dvd n}"
 | 
|
2097  | 
proof (rule Max_eqI[THEN sym], simp_all)  | 
|
2098  | 
  show "finite {d. d dvd m \<and> d dvd n}"
 | 
|
2099  | 
by (simp add: \<open>n > 0\<close>)  | 
|
2100  | 
show "\<And>y. y dvd m \<and> y dvd n \<Longrightarrow> y \<le> gcd m n"  | 
|
2101  | 
by (simp add: \<open>n > 0\<close> dvd_imp_le)  | 
|
2102  | 
qed  | 
|
2103  | 
||
2104  | 
lemma gcd_is_Max_divisors_int:  | 
|
2105  | 
fixes m n :: int  | 
|
2106  | 
  assumes "n \<noteq> 0" shows "gcd m n = Max {d. d dvd m \<and> d dvd n}"
 | 
|
2107  | 
proof (rule Max_eqI[THEN sym], simp_all)  | 
|
2108  | 
  show "finite {d. d dvd m \<and> d dvd n}"
 | 
|
2109  | 
by (simp add: \<open>n \<noteq> 0\<close>)  | 
|
2110  | 
show "\<And>y. y dvd m \<and> y dvd n \<Longrightarrow> y \<le> gcd m n"  | 
|
2111  | 
by (simp add: \<open>n \<noteq> 0\<close> zdvd_imp_le)  | 
|
2112  | 
qed  | 
|
| 63489 | 2113  | 
|
2114  | 
lemma gcd_code_int [code]: "gcd k l = \<bar>if l = 0 then k else gcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"  | 
|
2115  | 
for k l :: int  | 
|
| 67118 | 2116  | 
using gcd_red_int [of "\<bar>k\<bar>" "\<bar>l\<bar>"] by simp  | 
| 
34030
 
829eb528b226
resorted code equations from "old" number theory version
 
haftmann 
parents: 
33946 
diff
changeset
 | 
2117  | 
|
| 67051 | 2118  | 
lemma coprime_Suc_left_nat [simp]:  | 
2119  | 
"coprime (Suc n) n"  | 
|
2120  | 
using coprime_add_one_left [of n] by simp  | 
|
2121  | 
||
2122  | 
lemma coprime_Suc_right_nat [simp]:  | 
|
2123  | 
"coprime n (Suc n)"  | 
|
2124  | 
using coprime_Suc_left_nat [of n] by (simp add: ac_simps)  | 
|
2125  | 
||
2126  | 
lemma coprime_diff_one_left_nat [simp]:  | 
|
2127  | 
"coprime (n - 1) n" if "n > 0" for n :: nat  | 
|
2128  | 
using that coprime_Suc_right_nat [of "n - 1"] by simp  | 
|
2129  | 
||
2130  | 
lemma coprime_diff_one_right_nat [simp]:  | 
|
2131  | 
"coprime n (n - 1)" if "n > 0" for n :: nat  | 
|
2132  | 
using that coprime_diff_one_left_nat [of n] by (simp add: ac_simps)  | 
|
2133  | 
||
2134  | 
lemma coprime_crossproduct_nat:  | 
|
2135  | 
fixes a b c d :: nat  | 
|
2136  | 
assumes "coprime a d" and "coprime b c"  | 
|
2137  | 
shows "a * c = b * d \<longleftrightarrow> a = b \<and> c = d"  | 
|
2138  | 
using assms coprime_crossproduct [of a d b c] by simp  | 
|
2139  | 
||
2140  | 
lemma coprime_crossproduct_int:  | 
|
2141  | 
fixes a b c d :: int  | 
|
2142  | 
assumes "coprime a d" and "coprime b c"  | 
|
2143  | 
shows "\<bar>a\<bar> * \<bar>c\<bar> = \<bar>b\<bar> * \<bar>d\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>b\<bar> \<and> \<bar>c\<bar> = \<bar>d\<bar>"  | 
|
2144  | 
using assms coprime_crossproduct [of a d b c] by simp  | 
|
| 31706 | 2145  | 
|
2146  | 
||
| 60758 | 2147  | 
subsection \<open>Bezout's theorem\<close>  | 
| 31706 | 2148  | 
|
| 63489 | 2149  | 
text \<open>  | 
2150  | 
Function \<open>bezw\<close> returns a pair of witnesses to Bezout's theorem --  | 
|
2151  | 
see the theorems that follow the definition.  | 
|
2152  | 
\<close>  | 
|
2153  | 
||
2154  | 
fun bezw :: "nat \<Rightarrow> nat \<Rightarrow> int * int"  | 
|
2155  | 
where "bezw x y =  | 
|
2156  | 
(if y = 0 then (1, 0)  | 
|
2157  | 
else  | 
|
| 31706 | 2158  | 
(snd (bezw y (x mod y)),  | 
2159  | 
fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y)))"  | 
|
2160  | 
||
| 63489 | 2161  | 
lemma bezw_0 [simp]: "bezw x 0 = (1, 0)"  | 
2162  | 
by simp  | 
|
2163  | 
||
2164  | 
lemma bezw_non_0:  | 
|
2165  | 
"y > 0 \<Longrightarrow> bezw x y =  | 
|
2166  | 
(snd (bezw y (x mod y)), fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y))"  | 
|
| 31706 | 2167  | 
by simp  | 
2168  | 
||
2169  | 
declare bezw.simps [simp del]  | 
|
2170  | 
||
| 68708 | 2171  | 
|
2172  | 
lemma bezw_aux: "int (gcd x y) = fst (bezw x y) * int x + snd (bezw x y) * int y"  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31814 
diff
changeset
 | 
2173  | 
proof (induct x y rule: gcd_nat_induct)  | 
| 68708 | 2174  | 
case (step m n)  | 
2175  | 
then have "fst (bezw m n) * int m + snd (bezw m n) * int n - int (gcd m n)  | 
|
2176  | 
= int m * snd (bezw n (m mod n)) -  | 
|
2177  | 
(int (m mod n) * snd (bezw n (m mod n)) + int n * (int (m div n) * snd (bezw n (m mod n))))"  | 
|
2178  | 
by (simp add: bezw_non_0 gcd_non_0_nat field_simps)  | 
|
2179  | 
also have "\<dots> = int m * snd (bezw n (m mod n)) - (int (m mod n) + int (n * (m div n))) * snd (bezw n (m mod n))"  | 
|
2180  | 
by (simp add: distrib_right)  | 
|
2181  | 
also have "\<dots> = 0"  | 
|
2182  | 
by (metis cancel_comm_monoid_add_class.diff_cancel mod_mult_div_eq of_nat_add)  | 
|
2183  | 
finally show ?case  | 
|
2184  | 
by simp  | 
|
2185  | 
qed auto  | 
|
2186  | 
||
| 31706 | 2187  | 
|
| 63489 | 2188  | 
lemma bezout_int: "\<exists>u v. u * x + v * y = gcd x y"  | 
2189  | 
for x y :: int  | 
|
| 31706 | 2190  | 
proof -  | 
| 63489 | 2191  | 
have aux: "x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> \<exists>u v. u * x + v * y = gcd x y" for x y :: int  | 
| 31706 | 2192  | 
apply (rule_tac x = "fst (bezw (nat x) (nat y))" in exI)  | 
2193  | 
apply (rule_tac x = "snd (bezw (nat x) (nat y))" in exI)  | 
|
| 68708 | 2194  | 
by (simp add: bezw_aux gcd_int_def)  | 
| 63489 | 2195  | 
consider "x \<ge> 0" "y \<ge> 0" | "x \<ge> 0" "y \<le> 0" | "x \<le> 0" "y \<ge> 0" | "x \<le> 0" "y \<le> 0"  | 
| 68708 | 2196  | 
using linear by blast  | 
| 63489 | 2197  | 
then show ?thesis  | 
2198  | 
proof cases  | 
|
2199  | 
case 1  | 
|
2200  | 
then show ?thesis by (rule aux)  | 
|
2201  | 
next  | 
|
2202  | 
case 2  | 
|
2203  | 
then show ?thesis  | 
|
| 68708 | 2204  | 
using aux [of x "-y"]  | 
2205  | 
by (metis gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le)  | 
|
| 63489 | 2206  | 
next  | 
2207  | 
case 3  | 
|
2208  | 
then show ?thesis  | 
|
| 68708 | 2209  | 
using aux [of "-x" y]  | 
2210  | 
by (metis gcd.commute gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le)  | 
|
| 63489 | 2211  | 
next  | 
2212  | 
case 4  | 
|
2213  | 
then show ?thesis  | 
|
| 68708 | 2214  | 
using aux [of "-x" "-y"]  | 
2215  | 
by (metis diff_0 diff_ge_0_iff_ge gcd_neg1_int gcd_neg2_int mult.commute mult_minus_right)  | 
|
| 63489 | 2216  | 
qed  | 
| 31706 | 2217  | 
qed  | 
2218  | 
||
| 63489 | 2219  | 
|
2220  | 
text \<open>Versions of Bezout for \<open>nat\<close>, by Amine Chaieb.\<close>  | 
|
| 31706 | 2221  | 
|
| 68708 | 2222  | 
lemma Euclid_induct [case_names swap zero add]:  | 
| 63489 | 2223  | 
fixes P :: "nat \<Rightarrow> nat \<Rightarrow> bool"  | 
| 68708 | 2224  | 
assumes c: "\<And>a b. P a b \<longleftrightarrow> P b a"  | 
2225  | 
and z: "\<And>a. P a 0"  | 
|
2226  | 
and add: "\<And>a b. P a b \<longrightarrow> P a (a + b)"  | 
|
| 
27669
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2227  | 
shows "P a b"  | 
| 63489 | 2228  | 
proof (induct "a + b" arbitrary: a b rule: less_induct)  | 
| 34915 | 2229  | 
case less  | 
| 63489 | 2230  | 
consider (eq) "a = b" | (lt) "a < b" "a + b - a < a + b" | "b = 0" | "b + a - b < a + b"  | 
2231  | 
by arith  | 
|
2232  | 
show ?case  | 
|
2233  | 
proof (cases a b rule: linorder_cases)  | 
|
2234  | 
case equal  | 
|
2235  | 
with add [rule_format, OF z [rule_format, of a]] show ?thesis by simp  | 
|
2236  | 
next  | 
|
2237  | 
case lt: less  | 
|
2238  | 
then consider "a = 0" | "a + b - a < a + b" by arith  | 
|
2239  | 
then show ?thesis  | 
|
2240  | 
proof cases  | 
|
2241  | 
case 1  | 
|
2242  | 
with z c show ?thesis by blast  | 
|
2243  | 
next  | 
|
2244  | 
case 2  | 
|
2245  | 
also have *: "a + b - a = a + (b - a)" using lt by arith  | 
|
| 34915 | 2246  | 
finally have "a + (b - a) < a + b" .  | 
| 63489 | 2247  | 
then have "P a (a + (b - a))" by (rule add [rule_format, OF less])  | 
2248  | 
then show ?thesis by (simp add: *[symmetric])  | 
|
2249  | 
qed  | 
|
2250  | 
next  | 
|
2251  | 
case gt: greater  | 
|
2252  | 
then consider "b = 0" | "b + a - b < a + b" by arith  | 
|
2253  | 
then show ?thesis  | 
|
2254  | 
proof cases  | 
|
2255  | 
case 1  | 
|
2256  | 
with z c show ?thesis by blast  | 
|
2257  | 
next  | 
|
2258  | 
case 2  | 
|
2259  | 
also have *: "b + a - b = b + (a - b)" using gt by arith  | 
|
| 34915 | 2260  | 
finally have "b + (a - b) < a + b" .  | 
| 63489 | 2261  | 
then have "P b (b + (a - b))" by (rule add [rule_format, OF less])  | 
2262  | 
then have "P b a" by (simp add: *[symmetric])  | 
|
2263  | 
with c show ?thesis by blast  | 
|
2264  | 
qed  | 
|
2265  | 
qed  | 
|
| 
27669
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2266  | 
qed  | 
| 
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2267  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31814 
diff
changeset
 | 
2268  | 
lemma bezout_lemma_nat:  | 
| 68708 | 2269  | 
fixes d::nat  | 
2270  | 
shows "\<lbrakk>d dvd a; d dvd b; a * x = b * y + d \<or> b * x = a * y + d\<rbrakk>  | 
|
2271  | 
\<Longrightarrow> \<exists>x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"  | 
|
| 31706 | 2272  | 
apply auto  | 
| 68708 | 2273  | 
apply (metis add_mult_distrib2 left_add_mult_distrib)  | 
2274  | 
apply (rule_tac x=x in exI)  | 
|
2275  | 
by (metis add_mult_distrib2 mult.commute add.assoc)  | 
|
2276  | 
||
2277  | 
lemma bezout_add_nat:  | 
|
2278  | 
"\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"  | 
|
2279  | 
proof (induct a b rule: Euclid_induct)  | 
|
2280  | 
case (swap a b)  | 
|
2281  | 
then show ?case  | 
|
2282  | 
by blast  | 
|
2283  | 
next  | 
|
2284  | 
case (zero a)  | 
|
2285  | 
then show ?case  | 
|
2286  | 
by fastforce  | 
|
2287  | 
next  | 
|
2288  | 
case (add a b)  | 
|
2289  | 
then show ?case  | 
|
2290  | 
by (meson bezout_lemma_nat)  | 
|
2291  | 
qed  | 
|
2292  | 
||
2293  | 
lemma bezout1_nat: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"  | 
|
2294  | 
using bezout_add_nat[of a b] by (metis add_diff_cancel_left')  | 
|
| 63489 | 2295  | 
|
2296  | 
lemma bezout_add_strong_nat:  | 
|
2297  | 
fixes a b :: nat  | 
|
2298  | 
assumes a: "a \<noteq> 0"  | 
|
| 
27669
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2299  | 
shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"  | 
| 63489 | 2300  | 
proof -  | 
2301  | 
consider d x y where "d dvd a" "d dvd b" "a * x = b * y + d"  | 
|
| 68708 | 2302  | 
| d x y where "d dvd a" "d dvd b" "b * x = a * y + d"  | 
| 63489 | 2303  | 
using bezout_add_nat [of a b] by blast  | 
2304  | 
then show ?thesis  | 
|
2305  | 
proof cases  | 
|
2306  | 
case 1  | 
|
2307  | 
then show ?thesis by blast  | 
|
2308  | 
next  | 
|
2309  | 
case H: 2  | 
|
2310  | 
show ?thesis  | 
|
2311  | 
proof (cases "b = 0")  | 
|
2312  | 
case True  | 
|
2313  | 
with H show ?thesis by simp  | 
|
2314  | 
next  | 
|
2315  | 
case False  | 
|
2316  | 
then have bp: "b > 0" by simp  | 
|
2317  | 
with dvd_imp_le [OF H(2)] consider "d = b" | "d < b"  | 
|
2318  | 
by atomize_elim auto  | 
|
2319  | 
then show ?thesis  | 
|
2320  | 
proof cases  | 
|
2321  | 
case 1  | 
|
2322  | 
with a H show ?thesis  | 
|
| 68708 | 2323  | 
by (metis Suc_pred add.commute mult.commute mult_Suc_right neq0_conv)  | 
| 63489 | 2324  | 
next  | 
2325  | 
case 2  | 
|
2326  | 
show ?thesis  | 
|
2327  | 
proof (cases "x = 0")  | 
|
2328  | 
case True  | 
|
2329  | 
with a H show ?thesis by simp  | 
|
2330  | 
next  | 
|
2331  | 
case x0: False  | 
|
2332  | 
then have xp: "x > 0" by simp  | 
|
2333  | 
from \<open>d < b\<close> have "d \<le> b - 1" by simp  | 
|
2334  | 
then have "d * b \<le> b * (b - 1)" by simp  | 
|
2335  | 
with xp mult_mono[of "1" "x" "d * b" "b * (b - 1)"]  | 
|
2336  | 
have dble: "d * b \<le> x * b * (b - 1)" using bp by simp  | 
|
2337  | 
from H(3) have "d + (b - 1) * (b * x) = d + (b - 1) * (a * y + d)"  | 
|
| 31706 | 2338  | 
by simp  | 
| 63489 | 2339  | 
then have "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
56218 
diff
changeset
 | 
2340  | 
by (simp only: mult.assoc distrib_left)  | 
| 63489 | 2341  | 
then have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x * b * (b - 1)"  | 
| 31706 | 2342  | 
by algebra  | 
| 63489 | 2343  | 
then have "a * ((b - 1) * y) = d + x * b * (b - 1) - d * b"  | 
2344  | 
using bp by simp  | 
|
2345  | 
then have "a * ((b - 1) * y) = d + (x * b * (b - 1) - d * b)"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32879 
diff
changeset
 | 
2346  | 
by (simp only: diff_add_assoc[OF dble, of d, symmetric])  | 
| 63489 | 2347  | 
then have "a * ((b - 1) * y) = b * (x * (b - 1) - d) + d"  | 
| 59008 | 2348  | 
by (simp only: diff_mult_distrib2 ac_simps)  | 
| 63489 | 2349  | 
with H(1,2) show ?thesis  | 
| 68708 | 2350  | 
by blast  | 
| 63489 | 2351  | 
qed  | 
2352  | 
qed  | 
|
2353  | 
qed  | 
|
2354  | 
qed  | 
|
| 
27669
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2355  | 
qed  | 
| 
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2356  | 
|
| 63489 | 2357  | 
lemma bezout_nat:  | 
2358  | 
fixes a :: nat  | 
|
2359  | 
assumes a: "a \<noteq> 0"  | 
|
| 
27669
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2360  | 
shows "\<exists>x y. a * x = b * y + gcd a b"  | 
| 63489 | 2361  | 
proof -  | 
2362  | 
obtain d x y where d: "d dvd a" "d dvd b" and eq: "a * x = b * y + d"  | 
|
2363  | 
using bezout_add_strong_nat [OF a, of b] by blast  | 
|
2364  | 
from d have "d dvd gcd a b"  | 
|
2365  | 
by simp  | 
|
2366  | 
then obtain k where k: "gcd a b = d * k"  | 
|
2367  | 
unfolding dvd_def by blast  | 
|
2368  | 
from eq have "a * x * k = (b * y + d) * k"  | 
|
2369  | 
by auto  | 
|
2370  | 
then have "a * (x * k) = b * (y * k) + gcd a b"  | 
|
2371  | 
by (algebra add: k)  | 
|
2372  | 
then show ?thesis  | 
|
2373  | 
by blast  | 
|
| 
27669
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2374  | 
qed  | 
| 
 
4b1642284dd7
Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
 
chaieb 
parents: 
27651 
diff
changeset
 | 
2375  | 
|
| 31706 | 2376  | 
|
| 69593 | 2377  | 
subsection \<open>LCM properties on \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close>\<close>  | 
| 63489 | 2378  | 
|
2379  | 
lemma lcm_altdef_int [code]: "lcm a b = \<bar>a\<bar> * \<bar>b\<bar> div gcd a b"  | 
|
2380  | 
for a b :: int  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2381  | 
by (simp add: abs_mult lcm_gcd abs_div)  | 
| 67118 | 2382  | 
|
| 63489 | 2383  | 
lemma prod_gcd_lcm_nat: "m * n = gcd m n * lcm m n"  | 
2384  | 
for m n :: nat  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2385  | 
by (simp add: lcm_gcd)  | 
| 31706 | 2386  | 
|
| 63489 | 2387  | 
lemma prod_gcd_lcm_int: "\<bar>m\<bar> * \<bar>n\<bar> = gcd m n * lcm m n"  | 
2388  | 
for m n :: int  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2389  | 
by (simp add: lcm_gcd abs_div abs_mult)  | 
| 63489 | 2390  | 
|
2391  | 
lemma lcm_pos_nat: "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> lcm m n > 0"  | 
|
2392  | 
for m n :: nat  | 
|
| 67118 | 2393  | 
using lcm_eq_0_iff [of m n] by auto  | 
| 63489 | 2394  | 
|
2395  | 
lemma lcm_pos_int: "m \<noteq> 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> lcm m n > 0"  | 
|
2396  | 
for m n :: int  | 
|
| 67118 | 2397  | 
by (simp add: less_le lcm_eq_0_iff)  | 
| 
23687
 
06884f7ffb18
extended - convers now basic lcm properties also
 
haftmann 
parents: 
23431 
diff
changeset
 | 
2398  | 
|
| 63489 | 2399  | 
lemma dvd_pos_nat: "n > 0 \<Longrightarrow> m dvd n \<Longrightarrow> m > 0" (* FIXME move *)  | 
2400  | 
for m n :: nat  | 
|
| 68708 | 2401  | 
by auto  | 
| 63489 | 2402  | 
|
2403  | 
lemma lcm_unique_nat:  | 
|
2404  | 
"a dvd d \<and> b dvd d \<and> (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"  | 
|
2405  | 
for a b d :: nat  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2406  | 
by (auto intro: dvd_antisym lcm_least)  | 
| 
27568
 
9949dc7a24de
Theorem names as in IntPrimes.thy, also several theorems moved from there
 
chaieb 
parents: 
27556 
diff
changeset
 | 
2407  | 
|
| 63489 | 2408  | 
lemma lcm_unique_int:  | 
2409  | 
"d \<ge> 0 \<and> a dvd d \<and> b dvd d \<and> (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"  | 
|
2410  | 
for a b d :: int  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2411  | 
using lcm_least zdvd_antisym_nonneg by auto  | 
| 
34973
 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 
haftmann 
parents: 
34915 
diff
changeset
 | 
2412  | 
|
| 63489 | 2413  | 
lemma lcm_proj2_if_dvd_nat [simp]: "x dvd y \<Longrightarrow> lcm x y = y"  | 
2414  | 
for x y :: nat  | 
|
| 68708 | 2415  | 
by (simp add: lcm_proj2_if_dvd)  | 
| 63489 | 2416  | 
|
2417  | 
lemma lcm_proj2_if_dvd_int [simp]: "x dvd y \<Longrightarrow> lcm x y = \<bar>y\<bar>"  | 
|
2418  | 
for x y :: int  | 
|
| 68708 | 2419  | 
by (simp add: lcm_proj2_if_dvd)  | 
| 63489 | 2420  | 
|
2421  | 
lemma lcm_proj1_if_dvd_nat [simp]: "x dvd y \<Longrightarrow> lcm y x = y"  | 
|
2422  | 
for x y :: nat  | 
|
2423  | 
by (subst lcm.commute) (erule lcm_proj2_if_dvd_nat)  | 
|
2424  | 
||
2425  | 
lemma lcm_proj1_if_dvd_int [simp]: "x dvd y \<Longrightarrow> lcm y x = \<bar>y\<bar>"  | 
|
2426  | 
for x y :: int  | 
|
2427  | 
by (subst lcm.commute) (erule lcm_proj2_if_dvd_int)  | 
|
2428  | 
||
2429  | 
lemma lcm_proj1_iff_nat [simp]: "lcm m n = m \<longleftrightarrow> n dvd m"  | 
|
2430  | 
for m n :: nat  | 
|
2431  | 
by (metis lcm_proj1_if_dvd_nat lcm_unique_nat)  | 
|
2432  | 
||
2433  | 
lemma lcm_proj2_iff_nat [simp]: "lcm m n = n \<longleftrightarrow> m dvd n"  | 
|
2434  | 
for m n :: nat  | 
|
2435  | 
by (metis lcm_proj2_if_dvd_nat lcm_unique_nat)  | 
|
2436  | 
||
2437  | 
lemma lcm_proj1_iff_int [simp]: "lcm m n = \<bar>m\<bar> \<longleftrightarrow> n dvd m"  | 
|
2438  | 
for m n :: int  | 
|
2439  | 
by (metis dvd_abs_iff lcm_proj1_if_dvd_int lcm_unique_int)  | 
|
2440  | 
||
2441  | 
lemma lcm_proj2_iff_int [simp]: "lcm m n = \<bar>n\<bar> \<longleftrightarrow> m dvd n"  | 
|
2442  | 
for m n :: int  | 
|
2443  | 
by (metis dvd_abs_iff lcm_proj2_if_dvd_int lcm_unique_int)  | 
|
2444  | 
||
2445  | 
lemma lcm_1_iff_nat [simp]: "lcm m n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = Suc 0"  | 
|
2446  | 
for m n :: nat  | 
|
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2447  | 
using lcm_eq_1_iff [of m n] by simp  | 
| 63489 | 2448  | 
|
2449  | 
lemma lcm_1_iff_int [simp]: "lcm m n = 1 \<longleftrightarrow> (m = 1 \<or> m = -1) \<and> (n = 1 \<or> n = -1)"  | 
|
2450  | 
for m n :: int  | 
|
| 61913 | 2451  | 
by auto  | 
| 31995 | 2452  | 
|
| 
34030
 
829eb528b226
resorted code equations from "old" number theory version
 
haftmann 
parents: 
33946 
diff
changeset
 | 
2453  | 
|
| 69593 | 2454  | 
subsection \<open>The complete divisibility lattice on \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close>\<close>  | 
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2455  | 
|
| 63489 | 2456  | 
text \<open>  | 
2457  | 
Lifting \<open>gcd\<close> and \<open>lcm\<close> to sets (\<open>Gcd\<close> / \<open>Lcm\<close>).  | 
|
2458  | 
\<open>Gcd\<close> is defined via \<open>Lcm\<close> to facilitate the proof that we have a complete lattice.  | 
|
| 60758 | 2459  | 
\<close>  | 
| 45264 | 2460  | 
|
| 62345 | 2461  | 
instantiation nat :: semiring_Gcd  | 
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2462  | 
begin  | 
| 
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2463  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2464  | 
interpretation semilattice_neutr_set lcm "1::nat"  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2465  | 
by standard simp_all  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54489 
diff
changeset
 | 
2466  | 
|
| 63489 | 2467  | 
definition "Lcm M = (if finite M then F M else 0)" for M :: "nat set"  | 
2468  | 
||
2469  | 
lemma Lcm_nat_empty: "Lcm {} = (1::nat)"
 | 
|
| 60690 | 2470  | 
by (simp add: Lcm_nat_def del: One_nat_def)  | 
| 51489 | 2471  | 
|
| 63489 | 2472  | 
lemma Lcm_nat_insert: "Lcm (insert n M) = lcm n (Lcm M)" for n :: nat  | 
| 68708 | 2473  | 
by (cases "finite M") (auto simp: Lcm_nat_def simp del: One_nat_def)  | 
| 61929 | 2474  | 
|
| 63489 | 2475  | 
lemma Lcm_nat_infinite: "infinite M \<Longrightarrow> Lcm M = 0" for M :: "nat set"  | 
| 61929 | 2476  | 
by (simp add: Lcm_nat_def)  | 
2477  | 
||
2478  | 
lemma dvd_Lcm_nat [simp]:  | 
|
2479  | 
fixes M :: "nat set"  | 
|
2480  | 
assumes "m \<in> M"  | 
|
2481  | 
shows "m dvd Lcm M"  | 
|
2482  | 
proof -  | 
|
| 63489 | 2483  | 
from assms have "insert m M = M"  | 
2484  | 
by auto  | 
|
| 61929 | 2485  | 
moreover have "m dvd Lcm (insert m M)"  | 
2486  | 
by (simp add: Lcm_nat_insert)  | 
|
| 63489 | 2487  | 
ultimately show ?thesis  | 
2488  | 
by simp  | 
|
| 61929 | 2489  | 
qed  | 
2490  | 
||
2491  | 
lemma Lcm_dvd_nat [simp]:  | 
|
2492  | 
fixes M :: "nat set"  | 
|
2493  | 
assumes "\<forall>m\<in>M. m dvd n"  | 
|
2494  | 
shows "Lcm M dvd n"  | 
|
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2495  | 
proof (cases "n > 0")  | 
| 63489 | 2496  | 
case False  | 
2497  | 
then show ?thesis by simp  | 
|
| 61929 | 2498  | 
next  | 
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2499  | 
case True  | 
| 63489 | 2500  | 
  then have "finite {d. d dvd n}"
 | 
2501  | 
by (rule finite_divisors_nat)  | 
|
2502  | 
  moreover have "M \<subseteq> {d. d dvd n}"
 | 
|
2503  | 
using assms by fast  | 
|
2504  | 
ultimately have "finite M"  | 
|
2505  | 
by (rule rev_finite_subset)  | 
|
2506  | 
then show ?thesis  | 
|
2507  | 
using assms by (induct M) (simp_all add: Lcm_nat_empty Lcm_nat_insert)  | 
|
| 61929 | 2508  | 
qed  | 
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2509  | 
|
| 63489 | 2510  | 
definition "Gcd M = Lcm {d. \<forall>m\<in>M. d dvd m}" for M :: "nat set"
 | 
2511  | 
||
2512  | 
instance  | 
|
2513  | 
proof  | 
|
2514  | 
fix N :: "nat set"  | 
|
2515  | 
fix n :: nat  | 
|
2516  | 
show "Gcd N dvd n" if "n \<in> N"  | 
|
| 68708 | 2517  | 
using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def)  | 
| 63489 | 2518  | 
show "n dvd Gcd N" if "\<And>m. m \<in> N \<Longrightarrow> n dvd m"  | 
| 68708 | 2519  | 
using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def)  | 
| 63489 | 2520  | 
show "n dvd Lcm N" if "n \<in> N"  | 
2521  | 
using that by (induct N rule: infinite_finite_induct) auto  | 
|
2522  | 
show "Lcm N dvd n" if "\<And>m. m \<in> N \<Longrightarrow> m dvd n"  | 
|
2523  | 
using that by (induct N rule: infinite_finite_induct) auto  | 
|
2524  | 
show "normalize (Gcd N) = Gcd N" and "normalize (Lcm N) = Lcm N"  | 
|
2525  | 
by simp_all  | 
|
2526  | 
qed  | 
|
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2527  | 
|
| 62345 | 2528  | 
end  | 
| 61913 | 2529  | 
|
| 63489 | 2530  | 
lemma Gcd_nat_eq_one: "1 \<in> N \<Longrightarrow> Gcd N = 1"  | 
2531  | 
for N :: "nat set"  | 
|
| 62346 | 2532  | 
by (rule Gcd_eq_1_I) auto  | 
2533  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2534  | 
instance nat :: semiring_gcd_mult_normalize  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2535  | 
by intro_classes (auto simp: unit_factor_nat_def)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2536  | 
|
| 63489 | 2537  | 
|
2538  | 
text \<open>Alternative characterizations of Gcd:\<close>  | 
|
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2539  | 
|
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2540  | 
lemma Gcd_eq_Max:  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2541  | 
fixes M :: "nat set"  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2542  | 
  assumes "finite (M::nat set)" and "M \<noteq> {}" and "0 \<notin> M"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2543  | 
  shows "Gcd M = Max (\<Inter>m\<in>M. {d. d dvd m})"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2544  | 
proof (rule antisym)  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2545  | 
from assms obtain m where "m \<in> M" and "m > 0"  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2546  | 
by auto  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2547  | 
  from \<open>m > 0\<close> have "finite {d. d dvd m}"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2548  | 
by (blast intro: finite_divisors_nat)  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2549  | 
  with \<open>m \<in> M\<close> have fin: "finite (\<Inter>m\<in>M. {d. d dvd m})"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2550  | 
by blast  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2551  | 
  from fin show "Gcd M \<le> Max (\<Inter>m\<in>M. {d. d dvd m})"
 | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2552  | 
by (auto intro: Max_ge Gcd_dvd)  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2553  | 
  from fin show "Max (\<Inter>m\<in>M. {d. d dvd m}) \<le> Gcd M"
 | 
| 68708 | 2554  | 
proof (rule Max.boundedI, simp_all)  | 
2555  | 
    show "(\<Inter>m\<in>M. {d. d dvd m}) \<noteq> {}"
 | 
|
2556  | 
by auto  | 
|
2557  | 
show "\<And>a. \<forall>x\<in>M. a dvd x \<Longrightarrow> a \<le> Gcd M"  | 
|
2558  | 
by (meson Gcd_dvd Gcd_greatest \<open>0 < m\<close> \<open>m \<in> M\<close> dvd_imp_le dvd_pos_nat)  | 
|
2559  | 
qed  | 
|
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2560  | 
qed  | 
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2561  | 
|
| 63489 | 2562  | 
lemma Gcd_remove0_nat: "finite M \<Longrightarrow> Gcd M = Gcd (M - {0})"
 | 
2563  | 
for M :: "nat set"  | 
|
| 68708 | 2564  | 
proof (induct pred: finite)  | 
2565  | 
case (insert x M)  | 
|
2566  | 
then show ?case  | 
|
2567  | 
by (simp add: insert_Diff_if)  | 
|
2568  | 
qed auto  | 
|
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2569  | 
|
| 
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2570  | 
lemma Lcm_in_lcm_closed_set_nat:  | 
| 68708 | 2571  | 
fixes M :: "nat set"  | 
2572  | 
  assumes "finite M" "M \<noteq> {}" "\<And>m n. \<lbrakk>m \<in> M; n \<in> M\<rbrakk> \<Longrightarrow> lcm m n \<in> M"
 | 
|
2573  | 
shows "Lcm M \<in> M"  | 
|
2574  | 
using assms  | 
|
2575  | 
proof (induction M rule: finite_linorder_min_induct)  | 
|
2576  | 
case (insert x M)  | 
|
2577  | 
then have "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> lcm m n \<in> M"  | 
|
2578  | 
by (metis dvd_lcm1 gr0I insert_iff lcm_pos_nat nat_dvd_not_less)  | 
|
2579  | 
with insert show ?case  | 
|
2580  | 
by simp (metis Lcm_nat_empty One_nat_def dvd_1_left dvd_lcm2)  | 
|
2581  | 
qed auto  | 
|
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2582  | 
|
| 
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2583  | 
lemma Lcm_eq_Max_nat:  | 
| 68708 | 2584  | 
fixes M :: "nat set"  | 
2585  | 
  assumes M: "finite M" "M \<noteq> {}" "0 \<notin> M" and lcm: "\<And>m n. \<lbrakk>m \<in> M; n \<in> M\<rbrakk> \<Longrightarrow> lcm m n \<in> M"
 | 
|
2586  | 
shows "Lcm M = Max M"  | 
|
2587  | 
proof (rule antisym)  | 
|
2588  | 
show "Lcm M \<le> Max M"  | 
|
2589  | 
    by (simp add: Lcm_in_lcm_closed_set_nat \<open>finite M\<close> \<open>M \<noteq> {}\<close> lcm)
 | 
|
2590  | 
show "Max M \<le> Lcm M"  | 
|
2591  | 
by (meson Lcm_0_iff Max_in M dvd_Lcm dvd_imp_le le_0_eq not_le)  | 
|
2592  | 
qed  | 
|
| 
32112
 
6da9c2a49fed
Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
 
nipkow 
parents: 
32111 
diff
changeset
 | 
2593  | 
|
| 34222 | 2594  | 
lemma mult_inj_if_coprime_nat:  | 
| 68708 | 2595  | 
"inj_on f A \<Longrightarrow> inj_on g B \<Longrightarrow> (\<And>a b. \<lbrakk>a\<in>A; b\<in>B\<rbrakk> \<Longrightarrow> coprime (f a) (g b)) \<Longrightarrow>  | 
| 63489 | 2596  | 
inj_on (\<lambda>(a, b). f a * g b) (A \<times> B)"  | 
2597  | 
for f :: "'a \<Rightarrow> nat" and g :: "'b \<Rightarrow> nat"  | 
|
| 68708 | 2598  | 
by (auto simp: inj_on_def coprime_crossproduct_nat simp del: One_nat_def)  | 
| 34222 | 2599  | 
|
| 63489 | 2600  | 
|
2601  | 
subsubsection \<open>Setwise GCD and LCM for integers\<close>  | 
|
| 45264 | 2602  | 
|
| 67118 | 2603  | 
instantiation int :: Gcd  | 
| 45264 | 2604  | 
begin  | 
2605  | 
||
| 67118 | 2606  | 
definition Gcd_int :: "int set \<Rightarrow> int"  | 
2607  | 
where "Gcd K = int (GCD k\<in>K. (nat \<circ> abs) k)"  | 
|
2608  | 
||
2609  | 
definition Lcm_int :: "int set \<Rightarrow> int"  | 
|
2610  | 
where "Lcm K = int (LCM k\<in>K. (nat \<circ> abs) k)"  | 
|
2611  | 
||
2612  | 
instance ..  | 
|
| 62345 | 2613  | 
|
2614  | 
end  | 
|
2615  | 
||
| 67118 | 2616  | 
lemma Gcd_int_eq [simp]:  | 
2617  | 
"(GCD n\<in>N. int n) = int (Gcd N)"  | 
|
2618  | 
by (simp add: Gcd_int_def image_image)  | 
|
2619  | 
||
2620  | 
lemma Gcd_nat_abs_eq [simp]:  | 
|
2621  | 
"(GCD k\<in>K. nat \<bar>k\<bar>) = nat (Gcd K)"  | 
|
2622  | 
by (simp add: Gcd_int_def)  | 
|
2623  | 
||
2624  | 
lemma abs_Gcd_eq [simp]:  | 
|
2625  | 
"\<bar>Gcd K\<bar> = Gcd K" for K :: "int set"  | 
|
2626  | 
by (simp only: Gcd_int_def)  | 
|
2627  | 
||
2628  | 
lemma Gcd_int_greater_eq_0 [simp]:  | 
|
2629  | 
"Gcd K \<ge> 0"  | 
|
| 63489 | 2630  | 
for K :: "int set"  | 
| 67118 | 2631  | 
using abs_ge_zero [of "Gcd K"] by simp  | 
2632  | 
||
2633  | 
lemma Gcd_abs_eq [simp]:  | 
|
2634  | 
"(GCD k\<in>K. \<bar>k\<bar>) = Gcd K"  | 
|
| 63489 | 2635  | 
for K :: "int set"  | 
| 67118 | 2636  | 
by (simp only: Gcd_int_def image_image) simp  | 
2637  | 
||
2638  | 
lemma Lcm_int_eq [simp]:  | 
|
2639  | 
"(LCM n\<in>N. int n) = int (Lcm N)"  | 
|
2640  | 
by (simp add: Lcm_int_def image_image)  | 
|
2641  | 
||
2642  | 
lemma Lcm_nat_abs_eq [simp]:  | 
|
2643  | 
"(LCM k\<in>K. nat \<bar>k\<bar>) = nat (Lcm K)"  | 
|
2644  | 
by (simp add: Lcm_int_def)  | 
|
2645  | 
||
2646  | 
lemma abs_Lcm_eq [simp]:  | 
|
2647  | 
"\<bar>Lcm K\<bar> = Lcm K" for K :: "int set"  | 
|
2648  | 
by (simp only: Lcm_int_def)  | 
|
2649  | 
||
2650  | 
lemma Lcm_int_greater_eq_0 [simp]:  | 
|
2651  | 
"Lcm K \<ge> 0"  | 
|
2652  | 
for K :: "int set"  | 
|
2653  | 
using abs_ge_zero [of "Lcm K"] by simp  | 
|
2654  | 
||
2655  | 
lemma Lcm_abs_eq [simp]:  | 
|
2656  | 
"(LCM k\<in>K. \<bar>k\<bar>) = Lcm K"  | 
|
2657  | 
for K :: "int set"  | 
|
2658  | 
by (simp only: Lcm_int_def image_image) simp  | 
|
2659  | 
||
2660  | 
instance int :: semiring_Gcd  | 
|
2661  | 
proof  | 
|
2662  | 
fix K :: "int set" and k :: int  | 
|
2663  | 
show "Gcd K dvd k" and "k dvd Lcm K" if "k \<in> K"  | 
|
2664  | 
using that Gcd_dvd [of "nat \<bar>k\<bar>" "(nat \<circ> abs) ` K"]  | 
|
2665  | 
dvd_Lcm [of "nat \<bar>k\<bar>" "(nat \<circ> abs) ` K"]  | 
|
2666  | 
by (simp_all add: comp_def)  | 
|
2667  | 
show "k dvd Gcd K" if "\<And>l. l \<in> K \<Longrightarrow> k dvd l"  | 
|
2668  | 
proof -  | 
|
2669  | 
have "nat \<bar>k\<bar> dvd (GCD k\<in>K. nat \<bar>k\<bar>)"  | 
|
2670  | 
by (rule Gcd_greatest) (use that in auto)  | 
|
2671  | 
then show ?thesis by simp  | 
|
2672  | 
qed  | 
|
2673  | 
show "Lcm K dvd k" if "\<And>l. l \<in> K \<Longrightarrow> l dvd k"  | 
|
2674  | 
proof -  | 
|
2675  | 
have "(LCM k\<in>K. nat \<bar>k\<bar>) dvd nat \<bar>k\<bar>"  | 
|
2676  | 
by (rule Lcm_least) (use that in auto)  | 
|
2677  | 
then show ?thesis by simp  | 
|
2678  | 
qed  | 
|
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2679  | 
qed (simp_all add: sgn_mult)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2680  | 
|
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2681  | 
instance int :: semiring_gcd_mult_normalize  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69906 
diff
changeset
 | 
2682  | 
by intro_classes (auto simp: sgn_mult)  | 
| 62346 | 2683  | 
|
| 62345 | 2684  | 
|
| 69593 | 2685  | 
subsection \<open>GCD and LCM on \<^typ>\<open>integer\<close>\<close>  | 
| 62345 | 2686  | 
|
2687  | 
instantiation integer :: gcd  | 
|
2688  | 
begin  | 
|
2689  | 
||
2690  | 
context  | 
|
2691  | 
includes integer.lifting  | 
|
2692  | 
begin  | 
|
2693  | 
||
| 63489 | 2694  | 
lift_definition gcd_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer" is gcd .  | 
2695  | 
||
2696  | 
lift_definition lcm_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer" is lcm .  | 
|
| 62345 | 2697  | 
|
2698  | 
end  | 
|
| 63489 | 2699  | 
|
| 45264 | 2700  | 
instance ..  | 
| 60686 | 2701  | 
|
| 21256 | 2702  | 
end  | 
| 45264 | 2703  | 
|
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2704  | 
lifting_update integer.lifting  | 
| 
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2705  | 
lifting_forget integer.lifting  | 
| 
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2706  | 
|
| 62345 | 2707  | 
context  | 
2708  | 
includes integer.lifting  | 
|
2709  | 
begin  | 
|
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2710  | 
|
| 63489 | 2711  | 
lemma gcd_code_integer [code]: "gcd k l = \<bar>if l = (0::integer) then k else gcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"  | 
| 62345 | 2712  | 
by transfer (fact gcd_code_int)  | 
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2713  | 
|
| 63489 | 2714  | 
lemma lcm_code_integer [code]: "lcm a b = \<bar>a\<bar> * \<bar>b\<bar> div gcd a b"  | 
2715  | 
for a b :: integer  | 
|
| 62345 | 2716  | 
by transfer (fact lcm_altdef_int)  | 
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2717  | 
|
| 
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2718  | 
end  | 
| 
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2719  | 
|
| 63489 | 2720  | 
code_printing  | 
2721  | 
constant "gcd :: integer \<Rightarrow> _" \<rightharpoonup>  | 
|
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69785 
diff
changeset
 | 
2722  | 
(OCaml) "!(fun k l -> if Z.equal k Z.zero then/ Z.abs l else if Z.equal/ l Z.zero then Z.abs k else Z.gcd k l)"  | 
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2723  | 
and (Haskell) "Prelude.gcd"  | 
| 
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2724  | 
and (Scala) "_.gcd'((_)')"  | 
| 61975 | 2725  | 
\<comment> \<open>There is no gcd operation in the SML standard library, so no code setup for SML\<close>  | 
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2726  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2727  | 
text \<open>Some code equations\<close>  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2728  | 
|
| 64850 | 2729  | 
lemmas Gcd_nat_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = nat]  | 
2730  | 
lemmas Lcm_nat_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = nat]  | 
|
2731  | 
lemmas Gcd_int_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = int]  | 
|
2732  | 
lemmas Lcm_int_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = int]  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2733  | 
|
| 63489 | 2734  | 
text \<open>Fact aliases.\<close>  | 
2735  | 
||
2736  | 
lemma lcm_0_iff_nat [simp]: "lcm m n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"  | 
|
2737  | 
for m n :: nat  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2738  | 
by (fact lcm_eq_0_iff)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2739  | 
|
| 63489 | 2740  | 
lemma lcm_0_iff_int [simp]: "lcm m n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"  | 
2741  | 
for m n :: int  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2742  | 
by (fact lcm_eq_0_iff)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2743  | 
|
| 63489 | 2744  | 
lemma dvd_lcm_I1_nat [simp]: "k dvd m \<Longrightarrow> k dvd lcm m n"  | 
2745  | 
for k m n :: nat  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2746  | 
by (fact dvd_lcmI1)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2747  | 
|
| 63489 | 2748  | 
lemma dvd_lcm_I2_nat [simp]: "k dvd n \<Longrightarrow> k dvd lcm m n"  | 
2749  | 
for k m n :: nat  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2750  | 
by (fact dvd_lcmI2)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2751  | 
|
| 63489 | 2752  | 
lemma dvd_lcm_I1_int [simp]: "i dvd m \<Longrightarrow> i dvd lcm m n"  | 
2753  | 
for i m n :: int  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2754  | 
by (fact dvd_lcmI1)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2755  | 
|
| 63489 | 2756  | 
lemma dvd_lcm_I2_int [simp]: "i dvd n \<Longrightarrow> i dvd lcm m n"  | 
2757  | 
for i m n :: int  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2758  | 
by (fact dvd_lcmI2)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2759  | 
|
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2760  | 
lemmas Gcd_dvd_nat [simp] = Gcd_dvd [where ?'a = nat]  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2761  | 
lemmas Gcd_dvd_int [simp] = Gcd_dvd [where ?'a = int]  | 
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2762  | 
lemmas Gcd_greatest_nat [simp] = Gcd_greatest [where ?'a = nat]  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2763  | 
lemmas Gcd_greatest_int [simp] = Gcd_greatest [where ?'a = int]  | 
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2764  | 
|
| 63489 | 2765  | 
lemma dvd_Lcm_int [simp]: "m \<in> M \<Longrightarrow> m dvd Lcm M"  | 
2766  | 
for M :: "int set"  | 
|
2767  | 
by (fact dvd_Lcm)  | 
|
2768  | 
||
2769  | 
lemma gcd_neg_numeral_1_int [simp]: "gcd (- numeral n :: int) x = gcd (numeral n) x"  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2770  | 
by (fact gcd_neg1_int)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2771  | 
|
| 63489 | 2772  | 
lemma gcd_neg_numeral_2_int [simp]: "gcd x (- numeral n :: int) = gcd x (numeral n)"  | 
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2773  | 
by (fact gcd_neg2_int)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2774  | 
|
| 63489 | 2775  | 
lemma gcd_proj1_if_dvd_nat [simp]: "x dvd y \<Longrightarrow> gcd x y = x"  | 
2776  | 
for x y :: nat  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2777  | 
by (fact gcd_nat.absorb1)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2778  | 
|
| 63489 | 2779  | 
lemma gcd_proj2_if_dvd_nat [simp]: "y dvd x \<Longrightarrow> gcd x y = y"  | 
2780  | 
for x y :: nat  | 
|
| 
62344
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2781  | 
by (fact gcd_nat.absorb2)  | 
| 
 
759d684c0e60
pulled out legacy aliasses and infamous dvd interpretations into theory appendix
 
haftmann 
parents: 
62343 
diff
changeset
 | 
2782  | 
|
| 
62353
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2783  | 
lemmas Lcm_eq_0_I_nat [simp] = Lcm_eq_0_I [where ?'a = nat]  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2784  | 
lemmas Lcm_0_iff_nat [simp] = Lcm_0_iff [where ?'a = nat]  | 
| 
 
7f927120b5a2
dropped various legacy fact bindings and tuned proofs
 
haftmann 
parents: 
62350 
diff
changeset
 | 
2785  | 
lemmas Lcm_least_int [simp] = Lcm_least [where ?'a = int]  | 
| 62345 | 2786  | 
|
| 
61856
 
4b1b85f38944
add gcd instance for integer and serialisation to target language operations
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
2787  | 
end  |