src/HOL/GCD.thy
author Manuel Eberl <eberlm@in.tum.de>
Fri, 01 Jul 2016 08:35:15 +0200
changeset 63359 99b51ba8da1c
parent 63145 703edebd1d92
child 63489 cd540c8031a4
permissions -rw-r--r--
More lemmas on Gcd/Lcm
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32479
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
     1
(*  Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
     2
                Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     3
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     4
32479
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
     5
This file deals with the functions gcd and lcm.  Definitions and
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
     6
lemmas are proved uniformly for the natural numbers and integers.
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     7
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     8
This file combines and revises a number of prior developments.
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     9
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    10
The original theories "GCD" and "Primes" were by Christophe Tabacznyj
58623
2db1df2c8467 more bibtex entries;
wenzelm
parents: 57514
diff changeset
    11
and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    12
gcd, lcm, and prime for the natural numbers.
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 theory "IntPrimes" was by Thomas M. Rasmussen, and
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    15
extended gcd, lcm, primes to the integers. Amine Chaieb provided
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    16
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
    17
of results to "Primes" and "GCD". IntPrimes also defined and developed
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    18
the congruence relations on the integers. The notion was extended to
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
    19
the natural numbers by Chaieb.
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    20
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
    21
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
    22
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
    23
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
    24
Tobias Nipkow cleaned up a lot.
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    25
*)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    26
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    27
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
    28
section \<open>Greatest common divisor and least common multiple\<close>
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
theory GCD
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
    31
imports Main
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    32
begin
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    33
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
    34
subsection \<open>Abstract GCD and LCM\<close>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    35
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
    36
class gcd = zero + one + dvd +
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 37770
diff changeset
    37
  fixes gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
efa734d9b221 eliminated global prems;
wenzelm
parents: 37770
diff changeset
    38
    and lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    39
begin
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    40
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    41
abbreviation coprime :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    42
  where "coprime x y \<equiv> gcd x y = 1"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    43
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    44
end
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    45
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    46
class Gcd = gcd +
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    47
  fixes Gcd :: "'a set \<Rightarrow> 'a"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    48
    and Lcm :: "'a set \<Rightarrow> 'a"
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    49
begin
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    50
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    51
abbreviation GREATEST_COMMON_DIVISOR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    52
where
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    53
  "GREATEST_COMMON_DIVISOR A f \<equiv> Gcd (f ` A)"
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    54
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    55
abbreviation LEAST_COMMON_MULTIPLE :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    56
where
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    57
  "LEAST_COMMON_MULTIPLE A f \<equiv> Lcm (f ` A)"
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    58
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    59
end
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    60
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    61
syntax
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    62
  "_GCD1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3GCD _./ _)" [0, 10] 10)
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    63
  "_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
    64
  "_LCM1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3LCM _./ _)" [0, 10] 10)
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    65
  "_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
    66
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    67
translations
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    68
  "GCD x y. B"   \<rightleftharpoons> "GCD x. GCD y. B"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    69
  "GCD x. B"     \<rightleftharpoons> "CONST GREATEST_COMMON_DIVISOR CONST UNIV (\<lambda>x. B)"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    70
  "GCD x. B"     \<rightleftharpoons> "GCD x \<in> CONST UNIV. B"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    71
  "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
    72
  "LCM x y. B"   \<rightleftharpoons> "LCM x. LCM y. B"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    73
  "LCM x. B"     \<rightleftharpoons> "CONST LEAST_COMMON_MULTIPLE CONST UNIV (\<lambda>x. B)"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    74
  "LCM x. B"     \<rightleftharpoons> "LCM x \<in> CONST UNIV. B"
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    75
  "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
    76
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
    77
