src/HOL/Library/GCD.thy
author chaieb
Mon, 14 Jul 2008 16:13:51 +0200
changeset 27568 9949dc7a24de
parent 27556 292098f2efdf
child 27651 16a26996c30e
permissions -rw-r--r--
Theorem names as in IntPrimes.thy, also several theorems moved from there
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/GCD.thy
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     5
*)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     6
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     7
header {* The Greatest Common Divisor *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     8
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     9
theory GCD
27487
c8a6ce181805 absolute imports of HOL/*.thy theories
haftmann
parents: 27368
diff changeset
    10
imports Plain "~~/src/HOL/Presburger"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    11
begin
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    12
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    13
text {*
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    14
  See \cite{davenport92}. \bigskip
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    15
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    16
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    17
subsection {* Specification of GCD on nats *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    18
21263
wenzelm
parents: 21256
diff changeset
    19
definition
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    20
  is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    21
  [code func del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    22
    (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    23
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    24
text {* Uniqueness *}
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    25
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    26
lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    27
  by (simp add: is_gcd_def) (blast intro: dvd_anti_sym)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    28
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    29
text {* Connection to divides relation *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    30
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    31
lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    32
  by (auto simp add: is_gcd_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    33
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    34
text {* Commutativity *}
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    35
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    36
lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    37
  by (auto simp add: is_gcd_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    38
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    39
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    40
subsection {* GCD on nat by Euclid's algorithm *}
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    41
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
    42
fun
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
    43
  gcd  :: "nat => nat => nat"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
    44
where
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    45
  "gcd m n = (if n = 0 then m else gcd n (m mod n))"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    46
lemma gcd_induct [case_names "0" rec]:
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    47
  fixes m n :: nat
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    48
  assumes "\<And>m. P m 0"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    49
    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
    50
  shows "P m n"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    51
proof (induct m n rule: gcd.induct)
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    52
  case (1 m n) with assms show ?case by (cases "n = 0") simp_all
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    53
qed
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    54
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    55
lemma gcd_0 [simp]: "gcd m 0 = m"
21263
wenzelm
parents: 21256
diff changeset
    56
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    57
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    58
lemma gcd_0_left [simp]: "gcd 0 m = m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    59
  by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    60
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    61
lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    62
  by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    63
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    64
lemma gcd_1 [simp]: "gcd m (Suc 0) = 1"
21263
wenzelm
parents: 21256
diff changeset
    65
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    66
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    67
declare gcd.simps [simp del]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    68
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    69
text {*
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    70
  \medskip @{term "gcd m n"} divides @{text m} and @{text n}.  The
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    71
  conjunctions don't seem provable separately.
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    72
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    73
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    74
lemma gcd_dvd1 [iff]: "gcd m n dvd m"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    75
  and gcd_dvd2 [iff]: "gcd m n dvd n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    76
  apply (induct m n rule: gcd_induct)
21263
wenzelm
parents: 21256
diff changeset
    77
     apply (simp_all add: gcd_non_0)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    78
  apply (blast dest: dvd_mod_imp_dvd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    79
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    80
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    81
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    82
  \medskip Maximality: for all @{term m}, @{term n}, @{term k}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    83
  naturals, if @{term k} divides @{term m} and @{term k} divides
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    84
  @{term n} then @{term k} divides @{term "gcd m n"}.
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    85
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    86
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    87
lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
21263
wenzelm
parents: 21256
diff changeset
    88
  by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    89
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    90
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    91
  \medskip Function gcd yields the Greatest Common Divisor.
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    92
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    93
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    94
lemma is_gcd: "is_gcd m n (gcd m n) "
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    95
  by (simp add: is_gcd_def gcd_greatest)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    96
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    97
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    98
subsection {* Derived laws for GCD *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    99
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   100
lemma gcd_greatest_iff [iff]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   101
  by (blast intro!: gcd_greatest intro: dvd_trans)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   102
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   103
lemma gcd_zero: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   104
  by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   105
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   106
lemma gcd_commute: "gcd m n = gcd n m"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   107
  apply (rule is_gcd_unique)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   108
   apply (rule is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   109
  apply (subst is_gcd_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   110
  apply (simp add: is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   111
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   112
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   113
lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   114
  apply (rule is_gcd_unique)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   115
   apply (rule is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   116
  apply (simp add: is_gcd_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   117
  apply (blast intro: dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   118
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   119
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   120
lemma gcd_1_left [simp]: "gcd (Suc 0) m = 1"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   121
  by (simp add: gcd_commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   122
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   123
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   124
  \medskip Multiplication laws
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   125
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   126
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   127
lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   128
    -- {* \cite[page 27]{davenport92} *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   129
  apply (induct m n rule: gcd_induct)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   130
   apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   131
  apply (case_tac "k = 0")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   132
   apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   133
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   134
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   135
lemma gcd_mult [simp]: "gcd k (k * n) = k"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   136
  apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   137
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   138
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   139
lemma gcd_self [simp]: "gcd k k = k"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   140
  apply (rule gcd_mult [of k 1, simplified])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   141
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   142
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   143
lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   144
  apply (insert gcd_mult_distrib2 [of m k n])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   145
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   146
  apply (erule_tac t = m in ssubst)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   147
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   148
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   149
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   150
lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   151
  apply (blast intro: relprime_dvd_mult dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   152
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   153
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   154
lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   155
  apply (rule dvd_anti_sym)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   156
   apply (rule gcd_greatest)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   157
    apply (rule_tac n = k in relprime_dvd_mult)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   158
     apply (simp add: gcd_assoc)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   159
     apply (simp add: gcd_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   160
    apply (simp_all add: mult_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   161
  apply (blast intro: dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   162
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   163
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   164
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   165
text {* \medskip Addition laws *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   166
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   167
lemma gcd_add1 [simp]: "gcd (m + n) n = gcd m n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   168
  apply (case_tac "n = 0")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   169
   apply (simp_all add: gcd_non_0)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   170
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   171
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   172
lemma gcd_add2 [simp]: "gcd m (m + n) = gcd m n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   173
proof -
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   174
  have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   175
  also have "... = gcd (n + m) m" by (simp add: add_commute)
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   176
  also have "... = gcd n m" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   177
  also have  "... = gcd m n" by (rule gcd_commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   178
  finally show ?thesis .
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   179
qed
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   180
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   181
lemma gcd_add2' [simp]: "gcd m (n + m) = gcd m n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   182
  apply (subst add_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   183
  apply (rule gcd_add2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   184
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   185
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   186
lemma gcd_add_mult: "gcd m (k * m + n) = gcd m n"
21263
wenzelm
parents: 21256
diff changeset
   187
  by (induct k) (simp_all add: add_assoc)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   188
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   189
lemma gcd_dvd_prod: "gcd m n dvd m * n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   190
  using mult_dvd_mono [of 1] by auto
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   191
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   192
text {*
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   193
  \medskip Division by gcd yields rrelatively primes.
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   194
*}
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   195
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   196
lemma div_gcd_relprime:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   197
  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   198
  shows "gcd (a div gcd a b) (b div gcd a b) = 1"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   199
proof -
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   200
  let ?g = "gcd a b"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   201
  let ?a' = "a div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   202
  let ?b' = "b div ?g"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   203
  let ?g' = "gcd ?a' ?b'"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   204
  have dvdg: "?g dvd a" "?g dvd b" by simp_all
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   205
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   206
  from dvdg dvdg' obtain ka kb ka' kb' where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   207
      kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   208
    unfolding dvd_def by blast
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   209
  then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   210
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   211
    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   212
      dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   213
  have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   214
  then have gp: "?g > 0" by simp
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   215
  from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   216
  with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   217
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   218
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   219
subsection {* LCM defined by GCD *}
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   220
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   221
definition
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   222
  lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   223
where
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   224
  lcm_def: "lcm m n = m * n div gcd m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   225
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   226
lemma prod_gcd_lcm:
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   227
  "m * n = gcd m n * lcm m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   228
  unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod])
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   229
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   230
lemma lcm_0 [simp]: "lcm m 0 = 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   231
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   232
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   233
lemma lcm_1 [simp]: "lcm m 1 = m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   234
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   235
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   236
lemma lcm_0_left [simp]: "lcm 0 n = 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   237
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   238
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   239
lemma lcm_1_left [simp]: "lcm 1 m = m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   240
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   241
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   242
lemma dvd_pos:
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   243
  fixes n m :: nat
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   244
  assumes "n > 0" and "m dvd n"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   245
  shows "m > 0"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   246
using assms by (cases m) auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   247
23951
b188cac107ad renamed lcm_lowest to lcm_least
haftmann
parents: 23687
diff changeset
   248
lemma lcm_least:
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   249
  assumes "m dvd k" and "n dvd k"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   250
  shows "lcm m n dvd k"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   251
proof (cases k)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   252
  case 0 then show ?thesis by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   253
next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   254
  case (Suc _) then have pos_k: "k > 0" by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   255
  from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   256
  with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   257
  from assms obtain p where k_m: "k = m * p" using dvd_def by blast
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   258
  from assms obtain q where k_n: "k = n * q" using dvd_def by blast
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   259
  from pos_k k_m have pos_p: "p > 0" by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   260
  from pos_k k_n have pos_q: "q > 0" by auto
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   261
  have "k * k * gcd q p = k * gcd (k * q) (k * p)"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   262
    by (simp add: mult_ac gcd_mult_distrib2)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   263
  also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   264
    by (simp add: k_m [symmetric] k_n [symmetric])
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   265
  also have "\<dots> = k * p * q * gcd m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   266
    by (simp add: mult_ac gcd_mult_distrib2)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   267
  finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   268
    by (simp only: k_m [symmetric] k_n [symmetric])
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   269
  then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   270
    by (simp add: mult_ac)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   271
  with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   272
    by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   273
  with prod_gcd_lcm [of m n]
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   274
  have "lcm m n * gcd q p * gcd m n = k * gcd m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   275
    by (simp add: mult_ac)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   276
  with pos_gcd have "lcm m n * gcd q p = k" by simp
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   277
  then show ?thesis using dvd_def by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   278
qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   279
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   280
lemma lcm_dvd1 [iff]:
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   281
  "m dvd lcm m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   282
proof (cases m)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   283
  case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   284
next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   285
  case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   286
  then have mpos: "m > 0" by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   287
  show ?thesis
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   288
  proof (cases n)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   289
    case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   290
  next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   291
    case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   292
    then have npos: "n > 0" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   293
    have "gcd m n dvd n" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   294
    then obtain k where "n = gcd m n * k" using dvd_def by auto
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   295
    then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac)
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   296
    also have "\<dots> = m * k" using mpos npos gcd_zero by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   297
    finally show ?thesis by (simp add: lcm_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   298
  qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   299
qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   300
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   301
lemma lcm_dvd2 [iff]: 
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   302
  "n dvd lcm m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   303
proof (cases n)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   304
  case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   305
next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   306
  case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   307
  then have npos: "n > 0" by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   308
  show ?thesis
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   309
  proof (cases m)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   310
    case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   311
  next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   312
    case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   313
    then have mpos: "m > 0" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   314
    have "gcd m n dvd m" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   315
    then obtain k where "m = gcd m n * k" using dvd_def by auto
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   316
    then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac)
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   317
    also have "\<dots> = n * k" using mpos npos gcd_zero by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   318
    finally show ?thesis by (simp add: lcm_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   319
  qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   320
qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   321
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   322
lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   323
  by (simp add: gcd_commute)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   324
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   325
lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   326
  apply (subgoal_tac "n = m + (n - m)")
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   327
   apply (erule ssubst, rule gcd_add1_eq, simp)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   328
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   329
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   330
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   331
subsection {* GCD and LCM on integers *}
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   332
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   333
definition
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   334
  zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   335
  "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   336
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   337
lemma zgcd_zdvd1 [iff,simp]: "zgcd i j dvd i"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   338
  by (simp add: zgcd_def int_dvd_iff)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   339
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   340
lemma zgcd_zdvd2 [iff,simp]: "zgcd i j dvd j"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   341
  by (simp add: zgcd_def int_dvd_iff)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   342
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   343
lemma zgcd_pos: "zgcd i j \<ge> 0"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   344
  by (simp add: zgcd_def)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   345
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   346
lemma zgcd0 [simp]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   347
  by (simp add: zgcd_def gcd_zero) arith
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   348
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   349
lemma zgcd_commute: "zgcd i j = zgcd j i"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   350
  unfolding zgcd_def by (simp add: gcd_commute)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   351
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   352
lemma zgcd_zminus [simp]: "zgcd (- i) j = zgcd i j"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   353
  unfolding zgcd_def by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   354
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   355
lemma zgcd_zminus2 [simp]: "zgcd i (- j) = zgcd i j"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   356
  unfolding zgcd_def by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   357
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   358
lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   359
  unfolding zgcd_def
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   360
proof -
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   361
  assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   362
  then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   363
  from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   364
  have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   365
    unfolding dvd_def
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   366
    by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   367
  from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   368
    unfolding dvd_def by blast
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   369
  from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23365
diff changeset
   370
  then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   371
  then show ?thesis
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   372
    apply (subst zdvd_abs1 [symmetric])
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   373
    apply (subst zdvd_abs2 [symmetric])
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   374
    apply (unfold dvd_def)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   375
    apply (rule_tac x = "int h'" in exI, simp)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   376
    done
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   377
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   378
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   379
lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   380
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   381
lemma zgcd_greatest:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   382
  assumes "k dvd m" and "k dvd n"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   383
  shows "k dvd zgcd m n"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   384
proof -
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   385
  let ?k' = "nat \<bar>k\<bar>"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   386
  let ?m' = "nat \<bar>m\<bar>"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   387
  let ?n' = "nat \<bar>n\<bar>"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   388
  from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   389
    unfolding zdvd_int by (simp_all only: int_nat_abs zdvd_abs1 zdvd_abs2)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   390
  from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   391
    unfolding zgcd_def by (simp only: zdvd_int)
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   392
  then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   393
  then show "k dvd zgcd m n" by (simp add: zdvd_abs1)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   394
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   395
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   396
lemma div_zgcd_relprime:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   397
  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   398
  shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   399
proof -
25112
98824cc791c0 fixed proofs
chaieb
parents: 23994
diff changeset
   400
  from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith 
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   401
  let ?g = "zgcd a b"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   402
  let ?a' = "a div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   403
  let ?b' = "b div ?g"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   404
  let ?g' = "zgcd ?a' ?b'"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   405
  have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   406
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   407
  from dvdg dvdg' obtain ka kb ka' kb' where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   408
   kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   409
    unfolding dvd_def by blast
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   410
  then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   411
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   412
    by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   413
      zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   414
  have "?g \<noteq> 0" using nz by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   415
  then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   416
  from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   417
  with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   418
  with zgcd_pos show "?g' = 1" by simp
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   419
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   420
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   421
    (* IntPrimes stuff *)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   422
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   423
lemma zgcd_0 [simp]: "zgcd m 0 = abs m"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   424
  by (simp add: zgcd_def abs_if)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   425
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   426
lemma zgcd_0_left [simp]: "zgcd 0 m = abs m"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   427
  by (simp add: zgcd_def abs_if)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   428
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   429
lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   430
  apply (frule_tac b = n and a = m in pos_mod_sign)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   431
  apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   432
  apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   433
  apply (frule_tac a = m in pos_mod_bound)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   434
  apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   435
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   436
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   437
lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   438
  apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   439
  apply (auto simp add: linorder_neq_iff zgcd_non_0)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   440
  apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   441
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   442
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   443
lemma zgcd_1 [simp]: "zgcd m 1 = 1"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   444
  by (simp add: zgcd_def abs_if)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   445
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   446
lemma zgcd_0_1_iff [simp]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   447
  by (simp add: zgcd_def abs_if)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   448
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   449
lemma zgcd_greatest_iff: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   450
  by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   451
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   452
lemma zgcd_1_left [simp]: "zgcd 1 m = 1"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   453
  by (simp add: zgcd_def gcd_1_left)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   454
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   455
lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   456
  by (simp add: zgcd_def gcd_assoc)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   457
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   458
lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   459
  apply (rule zgcd_commute [THEN trans])
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   460
  apply (rule zgcd_assoc [THEN trans])
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   461
  apply (rule zgcd_commute [THEN arg_cong])
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   462
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   463
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   464
lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   465
  -- {* addition is an AC-operator *}
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   466
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   467
lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   468
  by (simp del: minus_mult_right [symmetric]
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   469
      add: minus_mult_right nat_mult_distrib zgcd_def abs_if
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   470
          mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   471
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   472
lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   473
  by (simp add: abs_if zgcd_zmult_distrib2)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   474
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   475
lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   476
  by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   477
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   478
lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   479
  by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   480
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   481
lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k"
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   482
  by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   483
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   484
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   485
definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))"
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   486
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   487
lemma dvd_zlcm_self1[simp]: "i dvd zlcm i j"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   488
by(simp add:zlcm_def dvd_int_iff)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   489
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   490
lemma dvd_zlcm_self2[simp]: "j dvd zlcm i j"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   491
by(simp add:zlcm_def dvd_int_iff)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   492
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   493
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   494
lemma dvd_imp_dvd_zlcm1:
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   495
  assumes "k dvd i" shows "k dvd (zlcm i j)"
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   496
proof -
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   497
  have "nat(abs k) dvd nat(abs i)" using `k dvd i`
23994
3ddc90d1eda1 removed redundant ilcm_dvd1 ilcm_dvd2 zvdd_abs1
chaieb
parents: 23983
diff changeset
   498
    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   499
  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   500
qed
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   501
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   502
lemma dvd_imp_dvd_zlcm2:
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   503
  assumes "k dvd j" shows "k dvd (zlcm i j)"
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   504
proof -
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   505
  have "nat(abs k) dvd nat(abs j)" using `k dvd j`
23994
3ddc90d1eda1 removed redundant ilcm_dvd1 ilcm_dvd2 zvdd_abs1
chaieb
parents: 23983
diff changeset
   506
    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric] zdvd_abs1)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   507
  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   508
qed
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   509
23994
3ddc90d1eda1 removed redundant ilcm_dvd1 ilcm_dvd2 zvdd_abs1
chaieb
parents: 23983
diff changeset
   510
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   511
lemma zdvd_self_abs1: "(d::int) dvd (abs d)"
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   512
by (case_tac "d <0", simp_all)
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   513
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   514
lemma zdvd_self_abs2: "(abs (d::int)) dvd d"
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   515
by (case_tac "d<0", simp_all)
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   516
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   517
(* lcm a b is positive for positive a and b *)
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   518
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   519
lemma lcm_pos: 
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   520
  assumes mpos: "m > 0"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   521
  and npos: "n>0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   522
  shows "lcm m n > 0"
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   523
proof(rule ccontr, simp add: lcm_def gcd_zero)
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   524
assume h:"m*n div gcd m n = 0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   525
from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   526
hence gcdp: "gcd m n > 0" by simp
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   527
with h
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   528
have "m*n < gcd m n"
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   529
  by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   530
moreover 
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   531
have "gcd m n dvd m" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   532
 with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   533
 with npos have t1:"gcd m n *n \<le> m*n" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   534
 have "gcd m n \<le> gcd m n*n" using npos by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   535
 with t1 have "gcd m n \<le> m*n" by arith
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   536
ultimately show "False" by simp
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   537
qed
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   538
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   539
lemma zlcm_pos: 
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   540
  assumes anz: "a \<noteq> 0"
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   541
  and bnz: "b \<noteq> 0" 
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   542
  shows "0 < zlcm a b"
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   543
proof-
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   544
  let ?na = "nat (abs a)"
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   545
  let ?nb = "nat (abs b)"
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   546
  have nap: "?na >0" using anz by simp
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   547
  have nbp: "?nb >0" using bnz by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   548
  have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   549
  thus ?thesis by (simp add: zlcm_def)
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   550
qed
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   551
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   552
end