# HG changeset patch # User haftmann # Date 1455742316 -3600 # Node ID 97f2ed2404314092d21d525f6d75581ad5aa8906 # Parent e66d7841d5a2dba24690e14ace468e1713ca0330 more theorems concerning gcd/lcm/Gcd/Lcm diff -r e66d7841d5a2 -r 97f2ed240431 src/HOL/GCD.thy --- a/src/HOL/GCD.thy Wed Feb 17 21:51:56 2016 +0100 +++ b/src/HOL/GCD.thy Wed Feb 17 21:51:56 2016 +0100 @@ -582,6 +582,20 @@ by (auto intro: associated_eqI) qed +lemma Gcd_eqI: + assumes "normalize a = a" + assumes "\b. b \ A \ a dvd b" + and "\c. (\b. b \ A \ c dvd b) \ c dvd a" + shows "Gcd A = a" + using assms by (blast intro: associated_eqI Gcd_greatest Gcd_dvd normalize_Gcd) + +lemma Lcm_eqI: + assumes "normalize a = a" + assumes "\b. b \ A \ b dvd a" + and "\c. (\b. b \ A \ b dvd c) \ a dvd c" + shows "Lcm A = a" + using assms by (blast intro: associated_eqI Lcm_least dvd_Lcm normalize_Lcm) + end @@ -656,6 +670,14 @@ (* specific to int *) +lemma gcd_eq_int_iff: + "gcd k l = int n \ gcd (nat \k\) (nat \l\) = n" + by (simp add: gcd_int_def) + +lemma lcm_eq_int_iff: + "lcm k l = int n \ lcm (nat \k\) (nat \l\) = n" + by (simp add: lcm_int_def) + lemma gcd_neg1_int [simp]: "gcd (-x::int) y = gcd x y" by (simp add: gcd_int_def) @@ -1874,6 +1896,10 @@ end +lemma Gcd_nat_eq_one: + "1 \ N \ Gcd N = (1::nat)" + by (rule Gcd_eq_1_I) auto + text\Alternative characterizations of Gcd:\ lemma Gcd_eq_Max: "finite(M::nat set) \ M \ {} \ 0 \ M \ Gcd M = Max(\m\M. {d. d dvd m})" @@ -1940,10 +1966,10 @@ begin definition - "Lcm M = int (Lcm (nat ` abs ` M))" + "Lcm M = int (Lcm m\M. (nat \ abs) m)" definition - "Gcd M = int (Gcd (nat ` abs ` M))" + "Gcd M = int (Gcd m\M. (nat \ abs) m)" instance by standard (auto intro!: Gcd_dvd Gcd_greatest simp add: Gcd_int_def @@ -1951,6 +1977,24 @@ end +lemma abs_Gcd [simp]: + fixes K :: "int set" + shows "\Gcd K\ = Gcd K" + using normalize_Gcd [of K] by simp + +lemma abs_Lcm [simp]: + fixes K :: "int set" + shows "\Lcm K\ = Lcm K" + using normalize_Lcm [of K] by simp + +lemma Gcm_eq_int_iff: + "Gcd K = int n \ Gcd ((nat \ abs) ` K) = n" + by (simp add: Gcd_int_def comp_def image_image) + +lemma Lcm_eq_int_iff: + "Lcm K = int n \ Lcm ((nat \ abs) ` K) = n" + by (simp add: Lcm_int_def comp_def image_image) + subsection \GCD and LCM on @{typ integer}\