src/HOL/Old_Number_Theory/Legacy_GCD.thy
author wenzelm
Sat, 10 Oct 2015 16:26:23 +0200
changeset 61382 efac889fccbc
parent 58889 5b7a9633cfa8
child 61762 d50b993b4fb9
permissions -rw-r--r--
isabelle update_cartouches;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
     1
(*  Title:      HOL/Old_Number_Theory/Legacy_GCD.thy
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     2
    Author:     Christophe Tabacznyj and Lawrence C Paulson
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     3
    Copyright   1996  University of Cambridge
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     4
*)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     5
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>The Greatest Common Divisor\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     7
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     8
theory Legacy_GCD
30738
0842e906300c normalized imports
haftmann
parents: 30242
diff changeset
     9
imports Main
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    10
begin
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    11
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    12
text \<open>
58623
2db1df2c8467 more bibtex entries;
wenzelm
parents: 57514
diff changeset
    13
  See @{cite davenport92}. \bigskip
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    14
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    15
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    16
subsection \<open>Specification of GCD on nats\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    17
21263
wenzelm
parents: 21256
diff changeset
    18
definition
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    19
  is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- \<open>@{term gcd} as a relation\<close>
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 34915
diff changeset
    20
  "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
    21
    (\<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
    22
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    23
text \<open>Uniqueness\<close>
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    24
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    25
lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
    26
  by (simp add: is_gcd_def) (blast intro: dvd_antisym)
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    27
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    28
text \<open>Connection to divides relation\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    29
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    30
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
    31
  by (auto simp add: is_gcd_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    32
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    33
text \<open>Commutativity\<close>
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    34
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    35
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
    36
  by (auto simp add: is_gcd_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    37
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    38
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    39
subsection \<open>GCD on nat by Euclid's algorithm\<close>
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    40
38159
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
    41
fun gcd :: "nat => nat => nat"
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
    42
  where "gcd m n = (if n = 0 then m else gcd n (m mod n))"
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
    43
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    44
lemma gcd_induct [case_names "0" rec]:
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    45
  fixes m n :: nat
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    46
  assumes "\<And>m. P m 0"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    47
    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
    48
  shows "P m n"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    49
proof (induct m n rule: gcd.induct)
38159
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
    50
  case (1 m n)
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
    51
  with assms show ?case by (cases "n = 0") simp_all
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    52
qed
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    53
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
    54
lemma gcd_0 [simp, algebra]: "gcd m 0 = m"
21263
wenzelm
parents: 21256
diff changeset
    55
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    56
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
    57
lemma gcd_0_left [simp,algebra]: "gcd 0 m = m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    58
  by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    59
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    60
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
    61
  by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    62
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
    63
lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = Suc 0"
21263
wenzelm
parents: 21256
diff changeset
    64
  by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    65
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
    66
lemma nat_gcd_1_right [simp, algebra]: "gcd m 1 = 1"
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
    67
  unfolding One_nat_def by (rule gcd_1)
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
    68
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    69
declare gcd.simps [simp del]
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    70
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    71
text \<open>
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    72
  \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
    73
  conjunctions don't seem provable separately.
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    74
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    75
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
    76
lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
    77
  and gcd_dvd2 [iff, algebra]: "gcd m n dvd n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    78
  apply (induct m n rule: gcd_induct)
21263
wenzelm
parents: 21256
diff changeset
    79
     apply (simp_all add: gcd_non_0)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    80
  apply (blast dest: dvd_mod_imp_dvd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    81
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    82
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    83
text \<open>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    84
  \medskip Maximality: for all @{term m}, @{term n}, @{term k}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    85
  naturals, if @{term k} divides @{term m} and @{term k} divides
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    86
  @{term n} then @{term k} divides @{term "gcd m n"}.
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    87
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    88
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    89
lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
21263
wenzelm
parents: 21256
diff changeset
    90
  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
    91
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    92
text \<open>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    93
  \medskip Function gcd yields the Greatest Common Divisor.
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    94
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    95
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
    96
lemma is_gcd: "is_gcd m n (gcd m n) "
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
    97
  by (simp add: is_gcd_def gcd_greatest)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    98
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    99
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   100
subsection \<open>Derived laws for GCD\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   101
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   102
lemma gcd_greatest_iff [iff, algebra]: "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
   103
  by (blast intro!: gcd_greatest intro: dvd_trans)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   104
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   105
lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   106
  by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   107
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   108
lemma gcd_commute: "gcd m n = gcd n m"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   109
  apply (rule is_gcd_unique)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   110
   apply (rule is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   111
  apply (subst is_gcd_commute)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   112
  apply (simp add: is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   113
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   114
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   115
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
   116
  apply (rule is_gcd_unique)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   117
   apply (rule is_gcd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   118
  apply (simp add: is_gcd_def)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   119
  apply (blast intro: dvd_trans)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   120
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   121
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
   122
lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = Suc 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   123
  by (simp add: gcd_commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   124
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
   125
lemma nat_gcd_1_left [simp, algebra]: "gcd 1 m = 1"
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
   126
  unfolding One_nat_def by (rule gcd_1_left)
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 30042
diff changeset
   127
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   128
text \<open>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   129
  \medskip Multiplication laws
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   130
\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   131
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   132
lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   133
    -- \<open>@{cite \<open>page 27\<close> davenport92}\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   134
  apply (induct m n rule: gcd_induct)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   135
   apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   136
  apply (case_tac "k = 0")
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 44821
diff changeset
   137
   apply (simp_all add: gcd_non_0)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   138
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   139
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   140
lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   141
  apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
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
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   144
lemma gcd_self [simp, algebra]: "gcd k k = k"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   145
  apply (rule gcd_mult [of k 1, simplified])
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
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   148
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
   149
  apply (insert gcd_mult_distrib2 [of m k n])
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   150
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   151
  apply (erule_tac t = m in ssubst)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   152
  apply simp
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   153
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   154
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   155
lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27568
diff changeset
   156
  by (auto intro: relprime_dvd_mult dvd_mult2)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   157
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   158
lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   159
  apply (rule dvd_antisym)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   160
   apply (rule gcd_greatest)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   161
    apply (rule_tac n = k in relprime_dvd_mult)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   162
     apply (simp add: gcd_assoc)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   163
     apply (simp add: gcd_commute)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   164
    apply (simp_all add: mult.commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   165
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   166
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   167
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   168
text \<open>\medskip Addition laws\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   169
27676
55676111ed69 (re-)added simp rules for (_ + _) div/mod _
haftmann
parents: 27669
diff changeset
   170
lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n"
55676111ed69 (re-)added simp rules for (_ + _) div/mod _
haftmann
parents: 27669
diff changeset
   171
  by (cases "n = 0") (auto simp add: gcd_non_0)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   172
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   173
lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   174
proof -
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   175
  have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   176
  also have "... = gcd (n + m) m" by (simp add: add.commute)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   177
  also have "... = gcd n m" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   178
  also have  "... = gcd m n" by (rule gcd_commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   179
  finally show ?thesis .
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   180
qed
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   181
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   182
lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   183
  apply (subst add.commute)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   184
  apply (rule gcd_add2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   185
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   186
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   187
lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   188
  by (induct k) (simp_all add: add.assoc)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   189
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   190
lemma gcd_dvd_prod: "gcd m n dvd m * n" 
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   191
  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
   192
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   193
text \<open>
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   194
  \medskip Division by gcd yields rrelatively primes.
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   195
\<close>
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   196
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   197
lemma div_gcd_relprime:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   198
  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   199
  shows "gcd (a div gcd a b) (b div gcd a b) = 1"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   200
proof -
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   201
  let ?g = "gcd a b"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   202
  let ?a' = "a div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   203
  let ?b' = "b div ?g"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   204
  let ?g' = "gcd ?a' ?b'"
22027
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
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   206
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   207
  from dvdg dvdg' obtain ka kb ka' kb' where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   208
      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
   209
    unfolding dvd_def by blast
58834
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58623
diff changeset
   210
  from this(3-4) [symmetric] have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'"
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58623
diff changeset
   211
    by (simp_all only: ac_simps mult.left_commute [of _ "gcd a b"])
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   212
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   213
    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   214
      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
   215
  have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   216
  then have gp: "?g > 0" by simp
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   217
  from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   218
  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
   219
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   220
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   221
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   222
lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   223
proof(auto)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   224
  assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   225
  from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   226
  have th: "gcd a b dvd d" by blast
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   227
  from dvd_antisym[OF th gcd_greatest[OF H(1,2)]]  show "d = gcd a b" by blast 
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   228
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   229
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   230
lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   231
  shows "gcd x y = gcd u v"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   232
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   233
  from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   234
  with gcd_unique[of "gcd u v" x y]  show ?thesis by auto
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   235
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   236
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   237
lemma ind_euclid:
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   238
  assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   239
  and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)"
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   240
  shows "P a b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   241
proof(induct "a + b" arbitrary: a b rule: less_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   242
  case less
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   243
  have "a = b \<or> a < b \<or> b < a" by arith
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   244
  moreover {assume eq: "a= b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   245
    from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   246
    by simp}
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   247
  moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   248
  {assume lt: "a < b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   249
    hence "a + b - a < a + b \<or> a = 0" by arith
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   250
    moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   251
    {assume "a =0" with z c have "P a b" by blast }
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   252
    moreover
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   253
    {assume "a + b - a < a + b"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   254
      also have th0: "a + b - a = a + (b - a)" using lt by arith
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   255
      finally have "a + (b - a) < a + b" .
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   256
      then have "P a (a + (b - a))" by (rule add[rule_format, OF less])
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   257
      then have "P a b" by (simp add: th0[symmetric])}
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   258
    ultimately have "P a b" by blast}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   259
  moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   260
  {assume lt: "a > b"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   261
    hence "b + a - b < a + b \<or> b = 0" by arith
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   262
    moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   263
    {assume "b =0" with z c have "P a b" by blast }
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   264
    moreover
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   265
    {assume "b + a - b < a + b"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   266
      also have th0: "b + a - b = b + (a - b)" using lt by arith
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   267
      finally have "b + (a - b) < a + b" .
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   268
      then have "P b (b + (a - b))" by (rule add[rule_format, OF less])
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33657
diff changeset
   269
      then have "P b a" by (simp add: th0[symmetric])
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   270
      hence "P a b" using c by blast }
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   271
    ultimately have "P a b" by blast}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   272
ultimately  show "P a b" by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   273
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   274
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   275
lemma bezout_lemma: 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   276
  assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   277
  shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   278
using ex
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   279
apply clarsimp
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   280
apply (rule_tac x="d" in exI, simp)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   281
apply (case_tac "a * x = b * y + d" , simp_all)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   282
apply (rule_tac x="x + y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   283
apply (rule_tac x="y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   284
apply algebra
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   285
apply (rule_tac x="x" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   286
apply (rule_tac x="x + y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   287
apply algebra
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   288
done
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   289
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   290
lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   291
apply(induct a b rule: ind_euclid)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   292
apply blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   293
apply clarify
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   294
apply (rule_tac x="a" in exI, simp)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   295
apply clarsimp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   296
apply (rule_tac x="d" in exI)
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   297
apply (case_tac "a * x = b * y + d", simp_all)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   298
apply (rule_tac x="x+y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   299
apply (rule_tac x="y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   300
apply algebra
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   301
apply (rule_tac x="x" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   302
apply (rule_tac x="x+y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   303
apply algebra
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   304
done
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   305
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   306
lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   307
using bezout_add[of a b]
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   308
apply clarsimp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   309
apply (rule_tac x="d" in exI, simp)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   310
apply (rule_tac x="x" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   311
apply (rule_tac x="y" in exI)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   312
apply auto
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   313
done
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   314
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   315
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   316
text \<open>We can get a stronger version with a nonzeroness assumption.\<close>
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   317
lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   318
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   319
lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   320
  shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   321
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   322
  from nz have ap: "a > 0" by simp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   323
 from bezout_add[of a b] 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   324
 have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   325
 moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   326
 {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   327
   from H have ?thesis by blast }
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   328
 moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   329
 {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   330
   {assume b0: "b = 0" with H  have ?thesis by simp}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   331
   moreover 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   332
   {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   333
     from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   334
     moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   335
     {assume db: "d=b"
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   336
       from nz H db have ?thesis apply simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   337
         apply (rule exI[where x = b], simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   338
         apply (rule exI[where x = b])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   339
        by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   340
    moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   341
    {assume db: "d < b" 
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   342
        {assume "x=0" hence ?thesis using nz H by simp }
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   343
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   344
        {assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   345
          
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   346
          from db have "d \<le> b - 1" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   347
          hence "d*b \<le> b*(b - 1)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   348
          with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   349
          have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   350
          from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   351
          hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   352
          hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   353
            by (simp only: diff_add_assoc[OF dble, of d, symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   354
          hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
58834
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58623
diff changeset
   355
            by (simp only: diff_mult_distrib2 ac_simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   356
          hence ?thesis using H(1,2)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   357
            apply -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   358
            apply (rule exI[where x=d], simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   359
            apply (rule exI[where x="(b - 1) * y"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   360
            by (rule exI[where x="x*(b - 1) - d"], simp)}
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32479
diff changeset
   361
        ultimately have ?thesis by blast}
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   362
    ultimately have ?thesis by blast}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   363
  ultimately have ?thesis by blast}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   364
 ultimately show ?thesis by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   365
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   366
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   367
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   368
lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   369
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   370
  let ?g = "gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   371
  from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   372
  from d(1,2) have "d dvd ?g" by simp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   373
  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   374
  from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   375
  hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k" 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   376
    by (algebra add: diff_mult_distrib)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   377
  hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g" 
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   378
    by (simp add: k mult.assoc)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   379
  thus ?thesis by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   380
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   381
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   382
lemma bezout_gcd_strong: assumes a: "a \<noteq> 0" 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   383
  shows "\<exists>x y. a * x = b * y + gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   384
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   385
  let ?g = "gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   386
  from bezout_add_strong[OF a, of b]
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   387
  obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   388
  from d(1,2) have "d dvd ?g" by simp
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   389
  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   390
  from d(3) have "a * x * k = (b * y + d) *k " by algebra
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   391
  hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   392
  thus ?thesis by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   393
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   394
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   395
lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   396
by(simp add: gcd_mult_distrib2 mult.commute)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   397
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   398
lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   399
  (is "?lhs \<longleftrightarrow> ?rhs")
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   400
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   401
  let ?g = "gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   402
  {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   403
    from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   404
      by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   405
    hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   406
    hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k" 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   407
      by (simp only: diff_mult_distrib)
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   408
    hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   409
      by (simp add: k[symmetric] mult.assoc)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   410
    hence ?lhs by blast}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   411
  moreover
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   412
  {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   413
    have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   414
      using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31706
diff changeset
   415
    from dvd_diff_nat[OF dv(1,2)] dvd_diff_nat[OF dv(3,4)] H
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   416
    have ?rhs by auto}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   417
  ultimately show ?thesis by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   418
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   419
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   420
lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   421
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   422
  let ?g = "gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   423
    have dv: "?g dvd a*x" "?g dvd b * y" 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   424
      using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   425
    from dvd_add[OF dv] H
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   426
    show ?thesis by auto
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   427
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   428
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   429
lemma gcd_mult': "gcd b (a * b) = b"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 47162
diff changeset
   430
by (simp add: mult.commute[of a b]) 
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   431
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   432
lemma gcd_add: "gcd(a + b) b = gcd a b" 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   433
  "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   434
by (simp_all add: gcd_commute)
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   435
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   436
lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   437
proof-
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   438
  {fix a b assume H: "b \<le> (a::nat)"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   439
    hence th: "a - b + b = a" by arith
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   440
    from gcd_add(1)[of "a - b" b] th  have "gcd(a - b) b = gcd a b" by simp}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   441
  note th = this
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   442
{
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   443
  assume ab: "b \<le> a"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   444
  from th[OF ab] show "gcd (a - b)  b = gcd a b" by blast
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   445
next
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   446
  assume ab: "a \<le> b"
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   447
  from th[OF ab] show "gcd a (b - a) = gcd a b" 
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   448
    by (simp add: gcd_commute)}
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   449
qed
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   450
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   451
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   452
subsection \<open>LCM defined by GCD\<close>
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   453
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   454
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   455
definition
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   456
  lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   457
where
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   458
  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
   459
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   460
lemma prod_gcd_lcm:
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   461
  "m * n = gcd m n * lcm m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   462
  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
   463
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   464
lemma lcm_0 [simp]: "lcm m 0 = 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   465
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   466
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   467
lemma lcm_1 [simp]: "lcm m 1 = m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   468
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   469
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   470
lemma lcm_0_left [simp]: "lcm 0 n = 0"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   471
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   472
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   473
lemma lcm_1_left [simp]: "lcm 1 m = m"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   474
  unfolding lcm_def by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   475
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   476
lemma dvd_pos:
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   477
  fixes n m :: nat
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   478
  assumes "n > 0" and "m dvd n"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   479
  shows "m > 0"
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   480
using assms by (cases m) auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   481
23951
b188cac107ad renamed lcm_lowest to lcm_least
haftmann
parents: 23687
diff changeset
   482
lemma lcm_least:
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   483
  assumes "m dvd k" and "n dvd k"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   484
  shows "lcm m n dvd k"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   485
proof (cases k)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   486
  case 0 then show ?thesis by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   487
next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   488
  case (Suc _) then have pos_k: "k > 0" by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   489
  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
   490
  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
   491
  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
   492
  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
   493
  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
   494
  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
   495
  have "k * k * gcd q p = k * gcd (k * q) (k * p)"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   496
    by (simp add: ac_simps gcd_mult_distrib2)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   497
  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
   498
    by (simp add: k_m [symmetric] k_n [symmetric])
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   499
  also have "\<dots> = k * p * q * gcd m n"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   500
    by (simp add: ac_simps gcd_mult_distrib2)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   501
  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
   502
    by (simp only: k_m [symmetric] k_n [symmetric])
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   503
  then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   504
    by (simp add: ac_simps)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   505
  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
   506
    by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   507
  with prod_gcd_lcm [of m n]
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   508
  have "lcm m n * gcd q p * gcd m n = k * gcd m n"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   509
    by (simp add: ac_simps)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   510
  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
   511
  then show ?thesis using dvd_def by auto
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   512
qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   513
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   514
lemma lcm_dvd1 [iff]:
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   515
  "m dvd lcm m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   516
proof (cases m)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   517
  case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   518
next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   519
  case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   520
  then have mpos: "m > 0" by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   521
  show ?thesis
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   522
  proof (cases n)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   523
    case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   524
  next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   525
    case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   526
    then have npos: "n > 0" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   527
    have "gcd m n dvd n" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   528
    then obtain k where "n = gcd m n * k" using dvd_def by auto
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   529
    then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: ac_simps)
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   530
    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
   531
    finally show ?thesis by (simp add: lcm_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   532
  qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   533
qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   534
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   535
lemma lcm_dvd2 [iff]: 
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   536
  "n dvd lcm m n"
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   537
proof (cases n)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   538
  case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   539
next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   540
  case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   541
  then have npos: "n > 0" by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   542
  show ?thesis
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   543
  proof (cases m)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   544
    case 0 then show ?thesis by simp
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   545
  next
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   546
    case (Suc _)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   547
    then have mpos: "m > 0" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   548
    have "gcd m n dvd m" by simp
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   549
    then obtain k where "m = gcd m n * k" using dvd_def by auto
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   550
    then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: ac_simps)
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   551
    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
   552
    finally show ?thesis by (simp add: lcm_def)
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   553
  qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   554
qed
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   555
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   556
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
   557
  by (simp add: gcd_commute)
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   558
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   559
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
   560
  apply (subgoal_tac "n = m + (n - m)")
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   561
  apply (erule ssubst, rule gcd_add1_eq, simp)  
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   562
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   563
23687
06884f7ffb18 extended - convers now basic lcm properties also
haftmann
parents: 23431
diff changeset
   564
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   565
subsection \<open>GCD and LCM on integers\<close>
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   566
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   567
definition
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   568
  zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   569
  "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   570
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   571
lemma zgcd_zdvd1 [iff, algebra]: "zgcd i j dvd i"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   572
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
   573
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   574
lemma zgcd_zdvd2 [iff, algebra]: "zgcd i j dvd j"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   575
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
   576
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   577
lemma zgcd_pos: "zgcd i j \<ge> 0"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   578
by (simp add: zgcd_def)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   579
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   580
lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   581
by (simp add: zgcd_def gcd_zero)
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   582
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   583
lemma zgcd_commute: "zgcd i j = zgcd j i"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   584
unfolding zgcd_def by (simp add: gcd_commute)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   585
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   586
lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   587
unfolding zgcd_def by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   588
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   589
lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j"
29700
22faf21db3df added some simp rules
nipkow
parents: 29655
diff changeset
   590
unfolding zgcd_def by simp
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   591
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   592
  (* should be solved by algebra*)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   593
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
   594
  unfolding zgcd_def
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   595
proof -
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   596
  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
   597
  then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   598
  from \<open>i dvd k * j\<close> 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
   599
  have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   600
    unfolding dvd_def
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   601
    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
   602
  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
   603
    unfolding dvd_def by blast
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   604
  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
   605
  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
   606
  then show ?thesis
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29700
diff changeset
   607
    apply (subst abs_dvd_iff [symmetric])
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29700
diff changeset
   608
    apply (subst dvd_abs_iff [symmetric])
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   609
    apply (unfold dvd_def)
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   610
    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
   611
    done
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   612
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   613
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   614
lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   615
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   616
lemma zgcd_greatest:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   617
  assumes "k dvd m" and "k dvd n"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   618
  shows "k dvd zgcd m n"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   619
proof -
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   620
  let ?k' = "nat \<bar>k\<bar>"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   621
  let ?m' = "nat \<bar>m\<bar>"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   622
  let ?n' = "nat \<bar>n\<bar>"
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   623
  from \<open>k dvd m\<close> and \<open>k dvd n\<close> have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29700
diff changeset
   624
    unfolding zdvd_int by (simp_all only: int_nat_abs abs_dvd_iff dvd_abs_iff)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   625
  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
   626
    unfolding zgcd_def by (simp only: zdvd_int)
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   627
  then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29700
diff changeset
   628
  then show "k dvd zgcd m n" by simp
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   629
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   630
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   631
lemma div_zgcd_relprime:
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   632
  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   633
  shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   634
proof -
25112
98824cc791c0 fixed proofs
chaieb
parents: 23994
diff changeset
   635
  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
   636
  let ?g = "zgcd a b"
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   637
  let ?a' = "a div ?g"
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   638
  let ?b' = "b div ?g"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   639
  let ?g' = "zgcd ?a' ?b'"
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   640
  have dvdg: "?g dvd a" "?g dvd b" by simp_all
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   641
  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   642
  from dvdg dvdg' obtain ka kb ka' kb' where
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   643
   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
   644
    unfolding dvd_def by blast
58834
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58623
diff changeset
   645
  from this(3-4) [symmetric] have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'"
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58623
diff changeset
   646
    by (simp_all only: ac_simps mult.left_commute [of _ "zgcd a b"])
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   647
  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
47162
9d7d919b9fd8 remove redundant lemma
huffman
parents: 45270
diff changeset
   648
    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
9d7d919b9fd8 remove redundant lemma
huffman
parents: 45270
diff changeset
   649
      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
   650
  have "?g \<noteq> 0" using nz by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   651
  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
   652
  from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   653
  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
   654
  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
   655
qed
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   656
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   657
lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   658
  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
   659
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   660
lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   661
  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
   662
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   663
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
   664
  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
   665
  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
   666
  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
   667
  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
   668
  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
   669
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   670
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   671
lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   672
  apply (cases "n = 0", simp)
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   673
  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
   674
  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
   675
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   676
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   677
lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   678
  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
   679
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   680
lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   681
  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
   682
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   683
lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   684
  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
   685
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   686
lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1"
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 38159
diff changeset
   687
  by (simp add: zgcd_def)
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   688
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   689
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
   690
  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
   691
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   692
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
   693
  apply (rule zgcd_commute [THEN trans])
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   694
  apply (rule zgcd_assoc [THEN trans])
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   695
  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
   696
  done
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   697
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   698
lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   699
  -- \<open>addition is an AC-operator\<close>
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   700
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   701
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
   702
  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
   703
      add: minus_mult_right nat_mult_distrib zgcd_def abs_if
44821
a92f65e174cf avoid using legacy theorem names
huffman
parents: 41541
diff changeset
   704
          mult_less_0_iff gcd_mult_distrib2 [symmetric] of_nat_mult)
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   705
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   706
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
   707
  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
   708
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   709
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
   710
  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
   711
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   712
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
   713
  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
   714
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   715
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
   716
  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
   717
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   718
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   719
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
   720
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   721
lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   722
by(simp add:zlcm_def dvd_int_iff)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   723
27669
4b1642284dd7 Tuned and simplified proofs; Rules added to presburger's and algebra's context; moved Bezout theorems from Primes.thy
chaieb
parents: 27651
diff changeset
   724
lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   725
by(simp add:zlcm_def dvd_int_iff)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   726
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   727
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   728
lemma dvd_imp_dvd_zlcm1:
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   729
  assumes "k dvd i" shows "k dvd (zlcm i j)"
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   730
proof -
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   731
  have "nat(abs k) dvd nat(abs i)" using \<open>k dvd i\<close>
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29700
diff changeset
   732
    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   733
  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   734
qed
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   735
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   736
lemma dvd_imp_dvd_zlcm2:
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   737
  assumes "k dvd j" shows "k dvd (zlcm i j)"
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   738
proof -
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   739
  have "nat(abs k) dvd nat(abs j)" using \<open>k dvd j\<close>
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29700
diff changeset
   740
    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   741
  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   742
qed
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   743
23994
3ddc90d1eda1 removed redundant ilcm_dvd1 ilcm_dvd2 zvdd_abs1
chaieb
parents: 23983
diff changeset
   744
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   745
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
   746
by (case_tac "d <0", simp_all)
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   747
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   748
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
   749
by (case_tac "d<0", simp_all)
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   750
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   751
(* 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
   752
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   753
lemma lcm_pos: 
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   754
  assumes mpos: "m > 0"
38159
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   755
    and npos: "n>0"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   756
  shows "lcm m n > 0"
38159
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   757
proof (rule ccontr, simp add: lcm_def gcd_zero)
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   758
  assume h:"m*n div gcd m n = 0"
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   759
  from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   760
  hence gcdp: "gcd m n > 0" by simp
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   761
  with h
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   762
  have "m*n < gcd m n"
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   763
    by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   764
  moreover 
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   765
  have "gcd m n dvd m" by simp
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   766
  with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   767
  with npos have t1:"gcd m n *n \<le> m*n" by simp
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   768
  have "gcd m n \<le> gcd m n*n" using npos by simp
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   769
  with t1 have "gcd m n \<le> m*n" by arith
e9b4835a54ee modernized specifications;
wenzelm
parents: 37765
diff changeset
   770
  ultimately show "False" by simp
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   771
qed
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   772
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   773
lemma zlcm_pos: 
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   774
  assumes anz: "a \<noteq> 0"
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   775
  and bnz: "b \<noteq> 0" 
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   776
  shows "0 < zlcm a b"
23244
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   777
proof-
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   778
  let ?na = "nat (abs a)"
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   779
  let ?nb = "nat (abs b)"
23983
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   780
  have nap: "?na >0" using anz by simp
79dc793bec43 Added lemmas
nipkow
parents: 23951
diff changeset
   781
  have nbp: "?nb >0" using bnz by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   782
  have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27487
diff changeset
   783
  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
   784
qed
1630951f0512 added lcm, ilcm (lcm for integers) and some lemmas about them;
chaieb
parents: 22367
diff changeset
   785
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 27676
diff changeset
   786
lemma zgcd_code [code]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27568
diff changeset
   787
  "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27568
diff changeset
   788
  by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27568
diff changeset
   789
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   790
end