src/HOL/Library/GCD.thy
author nipkow
Fri, 01 Jun 2007 22:09:16 +0200
changeset 23192 ec73b9707d48
parent 22367 6860f09242bf
child 23244 1630951f0512
permissions -rw-r--r--
Moved list comprehension into List
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
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    10
imports Main
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 {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    14
  See \cite{davenport92}.
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    15
  \bigskip
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    16
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    17
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    18
consts
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    19
  gcd  :: "nat \<times> nat => nat"  -- {* Euclid's algorithm *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    20
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    21
recdef gcd  "measure ((\<lambda>(m, n). n) :: nat \<times> nat => nat)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    22
  "gcd (m, n) = (if n = 0 then m else gcd (n, m mod n))"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    23
21263
wenzelm
parents: 21256
diff changeset
    24
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21263
diff changeset
    25
  is_gcd :: "nat => nat => nat => bool" where -- {* @{term gcd} as a relation *}
21263
wenzelm
parents: 21256
diff changeset
    26
  "is_gcd p m n = (p dvd m \<and> p dvd n \<and>
wenzelm
parents: 21256
diff changeset
    27
    (\<forall>d. d dvd m \<and> d dvd n --> d dvd p))"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    28
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
lemma gcd_induct:
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    31
  "(!!m. P m 0) ==>
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    32
    (!!m n. 0 < n ==> P n (m mod n) ==> P m n)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    33
  ==> P (m::nat) (n::nat)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    34
  apply (induct m n rule: gcd.induct)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    35
  apply (case_tac "n = 0")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    36
   apply simp_all
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    37
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    38
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    39
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    40
lemma gcd_0 [simp]: "gcd (m, 0) = m"
21263
wenzelm
parents: 21256
diff changeset
    41
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    42
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    43
lemma gcd_non_0: "0 < n ==> gcd (m, n) = gcd (n, m mod n)"
21263
wenzelm
parents: 21256
diff changeset
    44
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    45
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    46
declare gcd.simps [simp del]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    47
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    48
lemma gcd_1 [simp]: "gcd (m, Suc 0) = 1"
21263
wenzelm
parents: 21256
diff changeset
    49
  by (simp add: gcd_non_0)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    50
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    51
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    52
  \medskip @{term "gcd (m, n)"} divides @{text m} and @{text n}.  The
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    53
  conjunctions don't seem provable separately.
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    54
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    55
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    56
lemma gcd_dvd1 [iff]: "gcd (m, n) dvd m"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    57
  and gcd_dvd2 [iff]: "gcd (m, n) dvd n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    58
  apply (induct m n rule: gcd_induct)
21263
wenzelm
parents: 21256
diff changeset
    59
     apply (simp_all add: gcd_non_0)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    60
  apply (blast dest: dvd_mod_imp_dvd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    61
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    62
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    63
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    64
  \medskip Maximality: for all @{term m}, @{term n}, @{term k}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    65
  naturals, if @{term k} divides @{term m} and @{term k} divides
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    66
  @{term n} then @{term k} divides @{term "gcd (m, n)"}.
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    67
*}
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
lemma gcd_greatest: "k dvd m ==> k dvd n ==> k dvd gcd (m, n)"
21263
wenzelm
parents: 21256
diff changeset
    70
  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
    71
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    72
lemma gcd_greatest_iff [iff]: "(k dvd gcd (m, n)) = (k dvd m \<and> k dvd n)"
21263
wenzelm
parents: 21256
diff changeset
    73
  by (blast intro!: gcd_greatest intro: dvd_trans)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    74
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    75
lemma gcd_zero: "(gcd (m, n) = 0) = (m = 0 \<and> n = 0)"
21263
wenzelm
parents: 21256
diff changeset
    76
  by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    77
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    78
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    79
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    80
  \medskip Function gcd yields the Greatest Common Divisor.
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    81
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    82
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    83
lemma is_gcd: "is_gcd (gcd (m, n)) m n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    84
  apply (simp add: is_gcd_def gcd_greatest)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    85
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    86
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    87
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    88
  \medskip Uniqueness of GCDs.
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
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    91
lemma is_gcd_unique: "is_gcd m a b ==> is_gcd n a b ==> m = n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    92
  apply (simp add: is_gcd_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    93
  apply (blast intro: dvd_anti_sym)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    94
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    95
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    96
lemma is_gcd_dvd: "is_gcd m a b ==> k dvd a ==> k dvd b ==> k dvd m"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    97
  apply (auto simp add: is_gcd_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    98
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    99
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   100
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   101
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   102
  \medskip Commutativity
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   103
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   104
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   105
lemma is_gcd_commute: "is_gcd k m n = is_gcd k n m"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   106
  apply (auto simp add: is_gcd_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   107
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   108
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   109
lemma gcd_commute: "gcd (m, n) = gcd (n, m)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   110
  apply (rule is_gcd_unique)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   111
   apply (rule is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   112
  apply (subst is_gcd_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   113
  apply (simp add: is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   114
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   115
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   116
lemma gcd_assoc: "gcd (gcd (k, m), n) = gcd (k, gcd (m, n))"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   117
  apply (rule is_gcd_unique)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   118
   apply (rule is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   119
  apply (simp add: is_gcd_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   120
  apply (blast intro: dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   121
  done
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
lemma gcd_0_left [simp]: "gcd (0, m) = m"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   124
  apply (simp add: gcd_commute [of 0])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   125
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   126
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   127
lemma gcd_1_left [simp]: "gcd (Suc 0, m) = 1"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   128
  apply (simp add: gcd_commute [of "Suc 0"])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   129
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   130
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   131
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   132
text {*
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   133
  \medskip Multiplication laws
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   134
*}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   135
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   136
lemma gcd_mult_distrib2: "k * gcd (m, n) = gcd (k * m, k * n)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   137
    -- {* \cite[page 27]{davenport92} *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   138
  apply (induct m n rule: gcd_induct)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   139
   apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   140
  apply (case_tac "k = 0")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   141
   apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   142
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   143
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   144
lemma gcd_mult [simp]: "gcd (k, k * n) = k"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   145
  apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   146
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   147
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   148
lemma gcd_self [simp]: "gcd (k, k) = k"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   149
  apply (rule gcd_mult [of k 1, simplified])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   150
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   151
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   152
lemma relprime_dvd_mult: "gcd (k, n) = 1 ==> k dvd m * n ==> k dvd m"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   153
  apply (insert gcd_mult_distrib2 [of m k n])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   154
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   155
  apply (erule_tac t = m in ssubst)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   156
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   157
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   158
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   159
lemma relprime_dvd_mult_iff: "gcd (k, n) = 1 ==> (k dvd m * n) = (k dvd m)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   160
  apply (blast intro: relprime_dvd_mult dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   161
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   162
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   163
lemma gcd_mult_cancel: "gcd (k, n) = 1 ==> gcd (k * m, n) = gcd (m, n)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   164
  apply (rule dvd_anti_sym)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   165
   apply (rule gcd_greatest)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   166
    apply (rule_tac n = k in relprime_dvd_mult)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   167
     apply (simp add: gcd_assoc)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   168
     apply (simp add: gcd_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   169
    apply (simp_all add: mult_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   170
  apply (blast intro: dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   171
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   172
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   173
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   174
text {* \medskip Addition laws *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   175
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   176
lemma gcd_add1 [simp]: "gcd (m + n, n) = gcd (m, n)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   177
  apply (case_tac "n = 0")
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   178
   apply (simp_all add: gcd_non_0)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   179
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   180
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   181
lemma gcd_add2 [simp]: "gcd (m, m + n) = gcd (m, n)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   182
proof -
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   183
  have "gcd (m, m + n) = gcd (m + n, m)" by (rule gcd_commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   184
  also have "... = gcd (n + m, m)" by (simp add: add_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   185
  also have "... = gcd (n, m)" by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   186
  also have  "... = gcd (m, n)" by (rule gcd_commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   187
  finally show ?thesis .
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   188
qed
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   189
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   190
lemma gcd_add2' [simp]: "gcd (m, n + m) = gcd (m, n)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   191
  apply (subst add_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   192
  apply (rule gcd_add2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   193
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   194
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   195
lemma gcd_add_mult: "gcd (m, k * m + n) = gcd (m, n)"
21263
wenzelm
parents: 21256
diff changeset
   196
  by (induct k) (simp_all add: add_assoc)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   197
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   198
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   199
text {*
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   200
  \medskip Division by gcd yields rrelatively primes.
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   201
*}
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   202
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   203
lemma div_gcd_relprime:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   204
  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   205
  shows "gcd (a div gcd(a,b), b div gcd(a,b)) = 1"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   206
proof -
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   207
  let ?g = "gcd (a, b)"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   208
  let ?a' = "a div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   209
  let ?b' = "b div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   210
  let ?g' = "gcd (?a', ?b')"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   211
  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
   212
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   213
  from dvdg dvdg' obtain ka kb ka' kb' where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   214
      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
   215
    unfolding dvd_def by blast
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   216
  then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   217
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   218
    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   219
      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
   220
  have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   221
  then have gp: "?g > 0" by simp
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   222
  from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   223
  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
   224
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   225
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   226
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   227
text {*
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   228
  \medskip Gcd on integers.
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   229
*}
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   230
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   231
definition
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   232
  igcd :: "int \<Rightarrow> int \<Rightarrow> int" where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   233
  "igcd i j = int (gcd (nat (abs i), nat (abs j)))"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   234
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   235
lemma igcd_dvd1 [simp]: "igcd i j dvd i"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   236
  by (simp add: igcd_def int_dvd_iff)
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   237
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   238
lemma igcd_dvd2 [simp]: "igcd i j dvd j"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   239
  by (simp add: igcd_def int_dvd_iff)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   240
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   241
lemma igcd_pos: "igcd i j \<ge> 0"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   242
  by (simp add: igcd_def)
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   243
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   244
lemma igcd0 [simp]: "(igcd i j = 0) = (i = 0 \<and> j = 0)"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   245
  by (simp add: igcd_def gcd_zero) arith
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   246
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   247
lemma igcd_commute: "igcd i j = igcd j i"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   248
  unfolding igcd_def by (simp add: gcd_commute)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   249
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   250
lemma igcd_neg1 [simp]: "igcd (- i) j = igcd i j"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   251
  unfolding igcd_def by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   252
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   253
lemma igcd_neg2 [simp]: "igcd i (- j) = igcd i j"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   254
  unfolding igcd_def by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   255
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   256
lemma zrelprime_dvd_mult: "igcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   257
  unfolding igcd_def
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   258
proof -
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   259
  assume "int (gcd (nat \<bar>i\<bar>, nat \<bar>j\<bar>)) = 1" "i dvd k * j"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   260
  then have g: "gcd (nat \<bar>i\<bar>, nat \<bar>j\<bar>) = 1" by simp
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   261
  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
   262
  have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   263
    unfolding dvd_def
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   264
    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
   265
  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
   266
    unfolding dvd_def by blast
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   267
  from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   268
  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
   269
  then show ?thesis
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   270
    apply (subst zdvd_abs1 [symmetric])
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   271
    apply (subst zdvd_abs2 [symmetric])
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   272
    apply (unfold dvd_def)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   273
    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
   274
    done
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   275
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   276
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   277
lemma int_nat_abs: "int (nat (abs x)) = abs x"  by arith
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   278
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   279
lemma igcd_greatest:
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   280
  assumes "k dvd m" and "k dvd n"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   281
  shows "k dvd igcd m n"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   282
proof -
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   283
  let ?k' = "nat \<bar>k\<bar>"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   284
  let ?m' = "nat \<bar>m\<bar>"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   285
  let ?n' = "nat \<bar>n\<bar>"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   286
  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
   287
    unfolding zdvd_int by (simp_all only: int_nat_abs zdvd_abs1 zdvd_abs2)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   288
  from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd igcd m n"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   289
    unfolding igcd_def by (simp only: zdvd_int)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   290
  then have "\<bar>k\<bar> dvd igcd m n" by (simp only: int_nat_abs)
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   291
  then show "k dvd igcd m n" by (simp add: zdvd_abs1)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   292
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   293
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   294
lemma div_igcd_relprime:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   295
  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   296
  shows "igcd (a div (igcd a b)) (b div (igcd a b)) = 1"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   297
proof -
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   298
  from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by simp
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   299
  let ?g = "igcd a b"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   300
  let ?a' = "a div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   301
  let ?b' = "b div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   302
  let ?g' = "igcd ?a' ?b'"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   303
  have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: igcd_dvd1 igcd_dvd2)
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   304
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: igcd_dvd1 igcd_dvd2)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   305
  from dvdg dvdg' obtain ka kb ka' kb' where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   306
   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
   307
    unfolding dvd_def by blast
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   308
  then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   309
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   310
    by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   311
      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
   312
  have "?g \<noteq> 0" using nz by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   313
  then have gp: "?g \<noteq> 0" using igcd_pos[where i="a" and j="b"] by arith
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   314
  from igcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   315
  with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   316
  with igcd_pos show "?g' = 1" by simp
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   317
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   318
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   319
end