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