print_translation \<open>
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
    78
  [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
    79
    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
    80
\<close> \<comment> \<open>to avoid eta-contraction of body\<close>
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    81
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    82
class semiring_gcd = normalization_semidom + gcd +
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
    83
  assumes gcd_dvd1 [iff]: "gcd a b dvd a"
59977
ad2d1cd53877 eliminated hard tabs;
wenzelm
parents: 59807
diff changeset
    84
    and gcd_dvd2 [iff]: "gcd a b dvd b"
ad2d1cd53877 eliminated hard tabs;
wenzelm
parents: 59807
diff changeset
    85
    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
    86
    and normalize_gcd [simp]: "normalize (gcd a b) = gcd a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    87
    and lcm_gcd: "lcm a b = normalize (a * b) div gcd a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    88
begin    
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    89
60689
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    90
lemma gcd_greatest_iff [simp]:
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    91
  "a dvd gcd b c \<longleftrightarrow> a dvd b \<and> a dvd c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    92
  by (blast intro!: gcd_greatest intro: dvd_trans)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
    93
60689
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    94
lemma gcd_dvdI1:
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    95
  "a dvd c \<Longrightarrow> gcd a b dvd c"
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    96
  by (rule dvd_trans) (rule gcd_dvd1)
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    97
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    98
lemma gcd_dvdI2:
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
    99
  "b dvd c \<Longrightarrow> gcd a b dvd c"
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   100
  by (rule dvd_trans) (rule gcd_dvd2)
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   101
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   102
lemma dvd_gcdD1:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   103
  "a dvd gcd b c \<Longrightarrow> a dvd b"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   104
  using gcd_dvd1 [of b c] by (blast intro: dvd_trans)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   105
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   106
lemma dvd_gcdD2:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   107
  "a dvd gcd b c \<Longrightarrow> a dvd c"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   108
  using gcd_dvd2 [of b c] by (blast intro: dvd_trans)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   109
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   110
lemma gcd_0_left [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   111
  "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
   112
  by (rule associated_eqI) simp_all
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   113
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   114
lemma gcd_0_right [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   115
  "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
   116
  by (rule associated_eqI) simp_all
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   117
  
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   118
lemma gcd_eq_0_iff [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   119
  "gcd a b = 0 \<longleftrightarrow> a = 0 \<and> b = 0" (is "?P \<longleftrightarrow> ?Q")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   120
proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   121
  assume ?P then have "0 dvd gcd a b" by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   122
  then have "0 dvd a" and "0 dvd b" by (blast intro: dvd_trans)+
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   123
  then show ?Q by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   124
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   125
  assume ?Q then show ?P by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   126
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   127
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   128
lemma unit_factor_gcd:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   129
  "unit_factor (gcd a b) = (if a = 0 \<and> b = 0 then 0 else 1)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   130
proof (cases "gcd a b = 0")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   131
  case True then show ?thesis by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   132
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   133
  case False
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   134
  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
   135
    by (rule unit_factor_mult_normalize)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   136
  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
   137
    by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   138
  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
   139
    by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   140
  with False show ?thesis by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   141
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   142
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60689
diff changeset
   143
lemma is_unit_gcd [simp]:
a9e45c9588c3 tuned facts
haftmann
parents: 60689
diff changeset
   144
  "is_unit (gcd a b) \<longleftrightarrow> coprime a b"
a9e45c9588c3 tuned facts
haftmann
parents: 60689
diff changeset
   145
  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
   146
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   147
sublocale gcd: abel_semigroup gcd
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   148
proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   149
  fix a b c
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   150
  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
   151
    by (rule associated_eqI) simp_all
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   152
  from gcd_dvd1 have "gcd (gcd a b) c dvd a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   153
    by (rule dvd_trans) simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   154
  moreover from gcd_dvd1 have "gcd (gcd a b) c dvd b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   155
    by (rule dvd_trans) simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   156
  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
   157
    by (auto intro!: gcd_greatest)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   158
  from gcd_dvd2 have "gcd a (gcd b c) dvd b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   159
    by (rule dvd_trans) simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   160
  moreover from gcd_dvd2 have "gcd a (gcd b c) dvd c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   161
    by (rule dvd_trans) simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   162
  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
   163
    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
   164
  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
   165
    by (rule associated_eqI) simp_all
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   166
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   167
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   168
lemma gcd_self [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   169
  "gcd a a = normalize a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   170
proof -
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   171
  have "a dvd gcd a a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   172
    by (rule gcd_greatest) simp_all
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   173
  then show ?thesis
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60687
diff changeset
   174
    by (auto intro: associated_eqI)
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   175
qed
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   176
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   177
lemma gcd_left_idem [simp]:
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   178
  "gcd a (gcd a b) = gcd a b"
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   179
  by (auto intro: associated_eqI)
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   180
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   181
lemma gcd_right_idem [simp]:
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   182
  "gcd (gcd a b) b = gcd a b"
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   183
  unfolding gcd.commute [of a] gcd.commute [of "gcd b a"] by simp
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   184
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   185
lemma coprime_1_left [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   186
  "coprime 1 a"
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60687
diff changeset
   187
  by (rule associated_eqI) simp_all
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   188
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   189
lemma coprime_1_right [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   190
  "coprime a 1"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   191
  using coprime_1_left [of a] by (simp add: ac_simps)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   192
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   193
lemma gcd_mult_left:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   194
  "gcd (c * a) (c * b) = normalize c * gcd a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   195
proof (cases "c = 0")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   196
  case True then show ?thesis by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   197
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   198
  case False
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   199
  then have "c * gcd a b dvd gcd (c * a) (c * b)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   200
    by (auto intro: gcd_greatest)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   201
  moreover from calculation False have "gcd (c * a) (c * b) dvd c * gcd a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   202
    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
   203
  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
   204
    by (auto intro: associated_eqI)
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   205
  then show ?thesis by (simp add: normalize_mult)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   206
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   207
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   208
lemma gcd_mult_right:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   209
  "gcd (a * c) (b * c) = gcd b a * normalize c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   210
  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
   211
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   212
lemma mult_gcd_left:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   213
  "c * gcd a b = unit_factor c * gcd (c * a) (c * b)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   214
  by (simp add: gcd_mult_left mult.assoc [symmetric])
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   215
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   216
lemma mult_gcd_right:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   217
  "gcd a b * c = gcd (a * c) (b * c) * unit_factor c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   218
  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
   219
60689
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   220
lemma dvd_lcm1 [iff]:
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   221
  "a dvd lcm a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   222
proof -
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   223
  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
   224
    by (simp add: lcm_gcd normalize_mult div_mult_swap)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   225
  then show ?thesis
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   226
    by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   227
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   228
  
60689
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   229
lemma dvd_lcm2 [iff]:
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   230
  "b dvd lcm a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   231
proof -
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   232
  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
   233
    by (simp add: lcm_gcd normalize_mult div_mult_swap ac_simps)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   234
  then show ?thesis
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   235
    by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   236
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   237
60689
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   238
lemma dvd_lcmI1:
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   239
  "a dvd b \<Longrightarrow> a dvd lcm b c"
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   240
  by (rule dvd_trans) (assumption, blast) 
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   241
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   242
lemma dvd_lcmI2:
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   243
  "a dvd c \<Longrightarrow> a dvd lcm b c"
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   244
  by (rule dvd_trans) (assumption, blast)
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   245
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   246
lemma lcm_dvdD1:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   247
  "lcm a b dvd c \<Longrightarrow> a dvd c"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   248
  using dvd_lcm1 [of a b] by (blast intro: dvd_trans)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   249
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   250
lemma lcm_dvdD2:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   251
  "lcm a b dvd c \<Longrightarrow> b dvd c"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   252
  using dvd_lcm2 [of a b] by (blast intro: dvd_trans)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   253
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   254
lemma lcm_least:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   255
  assumes "a dvd c" and "b dvd c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   256
  shows "lcm a b dvd c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   257
proof (cases "c = 0")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   258
  case True then show ?thesis by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   259
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   260
  case False then have U: "is_unit (unit_factor c)" by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   261
  show ?thesis
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   262
  proof (cases "gcd a b = 0")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   263
    case True with assms show ?thesis by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   264
  next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   265
    case False then have "a \<noteq> 0 \<or> b \<noteq> 0" by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   266
    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
   267
      by (simp_all add: mult_dvd_mono)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   268
    then have "normalize (a * b) dvd gcd (a * c) (b * c)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   269
      by (auto intro: gcd_greatest simp add: ac_simps)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   270
    then have "normalize (a * b) dvd gcd (a * c) (b * c) * unit_factor c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   271
      using U by (simp add: dvd_mult_unit_iff)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   272
    then have "normalize (a * b) dvd gcd a b * c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   273
      by (simp add: mult_gcd_right [of a b c])
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   274
    then have "normalize (a * b) div gcd a b dvd c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   275
      using False by (simp add: div_dvd_iff_mult ac_simps)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   276
    then show ?thesis by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   277
  qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   278
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   279
60689
8a2d7c04d8c0 more cautious use of [iff] declarations
haftmann
parents: 60688
diff changeset
   280
lemma lcm_least_iff [simp]:
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   281
  "lcm a b dvd c \<longleftrightarrow> a dvd c \<and> b dvd c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   282
  by (blast intro!: lcm_least intro: dvd_trans)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   283
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   284
lemma normalize_lcm [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   285
  "normalize (lcm a b) = lcm a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   286
  by (simp add: lcm_gcd dvd_normalize_div)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   287
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   288
lemma lcm_0_left [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   289
  "lcm 0 a = 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   290
  by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   291
  
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   292
lemma lcm_0_right [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   293
  "lcm a 0 = 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   294
  by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   295
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   296
lemma lcm_eq_0_iff:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   297
  "lcm a b = 0 \<longleftrightarrow> a = 0 \<or> b = 0" (is "?P \<longleftrightarrow> ?Q")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   298
proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   299
  assume ?P then have "0 dvd lcm a b" by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   300
  then have "0 dvd normalize (a * b) div gcd a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   301
    by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   302
  then have "0 * gcd a b dvd normalize (a * b)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   303
    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
   304
  then have "normalize (a * b) = 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   305
    by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   306
  then show ?Q by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   307
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   308
  assume ?Q then show ?P by auto
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   309
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   310
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   311
lemma lcm_eq_1_iff [simp]:
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   312
  "lcm a b = 1 \<longleftrightarrow> is_unit a \<and> is_unit b"
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   313
  by (auto intro: associated_eqI)
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   314
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   315
lemma unit_factor_lcm :
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   316
  "unit_factor (lcm a b) = (if a = 0 \<or> b = 0 then 0 else 1)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   317
  by (simp add: unit_factor_gcd dvd_unit_factor_div lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   318
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   319
sublocale lcm: abel_semigroup lcm
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   320
proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   321
  fix a b c
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   322
  show "lcm a b = lcm b a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   323
    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
   324
  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
   325
    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
   326
    by (auto intro: lcm_least
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   327
      dvd_trans [of b "lcm b c" "lcm a (lcm b c)"]
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   328
      dvd_trans [of c "lcm b c" "lcm a (lcm b c)"]
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   329
      dvd_trans [of a "lcm a b" "lcm (lcm a b) c"]
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   330
      dvd_trans [of b "lcm a b" "lcm (lcm a b) c"])
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   331
  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
   332
    by (rule associated_eqI) simp_all
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   333
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   334
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   335
lemma lcm_self [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   336
  "lcm a a = normalize a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   337
proof -
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   338
  have "lcm a a dvd a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   339
    by (rule lcm_least) simp_all
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   340
  then show ?thesis
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60687
diff changeset
   341
    by (auto intro: associated_eqI)
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   342
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   343
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   344
lemma lcm_left_idem [simp]:
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   345
  "lcm a (lcm a b) = lcm a b"
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   346
  by (auto intro: associated_eqI)
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   347
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   348
lemma lcm_right_idem [simp]:
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   349
  "lcm (lcm a b) b = lcm a b"
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   350
  unfolding lcm.commute [of a] lcm.commute [of "lcm b a"] by simp
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
   351
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   352
lemma gcd_mult_lcm [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   353
  "gcd a b * lcm a b = normalize a * normalize b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   354
  by (simp add: lcm_gcd normalize_mult)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   355
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   356
lemma lcm_mult_gcd [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   357
  "lcm a b * gcd a b = normalize a * normalize b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   358
  using gcd_mult_lcm [of a b] by (simp add: ac_simps) 
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   359
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   360
lemma gcd_lcm:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   361
  assumes "a \<noteq> 0" and "b \<noteq> 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   362
  shows "gcd a b = normalize (a * b) div lcm a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   363
proof -
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   364
  from assms have "lcm a b \<noteq> 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   365
    by (simp add: lcm_eq_0_iff)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   366
  have "gcd a b * lcm a b = normalize a * normalize b" by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   367
  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
   368
    by (simp_all add: normalize_mult)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   369
  with \<open>lcm a b \<noteq> 0\<close> show ?thesis
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   370
    using nonzero_mult_divide_cancel_right [of "lcm a b" "gcd a b"] by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   371
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   372
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   373
lemma lcm_1_left [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   374
  "lcm 1 a = normalize a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   375
  by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   376
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   377
lemma lcm_1_right [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   378
  "lcm a 1 = normalize a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   379
  by (simp add: lcm_gcd)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   380
  
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   381
lemma lcm_mult_left:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   382
  "lcm (c * a) (c * b) = normalize c * lcm a b"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   383
  by (cases "c = 0")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   384
    (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
   385
      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
   386
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   387
lemma lcm_mult_right:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   388
  "lcm (a * c) (b * c) = lcm b a * normalize c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   389
  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
   390
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   391
lemma mult_lcm_left:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   392
  "c * lcm a b = unit_factor c * lcm (c * a) (c * b)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   393
  by (simp add: lcm_mult_left mult.assoc [symmetric])
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   394
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   395
lemma mult_lcm_right:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   396
  "lcm a b * c = lcm (a * c) (b * c) * unit_factor c"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   397
  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
   398
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   399
lemma gcdI:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   400
  assumes "c dvd a" and "c dvd b" and greatest: "\<And>d. d dvd a \<Longrightarrow> d dvd b \<Longrightarrow> d dvd c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   401
    and "normalize c = c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   402
  shows "c = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   403
  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
   404
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   405
lemma gcd_unique: "d dvd a \<and> d dvd b \<and> 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   406
    normalize d = d \<and>
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   407
    (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   408
  by rule (auto intro: gcdI simp: gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   409
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   410
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
   411
  using mult_dvd_mono [of 1] by auto
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   412
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   413
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
   414
  by (rule gcdI [symmetric]) simp_all
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   415
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   416
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
   417
  by (rule gcdI [symmetric]) simp_all
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   418
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   419
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
   420
proof
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   421
  assume A: "gcd m n = normalize m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   422
  show "m dvd n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   423
  proof (cases "m = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   424
    assume [simp]: "m \<noteq> 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   425
    from A have B: "m = gcd m n * unit_factor m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   426
      by (simp add: unit_eq_div2)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   427
    show ?thesis by (subst B, simp add: mult_unit_dvd_iff)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   428
  qed (insert A, simp)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   429
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   430
  assume "m dvd n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   431
  then show "gcd m n = normalize m" by (rule gcd_proj1_if_dvd)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   432
qed
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   433
  
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   434
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
   435
  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
   436
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   437
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
   438
  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
   439
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   440
lemma gcd_mult_distrib:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   441
  "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
   442
proof-
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   443
  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
   444
    by (simp add: gcd_mult_distrib')
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   445
  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
   446
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   447
  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
   448
    by (simp only: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   449
  then show ?thesis
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   450
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   451
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   452
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   453
lemma lcm_mult_unit1:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   454
  "is_unit a \<Longrightarrow> lcm (b * a) c = lcm b c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   455
  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
   456
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   457
lemma lcm_mult_unit2:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   458
  "is_unit a \<Longrightarrow> lcm b (c * a) = lcm b c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   459
  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
   460
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   461
lemma lcm_div_unit1:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   462
  "is_unit a \<Longrightarrow> lcm (b div a) c = lcm b c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   463
  by (erule is_unitE [of _ b]) (simp add: lcm_mult_unit1) 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   464
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   465
lemma lcm_div_unit2:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   466
  "is_unit a \<Longrightarrow> lcm b (c div a) = lcm b c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   467
  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
   468
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   469
lemma normalize_lcm_left [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   470
  "lcm (normalize a) b = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   471
proof (cases "a = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   472
  case True then show ?thesis
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   473
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   474
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   475
  case False then have "is_unit (unit_factor a)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   476
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   477
  moreover have "normalize a = a div unit_factor a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   478
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   479
  ultimately show ?thesis
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   480
    by (simp only: lcm_div_unit1)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   481
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   482
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   483
lemma normalize_lcm_right [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   484
  "lcm a (normalize b) = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   485
  using normalize_lcm_left [of b a] by (simp add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   486
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   487
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
   488
  apply (rule gcdI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   489
  apply simp_all
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   490
  apply (rule dvd_trans, rule gcd_dvd1, simp add: unit_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   491
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   492
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   493
lemma gcd_mult_unit2: "is_unit a \<Longrightarrow> gcd b (c * a) = gcd b c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   494
  by (subst gcd.commute, subst gcd_mult_unit1, assumption, rule gcd.commute)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   495
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   496
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
   497
  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
   498
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   499
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
   500
  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
   501
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   502
lemma normalize_gcd_left [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   503
  "gcd (normalize a) b = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   504
proof (cases "a = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   505
  case True then show ?thesis
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   506
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   507
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   508
  case False then have "is_unit (unit_factor a)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   509
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   510
  moreover have "normalize a = a div unit_factor a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   511
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   512
  ultimately show ?thesis
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   513
    by (simp only: gcd_div_unit1)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   514
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   515
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   516
lemma normalize_gcd_right [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   517
  "gcd a (normalize b) = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   518
  using normalize_gcd_left [of b a] by (simp add: ac_simps)
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 comp_fun_idem_gcd: "comp_fun_idem gcd"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   521
  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
   522
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   523
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
   524
  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
   525
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   526
lemma gcd_dvd_antisym:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   527
  "gcd a b dvd gcd c d \<Longrightarrow> gcd c d dvd gcd a b \<Longrightarrow> gcd a b = gcd c d"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   528
proof (rule gcdI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   529
  assume A: "gcd a b dvd gcd c d" and B: "gcd c d dvd gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   530
  have "gcd c d dvd c" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   531
  with A show "gcd a b dvd c" by (rule dvd_trans)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   532
  have "gcd c d dvd d" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   533
  with A show "gcd a b dvd d" by (rule dvd_trans)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   534
  show "normalize (gcd a b) = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   535
    by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   536
  fix l assume "l dvd c" and "l dvd d"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   537
  hence "l dvd gcd c d" by (rule gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   538
  from this and B show "l dvd gcd a b" by (rule dvd_trans)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   539
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   540
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   541
lemma coprime_dvd_mult:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   542
  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
   543
  shows "a dvd c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   544
proof (cases "c = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   545
  case True then show ?thesis by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   546
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   547
  case False
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   548
  then have unit: "is_unit (unit_factor c)" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   549
  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
   550
  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
   551
    by (simp add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   552
  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
   553
    by (simp add: dvd_mult_unit_iff unit)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   554
  ultimately show ?thesis by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   555
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   556
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   557
lemma coprime_dvd_mult_iff:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   558
  assumes "coprime a c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   559
  shows "a dvd b * c \<longleftrightarrow> a dvd b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   560
  using assms by (auto intro: coprime_dvd_mult)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   561
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   562
lemma gcd_mult_cancel:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   563
  "coprime c b \<Longrightarrow> gcd (c * a) b = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   564
  apply (rule associated_eqI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   565
  apply (rule gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   566
  apply (rule_tac b = c in coprime_dvd_mult)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   567
  apply (simp add: gcd.assoc)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   568
  apply (simp_all add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   569
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   570
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   571
lemma coprime_crossproduct:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   572
  fixes a b c d
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   573
  assumes "coprime a d" and "coprime b c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   574
  shows "normalize a * normalize c = normalize b * normalize d
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   575
    \<longleftrightarrow> normalize a = normalize b \<and> normalize c = normalize d" (is "?lhs \<longleftrightarrow> ?rhs")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   576
proof
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   577
  assume ?rhs then show ?lhs by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   578
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   579
  assume ?lhs
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   580
  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
   581
    by (auto intro: dvdI dest: sym)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   582
  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
   583
    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
   584
  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
   585
    by (auto intro: dvdI dest: sym)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   586
  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
   587
    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
   588
  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
   589
    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
   590
  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
   591
    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
   592
  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
   593
    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
   594
  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
   595
    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
   596
  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
   597
    by (rule associatedI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   598
  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
   599
    by (rule associatedI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   600
  ultimately show ?rhs ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   601
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   602
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   603
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
   604
  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
   605
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   606
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
   607
  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
   608
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   609
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
   610
  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
   611
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   612
lemma coprimeI: "(\<And>l. \<lbrakk>l dvd a; l dvd b\<rbrakk> \<Longrightarrow> l dvd 1) \<Longrightarrow> gcd a b = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   613
  by (rule sym, rule gcdI, simp_all)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   614
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   615
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
   616
  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
   617
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   618
lemma div_gcd_coprime:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   619
  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
   620
  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
   621
proof -
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   622
  let ?g = "gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   623
  let ?a' = "a div ?g"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   624
  let ?b' = "b div ?g"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   625
  let ?g' = "gcd ?a' ?b'"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   626
  have dvdg: "?g dvd a" "?g dvd b" by simp_all
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   627
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   628
  from dvdg dvdg' obtain ka kb ka' kb' where
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   629
      kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   630
    unfolding dvd_def by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   631
  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
   632
    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
   633
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   634
    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   635
      dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   636
  have "?g \<noteq> 0" using nz by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   637
  moreover from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   638
  thm dvd_mult_cancel_left
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   639
  ultimately show ?thesis using dvd_times_left_cancel_iff [of "gcd a b" _ 1] by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   640
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   641
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   642
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   643
lemma divides_mult:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   644
  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
   645
  shows "a * b dvd c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   646
proof-
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   647
  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
   648
  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
   649
    by (simp add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   650
  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
   651
    by (simp add: coprime_dvd_mult_iff)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   652
  then obtain a' where "b' = a * a'" ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   653
  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
   654
    by (simp add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   655
  then show ?thesis ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   656
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   657
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   658
lemma coprime_lmult:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   659
  assumes dab: "gcd d (a * b) = 1" 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   660
  shows "gcd d a = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   661
proof (rule coprimeI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   662
  fix l assume "l dvd d" and "l dvd a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   663
  hence "l dvd a * b" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   664
  with \<open>l dvd d\<close> and dab show "l dvd 1" by (auto intro: gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   665
qed
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_rmult:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   668
  assumes dab: "gcd d (a * b) = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   669
  shows "gcd d b = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   670
proof (rule coprimeI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   671
  fix l assume "l dvd d" and "l dvd b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   672
  hence "l dvd a * b" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   673
  with \<open>l dvd d\<close> and dab show "l dvd 1" by (auto intro: gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   674
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   675
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   676
lemma coprime_mult:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   677
  assumes da: "coprime d a" and db: "coprime d b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   678
  shows "coprime d (a * b)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   679
  apply (subst gcd.commute)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   680
  using da apply (subst gcd_mult_cancel)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   681
  apply (subst gcd.commute, assumption)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   682
  apply (subst gcd.commute, rule db)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   683
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   684
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   685
lemma coprime_mul_eq: "gcd d (a * b) = 1 \<longleftrightarrow> gcd d a = 1 \<and> gcd d b = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   686
  using coprime_rmult[of d a b] coprime_lmult[of d a b] coprime_mult[of d a b] by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   687
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   688
lemma gcd_coprime:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   689
  assumes c: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   690
  shows "gcd a' b' = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   691
proof -
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   692
  from c have "a \<noteq> 0 \<or> b \<noteq> 0" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   693
  with div_gcd_coprime have "gcd (a div gcd a b) (b div gcd a b) = 1" .
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   694
  also from assms have "a div gcd a b = a'" using dvd_div_eq_mult local.gcd_dvd1 by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   695
  also from assms have "b div gcd a b = b'" using dvd_div_eq_mult local.gcd_dvd1 by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   696
  finally show ?thesis .
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   697
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   698
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   699
lemma coprime_power:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   700
  assumes "0 < n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   701
  shows "gcd a (b ^ n) = 1 \<longleftrightarrow> gcd a b = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   702
using assms proof (induct n)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   703
  case (Suc n) then show ?case
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   704
    by (cases n) (simp_all add: coprime_mul_eq)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   705
qed simp
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_coprime_exists:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   708
  assumes nz: "gcd a b \<noteq> 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   709
  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
   710
  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
   711
  apply (rule_tac x = "b div gcd a b" in exI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   712
  apply (insert nz, auto intro: div_gcd_coprime)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   713
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   714
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   715
lemma coprime_exp:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   716
  "gcd d a = 1 \<Longrightarrow> gcd d (a^n) = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   717
  by (induct n, simp_all add: coprime_mult)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   718
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   719
lemma coprime_exp_left:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   720
  assumes "coprime a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   721
  shows "coprime (a ^ n) b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   722
  using assms by (induct n) (simp_all add: gcd_mult_cancel)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   723
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   724
lemma coprime_exp2:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   725
  assumes "coprime a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   726
  shows "coprime (a ^ n) (b ^ m)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   727
proof (rule coprime_exp_left)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   728
  from assms show "coprime a (b ^ m)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   729
    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
   730
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   731
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   732
lemma gcd_exp:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   733
  "gcd (a ^ n) (b ^ n) = gcd a b ^ n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   734
proof (cases "a = 0 \<and> b = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   735
  case True
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   736
  then show ?thesis by (cases n) simp_all
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   737
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   738
  case False
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   739
  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
   740
    using coprime_exp2[OF div_gcd_coprime[of a b], of n n, symmetric] by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   741
  then have "gcd a b ^ n = gcd a b ^ n * ..." by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   742
  also note gcd_mult_distrib
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   743
  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
   744
    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
   745
  also have "(gcd a b)^n * (a div gcd a b)^n = a^n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   746
    by (subst ac_simps, subst div_power, simp, rule dvd_div_mult_self, rule dvd_power_same, simp)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   747
  also have "(gcd a b)^n * (b div gcd a b)^n = b^n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   748
    by (subst ac_simps, subst div_power, simp, rule dvd_div_mult_self, rule dvd_power_same, simp)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   749
  finally show ?thesis by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   750
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   751
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   752
lemma coprime_common_divisor: 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   753
  "gcd a b = 1 \<Longrightarrow> a dvd a \<Longrightarrow> a dvd b \<Longrightarrow> is_unit a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   754
  apply (subgoal_tac "a dvd gcd a b")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   755
  apply simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   756
  apply (erule (1) gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   757
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   758
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   759
lemma division_decomp: 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   760
  assumes dc: "a dvd b * c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   761
  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
   762
proof (cases "gcd a b = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   763
  assume "gcd a b = 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   764
  hence "a = 0 \<and> b = 0" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   765
  hence "a = 0 * c \<and> 0 dvd b \<and> c dvd c" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   766
  then show ?thesis by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   767
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   768
  let ?d = "gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   769
  assume "?d \<noteq> 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   770
  from gcd_coprime_exists[OF this]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   771
    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
   772
    by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   773
  from ab'(1) have "a' dvd a" unfolding dvd_def by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   774
  with dc have "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   775
  from dc ab'(1,2) have "a'*?d dvd (b'*?d) * c" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   776
  hence "?d * a' dvd ?d * (b' * c)" by (simp add: mult_ac)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   777
  with \<open>?d \<noteq> 0\<close> have "a' dvd b' * c" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   778
  with coprime_dvd_mult[OF ab'(3)] 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   779
    have "a' dvd c" by (subst (asm) ac_simps, blast)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   780
  with ab'(1) have "a = ?d * a' \<and> ?d dvd b \<and> a' dvd c" by (simp add: mult_ac)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   781
  then show ?thesis by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   782
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   783
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   784
lemma pow_divs_pow:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   785
  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
   786
  shows "a dvd b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   787
proof (cases "gcd a b = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   788
  assume "gcd a b = 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   789
  then show ?thesis by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   790
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   791
  let ?d = "gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   792
  assume "?d \<noteq> 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   793
  from n obtain m where m: "n = Suc m" by (cases n, simp_all)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   794
  from \<open>?d \<noteq> 0\<close> have zn: "?d ^ n \<noteq> 0" by (rule power_not_zero)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   795
  from gcd_coprime_exists[OF \<open>?d \<noteq> 0\<close>]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   796
    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
   797
    by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   798
  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
   799
    by (simp add: ab'(1,2)[symmetric])
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   800
  hence "?d^n * a'^n dvd ?d^n * b'^n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   801
    by (simp only: power_mult_distrib ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   802
  with zn have "a'^n dvd b'^n" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   803
  hence "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by (simp add: m)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   804
  hence "a' dvd b'^m * b'" by (simp add: m ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   805
  with coprime_dvd_mult[OF coprime_exp[OF ab'(3), of m]]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   806
    have "a' dvd b'" by (subst (asm) ac_simps, blast)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   807
  hence "a'*?d dvd b'*?d" by (rule mult_dvd_mono, simp)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   808
  with ab'(1,2) show ?thesis by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   809
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   810
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   811
lemma pow_divs_eq [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   812
  "n \<noteq> 0 \<Longrightarrow> a ^ n dvd b ^ n \<longleftrightarrow> a dvd b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   813
  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
   814
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   815
lemma coprime_plus_one [simp]: "gcd (n + 1) n = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   816
  by (subst add_commute, simp)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   817
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   818
lemma setprod_coprime [rule_format]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   819
  "(\<forall>i\<in>A. gcd (f i) a = 1) \<longrightarrow> gcd (\<Prod>i\<in>A. f i) a = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   820
  apply (cases "finite A")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   821
  apply (induct set: finite)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   822
  apply (auto simp add: gcd_mult_cancel)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   823
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   824
  
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   825
lemma listprod_coprime:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   826
  "(\<And>x. x \<in> set xs \<Longrightarrow> coprime x y) \<Longrightarrow> coprime (listprod xs) y" 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   827
  by (induction xs) (simp_all add: gcd_mult_cancel)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   828
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   829
lemma coprime_divisors: 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   830
  assumes "d dvd a" "e dvd b" "gcd a b = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   831
  shows "gcd d e = 1" 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   832
proof -
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   833
  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
   834
    unfolding dvd_def by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   835
  with assms have "gcd (d * k) (e * l) = 1" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   836
  hence "gcd (d * k) e = 1" by (rule coprime_lmult)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   837
  also have "gcd (d * k) e = gcd e (d * k)" by (simp add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   838
  finally have "gcd e d = 1" by (rule coprime_lmult)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   839
  then show ?thesis by (simp add: ac_simps)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   840
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   841
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   842
lemma lcm_gcd_prod:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   843
  "lcm a b * gcd a b = normalize (a * b)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   844
  by (simp add: lcm_gcd)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   845
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   846
declare unit_factor_lcm [simp]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   847
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   848
lemma lcmI:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   849
  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
   850
    and "normalize c = c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   851
  shows "c = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   852
  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
   853
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   854
lemma gcd_dvd_lcm [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   855
  "gcd a b dvd lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   856
  using gcd_dvd2 by (rule dvd_lcmI2)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   857
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   858
lemmas lcm_0 = lcm_0_right
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   859
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   860
lemma lcm_unique:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   861
  "a dvd d \<and> b dvd d \<and> 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   862
  normalize d = d \<and>
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   863
  (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   864
  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
   865
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   866
lemma lcm_coprime:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   867
  "gcd a b = 1 \<Longrightarrow> lcm a b = normalize (a * b)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   868
  by (subst lcm_gcd) simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   869
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   870
lemma lcm_proj1_if_dvd: 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   871
  "b dvd a \<Longrightarrow> lcm a b = normalize a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   872
  by (cases "a = 0") (simp, rule sym, rule lcmI, simp_all)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   873
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   874
lemma lcm_proj2_if_dvd: 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   875
  "a dvd b \<Longrightarrow> lcm a b = normalize b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   876
  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
   877
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   878
lemma lcm_proj1_iff:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   879
  "lcm m n = normalize m \<longleftrightarrow> n dvd m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   880
proof
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   881
  assume A: "lcm m n = normalize m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   882
  show "n dvd m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   883
  proof (cases "m = 0")
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   884
    assume [simp]: "m \<noteq> 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   885
    from A have B: "m = lcm m n * unit_factor m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   886
      by (simp add: unit_eq_div2)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   887
    show ?thesis by (subst B, simp)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   888
  qed simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   889
next
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   890
  assume "n dvd m"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   891
  then show "lcm m n = normalize m" by (rule lcm_proj1_if_dvd)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   892
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   893
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   894
lemma lcm_proj2_iff:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   895
  "lcm m n = normalize n \<longleftrightarrow> m dvd n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   896
  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
   897
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   898
end
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   899
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   900
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
   901
begin
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   902
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   903
lemma coprime_minus_one: "coprime (n - 1) n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   904
  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
   905
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   906
lemma gcd_neg1 [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   907
  "gcd (-a) b = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   908
  by (rule sym, rule gcdI, simp_all add: gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   909
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   910
lemma gcd_neg2 [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   911
  "gcd a (-b) = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   912
  by (rule sym, rule gcdI, simp_all add: gcd_greatest)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   913
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   914
lemma gcd_neg_numeral_1 [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   915
  "gcd (- numeral n) a = gcd (numeral n) a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   916
  by (fact gcd_neg1)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   917
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   918
lemma gcd_neg_numeral_2 [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   919
  "gcd a (- numeral n) = gcd a (numeral n)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   920
  by (fact gcd_neg2)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   921
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   922
lemma gcd_diff1: "gcd (m - n) n = gcd m n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   923
  by (subst diff_conv_add_uminus, subst gcd_neg2[symmetric],  subst gcd_add1, simp)
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 gcd_diff2: "gcd (n - m) n = gcd m n"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   926
  by (subst gcd_neg1[symmetric], simp only: minus_diff_eq gcd_diff1)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   927
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   928
lemma lcm_neg1 [simp]: "lcm (-a) b = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   929
  by (rule sym, rule lcmI, simp_all add: lcm_least lcm_eq_0_iff)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   930
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   931
lemma lcm_neg2 [simp]: "lcm a (-b) = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   932
  by (rule sym, rule lcmI, simp_all add: lcm_least lcm_eq_0_iff)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   933
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   934
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
   935
  by (fact lcm_neg1)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   936
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   937
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
   938
  by (fact lcm_neg2)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   939
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
   940
end
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   941
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   942
class semiring_Gcd = semiring_gcd + Gcd +
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   943
  assumes Gcd_dvd: "a \<in> A \<Longrightarrow> Gcd A dvd a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   944
    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
   945
    and normalize_Gcd [simp]: "normalize (Gcd A) = Gcd A"
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   946
  assumes dvd_Lcm: "a \<in> A \<Longrightarrow> a dvd Lcm A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   947
    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
   948
    and normalize_Lcm [simp]: "normalize (Lcm A) = Lcm A"
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   949
begin
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   950
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   951
lemma Lcm_Gcd:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   952
  "Lcm A = Gcd {b. \<forall>a\<in>A. a dvd b}"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   953
  by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   954
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   955
lemma Gcd_Lcm:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   956
  "Gcd A = Lcm {b. \<forall>a\<in>A. b dvd a}"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   957
  by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   958
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   959
lemma Gcd_empty [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   960
  "Gcd {} = 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   961
  by (rule dvd_0_left, rule Gcd_greatest) simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
   962
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   963
lemma Lcm_empty [simp]:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   964
  "Lcm {} = 1"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   965
  by (auto intro: associated_eqI Lcm_least)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   966
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   967
lemma Gcd_insert [simp]:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   968
  "Gcd (insert a A) = gcd a (Gcd A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   969
proof -
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   970
  have "Gcd (insert a A) dvd gcd a (Gcd A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   971
    by (auto intro: Gcd_dvd Gcd_greatest)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   972
  moreover have "gcd a (Gcd A) dvd Gcd (insert a A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   973
  proof (rule Gcd_greatest)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   974
    fix b
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   975
    assume "b \<in> insert a A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   976
    then show "gcd a (Gcd A) dvd b"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   977
    proof
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   978
      assume "b = a" then show ?thesis by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   979
    next
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   980
      assume "b \<in> A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   981
      then have "Gcd A dvd b" by (rule Gcd_dvd)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   982
      moreover have "gcd a (Gcd A) dvd Gcd A" by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   983
      ultimately show ?thesis by (blast intro: dvd_trans)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   984
    qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   985
  qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   986
  ultimately show ?thesis
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   987
    by (auto intro: associated_eqI)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   988
qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   989
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   990
lemma Lcm_insert [simp]:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   991
  "Lcm (insert a A) = lcm a (Lcm A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   992
proof (rule sym)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   993
  have "lcm a (Lcm A) dvd Lcm (insert a A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   994
    by (auto intro: dvd_Lcm Lcm_least)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   995
  moreover have "Lcm (insert a A) dvd lcm a (Lcm A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   996
  proof (rule Lcm_least)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   997
    fix b
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   998
    assume "b \<in> insert a A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
   999
    then show "b dvd lcm a (Lcm A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1000
    proof
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1001
      assume "b = a" then show ?thesis by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1002
    next
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1003
      assume "b \<in> A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1004
      then have "b dvd Lcm A" by (rule dvd_Lcm)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1005
      moreover have "Lcm A dvd lcm a (Lcm A)" by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1006
      ultimately show ?thesis by (blast intro: dvd_trans)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1007
    qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1008
  qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1009
  ultimately show "lcm a (Lcm A) = Lcm (insert a A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1010
    by (rule associated_eqI) (simp_all add: lcm_eq_0_iff)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1011
qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1012
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1013
lemma LcmI:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1014
  assumes "\<And>a. a \<in> A \<Longrightarrow> a dvd b" and "\<And>c. (\<And>a. a \<in> A \<Longrightarrow> a dvd c) \<Longrightarrow> b dvd c"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1015
    and "normalize b = b" shows "b = Lcm A"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1016
  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
  1017
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1018
lemma Lcm_subset:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1019
  "A \<subseteq> B \<Longrightarrow> Lcm A dvd Lcm B"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1020
  by (blast intro: Lcm_least dvd_Lcm)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1021
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1022
lemma Lcm_Un:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1023
  "Lcm (A \<union> B) = lcm (Lcm A) (Lcm B)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1024
  apply (rule lcmI)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1025
  apply (blast intro: Lcm_subset)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1026
  apply (blast intro: Lcm_subset)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1027
  apply (intro Lcm_least ballI, elim UnE)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1028
  apply (rule dvd_trans, erule dvd_Lcm, assumption)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1029
  apply (rule dvd_trans, erule dvd_Lcm, assumption)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1030
  apply simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1031
  done
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1032
  
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1033
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1034
lemma Gcd_0_iff [simp]:
60687
33dbbcb6a8a3 eliminated some duplication
haftmann
parents: 60686
diff changeset
  1035
  "Gcd A = 0 \<longleftrightarrow> A \<subseteq> {0}" (is "?P \<longleftrightarrow> ?Q")
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1036
proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1037
  assume ?P
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1038
  show ?Q
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1039
  proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1040
    fix a
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1041
    assume "a \<in> A"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1042
    then have "Gcd A dvd a" by (rule Gcd_dvd)
60687
33dbbcb6a8a3 eliminated some duplication
haftmann
parents: 60686
diff changeset
  1043
    with \<open>?P\<close> have "a = 0" by simp
33dbbcb6a8a3 eliminated some duplication
haftmann
parents: 60686
diff changeset
  1044
    then show "a \<in> {0}" by simp
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1045
  qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1046
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1047
  assume ?Q
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1048
  have "0 dvd Gcd A"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1049
  proof (rule Gcd_greatest)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1050
    fix a
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1051
    assume "a \<in> A"
60687
33dbbcb6a8a3 eliminated some duplication
haftmann
parents: 60686
diff changeset
  1052
    with \<open>?Q\<close> have "a = 0" by auto
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1053
    then show "0 dvd a" by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1054
  qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1055
  then show ?P by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1056
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1057
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1058
lemma Lcm_1_iff [simp]:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1059
  "Lcm A = 1 \<longleftrightarrow> (\<forall>a\<in>A. is_unit a)" (is "?P \<longleftrightarrow> ?Q")
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1060
proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1061
  assume ?P
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1062
  show ?Q
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1063
  proof
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1064
    fix a
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1065
    assume "a \<in> A"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1066
    then have "a dvd Lcm A"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1067
      by (rule dvd_Lcm)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1068
    with \<open>?P\<close> show "is_unit a"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1069
      by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1070
  qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1071
next
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1072
  assume ?Q
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1073
  then have "is_unit (Lcm A)"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1074
    by (blast intro: Lcm_least)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1075
  then have "normalize (Lcm A) = 1"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1076
    by (rule is_unit_normalize)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1077
  then show ?P
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1078
    by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1079
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1080
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1081
lemma unit_factor_Lcm:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1082
  "unit_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1083
proof (cases "Lcm A = 0")
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1084
  case True then show ?thesis by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1085
next
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1086
  case False
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1087
  with unit_factor_normalize have "unit_factor (normalize (Lcm A)) = 1"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1088
    by blast
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1089
  with False show ?thesis
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1090
    by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1091
qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1092
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1093
lemma unit_factor_Gcd: "unit_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1094
proof -
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1095
  show "unit_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1096
    by (simp add: Gcd_Lcm unit_factor_Lcm)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1097
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1098
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1099
lemma GcdI:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1100
  assumes "\<And>a. a \<in> A \<Longrightarrow> b dvd a" and "\<And>c. (\<And>a. a \<in> A \<Longrightarrow> c dvd a) \<Longrightarrow> c dvd b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1101
    and "normalize b = b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1102
  shows "b = Gcd A"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1103
  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
  1104
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1105
lemma Gcd_eq_1_I:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1106
  assumes "is_unit a" and "a \<in> A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1107
  shows "Gcd A = 1"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1108
proof -
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1109
  from assms have "is_unit (Gcd A)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1110
    by (blast intro: Gcd_dvd dvd_unit_imp_unit)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1111
  then have "normalize (Gcd A) = 1"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1112
    by (rule is_unit_normalize)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1113
  then show ?thesis
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1114
    by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1115
qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1116
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1117
lemma Lcm_eq_0_I:
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1118
  assumes "0 \<in> A"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1119
  shows "Lcm A = 0"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1120
proof -
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1121
  from assms have "0 dvd Lcm A"
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1122
    by (rule dvd_Lcm)
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1123
  then show ?thesis
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1124
    by simp
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1125
qed
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1126
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1127
lemma Gcd_UNIV [simp]:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1128
  "Gcd UNIV = 1"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1129
  using dvd_refl by (rule Gcd_eq_1_I) simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1130
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1131
lemma Lcm_UNIV [simp]:
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1132
  "Lcm UNIV = 0"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1133
  by (rule Lcm_eq_0_I) simp
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1134
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1135
lemma Lcm_0_iff:
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1136
  assumes "finite A"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1137
  shows "Lcm A = 0 \<longleftrightarrow> 0 \<in> A"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1138
proof (cases "A = {}")
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1139
  case True then show ?thesis by simp
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1140
next
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1141
  case False with assms show ?thesis
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1142
    by (induct A rule: finite_ne_induct)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1143
      (auto simp add: lcm_eq_0_iff)
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1144
qed
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  1145
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1146
lemma Gcd_finite:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1147
  assumes "finite A"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1148
  shows "Gcd A = Finite_Set.fold gcd 0 A"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1149
  by (induct rule: finite.induct[OF \<open>finite A\<close>])
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1150
     (simp_all add: comp_fun_idem.fold_insert_idem[OF comp_fun_idem_gcd])
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1151
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1152
lemma Gcd_set [code_unfold]: "Gcd (set as) = foldl gcd 0 as"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1153
  by (simp add: Gcd_finite comp_fun_idem.fold_set_fold[OF comp_fun_idem_gcd] 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1154
                foldl_conv_fold gcd.commute)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1155
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1156
lemma Lcm_finite:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1157
  assumes "finite A"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1158
  shows "Lcm A = Finite_Set.fold lcm 1 A"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1159
  by (induct rule: finite.induct[OF \<open>finite A\<close>])
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1160
     (simp_all add: comp_fun_idem.fold_insert_idem[OF comp_fun_idem_lcm])
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1161
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1162
lemma Lcm_set [code_unfold]:
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1163
  "Lcm (set as) = foldl lcm 1 as"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1164
  by (simp add: Lcm_finite comp_fun_idem.fold_set_fold[OF comp_fun_idem_lcm] 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1165
                foldl_conv_fold lcm.commute)
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1166
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1167
lemma Gcd_image_normalize [simp]:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1168
  "Gcd (normalize ` A) = Gcd A"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1169
proof -
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1170
  have "Gcd (normalize ` A) dvd a" if "a \<in> A" for a
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1171
  proof -
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1172
    from that obtain B where "A = insert a B" by blast
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
  1173
    moreover have "gcd (normalize a) (Gcd (normalize ` B)) dvd normalize a"
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1174
      by (rule gcd_dvd1)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1175
    ultimately show "Gcd (normalize ` A) dvd a"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1176
      by simp
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1177
  qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1178
  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
  1179
    by (auto intro!: Gcd_greatest intro: Gcd_dvd)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1180
  then show ?thesis
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1181
    by (auto intro: associated_eqI)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1182
qed
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1183
62346
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1184
lemma Gcd_eqI:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1185
  assumes "normalize a = a"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1186
  assumes "\<And>b. b \<in> A \<Longrightarrow> a dvd b"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1187
    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
  1188
  shows "Gcd A = a"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1189
  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
  1190
63359
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1191
lemma dvd_GcdD:
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1192
  assumes "x dvd Gcd A" "y \<in> A"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1193
  shows   "x dvd y"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1194
  using assms Gcd_dvd dvd_trans by blast
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1195
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1196
lemma dvd_Gcd_iff:
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1197
  "x dvd Gcd A \<longleftrightarrow> (\<forall>y\<in>A. x dvd y)"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1198
  by (blast dest: dvd_GcdD intro: Gcd_greatest)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1199
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1200
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
  1201
proof (cases "c = 0")
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1202
  case [simp]: False
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1203
  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
  1204
    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
  1205
       (auto intro!: Gcd_greatest Gcd_dvd simp: mult.commute[of _ c])
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1206
  hence "Gcd (op * c ` A) dvd c * Gcd A"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1207
    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
  1208
  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
  1209
    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
  1210
  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
  1211
    by (simp add: dvd_mult_unit_iff)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1212
  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
  1213
  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
  1214
    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
  1215
  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
  1216
    by (rule associatedI)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1217
  thus ?thesis by (simp add: normalize_mult)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1218
qed auto
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1219
62346
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1220
lemma Lcm_eqI:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1221
  assumes "normalize a = a"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1222
  assumes "\<And>b. b \<in> A \<Longrightarrow> b dvd a"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1223
    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
  1224
  shows "Lcm A = a"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1225
  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
  1226
63359
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1227
lemma Lcm_dvdD:
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1228
  assumes "Lcm A dvd x" "y \<in> A"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1229
  shows   "y dvd x"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1230
  using assms dvd_Lcm dvd_trans by blast
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1231
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1232
lemma Lcm_dvd_iff:
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1233
  "Lcm A dvd x \<longleftrightarrow> (\<forall>y\<in>A. y dvd x)"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1234
  by (blast dest: Lcm_dvdD intro: Lcm_least)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1235
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1236
lemma Lcm_mult: 
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1237
  assumes "A \<noteq> {}"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1238
  shows   "Lcm (op * c ` A) = normalize c * Lcm A"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1239
proof (cases "c = 0")
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1240
  case True
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1241
  moreover from assms this have "op * c ` A = {0}" by auto
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1242
  ultimately show ?thesis by auto
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1243
next
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1244
  case [simp]: False
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1245
  from assms obtain x where x: "x \<in> A" by blast
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1246
  have "c dvd c * x" by simp
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1247
  also from x have "c * x dvd Lcm (op * c ` A)" by (intro dvd_Lcm) auto
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1248
  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
  1249
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1250
  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
  1251
    by (intro Lcm_least dvd_mult_imp_div)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1252
       (auto intro!: Lcm_least dvd_Lcm simp: mult.commute[of _ c])
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1253
  hence "c * Lcm A dvd Lcm (op * c ` A)"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1254
    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
  1255
  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
  1256
    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
  1257
  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
  1258
    by (simp add: mult_unit_dvd_iff)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1259
  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
  1260
  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
  1261
    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
  1262
  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
  1263
    by (rule associatedI)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1264
  thus ?thesis by (simp add: normalize_mult)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1265
qed
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63145
diff changeset
  1266
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1267
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1268
lemma Lcm_no_units:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1269
  "Lcm A = Lcm (A - {a. is_unit a})"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1270
proof -
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1271
  have "(A - {a. is_unit a}) \<union> {a\<in>A. is_unit a} = A" by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1272
  hence "Lcm A = lcm (Lcm (A - {a. is_unit a})) (Lcm {a\<in>A. is_unit a})"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1273
    by (simp add: Lcm_Un [symmetric])
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1274
  also have "Lcm {a\<in>A. is_unit a} = 1" by (simp add: Lcm_1_iff)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1275
  finally show ?thesis by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1276
qed
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1277
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1278
lemma Lcm_0_iff': "Lcm A = 0 \<longleftrightarrow> \<not>(\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l))"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1279
  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
  1280
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1281
lemma Lcm_no_multiple: "(\<forall>m. m \<noteq> 0 \<longrightarrow> (\<exists>a\<in>A. \<not>a dvd m)) \<Longrightarrow> Lcm A = 0"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1282
  by (auto simp: Lcm_0_iff')
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1283
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1284
lemma Lcm_singleton [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1285
  "Lcm {a} = normalize a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1286
  by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1287
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1288
lemma Lcm_2 [simp]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1289
  "Lcm {a,b} = lcm a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1290
  by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1291
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1292
lemma Lcm_coprime:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1293
  assumes "finite A" and "A \<noteq> {}" 
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1294
  assumes "\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> gcd a b = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1295
  shows "Lcm A = normalize (\<Prod>A)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1296
using assms proof (induct rule: finite_ne_induct)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1297
  case (insert a A)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1298
  have "Lcm (insert a A) = lcm a (Lcm A)" by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1299
  also from insert have "Lcm A = normalize (\<Prod>A)" by blast
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1300
  also have "lcm a \<dots> = lcm a (\<Prod>A)" by (cases "\<Prod>A = 0") (simp_all add: lcm_div_unit2)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1301
  also from insert have "gcd a (\<Prod>A) = 1" by (subst gcd.commute, intro setprod_coprime) auto
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1302
  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
  1303
    by (simp add: lcm_coprime)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1304
  finally show ?case .
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1305
qed simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1306
      
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1307
lemma Lcm_coprime':
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1308
  "card A \<noteq> 0 \<Longrightarrow> (\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> gcd a b = 1)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1309
    \<Longrightarrow> Lcm A = normalize (\<Prod>A)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1310
  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
  1311
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1312
lemma Gcd_1:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1313
  "1 \<in> A \<Longrightarrow> Gcd A = 1"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1314
  by (auto intro!: Gcd_eq_1_I)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1315
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1316
lemma Gcd_singleton [simp]: "Gcd {a} = normalize a"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1317
  by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1318
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1319
lemma Gcd_2 [simp]: "Gcd {a,b} = gcd a b"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1320
  by simp
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1321
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1322
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1323
definition pairwise_coprime where
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1324
  "pairwise_coprime A = (\<forall>x y. x \<in> A \<and> y \<in> A \<and> x \<noteq> y \<longrightarrow> coprime x y)"
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1325
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1326
lemma pairwise_coprimeI [intro?]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1327
  "(\<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
  1328
  by (simp add: pairwise_coprime_def)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1329
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1330
lemma pairwise_coprimeD:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1331
  "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
  1332
  by (simp add: pairwise_coprime_def)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1333
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1334
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
  1335
  by (force simp: pairwise_coprime_def)
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1336
62350
66a381d3f88f more sophisticated GCD syntax
haftmann
parents: 62349
diff changeset
  1337
end
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1338
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1339
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
  1340
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1341
instantiation nat :: gcd
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1342
begin
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1343
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1344
fun gcd_nat  :: "nat \<Rightarrow> nat \<Rightarrow> nat"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1345
where "gcd_nat x y =
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1346
  (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
  1347
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1348
definition lcm_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1349
where
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1350
  "lcm_nat x y = x * y div (gcd x y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1351
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1352
instance proof qed
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1353
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1354
end
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1355
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1356
instantiation int :: gcd
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1357
begin
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1358
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1359
definition gcd_int  :: "int \<Rightarrow> int \<Rightarrow> int"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1360
  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
  1361
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1362
definition lcm_int :: "int \<Rightarrow> int \<Rightarrow> int"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1363
  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
  1364
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1365
instance ..
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1366
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1367
end
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1368
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1369
text \<open>Transfer setup\<close>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1370
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1371
lemma transfer_nat_int_gcd:
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1372
  "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> gcd (nat x) (nat y) = nat (gcd x y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1373
  "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> lcm (nat x) (nat y) = nat (lcm x y)"
32479
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
  1374
  unfolding gcd_int_def lcm_int_def
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1375
  by auto
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1376
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1377
lemma transfer_nat_int_gcd_closures:
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1378
  "x >= (0::int) \<Longrightarrow> y >= 0 \<Longrightarrow> gcd x y >= 0"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1379
  "x >= (0::int) \<Longrightarrow> y >= 0 \<Longrightarrow> lcm x y >= 0"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1380
  by (auto simp add: gcd_int_def lcm_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1381
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35368
diff changeset
  1382
declare transfer_morphism_nat_int[transfer add return:
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1383
    transfer_nat_int_gcd transfer_nat_int_gcd_closures]
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1384
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1385
lemma transfer_int_nat_gcd:
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1386
  "gcd (int x) (int y) = int (gcd x y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1387
  "lcm (int x) (int y) = int (lcm x y)"
32479
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
  1388
  by (unfold gcd_int_def lcm_int_def, auto)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1389
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1390
lemma transfer_int_nat_gcd_closures:
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1391
  "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
  1392
  "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> lcm x y >= 0"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1393
  by (auto simp add: gcd_int_def lcm_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1394
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35368
diff changeset
  1395
declare transfer_morphism_int_nat[transfer add return:
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1396
    transfer_int_nat_gcd transfer_int_nat_gcd_closures]
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1397
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1398
lemma gcd_nat_induct:
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1399
  fixes m n :: nat
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1400
  assumes "\<And>m. P m 0"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1401
    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
  1402
  shows "P m n"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1403
  apply (rule gcd_nat.induct)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1404
  apply (case_tac "y = 0")
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1405
  using assms apply simp_all
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1406
done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1407
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1408
(* specific to int *)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1409
62346
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1410
lemma gcd_eq_int_iff:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1411
  "gcd k l = int n \<longleftrightarrow> gcd (nat \<bar>k\<bar>) (nat \<bar>l\<bar>) = n"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1412
  by (simp add: gcd_int_def)
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1413
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1414
lemma lcm_eq_int_iff:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1415
  "lcm k l = int n \<longleftrightarrow> lcm (nat \<bar>k\<bar>) (nat \<bar>l\<bar>) = n"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1416
  by (simp add: lcm_int_def)
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  1417
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1418
lemma gcd_neg1_int [simp]: "gcd (-x::int) y = gcd x y"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1419
  by (simp add: gcd_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1420
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1421
lemma gcd_neg2_int [simp]: "gcd (x::int) (-y) = gcd x y"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1422
  by (simp add: gcd_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1423
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1424
lemma abs_gcd_int [simp]: "\<bar>gcd (x::int) y\<bar> = gcd x y"
31813
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1425
by(simp add: gcd_int_def)
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1426
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1427
lemma gcd_abs_int: "gcd (x::int) y = gcd \<bar>x\<bar> \<bar>y\<bar>"
31813
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1428
by (simp add: gcd_int_def)
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1429
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1430
lemma gcd_abs1_int [simp]: "gcd \<bar>x\<bar> (y::int) = gcd x y"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1431
by (metis abs_idempotent gcd_abs_int)
31813
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1432
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1433
lemma gcd_abs2_int [simp]: "gcd x \<bar>y::int\<bar> = gcd x y"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1434
by (metis abs_idempotent gcd_abs_int)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1435
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1436
lemma gcd_cases_int:
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1437
  fixes x :: int and y
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1438
  assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (gcd x y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1439
      and "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (gcd x (-y))"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1440
      and "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (gcd (-x) y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1441
      and "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (gcd (-x) (-y))"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1442
  shows "P (gcd x y)"
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1443
  by (insert assms, auto, arith)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1444
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1445
lemma gcd_ge_0_int [simp]: "gcd (x::int) y >= 0"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1446
  by (simp add: gcd_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1447
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1448
lemma lcm_neg1_int: "lcm (-x::int) y = lcm x y"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1449
  by (simp add: lcm_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1450
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1451
lemma lcm_neg2_int: "lcm (x::int) (-y) = lcm x y"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1452
  by (simp add: lcm_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1453
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1454
lemma lcm_abs_int: "lcm (x::int) y = lcm \<bar>x\<bar> \<bar>y\<bar>"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1455
  by (simp add: lcm_int_def)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1456
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1457
lemma abs_lcm_int [simp]: "\<bar>lcm i j::int\<bar> = lcm i j"
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1458
  by (simp add:lcm_int_def)
31814
7c122634da81 lcm abs lemmas
nipkow
parents: 31813
diff changeset
  1459
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1460
lemma lcm_abs1_int [simp]: "lcm \<bar>x\<bar> (y::int) = lcm x y"
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1461
  by (metis abs_idempotent lcm_int_def)
31814
7c122634da81 lcm abs lemmas
nipkow
parents: 31813
diff changeset
  1462
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1463
lemma lcm_abs2_int [simp]: "lcm x \<bar>y::int\<bar> = lcm x y"
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1464
  by (metis abs_idempotent lcm_int_def)
31814
7c122634da81 lcm abs lemmas
nipkow
parents: 31813
diff changeset
  1465
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1466
lemma lcm_cases_int:
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1467
  fixes x :: int and y
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1468
  assumes "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm x y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1469
      and "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (lcm x (-y))"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1470
      and "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> P (lcm (-x) y)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1471
      and "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> P (lcm (-x) (-y))"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1472
  shows "P (lcm x y)"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 37770
diff changeset
  1473
  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
  1474
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1475
lemma lcm_ge_0_int [simp]: "lcm (x::int) y >= 0"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1476
  by (simp add: lcm_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1477
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54489
diff changeset
  1478
lemma gcd_0_nat: "gcd (x::nat) 0 = x"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1479
  by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1480
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1481
lemma gcd_0_int [simp]: "gcd (x::int) 0 = \<bar>x\<bar>"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1482
  by (unfold gcd_int_def, auto)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1483
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54489
diff changeset
  1484
lemma gcd_0_left_nat: "gcd 0 (x::nat) = x"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1485
  by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  1486
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1487
lemma gcd_0_left_int [simp]: "gcd 0 (x::int) = \<bar>x\<bar>"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1488
  by (unfold gcd_int_def, auto)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1489
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1490
lemma gcd_red_nat: "gcd (x::nat) y = gcd y (x mod y)"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1491
  by (case_tac "y = 0", auto)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1492
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1493
(* weaker, but useful for the simplifier *)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1494
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1495
lemma gcd_non_0_nat: "y \<noteq> (0::nat) \<Longrightarrow> gcd (x::nat) y = gcd y (x mod y)"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1496
  by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1497
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1498
lemma gcd_1_nat [simp]: "gcd (m::nat) 1 = 1"
21263
wenzelm
parents: 21256
diff changeset
  1499
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1500
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1501
lemma gcd_Suc_0 [simp]: "gcd (m::nat) (Suc 0) = Suc 0"
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60689
diff changeset
  1502
  by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1503
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1504
lemma gcd_1_int [simp]: "gcd (m::int) 1 = 1"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1505
  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
  1506
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1507
lemma gcd_idem_nat: "gcd (x::nat) x = x"
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  1508
by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1509
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1510
lemma gcd_idem_int: "gcd (x::int) x = \<bar>x\<bar>"
31813
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1511
by (auto simp add: gcd_int_def)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1512
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1513
declare gcd_nat.simps [simp del]
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1514
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1515
text \<open>
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  1516
  \medskip @{term "gcd m n"} divides \<open>m\<close> and \<open>n\<close>.  The
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1517
  conjunctions don't seem provable separately.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1518
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1519
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1520
instance nat :: semiring_gcd
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1521
proof
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1522
  fix m n :: nat
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1523
  show "gcd m n dvd m" and "gcd m n dvd n"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1524
  proof (induct m n rule: gcd_nat_induct)
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1525
    fix m n :: nat
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1526
    assume "gcd n (m mod n) dvd m mod n" and "gcd n (m mod n) dvd n"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1527
    then have "gcd n (m mod n) dvd m"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1528
      by (rule dvd_mod_imp_dvd)
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1529
    moreover assume "0 < n"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1530
    ultimately show "gcd m n dvd m"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1531
      by (simp add: gcd_non_0_nat)
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1532
  qed (simp_all add: gcd_0_nat gcd_non_0_nat)
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1533
next
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1534
  fix m n k :: nat
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1535
  assume "k dvd m" and "k dvd n"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1536
  then show "k dvd gcd m n"
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1537
    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
  1538
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
  1539
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1540
instance int :: ring_gcd
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1541
  by standard
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1542
    (simp_all add: dvd_int_unfold_dvd_nat gcd_int_def lcm_int_def zdiv_int nat_abs_mult_distrib [symmetric] lcm_gcd gcd_greatest)
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
  1543
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1544
lemma gcd_le1_nat [simp]: "a \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> a"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1545
  by (rule dvd_imp_le, auto)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1546
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1547
lemma gcd_le2_nat [simp]: "b \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> b"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1548
  by (rule dvd_imp_le, auto)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1549
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1550
lemma gcd_le1_int [simp]: "a > 0 \<Longrightarrow> gcd (a::int) b \<le> a"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1551
  by (rule zdvd_imp_le, auto)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1552
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1553
lemma gcd_le2_int [simp]: "b > 0 \<Longrightarrow> gcd (a::int) b \<le> b"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1554
  by (rule zdvd_imp_le, auto)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1555
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1556
lemma gcd_pos_nat [simp]: "(gcd (m::nat) n > 0) = (m ~= 0 | n ~= 0)"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1557
  by (insert gcd_eq_0_iff [of m n], arith)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1558
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1559
lemma gcd_pos_int [simp]: "(gcd (m::int) n > 0) = (m ~= 0 | n ~= 0)"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1560
  by (insert gcd_eq_0_iff [of m n], insert gcd_ge_0_int [of m n], arith)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1561
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1562
lemma gcd_unique_nat: "(d::nat) dvd a \<and> d dvd b \<and>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1563
    (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1564
  apply auto
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 33318
diff changeset
  1565
  apply (rule dvd_antisym)
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1566
  apply (erule (1) gcd_greatest)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1567
  apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1568
done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1569
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1570
lemma gcd_unique_int: "d >= 0 & (d::int) dvd a \<and> d dvd b \<and>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1571
    (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 33318
diff changeset
  1572
apply (case_tac "d = 0")
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 33318
diff changeset
  1573
 apply simp
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 33318
diff changeset
  1574
apply (rule iffI)
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 33318
diff changeset
  1575
 apply (rule zdvd_antisym_nonneg)
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1576
 apply (auto intro: gcd_greatest)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1577
done
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
  1578
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
  1579
interpretation gcd_nat:
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1580
  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
  1581
  by standard (auto simp add: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans)
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  1582
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1583
lemma gcd_proj1_if_dvd_int [simp]: "x dvd y \<Longrightarrow> gcd (x::int) y = \<bar>x\<bar>"
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54489
diff changeset
  1584
  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
  1585
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1586
lemma gcd_proj2_if_dvd_int [simp]: "y dvd x \<Longrightarrow> gcd (x::int) y = \<bar>y\<bar>"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1587
  by (metis gcd_proj1_if_dvd_int gcd.commute)
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  1588
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1589
text \<open>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1590
  \medskip Multiplication laws
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1591
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1592
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1593
lemma gcd_mult_distrib_nat: "(k::nat) * gcd m n = gcd (k * m) (k * n)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  1594
    \<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
  1595
  apply (induct m n rule: gcd_nat_induct)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1596
  apply simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1597
  apply (case_tac "k = 0")
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 45264
diff changeset
  1598
  apply (simp_all add: gcd_non_0_nat)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1599
done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1600
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1601
lemma gcd_mult_distrib_int: "\<bar>k::int\<bar> * gcd m n = gcd (k * m) (k * n)"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1602
  apply (subst (1 2) gcd_abs_int)
31813
4df828bbc411 gcd abs lemmas
nipkow
parents: 31798
diff changeset
  1603
  apply (subst (1 2) abs_mult)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1604
  apply (rule gcd_mult_distrib_nat [transferred])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1605
  apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1606
done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1607
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1608
lemma coprime_crossproduct_nat:
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1609
  fixes a b c d :: nat
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1610
  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
  1611
  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
  1612
  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
  1613
35368
19b340c3f1ff crossproduct coprimality lemmas
haftmann
parents: 35216
diff changeset
  1614
lemma coprime_crossproduct_int:
19b340c3f1ff crossproduct coprimality lemmas
haftmann
parents: 35216
diff changeset
  1615
  fixes a b c d :: int
19b340c3f1ff crossproduct coprimality lemmas
haftmann
parents: 35216
diff changeset
  1616
  assumes "coprime a d" and "coprime b c"
19b340c3f1ff crossproduct coprimality lemmas
haftmann
parents: 35216
diff changeset
  1617
  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
  1618
  using assms coprime_crossproduct [of a d b c] by simp
35368
19b340c3f1ff crossproduct coprimality lemmas
haftmann
parents: 35216
diff changeset
  1619
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1620
text \<open>\medskip Addition laws\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1621
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1622
(* to do: add the other variations? *)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1623
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1624
lemma gcd_diff1_nat: "(m::nat) >= n \<Longrightarrow> gcd (m - n) n = gcd m n"
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1625
  by (subst gcd_add1 [symmetric]) auto
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1626
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1627
lemma gcd_diff2_nat: "(n::nat) >= m \<Longrightarrow> gcd (n - m) n = gcd m n"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1628
  apply (subst gcd.commute)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1629
  apply (subst gcd_diff1_nat [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1630
  apply auto
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1631
  apply (subst gcd.commute)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1632
  apply (subst gcd_diff1_nat)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1633
  apply assumption
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1634
  apply (rule gcd.commute)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1635
  done
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1636
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1637
lemma gcd_non_0_int: "(y::int) > 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1638
  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
  1639
  apply (simp del: pos_mod_sign add: gcd_int_def abs_if nat_mod_distrib)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1640
  apply (auto simp add: gcd_non_0_nat nat_mod_distrib [symmetric]
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1641
    zmod_zminus1_eq_if)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1642
  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
  1643
  apply (subst (1 2) gcd.commute)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1644
  apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2_nat
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1645
    nat_le_eq_zle)
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1646
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  1647
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1648
lemma gcd_red_int: "gcd (x::int) y = gcd y (x mod y)"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1649
  apply (case_tac "y = 0")
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1650
  apply force
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1651
  apply (case_tac "y > 0")
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1652
  apply (subst gcd_non_0_int, auto)
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1653
  apply (insert gcd_non_0_int [of "-y" "-x"])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35028
diff changeset
  1654
  apply auto
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1655
done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1656
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1657
(* to do: differences, and all variations of addition rules
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1658
    as simplification rules for nat and int *)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1659
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1660
(* to do: add the three variations of these, and for ints? *)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1661
63145
703edebd1d92 isabelle update_cartouches -c -t;
wenzelm
parents: 63025
diff changeset
  1662
lemma finite_divisors_nat [simp]: \<comment> \<open>FIXME move\<close>
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1663
  fixes m :: nat
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1664
  assumes "m > 0" 
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1665
  shows "finite {d. d dvd m}"
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1666
proof-
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1667
  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
  1668
    by (auto dest: dvd_imp_le)
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1669
  then show ?thesis
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1670
    using finite_Collect_le_nat by (rule finite_subset)
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1671
qed
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1672
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1673
lemma finite_divisors_int [simp]:
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1674
  fixes i :: int
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1675
  assumes "i \<noteq> 0"
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1676
  shows "finite {d. d dvd i}"
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1677
proof -
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1678
  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
  1679
    by (auto simp: abs_if)
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1680
  then have "finite {d. \<bar>d\<bar> <= \<bar>i\<bar>}"
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1681
    by simp
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1682
  from finite_subset [OF _ this] show ?thesis using assms
60512
e0169291b31c tuned proofs -- slightly faster;
wenzelm
parents: 60357
diff changeset
  1683
    by (simp add: dvd_imp_le_int subset_iff)
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1684
qed
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1685
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1686
lemma Max_divisors_self_nat [simp]: "n\<noteq>0 \<Longrightarrow> Max{d::nat. d dvd n} = n"
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1687
apply(rule antisym)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44845
diff changeset
  1688
 apply (fastforce intro: Max_le_iff[THEN iffD2] simp: dvd_imp_le)
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1689
apply simp
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1690
done
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1691
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1692
lemma Max_divisors_self_int [simp]: "n\<noteq>0 \<Longrightarrow> Max{d::int. d dvd n} = \<bar>n\<bar>"
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1693
apply(rule antisym)
44278
1220ecb81e8f observe distinction between sets and predicates more properly
haftmann
parents: 42871
diff changeset
  1694
 apply(rule Max_le_iff [THEN iffD2])
1220ecb81e8f observe distinction between sets and predicates more properly
haftmann
parents: 42871
diff changeset
  1695
  apply (auto intro: abs_le_D1 dvd_imp_le_int)
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1696
done
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1697
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1698
lemma gcd_is_Max_divisors_nat:
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1699
  "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> gcd (m::nat) n = Max {d. d dvd m \<and> d dvd n}"
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1700
apply(rule Max_eqI[THEN sym])
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1701
  apply (metis finite_Collect_conjI finite_divisors_nat)
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1702
 apply simp
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1703
 apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le gcd_greatest_iff gcd_pos_nat)
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1704
apply simp
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1705
done
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1706
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1707
lemma gcd_is_Max_divisors_int:
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1708
  "m ~= 0 ==> n ~= 0 ==> gcd (m::int) n = (Max {d. d dvd m & d dvd n})"
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1709
apply(rule Max_eqI[THEN sym])
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  1710
  apply (metis finite_Collect_conjI finite_divisors_int)
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1711
 apply simp
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1712
 apply (metis gcd_greatest_iff gcd_pos_int zdvd_imp_le)
31734
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1713
apply simp
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1714
done
a4a79836d07b new lemmas
nipkow
parents: 31730
diff changeset
  1715
34030
829eb528b226 resorted code equations from "old" number theory version
haftmann
parents: 33946
diff changeset
  1716
lemma gcd_code_int [code]:
829eb528b226 resorted code equations from "old" number theory version
haftmann
parents: 33946
diff changeset
  1717
  "gcd k l = \<bar>if l = (0::int) then k else gcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
829eb528b226 resorted code equations from "old" number theory version
haftmann
parents: 33946
diff changeset
  1718
  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
  1719
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
  1720
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1721
subsection \<open>Coprimality\<close>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1722
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1723
lemma coprime_nat:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1724
  "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1725
  using coprime [of a b] by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1726
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1727
lemma coprime_Suc_0_nat:
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1728
  "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = Suc 0)"
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60689
diff changeset
  1729
  using coprime_nat by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1730
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1731
lemma coprime_int:
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  1732
  "coprime (a::int) b \<longleftrightarrow> (\<forall>d. d \<ge> 0 \<and> d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1733
  using gcd_unique_int [of 1 a b]
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1734
  apply clarsimp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1735
  apply (erule subst)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1736
  apply (rule iffI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1737
  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
  1738
  using abs_dvd_iff abs_ge_zero apply blast
59807
22bc39064290 prefer local fixes;
wenzelm
parents: 59667
diff changeset
  1739
  done
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1740
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1741
lemma pow_divides_eq_nat [simp]:
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  1742
  "n > 0 \<Longrightarrow> (a::nat) ^ 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
  1743
  using pow_divs_eq[of n] by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1744
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1745
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
  1746
  using coprime_plus_one[of n] by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1747
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1748
lemma coprime_minus_one_nat: "(n::nat) \<noteq> 0 \<Longrightarrow> coprime (n - 1) n"
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1749
  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
  1750
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 59977
diff changeset
  1751
lemma coprime_common_divisor_nat: 
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1752
  "coprime (a::nat) b \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> x = 1"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  1753
  by (metis gcd_greatest_iff nat_dvd_1_iff_1)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1754
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 59977
diff changeset
  1755
lemma coprime_common_divisor_int:
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  1756
  "coprime (a::int) b \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> \<bar>x\<bar> = 1"
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  1757
  using gcd_greatest_iff [of x a b] by auto
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1758
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1759
lemma invertible_coprime_nat: "(x::nat) * y mod m = 1 \<Longrightarrow> coprime x m"
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1760
by (metis coprime_lmult gcd_1_nat gcd.commute gcd_red_nat)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1761
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1762
lemma invertible_coprime_int: "(x::int) * y mod m = 1 \<Longrightarrow> coprime x m"
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  1763
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
  1764
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1765
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1766
subsection \<open>Bezout's theorem\<close>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1767
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1768
(* Function bezw returns a pair of witnesses to Bezout's theorem --
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1769
   see the theorems that follow the definition. *)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1770
fun
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1771
  bezw  :: "nat \<Rightarrow> nat \<Rightarrow> int * int"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1772
where
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1773
  "bezw x y =
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1774
  (if y = 0 then (1, 0) else
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1775
      (snd (bezw y (x mod y)),
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1776
       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
  1777
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1778
lemma bezw_0 [simp]: "bezw x 0 = (1, 0)" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1779
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1780
lemma bezw_non_0: "y > 0 \<Longrightarrow> bezw x y = (snd (bezw y (x mod y)),
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1781
       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
  1782
  by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1783
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1784
declare bezw.simps [simp del]
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1785
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1786
lemma bezw_aux [rule_format]:
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1787
    "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
  1788
proof (induct x y rule: gcd_nat_induct)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1789
  fix m :: nat
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1790
  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
  1791
    by auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1792
  next fix m :: nat and n
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1793
    assume ngt0: "n > 0" and
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1794
      ih: "fst (bezw n (m mod n)) * int n +
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1795
        snd (bezw n (m mod n)) * int (m mod n) =
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1796
        int (gcd n (m mod n))"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1797
    thus "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1798
      apply (simp add: bezw_non_0 gcd_non_0_nat)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1799
      apply (erule subst)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 35726
diff changeset
  1800
      apply (simp add: field_simps)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1801
      apply (subst mod_div_equality [of m n, symmetric])
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1802
      (* applying simp here undoes the last substitution!
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1803
         what is procedure cancel_div_mod? *)
58776
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 58770
diff changeset
  1804
      apply (simp only: NO_MATCH_def field_simps of_nat_add of_nat_mult)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1805
      done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1806
qed
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1807
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1808
lemma bezout_int:
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1809
  fixes x y
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1810
  shows "EX u v. u * (x::int) + v * y = gcd x y"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1811
proof -
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1812
  have bezout_aux: "!!x y. x \<ge> (0::int) \<Longrightarrow> y \<ge> 0 \<Longrightarrow>
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1813
      EX u v. u * x + v * y = gcd x y"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1814
    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
  1815
    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
  1816
    apply (unfold gcd_int_def)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1817
    apply simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1818
    apply (subst bezw_aux [symmetric])
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1819
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1820
    done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1821
  have "(x \<ge> 0 \<and> y \<ge> 0) | (x \<ge> 0 \<and> y \<le> 0) | (x \<le> 0 \<and> y \<ge> 0) |
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1822
      (x \<le> 0 \<and> y \<le> 0)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1823
    by auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1824
  moreover have "x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> ?thesis"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1825
    by (erule (1) bezout_aux)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1826
  moreover have "x >= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> ?thesis"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1827
    apply (insert bezout_aux [of x "-y"])
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1828
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1829
    apply (rule_tac x = u in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1830
    apply (rule_tac x = "-v" in exI)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1831
    apply (subst gcd_neg2_int [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1832
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1833
    done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1834
  moreover have "x <= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> ?thesis"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1835
    apply (insert bezout_aux [of "-x" y])
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1836
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1837
    apply (rule_tac x = "-u" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1838
    apply (rule_tac x = v in exI)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1839
    apply (subst gcd_neg1_int [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1840
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1841
    done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1842
  moreover have "x <= 0 \<Longrightarrow> y <= 0 \<Longrightarrow> ?thesis"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1843
    apply (insert bezout_aux [of "-x" "-y"])
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1844
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1845
    apply (rule_tac x = "-u" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1846
    apply (rule_tac x = "-v" in exI)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1847
    apply (subst gcd_neg1_int [symmetric])
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1848
    apply (subst gcd_neg2_int [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1849
    apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1850
    done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1851
  ultimately show ?thesis by blast
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1852
qed
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1853
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1854
text \<open>versions of Bezout for nat, by Amine Chaieb\<close>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1855
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1856
lemma ind_euclid:
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1857
  assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1858
  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
  1859
  shows "P a b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1860
proof(induct "a + b" arbitrary: a b rule: less_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1861
  case less
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
  1862
  have "a = b \<or> a < b \<or> b < a" by arith
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
  1863
  moreover {assume eq: "a= b"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1864
    from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1865
    by simp}
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
  1866
  moreover
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
  1867
  {assume lt: "a < b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1868
    hence "a + b - a < a + b \<or> a = 0" by arith
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
  1869
    moreover
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
  1870
    {assume "a =0" with z c have "P a b" by blast }
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
  1871
    moreover
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1872
    {assume "a + b - a < a + b"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1873
      also have th0: "a + b - a = a + (b - a)" using lt by arith
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1874
      finally have "a + (b - a) < a + b" .
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1875
      then have "P a (a + (b - a))" by (rule add[rule_format, OF less])
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1876
      then have "P a b" by (simp add: th0[symmetric])}
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
  1877
    ultimately have "P a b" by blast}
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
  1878
  moreover
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
  1879
  {assume lt: "a > b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1880
    hence "b + a - b < a + b \<or> b = 0" by arith
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
  1881
    moreover
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
  1882
    {assume "b =0" with z c have "P a b" by blast }
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
  1883
    moreover
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1884
    {assume "b + a - b < a + b"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1885
      also have th0: "b + a - b = b + (a - b)" using lt by arith
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1886
      finally have "b + (a - b) < a + b" .
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1887
      then have "P b (b + (a - b))" by (rule add[rule_format, OF less])
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 34223
diff changeset
  1888
      then have "P b a" by (simp add: th0[symmetric])
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
  1889
      hence "P a b" using c by blast }
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
  1890
    ultimately have "P a b" by blast}
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
  1891
ultimately  show "P a b" by blast
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
  1892
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
  1893
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1894
lemma bezout_lemma_nat:
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1895
  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
  1896
    (a * x = b * y + d \<or> b * x = a * y + d)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1897
  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
  1898
    (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
  1899
  using ex
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1900
  apply clarsimp
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35028
diff changeset
  1901
  apply (rule_tac x="d" in exI, simp)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1902
  apply (case_tac "a * x = b * y + d" , simp_all)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1903
  apply (rule_tac x="x + y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1904
  apply (rule_tac x="y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1905
  apply algebra
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1906
  apply (rule_tac x="x" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1907
  apply (rule_tac x="x + y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1908
  apply algebra
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
  1909
done
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
  1910
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1911
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
  1912
    (a * x = b * y + d \<or> b * x = a * y + d)"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1913
  apply(induct a b rule: ind_euclid)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1914
  apply blast
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1915
  apply clarify
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35028
diff changeset
  1916
  apply (rule_tac x="a" in exI, simp)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1917
  apply clarsimp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1918
  apply (rule_tac x="d" in exI)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35028
diff changeset
  1919
  apply (case_tac "a * x = b * y + d", simp_all)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1920
  apply (rule_tac x="x+y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1921
  apply (rule_tac x="y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1922
  apply algebra
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1923
  apply (rule_tac x="x" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1924
  apply (rule_tac x="x+y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1925
  apply algebra
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
  1926
done
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
  1927
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1928
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
  1929
    (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
  1930
  using bezout_add_nat[of a b]
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1931
  apply clarsimp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1932
  apply (rule_tac x="d" in exI, simp)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1933
  apply (rule_tac x="x" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1934
  apply (rule_tac x="y" in exI)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1935
  apply auto
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
  1936
done
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
  1937
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1938
lemma bezout_add_strong_nat: assumes nz: "a \<noteq> (0::nat)"
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
  1939
  shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
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
  1940
proof-
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1941
 from nz have ap: "a > 0" by simp
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1942
 from bezout_add_nat[of a b]
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1943
 have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or>
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1944
   (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" 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
  1945
 moreover
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1946
    {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1947
     from H have ?thesis 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
  1948
 moreover
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
  1949
 {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
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
  1950
   {assume b0: "b = 0" with H  have ?thesis by simp}
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1951
   moreover
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
  1952
   {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1953
     from b dvd_imp_le [OF H(2)] have "d < b \<or> d = b"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1954
       by auto
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
  1955
     moreover
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
  1956
     {assume db: "d=b"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 37770
diff changeset
  1957
       with nz H have ?thesis apply simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1958
         apply (rule exI[where x = b], simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1959
         apply (rule exI[where x = b])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1960
        by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
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
  1961
    moreover
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  1962
    {assume db: "d < b"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 37770
diff changeset
  1963
        {assume "x=0" hence ?thesis using nz H by simp }
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1964
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1965
        {assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1966
          from db have "d \<le> b - 1" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1967
          hence "d*b \<le> b*(b - 1)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1968
          with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1969
          have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1970
          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
  1971
            by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1972
          hence "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
  1973
            by (simp only: mult.assoc distrib_left)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1974
          hence "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
  1975
            by algebra
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1976
          hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1977
          hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1978
            by (simp only: diff_add_assoc[OF dble, of d, symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1979
          hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
59008
f61482b0f240 formally self-contained gcd type classes
haftmann
parents: 58889
diff changeset
  1980
            by (simp only: diff_mult_distrib2 ac_simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1981
          hence ?thesis using H(1,2)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1982
            apply -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1983
            apply (rule exI[where x=d], simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1984
            apply (rule exI[where x="(b - 1) * y"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1985
            by (rule exI[where x="x*(b - 1) - d"], simp)}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32879
diff changeset
  1986
        ultimately have ?thesis 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
  1987
    ultimately have ?thesis by blast}
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
  1988
  ultimately have ?thesis by blast}
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
  1989
 ultimately show ?thesis by blast
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
  1990
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
  1991
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1992
lemma bezout_nat: assumes a: "(a::nat) \<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
  1993
  shows "\<exists>x y. a * x = b * y + gcd a b"
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
  1994
proof-
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
  1995
  let ?g = "gcd a b"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  1996
  from bezout_add_strong_nat[OF a, of 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
  1997
  obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
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
  1998
  from d(1,2) have "d dvd ?g" by simp
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
  1999
  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2000
  from d(3) have "a * x * k = (b * y + d) *k " by auto
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
  2001
  hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
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
  2002
  thus ?thesis by blast
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
  2003
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
  2004
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2005
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2006
subsection \<open>LCM properties  on @{typ nat} and @{typ int}\<close>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2007
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  2008
lemma lcm_altdef_int [code]: "lcm (a::int) b = \<bar>a\<bar> * \<bar>b\<bar> div gcd a b"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2009
  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
  2010
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2011
lemma prod_gcd_lcm_nat: "(m::nat) * n = gcd m n * lcm m n"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2012
  unfolding lcm_nat_def
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  2013
  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
  2014
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  2015
lemma prod_gcd_lcm_int: "\<bar>m::int\<bar> * \<bar>n\<bar> = gcd m n * lcm m n"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2016
  unfolding lcm_int_def gcd_int_def
62348
9a5f43dac883 dropped various legacy fact bindings
haftmann
parents: 62347
diff changeset
  2017
  apply (subst of_nat_mult [symmetric])
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2018
  apply (subst prod_gcd_lcm_nat [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2019
  apply (subst nat_abs_mult_distrib [symmetric])
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2020
  apply (simp, simp add: abs_mult)
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2021
done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2022
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2023
lemma lcm_pos_nat:
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  2024
  "(m::nat) > 0 \<Longrightarrow> n>0 \<Longrightarrow> lcm m n > 0"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2025
by (metis gr0I mult_is_0 prod_gcd_lcm_nat)
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
  2026
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2027
lemma lcm_pos_int:
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  2028
  "(m::int) ~= 0 \<Longrightarrow> n ~= 0 \<Longrightarrow> lcm m n > 0"
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2029
  apply (subst lcm_abs_int)
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2030
  apply (rule lcm_pos_nat [transferred])
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  2031
  apply auto
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2032
  done
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  2033
63145
703edebd1d92 isabelle update_cartouches -c -t;
wenzelm
parents: 63025
diff changeset
  2034
lemma dvd_pos_nat: \<comment> \<open>FIXME move\<close>
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  2035
  fixes n m :: nat
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  2036
  assumes "n > 0" and "m dvd n"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
  2037
  shows "m > 0"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2038
  using assms by (cases m) auto
31729
b9299916d618 new lemmas and tuning
nipkow
parents: 31709
diff changeset
  2039
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2040
lemma lcm_unique_nat: "(a::nat) dvd d \<and> b dvd d \<and>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2041
    (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2042
  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
  2043
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2044
lemma lcm_unique_int: "d >= 0 \<and> (a::int) dvd d \<and> b dvd d \<and>
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2045
    (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2046
  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
  2047
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  2048
lemma lcm_proj2_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> lcm x y = y"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2049
  apply (rule sym)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2050
  apply (subst lcm_unique_nat [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2051
  apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2052
done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2053
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  2054
lemma lcm_proj2_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm x y = \<bar>y\<bar>"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2055
  apply (rule sym)
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31814
diff changeset
  2056
  apply (subst lcm_unique_int [symmetric])
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2057
  apply auto
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2058
done
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2059
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
  2060
lemma lcm_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> lcm y x = y"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2061
by (subst lcm.commute, erule lcm_proj2_if_dvd_nat)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2062
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  2063
lemma lcm_proj1_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm y x = \<bar>y\<bar>"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2064
by (subst lcm.commute, erule lcm_proj2_if_dvd_int)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
  2065
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2066
lemma lcm_proj1_iff_nat [simp]: "lcm m n = (m::nat) \<longleftrightarrow> n dvd m"
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2067
by (metis lcm_proj1_if_dvd_nat lcm_unique_nat)
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2068
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2069
lemma lcm_proj2_iff_nat [simp]: "lcm m n = (n::nat) \<longleftrightarrow> m dvd n"
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2070
by (metis lcm_proj2_if_dvd_nat lcm_unique_nat)
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2071
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2072
lemma lcm_proj1_iff_int [simp]: "lcm m n = \<bar>m::int\<bar> \<longleftrightarrow> n dvd m"
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2073
by (metis dvd_abs_iff lcm_proj1_if_dvd_int lcm_unique_int)
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2074
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2075
lemma lcm_proj2_iff_int [simp]: "lcm m n = \<bar>n::int\<bar> \<longleftrightarrow> m dvd n"
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31952
diff changeset
  2076
by (metis dvd_abs_iff lcm_proj2_if_dvd_int lcm_unique_int)
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
  2077
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2078
lemma lcm_1_iff_nat [simp]:
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2079
  "lcm (m::nat) n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = Suc 0"
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2080
  using lcm_eq_1_iff [of m n] by simp
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
  2081
  
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2082
lemma lcm_1_iff_int [simp]:
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2083
  "lcm (m::int) n = 1 \<longleftrightarrow> (m=1 \<or> m = -1) \<and> (n=1 \<or> n = -1)"
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
  2084
  by auto
31995
8f37cf60b885 more gcd/lcm lemmas
nipkow
parents: 31992
diff changeset
  2085
34030
829eb528b226 resorted code equations from "old" number theory version
haftmann
parents: 33946
diff changeset
  2086
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2087
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
  2088
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2089
text\<open>Lifting gcd and lcm to sets (Gcd/Lcm).
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2090
Gcd is defined via Lcm to facilitate the proof that we have a complete lattice.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2091
\<close>
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2092
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2093
instantiation nat :: semiring_Gcd
32112
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2094
begin
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2095
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2096
interpretation semilattice_neutr_set lcm "1::nat"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2097
  by standard simp_all
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54489
diff changeset
  2098
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2099
definition
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2100
  "Lcm (M::nat set) = (if finite M then F M else 0)"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49962
diff changeset
  2101
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49962
diff changeset
  2102
lemma Lcm_nat_empty:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49962
diff changeset
  2103
  "Lcm {} = (1::nat)"
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60689
diff changeset
  2104
  by (simp add: Lcm_nat_def del: One_nat_def)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49962
diff changeset
  2105
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49962
diff changeset
  2106
lemma Lcm_nat_insert:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 49962
diff changeset
  2107
  "Lcm (insert n M) = lcm (n::nat) (Lcm M)"
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2108
  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
  2109
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2110
lemma Lcm_nat_infinite:
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2111
  "infinite M \<Longrightarrow> Lcm M = (0::nat)"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2112
  by (simp add: Lcm_nat_def)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2113
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2114
lemma dvd_Lcm_nat [simp]:
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2115
  fixes M :: "nat set"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2116
  assumes "m \<in> M"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2117
  shows "m dvd Lcm M"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2118
proof -
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2119
  from assms have "insert m M = M" by auto
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2120
  moreover have "m dvd Lcm (insert m M)"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2121
    by (simp add: Lcm_nat_insert)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2122
  ultimately show ?thesis by simp
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2123
qed
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2124
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2125
lemma Lcm_dvd_nat [simp]:
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2126
  fixes M :: "nat set"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2127
  assumes "\<forall>m\<in>M. m dvd n"
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2128
  shows "Lcm M dvd n"
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2129
proof (cases "n > 0")
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2130
  case False then show ?thesis by simp
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2131
next
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2132
  case True
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2133
  then have "finite {d. d dvd n}" by (rule finite_divisors_nat)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2134
  moreover have "M \<subseteq> {d. d dvd n}" using assms by fast
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2135
  ultimately have "finite M" by (rule rev_finite_subset)
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2136
  then show ?thesis using assms
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2137
    by (induct M) (simp_all add: Lcm_nat_empty Lcm_nat_insert)
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2138
qed
32112
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2139
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2140
definition
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2141
  "Gcd (M::nat set) = Lcm {d. \<forall>m\<in>M. d dvd m}"
32112
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2142
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2143
instance proof
61929
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2144
  show "Gcd N dvd n" if "n \<in> N" for N and n :: nat
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2145
  using that by (induct N rule: infinite_finite_induct)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2146
    (auto simp add: Gcd_nat_def)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2147
  show "n dvd Gcd N" if "\<And>m. m \<in> N \<Longrightarrow> n dvd m" for N and n :: nat
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2148
  using that by (induct N rule: infinite_finite_induct)
b8e242e52c97 tuned proofs and augmented lemmas
haftmann
parents: 61913
diff changeset
  2149
    (auto simp add: Gcd_nat_def)
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2150
  show "n dvd Lcm N" if "n \<in> N" for N and n ::nat
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2151
  using that by (induct N rule: infinite_finite_induct)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2152
    auto
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2153
  show "Lcm N dvd n" if "\<And>m. m \<in> N \<Longrightarrow> m dvd n" for N and n ::nat
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2154
  using that by (induct N rule: infinite_finite_induct)
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2155
    auto
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2156
qed simp_all
32112
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2157
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2158
end
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
  2159
62346
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2160
lemma Gcd_nat_eq_one:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2161
  "1 \<in> N \<Longrightarrow> Gcd N = (1::nat)"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2162
  by (rule Gcd_eq_1_I) auto
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2163
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2164
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
  2165
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2166
lemma Gcd_eq_Max:
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2167
  fixes M :: "nat set"
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2168
  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
  2169
  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
  2170
proof (rule antisym)
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2171
  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
  2172
    by auto
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2173
  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
  2174
    by (blast intro: finite_divisors_nat)
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2175
  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
  2176
    by blast
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2177
  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
  2178
    by (auto intro: Max_ge Gcd_dvd)
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2179
  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
  2180
    apply (rule Max.boundedI)
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2181
    apply auto
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2182
    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
  2183
    done
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2184
qed
32112
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2185
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2186
lemma Gcd_remove0_nat: "finite M \<Longrightarrow> Gcd M = Gcd (M - {0::nat})"
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2187
apply(induct pred:finite)
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2188
 apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2189
apply(case_tac "x=0")
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2190
 apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2191
apply(subgoal_tac "insert x F - {0} = insert x (F - {0})")
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2192
 apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2193
apply blast
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2194
done
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2195
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2196
lemma Lcm_in_lcm_closed_set_nat:
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2197
  "finite M \<Longrightarrow> M \<noteq> {} \<Longrightarrow> ALL m n :: nat. m:M \<longrightarrow> n:M \<longrightarrow> lcm m n : M \<Longrightarrow> Lcm M : M"
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2198
apply(induct rule:finite_linorder_min_induct)
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2199
 apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2200
apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2201
apply(subgoal_tac "ALL m n :: nat. m:A \<longrightarrow> n:A \<longrightarrow> lcm m n : A")
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2202
 apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2203
 apply(case_tac "A={}")
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2204
  apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2205
 apply simp
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2206
apply (metis lcm_pos_nat lcm_unique_nat linorder_neq_iff nat_dvd_not_less not_less0)
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2207
done
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2208
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2209
lemma Lcm_eq_Max_nat:
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2210
  "finite M \<Longrightarrow> M \<noteq> {} \<Longrightarrow> 0 \<notin> M \<Longrightarrow> ALL m n :: nat. m:M \<longrightarrow> n:M \<longrightarrow> lcm m n : M \<Longrightarrow> Lcm M = Max M"
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2211
apply(rule antisym)
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2212
 apply(rule Max_ge, assumption)
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2213
 apply(erule (2) Lcm_in_lcm_closed_set_nat)
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2214
apply (auto simp add: not_le Lcm_0_iff dvd_imp_le leD le_neq_trans)
32112
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2215
done
6da9c2a49fed Made dvd/gcd/lcm a complete lattice by introducing Gcd/GCD/Lcm/LCM
nipkow
parents: 32111
diff changeset
  2216
34222
e33ee7369ecb added lemma
nipkow
parents: 34221
diff changeset
  2217
lemma mult_inj_if_coprime_nat:
e33ee7369ecb added lemma
nipkow
parents: 34221
diff changeset
  2218
  "inj_on f A \<Longrightarrow> inj_on g B \<Longrightarrow> ALL a:A. ALL b:B. coprime (f a) (g b)
e33ee7369ecb added lemma
nipkow
parents: 34221
diff changeset
  2219
   \<Longrightarrow> inj_on (%(a,b). f a * g b::nat) (A \<times> B)"
61913
58b153bfa737 tuned proofs and augmented some lemmas
haftmann
parents: 61856
diff changeset
  2220
  by (auto simp add: inj_on_def coprime_crossproduct_nat simp del: One_nat_def)
34222
e33ee7369ecb added lemma
nipkow
parents: 34221
diff changeset
  2221
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2222
text\<open>Nitpick:\<close>
34222
e33ee7369ecb added lemma
nipkow
parents: 34221
diff changeset
  2223
41792
ff3cb0c418b7 renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents: 41550
diff changeset
  2224
lemma gcd_eq_nitpick_gcd [nitpick_unfold]: "gcd x y = Nitpick.nat_gcd x y"
ff3cb0c418b7 renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents: 41550
diff changeset
  2225
by (induct x y rule: nat_gcd.induct)
ff3cb0c418b7 renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents: 41550
diff changeset
  2226
   (simp add: gcd_nat.simps Nitpick.nat_gcd.simps)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32960
diff changeset
  2227
41792
ff3cb0c418b7 renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents: 41550
diff changeset
  2228
lemma lcm_eq_nitpick_lcm [nitpick_unfold]: "lcm x y = Nitpick.nat_lcm x y"
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32960
diff changeset
  2229
by (simp only: lcm_nat_def Nitpick.nat_lcm_def gcd_eq_nitpick_gcd)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents: 32960
diff changeset
  2230
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54489
diff changeset
  2231
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2232
subsubsection \<open>Setwise gcd and lcm for integers\<close>
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2233
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2234
instantiation int :: semiring_Gcd
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2235
begin
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2236
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2237
definition
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
  2238
  "Lcm M = int (LCM m\<in>M. (nat \<circ> abs) m)"
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2239
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2240
definition
63025
92680537201f capitalized GCD and LCM syntax
haftmann
parents: 62429
diff changeset
  2241
  "Gcd M = int (GCD m\<in>M. (nat \<circ> abs) m)"
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2242
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2243
instance by standard
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2244
  (auto intro!: Gcd_dvd Gcd_greatest simp add: Gcd_int_def
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2245
    Lcm_int_def int_dvd_iff dvd_int_iff dvd_int_unfold_dvd_nat [symmetric])
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2246
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2247
end
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2248
62346
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2249
lemma abs_Gcd [simp]:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2250
  fixes K :: "int set"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2251
  shows "\<bar>Gcd K\<bar> = Gcd K"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2252
  using normalize_Gcd [of K] by simp
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2253
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2254
lemma abs_Lcm [simp]:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2255
  fixes K :: "int set"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2256
  shows "\<bar>Lcm K\<bar> = Lcm K"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2257
  using normalize_Lcm [of K] by simp
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2258
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2259
lemma Gcm_eq_int_iff:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2260
  "Gcd K = int n \<longleftrightarrow> Gcd ((nat \<circ> abs) ` K) = n"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2261
  by (simp add: Gcd_int_def comp_def image_image)
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2262
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2263
lemma Lcm_eq_int_iff:
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2264
  "Lcm K = int n \<longleftrightarrow> Lcm ((nat \<circ> abs) ` K) = n"
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2265
  by (simp add: Lcm_int_def comp_def image_image)
97f2ed240431 more theorems concerning gcd/lcm/Gcd/Lcm
haftmann
parents: 62345
diff changeset
  2266
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2267
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2268
subsection \<open>GCD and LCM on @{typ integer}\<close>
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2269
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2270
instantiation integer :: gcd
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2271
begin
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2272
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2273
context
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2274
  includes integer.lifting
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2275
begin
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2276
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2277
lift_definition gcd_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2278
  is gcd .
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2279
lift_definition lcm_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2280
  is lcm .
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2281
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2282
end
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2283
instance ..
60686
ea5bc46c11e6 more algebraic properties for gcd/lcm
haftmann
parents: 60597
diff changeset
  2284
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
  2285
end
45264
3b2c770f6631 merge Gcd/GCD and Lcm/LCM
huffman
parents: 44890
diff changeset
  2286
61856
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2287
lifting_update integer.lifting
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2288
lifting_forget integer.lifting
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2289
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2290
context
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2291
  includes integer.lifting
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2292
begin
61856
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2293
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2294
lemma gcd_code_integer [code]:
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2295
  "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
  2296
  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
  2297
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61929
diff changeset
  2298
lemma lcm_code_integer [code]: "lcm (a::integer) b = \<bar>a\<bar> * \<bar>b\<bar> div gcd a b"
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2299
  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
  2300
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2301
end
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2302
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2303
code_printing constant "gcd :: integer \<Rightarrow> _"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2304
  \<rightharpoonup> (OCaml) "Big'_int.gcd'_big'_int"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2305
  and (Haskell) "Prelude.gcd"
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2306
  and (Scala) "_.gcd'((_)')"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61954
diff changeset
  2307
  \<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
  2308
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2309
text \<open>Some code equations\<close>
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2310
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  2311
lemmas Lcm_set_nat [code, code_unfold] = Lcm_set[where ?'a = nat]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  2312
lemmas Gcd_set_nat [code] = Gcd_set[where ?'a = nat]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  2313
lemmas Lcm_set_int [code, code_unfold] = Lcm_set[where ?'a = int]
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62353
diff changeset
  2314
lemmas Gcd_set_int [code] = Gcd_set[where ?'a = int]
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2315
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2316
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2317
text \<open>Fact aliasses\<close>
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2318
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2319
lemma lcm_0_iff_nat [simp]: "lcm (m::nat) n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2320
  by (fact lcm_eq_0_iff)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2321
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2322
lemma lcm_0_iff_int [simp]: "lcm (m::int) n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2323
  by (fact lcm_eq_0_iff)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2324
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2325
lemma dvd_lcm_I1_nat [simp]: "(k::nat) dvd m \<Longrightarrow> k dvd lcm m n"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2326
  by (fact dvd_lcmI1)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2327
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2328
lemma dvd_lcm_I2_nat [simp]: "(k::nat) dvd n \<Longrightarrow> k dvd lcm m n"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2329
  by (fact dvd_lcmI2)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2330
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2331
lemma dvd_lcm_I1_int [simp]: "(i::int) dvd m \<Longrightarrow> i dvd lcm m n"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2332
  by (fact dvd_lcmI1)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2333
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2334
lemma dvd_lcm_I2_int [simp]: "(i::int) dvd n \<Longrightarrow> i dvd lcm m n"
62344
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2335
  by (fact dvd_lcmI2)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2336
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2337
lemma coprime_exp2_nat [intro]: "coprime (a::nat) b \<Longrightarrow> coprime (a^n) (b^m)"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2338
  by (fact coprime_exp2)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2339
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2340
lemma coprime_exp2_int [intro]: "coprime (a::int) b \<Longrightarrow> coprime (a^n) (b^m)"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2341
  by (fact coprime_exp2)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2342
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2343
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
  2344
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
  2345
lemmas Gcd_greatest_nat [simp] = Gcd_greatest [where ?'a = nat]
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2346
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
  2347
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2348
lemma dvd_Lcm_int [simp]:
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2349
  fixes M :: "int set" assumes "m \<in> M" shows "m dvd Lcm M"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2350
  using assms by (fact dvd_Lcm)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2351
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2352
lemma gcd_neg_numeral_1_int [simp]:
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2353
  "gcd (- numeral n :: int) x = gcd (numeral n) x"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2354
  by (fact gcd_neg1_int)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2355
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2356
lemma gcd_neg_numeral_2_int [simp]:
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2357
  "gcd x (- numeral n :: int) = gcd x (numeral n)"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2358
  by (fact gcd_neg2_int)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2359
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2360
lemma gcd_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> gcd x y = x"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2361
  by (fact gcd_nat.absorb1)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2362
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2363
lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \<Longrightarrow> gcd x y = y"
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2364
  by (fact gcd_nat.absorb2)
759d684c0e60 pulled out legacy aliasses and infamous dvd interpretations into theory appendix
haftmann
parents: 62343
diff changeset
  2365
62353
7f927120b5a2 dropped various legacy fact bindings and tuned proofs
haftmann
parents: 62350
diff changeset
  2366
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
  2367
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
  2368
lemmas Lcm_least_int [simp] = Lcm_least [where ?'a = int]
62345
e66d7841d5a2 further generalization and polishing
haftmann
parents: 62344
diff changeset
  2369
61856
4b1b85f38944 add gcd instance for integer and serialisation to target language operations
Andreas Lochbihler
parents: 61799
diff changeset
  2370
end