# HG changeset patch # User wenzelm # Date 1455748175 -3600 # Node ID ab76bd43c14a764d0037ba1316bffcdc249298da # Parent 7f927120b5a2f11fe728125e87e5a8cec4180216# Parent e307a410f46c66ac9261cbf4f2211c80f4fc469a merged diff -r e307a410f46c -r ab76bd43c14a CONTRIBUTORS --- a/CONTRIBUTORS Wed Feb 17 23:28:58 2016 +0100 +++ b/CONTRIBUTORS Wed Feb 17 23:29:35 2016 +0100 @@ -6,6 +6,10 @@ Contributions to this Isabelle version -------------------------------------- +* January 2016: Florian Haftmann + Abolition of compound operators INFIMUM and SUPREMUM + for complete lattices. + Contributions to Isabelle2016 ----------------------------- diff -r e307a410f46c -r ab76bd43c14a NEWS --- a/NEWS Wed Feb 17 23:28:58 2016 +0100 +++ b/NEWS Wed Feb 17 23:29:35 2016 +0100 @@ -32,6 +32,109 @@ pred_prod_apply ~> pred_prod_inject INCOMPATIBILITY. +* Compound constants INFIMUM and SUPREMUM are mere abbreviations now. +INCOMPATIBILITY. + +* Class semiring_Lcd merged into semiring_Gcd. INCOMPATIBILITY. + +* Library/Polynomial.thy contains also derivation of polynomials +but not gcd/lcm on polynomials over fields. This has been moved +to a separate theory Library/Polynomial_GCD_euclidean.thy, to +pave way for a possible future different type class instantiation +for polynomials over factorial rings. INCOMPATIBILITY. + +* Dropped various legacy fact bindings, whose replacements are often +of a more general type also: + lcm_left_commute_nat ~> lcm.left_commute + lcm_left_commute_int ~> lcm.left_commute + gcd_left_commute_nat ~> gcd.left_commute + gcd_left_commute_int ~> gcd.left_commute + gcd_greatest_iff_nat ~> gcd_greatest_iff + gcd_greatest_iff_int ~> gcd_greatest_iff + coprime_dvd_mult_nat ~> coprime_dvd_mult + coprime_dvd_mult_int ~> coprime_dvd_mult + zpower_numeral_even ~> power_numeral_even + gcd_mult_cancel_nat ~> gcd_mult_cancel + gcd_mult_cancel_int ~> gcd_mult_cancel + div_gcd_coprime_nat ~> div_gcd_coprime + div_gcd_coprime_int ~> div_gcd_coprime + zpower_numeral_odd ~> power_numeral_odd + zero_less_int_conv ~> of_nat_0_less_iff + gcd_greatest_nat ~> gcd_greatest + gcd_greatest_int ~> gcd_greatest + coprime_mult_nat ~> coprime_mult + coprime_mult_int ~> coprime_mult + lcm_commute_nat ~> lcm.commute + lcm_commute_int ~> lcm.commute + int_less_0_conv ~> of_nat_less_0_iff + gcd_commute_nat ~> gcd.commute + gcd_commute_int ~> gcd.commute + Gcd_insert_nat ~> Gcd_insert + Gcd_insert_int ~> Gcd_insert + of_int_int_eq ~> of_int_of_nat_eq + lcm_least_nat ~> lcm_least + lcm_least_int ~> lcm_least + lcm_assoc_nat ~> lcm.assoc + lcm_assoc_int ~> lcm.assoc + int_le_0_conv ~> of_nat_le_0_iff + int_eq_0_conv ~> of_nat_eq_0_iff + Gcd_empty_nat ~> Gcd_empty + Gcd_empty_int ~> Gcd_empty + gcd_assoc_nat ~> gcd.assoc + gcd_assoc_int ~> gcd.assoc + zero_zle_int ~> of_nat_0_le_iff + lcm_dvd2_nat ~> dvd_lcm2 + lcm_dvd2_int ~> dvd_lcm2 + lcm_dvd1_nat ~> dvd_lcm1 + lcm_dvd1_int ~> dvd_lcm1 + gcd_zero_nat ~> gcd_eq_0_iff + gcd_zero_int ~> gcd_eq_0_iff + gcd_dvd2_nat ~> gcd_dvd2 + gcd_dvd2_int ~> gcd_dvd2 + gcd_dvd1_nat ~> gcd_dvd1 + gcd_dvd1_int ~> gcd_dvd1 + int_numeral ~> of_nat_numeral + lcm_ac_nat ~> ac_simps + lcm_ac_int ~> ac_simps + gcd_ac_nat ~> ac_simps + gcd_ac_int ~> ac_simps + abs_int_eq ~> abs_of_nat + zless_int ~> of_nat_less_iff + zdiff_int ~> of_nat_diff + zadd_int ~> of_nat_add + int_mult ~> of_nat_mult + int_Suc ~> of_nat_Suc + inj_int ~> inj_of_nat + int_1 ~> of_nat_1 + int_0 ~> of_nat_0 + Lcm_empty_nat ~> Lcm_empty + Lcm_empty_int ~> Lcm_empty + Lcm_insert_nat ~> Lcm_insert + Lcm_insert_int ~> Lcm_insert + comp_fun_idem_gcd_nat ~> comp_fun_idem_gcd + comp_fun_idem_gcd_int ~> comp_fun_idem_gcd + comp_fun_idem_lcm_nat ~> comp_fun_idem_lcm + comp_fun_idem_lcm_int ~> comp_fun_idem_lcm + Lcm_eq_0 ~> Lcm_eq_0_I + Lcm0_iff ~> Lcm_0_iff + Lcm_dvd_int ~> Lcm_least + divides_mult_nat ~> divides_mult + divides_mult_int ~> divides_mult + lcm_0_nat ~> lcm_0_right + lcm_0_int ~> lcm_0_right + lcm_0_left_nat ~> lcm_0_left + lcm_0_left_int ~> lcm_0_left + dvd_gcd_D1_nat ~> dvd_gcdD1 + dvd_gcd_D1_int ~> dvd_gcdD1 + dvd_gcd_D2_nat ~> dvd_gcdD2 + dvd_gcd_D2_int ~> dvd_gcdD2 + coprime_dvd_mult_iff_nat ~> coprime_dvd_mult_iff + coprime_dvd_mult_iff_int ~> coprime_dvd_mult_iff + realpow_minus_mult ~> power_minus_mult + realpow_Suc_le_self ~> power_Suc_le_self + dvd_Gcd, dvd_Gcd_nat, dvd_Gcd_int removed in favour of Gcd_greatest +INCOMPATIBILITY. + *** System *** diff -r e307a410f46c -r ab76bd43c14a src/HOL/Algebra/AbelCoset.thy --- a/src/HOL/Algebra/AbelCoset.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Algebra/AbelCoset.thy Wed Feb 17 23:29:35 2016 +0100 @@ -258,7 +258,7 @@ from a_subgroup have Hcarr: "H \ carrier G" unfolding subgroup_def by simp from xcarr Hcarr show "(\h\H. {h \\<^bsub>G\<^esub> x}) = (\h\H. {x \\<^bsub>G\<^esub> h})" - using m_comm [simplified] by fast + using m_comm [simplified] by fastforce qed qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Algebra/Coset.thy Wed Feb 17 23:29:35 2016 +0100 @@ -517,7 +517,7 @@ by (auto simp add: set_mult_def subsetD) lemma (in group) subgroup_mult_id: "subgroup H G \ H <#> H = H" -apply (auto simp add: subgroup.m_closed set_mult_def Sigma_def image_def) +apply (auto simp add: subgroup.m_closed set_mult_def Sigma_def) apply (rule_tac x = x in bexI) apply (rule bexI [of _ "\"]) apply (auto simp add: subgroup.one_closed subgroup.subset [THEN subsetD]) @@ -932,7 +932,7 @@ obtain g where g: "g \ carrier G" and "X = kernel G H h #> g" by (auto simp add: FactGroup_def RCOSETS_def) - hence "h ` X = {h g}" by (auto simp add: kernel_def r_coset_def image_def g) + hence "h ` X = {h g}" by (auto simp add: kernel_def r_coset_def g intro!: imageI) thus ?thesis by (auto simp add: g) qed @@ -952,11 +952,11 @@ and Xsub: "X \ carrier G" and X'sub: "X' \ carrier G" by (force simp add: kernel_def r_coset_def image_def)+ hence "h ` (X <#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' - by (auto dest!: FactGroup_nonempty - simp add: set_mult_def image_eq_UN + by (auto dest!: FactGroup_nonempty intro!: image_eqI + simp add: set_mult_def subsetD [OF Xsub] subsetD [OF X'sub]) - thus "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" - by (simp add: all image_eq_UN FactGroup_nonempty X X') + then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" + by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) qed @@ -964,7 +964,7 @@ lemma (in group_hom) FactGroup_subset: "\g \ carrier G; g' \ carrier G; h g = h g'\ \ kernel G H h #> g \ kernel G H h #> g'" -apply (clarsimp simp add: kernel_def r_coset_def image_def) +apply (clarsimp simp add: kernel_def r_coset_def) apply (rename_tac y) apply (rule_tac x="y \ g \ inv g'" in exI) apply (simp add: G.m_assoc) @@ -985,7 +985,7 @@ by (force simp add: kernel_def r_coset_def image_def)+ assume "the_elem (h ` X) = the_elem (h ` X')" hence h: "h g = h g'" - by (simp add: image_eq_UN all FactGroup_nonempty X X') + by (simp add: all FactGroup_nonempty X X' the_elem_image_unique) show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX) qed @@ -1006,7 +1006,10 @@ hence "(\x\kernel G H h #> g. {h x}) = {y}" by (auto simp add: y kernel_def r_coset_def) with g show "y \ (\X. the_elem (h ` X)) ` carrier (G Mod kernel G H h)" - by (auto intro!: bexI simp add: FactGroup_def RCOSETS_def image_eq_UN) + apply (auto intro!: bexI image_eqI simp add: FactGroup_def RCOSETS_def) + apply (subst the_elem_image_unique) + apply auto + done qed qed @@ -1019,5 +1022,4 @@ by (simp add: iso_def FactGroup_hom FactGroup_inj_on bij_betw_def FactGroup_onto) - end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Algebra/Exponent.thy --- a/src/HOL/Algebra/Exponent.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Algebra/Exponent.thy Wed Feb 17 23:29:35 2016 +0100 @@ -21,10 +21,10 @@ text\Prime Theorems\ lemma prime_iff: - "(prime p) = (Suc 0 < p & (\a b. p dvd a*b --> (p dvd a) | (p dvd b)))" -apply (auto simp add: prime_gt_Suc_0_nat) -by (metis (full_types) One_nat_def Suc_lessD dvd.order_refl nat_dvd_not_less not_prime_eq_prod_nat) - + "prime p \ Suc 0 < p \ (\a b. p dvd a * b \ p dvd a \ p dvd b)" + by (auto simp add: prime_gt_Suc_0_nat) + (metis One_nat_def Suc_lessD dvd_refl nat_dvd_not_less not_prime_eq_prod_nat) + lemma zero_less_prime_power: fixes p::nat shows "prime p ==> 0 < p^a" by (force simp add: prime_iff) @@ -201,7 +201,7 @@ apply (drule less_imp_le [of a]) apply (drule le_imp_power_dvd) apply (drule_tac b = "p ^ r" in dvd_trans, assumption) -apply (metis diff_is_0_eq dvd_diffD1 gcd_dvd2_nat gcd_mult' gr0I le_extend_mult less_diff_conv nat_dvd_not_less mult.commute not_add_less2 xt1(10)) +apply (metis diff_is_0_eq dvd_diffD1 gcd_dvd2 gcd_mult' gr0I le_extend_mult less_diff_conv nat_dvd_not_less mult.commute not_add_less2 xt1(10)) done lemma p_fac_forw: "[| (m::nat) > 0; k>0; k < p^a; (p^r) dvd (p^a)* m - k |] diff -r e307a410f46c -r ab76bd43c14a src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Algebra/IntRing.thy Wed Feb 17 23:29:35 2016 +0100 @@ -251,7 +251,7 @@ then obtain x where "1 = x * int p" by best then have "\int p * x\ = 1" by (simp add: mult.commute) then show False - by (metis abs_of_nat int_1 of_nat_eq_iff abs_zmult_eq_1 one_not_prime_nat prime) + by (metis abs_of_nat of_nat_1 of_nat_eq_iff abs_zmult_eq_1 one_not_prime_nat prime) qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Archimedean_Field.thy Wed Feb 17 23:29:35 2016 +0100 @@ -575,11 +575,10 @@ lemma frac_unique_iff: fixes x :: "'a::floor_ceiling" - shows "(frac x) = a \ x - a \ \ \ 0 \ a \ a < 1" - apply (auto simp: Ints_def frac_def) - apply linarith - apply linarith - by (metis (no_types) add.commute add.left_neutral eq_diff_eq floor_add_of_int floor_unique of_int_0) + shows "frac x = a \ x - a \ \ \ 0 \ a \ a < 1" + apply (auto simp: Ints_def frac_def algebra_simps floor_unique) + apply linarith+ + done lemma frac_eq: "(frac x) = x \ 0 \ x \ x < 1" by (simp add: frac_unique_iff) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Auth/Guard/Proto.thy --- a/src/HOL/Auth/Guard/Proto.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Auth/Guard/Proto.thy Wed Feb 17 23:29:35 2016 +0100 @@ -56,7 +56,7 @@ Nonce n ~:parts (apm s `(msg `(fst R))) |] ==> (EX k. Nonce k:parts {X} & nonce s k = n)" apply (erule Nonce_apm, unfold wdef_def) -apply (drule_tac x=R in spec, drule_tac x=k in spec, clarsimp simp: image_eq_UN) +apply (drule_tac x=R in spec, drule_tac x=k in spec, clarsimp) apply (drule_tac x=x in bspec, simp) apply (drule_tac Y="msg x" and s=s in apm_parts, simp) by (blast dest: parts_parts) @@ -134,7 +134,7 @@ lemma ok_not_used: "[| Nonce n ~:used evs; ok evs R s; ALL x. x:fst R --> is_Says x |] ==> Nonce n ~:parts (apm s `(msg `(fst R)))" -apply (unfold ok_def, clarsimp simp: image_eq_UN) +apply (unfold ok_def, clarsimp) apply (drule_tac x=x in spec, drule_tac x=x in spec) by (auto simp: is_Says_def dest: Says_imp_spies not_used_not_spied parts_parts) @@ -188,10 +188,10 @@ apply (drule wdef_Nonce, simp+) apply (frule ok_not_used, simp+) apply (clarify, erule ok_is_Says, simp+) -apply (clarify, rule_tac x=k in exI, simp add: newn_def image_eq_UN) +apply (clarify, rule_tac x=k in exI, simp add: newn_def) apply (clarify, drule_tac Y="msg x" and s=s in apm_parts) apply (drule ok_not_used, simp+) -by (clarify, erule ok_is_Says, simp_all add: image_eq_UN) +by (clarify, erule ok_is_Says, simp_all) lemma fresh_rule: "[| evs' @ ev # evs:tr p; wdef p; Nonce n ~:used evs; Nonce n:parts {msg ev} |] ==> EX R s. R:p & ap' s R = ev" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Auth/Message.thy Wed Feb 17 23:29:35 2016 +0100 @@ -207,8 +207,16 @@ apply (erule parts.induct, blast+) done -lemma parts_UN [simp]: "parts(\x\A. H x) = (\x\A. parts(H x))" -by (intro equalityI parts_UN_subset1 parts_UN_subset2) +lemma parts_UN [simp]: + "parts (\x\A. H x) = (\x\A. parts (H x))" + by (intro equalityI parts_UN_subset1 parts_UN_subset2) + +lemma parts_image [simp]: + "parts (f ` A) = (\x\A. parts {f x})" + apply auto + apply (metis (mono_tags, hide_lams) image_iff parts_singleton) + apply (metis empty_subsetI image_eqI insert_absorb insert_subset parts_mono) + done text\Added to simplify arguments to parts, analz and synth. NOTE: the UN versions are no longer used!\ @@ -299,10 +307,7 @@ done lemma parts_image_Key [simp]: "parts (Key`N) = Key`N" -apply auto -apply (erule parts.induct, auto) -done - +by auto text\In any message, there is an upper bound N on its greatest nonce.\ lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Auth/Recur.thy --- a/src/HOL/Auth/Recur.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Auth/Recur.thy Wed Feb 17 23:29:35 2016 +0100 @@ -337,7 +337,7 @@ RB \ responses evs |] ==> (Key K \ parts{RB} | Key K \ analz H)" apply (erule rev_mp, erule responses.induct) -apply (simp_all del: image_insert +apply (simp_all del: image_insert parts_image add: analz_image_freshK_simps resp_analz_image_freshK_lemma) txt\Simplification using two distinct treatments of "image"\ apply (simp add: parts_insert2, blast) @@ -389,7 +389,7 @@ apply (erule respond.induct) apply (frule_tac [2] respond_imp_responses) apply (frule_tac [2] respond_imp_not_used) -apply (simp_all del: image_insert +apply (simp_all del: image_insert parts_image add: analz_image_freshK_simps split_ifs shrK_in_analz_respond resp_analz_image_freshK parts_insert2) txt\Base case of respond\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Auth/Smartcard/ShoupRubin.thy --- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Feb 17 23:29:35 2016 +0100 @@ -809,9 +809,9 @@ (*Because initState contains a set of nonces, this is needed for base case of analz_image_freshK*) -lemma analz_image_Key_Un_Nonce: "analz (Key`K \ Nonce`N) = Key`K \ Nonce`N" -apply auto -done +lemma analz_image_Key_Un_Nonce: + "analz (Key ` K \ Nonce ` N) = Key ` K \ Nonce ` N" + by (auto simp del: parts_image) method_setup sc_analz_freshK = \ Scan.succeed (fn ctxt => diff -r e307a410f46c -r ab76bd43c14a src/HOL/Auth/Smartcard/ShoupRubinBella.thy --- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Feb 17 23:29:35 2016 +0100 @@ -819,9 +819,9 @@ (*Because initState contains a set of nonces, this is needed for base case of analz_image_freshK*) -lemma analz_image_Key_Un_Nonce: "analz (Key`K \ Nonce`N) = Key`K \ Nonce`N" -apply auto -done +lemma analz_image_Key_Un_Nonce: + "analz (Key ` K \ Nonce ` N) = Key ` K \ Nonce ` N" + by (auto simp del: parts_image) method_setup sc_analz_freshK = \ Scan.succeed (fn ctxt => diff -r e307a410f46c -r ab76bd43c14a src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Wed Feb 17 23:29:35 2016 +0100 @@ -115,10 +115,7 @@ text\Added to extend initstate with set of nonces\ lemma parts_image_Nonce [simp]: "parts (Nonce`N) = Nonce`N" -apply auto -apply (erule parts.induct) -apply auto -done + by auto lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}" apply (unfold keysFor_def) diff -r e307a410f46c -r ab76bd43c14a src/HOL/BNF_Cardinal_Order_Relation.thy --- a/src/HOL/BNF_Cardinal_Order_Relation.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/BNF_Cardinal_Order_Relation.thy Wed Feb 17 23:29:35 2016 +0100 @@ -711,7 +711,7 @@ lemma card_of_UNION_Sigma: "|\i \ I. A i| \o |SIGMA i : I. A i|" -using Ex_inj_on_UNION_Sigma[of I A] card_of_ordLeq by blast +using Ex_inj_on_UNION_Sigma [of A I] card_of_ordLeq by blast lemma card_of_bool: assumes "a1 \ a2" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Binomial.thy --- a/src/HOL/Binomial.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Binomial.thy Wed Feb 17 23:29:35 2016 +0100 @@ -25,9 +25,19 @@ lemma fact_Suc_0 [simp]: "fact (Suc 0) = Suc 0" by simp -lemma of_nat_fact [simp]: "of_nat (fact n) = fact n" +lemma of_nat_fact [simp]: + "of_nat (fact n) = fact n" by (induct n) (auto simp add: algebra_simps of_nat_mult) +lemma of_int_fact [simp]: + "of_int (fact n) = fact n" +proof - + have "of_int (of_nat (fact n)) = fact n" + unfolding of_int_of_nat_eq by simp + then show ?thesis + by simp +qed + lemma fact_reduce: "n > 0 \ fact n = of_nat n * fact (n - 1)" by (cases n) auto @@ -1323,7 +1333,7 @@ also have "... = (fact (m+r+k) * fact (m+r)) div (fact r * (fact k * fact m) * fact (m+r))" apply (subst div_mult_div_if_dvd [symmetric]) apply (auto simp add: algebra_simps) - apply (metis fact_fact_dvd_fact dvd.order.trans nat_mult_dvd_cancel_disj) + apply (metis fact_fact_dvd_fact dvd_trans nat_mult_dvd_cancel_disj) done also have "... = (fact (m+r+k) div (fact k * fact (m+r)) * (fact (m+r) div (fact r * fact m)))" apply (subst div_mult_div_if_dvd) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Cardinals/Wellorder_Relation.thy --- a/src/HOL/Cardinals/Wellorder_Relation.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Cardinals/Wellorder_Relation.thy Wed Feb 17 23:29:35 2016 +0100 @@ -450,8 +450,7 @@ lemma ofilter_under_Union: "ofilter A \ A = \{under a| a. a \ A}" -using ofilter_under_UNION[of A] -by(unfold Union_eq, auto) +using ofilter_under_UNION [of A] by auto subsubsection {* Other properties *} diff -r e307a410f46c -r ab76bd43c14a src/HOL/Codegenerator_Test/Candidates.thy --- a/src/HOL/Codegenerator_Test/Candidates.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Codegenerator_Test/Candidates.thy Wed Feb 17 23:29:35 2016 +0100 @@ -8,6 +8,7 @@ Complex_Main "~~/src/HOL/Library/Library" "~~/src/HOL/Library/Sublist_Order" + "~~/src/HOL/Library/Polynomial_GCD_euclidean" "~~/src/HOL/Data_Structures/Tree_Map" "~~/src/HOL/Data_Structures/Tree_Set" "~~/src/HOL/Number_Theory/Eratosthenes" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Complete_Lattices.thy Wed Feb 17 23:29:35 2016 +0100 @@ -17,28 +17,25 @@ fixes Inf :: "'a set \ 'a" ("\_" [900] 900) begin -definition INFIMUM :: "'b set \ ('b \ 'a) \ 'a" where - INF_def: "INFIMUM A f = \(f ` A)" - -lemma Inf_image_eq [simp]: - "\(f ` A) = INFIMUM A f" - by (simp add: INF_def) +abbreviation INFIMUM :: "'b set \ ('b \ 'a) \ 'a" +where + "INFIMUM A f \ \(f ` A)" lemma INF_image [simp]: "INFIMUM (f ` A) g = INFIMUM A (g \ f)" - by (simp only: INF_def image_comp) + by (simp add: image_comp) lemma INF_identity_eq [simp]: "INFIMUM A (\x. x) = \A" - by (simp add: INF_def) + by simp lemma INF_id_eq [simp]: "INFIMUM A id = \A" - by (simp add: id_def) + by simp lemma INF_cong: "A = B \ (\x. x \ B \ C x = D x) \ INFIMUM A C = INFIMUM B D" - by (simp add: INF_def image_def) + by (simp add: image_def) lemma strong_INF_cong [cong]: "A = B \ (\x. x \ B =simp=> C x = D x) \ INFIMUM A C = INFIMUM B D" @@ -50,20 +47,17 @@ fixes Sup :: "'a set \ 'a" ("\_" [900] 900) begin -definition SUPREMUM :: "'b set \ ('b \ 'a) \ 'a" where - SUP_def: "SUPREMUM A f = \(f ` A)" - -lemma Sup_image_eq [simp]: - "\(f ` A) = SUPREMUM A f" - by (simp add: SUP_def) +abbreviation SUPREMUM :: "'b set \ ('b \ 'a) \ 'a" +where + "SUPREMUM A f \ \(f ` A)" lemma SUP_image [simp]: "SUPREMUM (f ` A) g = SUPREMUM A (g \ f)" - by (simp only: SUP_def image_comp) + by (simp add: image_comp) lemma SUP_identity_eq [simp]: "SUPREMUM A (\x. x) = \A" - by (simp add: SUP_def) + by simp lemma SUP_id_eq [simp]: "SUPREMUM A id = \A" @@ -71,7 +65,7 @@ lemma SUP_cong: "A = B \ (\x. x \ B \ C x = D x) \ SUPREMUM A C = SUPREMUM B D" - by (simp add: SUP_def image_def) + by (simp add: image_def) lemma strong_SUP_cong [cong]: "A = B \ (\x. x \ B =simp=> C x = D x) \ SUPREMUM A C = SUPREMUM B D" @@ -153,14 +147,6 @@ context complete_lattice begin -lemma INF_foundation_dual: - "Sup.SUPREMUM Inf = INFIMUM" - by (simp add: fun_eq_iff Sup.SUP_def) - -lemma SUP_foundation_dual: - "Inf.INFIMUM Sup = SUPREMUM" - by (simp add: fun_eq_iff Inf.INF_def) - lemma Sup_eqI: "(\y. y \ A \ y \ x) \ (\y. (\z. z \ A \ z \ y) \ x \ y) \ \A = x" by (blast intro: antisym Sup_least Sup_upper) @@ -217,19 +203,19 @@ by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) lemma INF_insert [simp]: "(\x\insert a A. f x) = f a \ INFIMUM A f" - unfolding INF_def Inf_insert by simp + by (simp cong del: strong_INF_cong) lemma Sup_insert [simp]: "\insert a A = a \ \A" by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) lemma SUP_insert [simp]: "(\x\insert a A. f x) = f a \ SUPREMUM A f" - unfolding SUP_def Sup_insert by simp + by (simp cong del: strong_SUP_cong) lemma INF_empty [simp]: "(\x\{}. f x) = \" - by (simp add: INF_def) + by (simp cong del: strong_INF_cong) lemma SUP_empty [simp]: "(\x\{}. f x) = \" - by (simp add: SUP_def) + by (simp cong del: strong_SUP_cong) lemma Inf_UNIV [simp]: "\UNIV = \" @@ -308,18 +294,18 @@ ultimately show ?thesis by (rule Sup_upper2) qed +lemma INF_eq: + assumes "\i. i \ A \ \j\B. f i \ g j" + assumes "\j. j \ B \ \i\A. g j \ f i" + shows "INFIMUM A f = INFIMUM B g" + by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+ + lemma SUP_eq: assumes "\i. i \ A \ \j\B. f i \ g j" assumes "\j. j \ B \ \i\A. g j \ f i" - shows "(\i\A. f i) = (\j\B. g j)" + shows "SUPREMUM A f = SUPREMUM B g" by (intro antisym SUP_least) (blast intro: SUP_upper2 dest: assms)+ -lemma INF_eq: - assumes "\i. i \ A \ \j\B. f i \ g j" - assumes "\j. j \ B \ \i\A. g j \ f i" - shows "(\i\A. f i) = (\j\B. g j)" - by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+ - lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" by (auto intro: Inf_greatest Inf_lower) @@ -498,24 +484,24 @@ lemma sup_INF: "a \ (\b\B. f b) = (\b\B. a \ f b)" - by (simp only: INF_def sup_Inf image_image) + unfolding sup_Inf by simp lemma inf_SUP: "a \ (\b\B. f b) = (\b\B. a \ f b)" - by (simp only: SUP_def inf_Sup image_image) + unfolding inf_Sup by simp lemma dual_complete_distrib_lattice: "class.complete_distrib_lattice Sup Inf sup (op \) (op >) inf \ \" apply (rule class.complete_distrib_lattice.intro) apply (fact dual_complete_lattice) apply (rule class.complete_distrib_lattice_axioms.intro) - apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf) + apply (simp_all add: inf_Sup sup_Inf) done subclass distrib_lattice proof fix a b c from sup_Inf have "a \ \{b, c} = (\d\{b, c}. a \ d)" . - then show "a \ b \ c = (a \ b) \ (a \ c)" by (simp add: INF_def) + then show "a \ b \ c = (a \ b) \ (a \ c)" by simp qed lemma Inf_sup: @@ -592,7 +578,7 @@ qed lemma uminus_INF: "- (\x\A. B x) = (\x\A. - B x)" - by (simp only: INF_def SUP_def uminus_Inf image_image) + by (simp add: uminus_Inf image_image) lemma uminus_Sup: "- (\A) = \(uminus ` A)" @@ -602,7 +588,7 @@ qed lemma uminus_SUP: "- (\x\A. B x) = (\x\A. - B x)" - by (simp only: INF_def SUP_def uminus_Sup image_image) + by (simp add: uminus_Sup image_image) end @@ -731,11 +717,11 @@ lemma INF_bool_eq [simp]: "INFIMUM = Ball" - by (simp add: fun_eq_iff INF_def) + by (simp add: fun_eq_iff) lemma SUP_bool_eq [simp]: "SUPREMUM = Bex" - by (simp add: fun_eq_iff SUP_def) + by (simp add: fun_eq_iff) instance bool :: complete_boolean_algebra proof qed (auto intro: bool_induct) @@ -788,8 +774,7 @@ using Sup_apply [of "f ` A"] by (simp add: comp_def) instance "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice proof -qed (auto simp add: INF_def SUP_def inf_Sup sup_Inf fun_eq_iff image_image - simp del: Inf_image_eq Sup_image_eq) +qed (auto simp add: inf_Sup sup_Inf fun_eq_iff image_image) instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. @@ -903,7 +888,7 @@ instance "set" :: (type) complete_boolean_algebra proof -qed (auto simp add: INF_def SUP_def Inf_set_def Sup_set_def image_def) +qed (auto simp add: Inf_set_def Sup_set_def image_def) subsubsection \Inter\ @@ -1011,22 +996,18 @@ "(\x\A. B x) = {y. \x\A. y \ B x}" by (auto intro!: INF_eqI) -lemma Inter_image_eq: - "\(B ` A) = (\x\A. B x)" - by (fact Inf_image_eq) - lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" using Inter_iff [of _ "B ` A"] by simp lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" - by (auto simp add: INF_def image_def) + by auto lemma INT_D [elim, Pure.elim]: "b \ (\x\A. B x) \ a \ A \ b \ B a" by auto lemma INT_E [elim]: "b \ (\x\A. B x) \ (b \ B a \ R) \ (a \ A \ R) \ R" \ \"Classical" elimination -- by the Excluded Middle on @{prop "a\A"}.\ - by (auto simp add: INF_def image_def) + by auto lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" by blast @@ -1198,10 +1179,6 @@ "x \ Set.bind P f \ x \ UNION P f " by (simp add: bind_UNION) -lemma Union_image_eq: - "\(B ` A) = (\x\A. B x)" - by (fact Sup_image_eq) - lemma Union_SetCompr_eq: "\{f x| x. P x} = {a. \x. P x \ a \ f x}" by blast @@ -1214,10 +1191,7 @@ by auto lemma UN_E [elim!]: "b \ (\x\A. B x) \ (\x. x\A \ b \ B x \ R) \ R" - by (auto simp add: SUP_def image_def) - -lemma image_eq_UN: "f ` A = (\x\A. {f x})" - by blast + by auto lemma UN_upper: "a \ A \ B a \ (\x\A. B x)" by (fact SUP_upper) @@ -1273,7 +1247,7 @@ by blast lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" - by (auto simp add: split_if_mem2) + by safe (auto simp add: split_if_mem2) lemma UN_bool_eq: "(\b. A b) = (A True \ A False)" by (fact SUP_UNIV_bool_expand) @@ -1355,7 +1329,7 @@ lemma inj_on_INTER: "I \ {} \ (\i. i \ I \ inj_on f (A i)) \ inj_on f (\i \ I. A i)" - unfolding inj_on_def by blast + unfolding inj_on_def by safe simp lemma inj_on_UNION_chain: assumes CH: "\ i j. \i \ I; j \ I\ \ A i \ A j \ A j \ A i" and @@ -1414,20 +1388,19 @@ lemma image_INT: "[| inj_on f C; ALL x:A. B x <= C; j:A |] ==> f ` (INTER A B) = (INT x:A. f ` B x)" -apply (simp add: inj_on_def, blast) -done + by (simp add: inj_on_def, auto) blast -(*Compare with image_INT: no use of inj_on, and if f is surjective then - it doesn't matter whether A is empty*) lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)" -apply (simp add: bij_def) -apply (simp add: inj_on_def surj_def, blast) -done + apply (simp add: bij_def) + apply (simp add: inj_on_def surj_def) + apply auto + apply blast + done lemma UNION_fun_upd: - "UNION J (A(i:=B)) = (UNION (J-{i}) A \ (if i\J then B else {}))" -by (auto split: if_splits) - + "UNION J (A(i := B)) = UNION (J - {i}) A \ (if i \ J then B else {})" + by (auto simp add: set_eq_iff) + subsubsection \Complement\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Wed Feb 17 23:29:35 2016 +0100 @@ -321,10 +321,10 @@ by (metis cSUP_upper le_less_trans) lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ INFIMUM (insert a A) f = inf (f a) (INFIMUM A f)" - by (metis cInf_insert Inf_image_eq image_insert image_is_empty) + by (metis cInf_insert image_insert image_is_empty) lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ SUPREMUM (insert a A) f = sup (f a) (SUPREMUM A f)" - by (metis cSup_insert Sup_image_eq image_insert image_is_empty) + by (metis cSup_insert image_insert image_is_empty) lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ INFIMUM A f \ INFIMUM B g" using cInf_mono [of "g ` B" "f ` A"] by auto diff -r e307a410f46c -r ab76bd43c14a src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Decision_Procs/Cooper.thy Wed Feb 17 23:29:35 2016 +0100 @@ -2442,9 +2442,9 @@ | num_of_term vs @{term "1::int"} = @{code C} (@{code int_of_integer} 1) | num_of_term vs @{term "- 1::int"} = @{code C} (@{code int_of_integer} (~ 1)) | num_of_term vs (@{term "numeral :: _ \ int"} $ t) = - @{code C} (@{code int_of_integer} (HOLogic.dest_num t)) + @{code C} (@{code int_of_integer} (HOLogic.dest_numeral t)) | num_of_term vs (@{term "- numeral :: _ \ int"} $ t) = - @{code C} (@{code int_of_integer} (~(HOLogic.dest_num t))) + @{code C} (@{code int_of_integer} (~(HOLogic.dest_numeral t))) | num_of_term vs (Bound i) = @{code Bound} (@{code nat_of_integer} i) | num_of_term vs (@{term "uminus :: int \ int"} $ t') = @{code Neg} (num_of_term vs t') | num_of_term vs (@{term "op + :: int \ int \ int"} $ t1 $ t2) = diff -r e307a410f46c -r ab76bd43c14a src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Wed Feb 17 23:29:35 2016 +0100 @@ -5547,17 +5547,17 @@ of @{code C} i => @{code Mul} (i, num_of_term vs t2) | _ => error "num_of_term: unsupported Multiplication") | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t')) = - mk_C (HOLogic.dest_num t') + mk_C (HOLogic.dest_numeral t') | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t')) = - mk_C (~ (HOLogic.dest_num t')) + mk_C (~ (HOLogic.dest_numeral t')) | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "floor :: real \ int"} $ t')) = @{code Floor} (num_of_term vs t') | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "ceiling :: real \ int"} $ t')) = @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) | num_of_term vs (@{term "numeral :: _ \ real"} $ t') = - mk_C (HOLogic.dest_num t') + mk_C (HOLogic.dest_numeral t') | num_of_term vs (@{term "- numeral :: _ \ real"} $ t') = - mk_C (~ (HOLogic.dest_num t')) + mk_C (~ (HOLogic.dest_numeral t')) | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); fun fm_of_term vs @{term True} = @{code T} @@ -5569,9 +5569,9 @@ | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs (@{term "op rdvd"} $ (@{term "of_int :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t1)) $ t2) = - mk_Dvd (HOLogic.dest_num t1, num_of_term vs t2) + mk_Dvd (HOLogic.dest_numeral t1, num_of_term vs t2) | fm_of_term vs (@{term "op rdvd"} $ (@{term "of_int :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t1)) $ t2) = - mk_Dvd (~ (HOLogic.dest_num t1), num_of_term vs t2) + mk_Dvd (~ (HOLogic.dest_numeral t1), num_of_term vs t2) | fm_of_term vs (@{term "op = :: bool \ bool \ bool"} $ t1 $ t2) = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) = diff -r e307a410f46c -r ab76bd43c14a src/HOL/Decision_Procs/Rat_Pair.thy --- a/src/HOL/Decision_Procs/Rat_Pair.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Decision_Procs/Rat_Pair.thy Wed Feb 17 23:29:35 2016 +0100 @@ -27,7 +27,7 @@ (let g = gcd a b in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))" -declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger] +declare gcd_dvd1[presburger] gcd_dvd2[presburger] lemma normNum_isnormNum [simp]: "isnormNum (normNum x)" proof - @@ -51,7 +51,7 @@ from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] ab have nz': "?a' \ 0" "?b' \ 0" by - (rule notI, simp)+ from ab have stupid: "a \ 0 \ b \ 0" by arith - from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" . + from div_gcd_coprime[OF stupid] have gp1: "?g' = 1" . from ab consider "b < 0" | "b > 0" by arith then show ?thesis proof cases @@ -142,7 +142,7 @@ lemma Ninv_normN[simp]: "isnormNum x \ isnormNum (Ninv x)" apply (simp add: Ninv_def isnormNum_def split_def) apply (cases "fst x = 0") - apply (auto simp add: gcd_commute_int) + apply (auto simp add: gcd.commute) done lemma isnormNum_int[simp]: "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \ 0 \ isnormNum (i)\<^sub>N" @@ -197,7 +197,7 @@ by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult) from \a \ 0\ \a' \ 0\ na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1" - by (simp_all add: x y isnormNum_def add: gcd_commute_int) + by (simp_all add: x y isnormNum_def add: gcd.commute) from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'" apply - apply algebra @@ -205,8 +205,8 @@ apply simp apply algebra done - from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)] - coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]] + from zdvd_antisym_abs[OF coprime_dvd_mult[OF gcd1(2) raw_dvd(2)] + coprime_dvd_mult[OF gcd1(4) raw_dvd(4)]] have eq1: "b = b'" using pos by arith with eq have "a = a'" using pos by simp with eq1 show ?thesis by (simp add: x y) @@ -551,7 +551,7 @@ qed lemma Nmul_commute: "isnormNum x \ isnormNum y \ x *\<^sub>N y = y *\<^sub>N x" - by (simp add: Nmul_def split_def Let_def gcd_commute_int mult.commute) + by (simp add: Nmul_def split_def Let_def gcd.commute mult.commute) lemma Nmul_assoc: assumes "SORT_CONSTRAINT('a::{field_char_0,field})" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Decision_Procs/cooper_tac.ML --- a/src/HOL/Decision_Procs/cooper_tac.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Decision_Procs/cooper_tac.ML Wed Feb 17 23:29:35 2016 +0100 @@ -62,13 +62,14 @@ val simpset1 = put_simpset HOL_basic_ss ctxt addsimps @{thms zdvd_int} @ [@{thm "of_nat_add"}, @{thm "of_nat_mult"}] @ - map (fn r => r RS sym) @{thms int_numeral int_int_eq zle_int zless_int} + map (fn r => r RS sym) @{thms of_nat_numeral [where ?'a = int] int_int_eq zle_int of_nat_less_iff [where ?'a = int]} |> Splitter.add_split @{thm zdiff_int_split} (*simp rules for elimination of int n*) val simpset2 = put_simpset HOL_basic_ss ctxt - addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, @{thm zero_le_numeral}, @{thm order_refl}(* FIXME: necessary? *), @{thm int_0}, @{thm int_1}] + addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, @{thm zero_le_numeral}, @{thm order_refl}(* FIXME: necessary? *), + @{thm of_nat_0 [where ?'a = int]}, @{thm of_nat_1 [where ?'a = int]}] |> fold Simplifier.add_cong @{thms conj_le_cong imp_le_cong} (* simp rules for elimination of abs *) val simpset3 = put_simpset HOL_basic_ss ctxt |> Splitter.add_split @{thm abs_split} diff -r e307a410f46c -r ab76bd43c14a src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Decision_Procs/mir_tac.ML Wed Feb 17 23:29:35 2016 +0100 @@ -88,13 +88,13 @@ (* Simp rules for changing (n::int) to int n *) val simpset1 = put_simpset HOL_basic_ss ctxt addsimps [@{thm "zdvd_int"}, @{thm "of_nat_add"}, @{thm "of_nat_mult"}] @ - map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm nat_numeral}] + map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "of_nat_less_iff"}, @{thm nat_numeral}] |> Splitter.add_split @{thm "zdiff_int_split"} (*simp rules for elimination of int n*) val simpset2 = put_simpset HOL_basic_ss ctxt addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm zero_le_numeral}, - @{thm "int_0"}, @{thm "int_1"}] + @{thm "of_nat_0"}, @{thm "of_nat_1"}] |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] (* simp rules for elimination of abs *) val ct = Thm.cterm_of ctxt (HOLogic.mk_Trueprop t) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Enum.thy --- a/src/HOL/Enum.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Enum.thy Wed Feb 17 23:29:35 2016 +0100 @@ -556,7 +556,7 @@ end instance finite_1 :: complete_distrib_lattice -by intro_classes(simp_all add: INF_def SUP_def) + by standard simp_all instance finite_1 :: complete_linorder .. @@ -679,7 +679,7 @@ end instance finite_2 :: complete_distrib_lattice -by(intro_classes)(auto simp add: INF_def SUP_def sup_finite_2_def inf_finite_2_def Inf_finite_2_def Sup_finite_2_def) + by standard (auto simp add: sup_finite_2_def inf_finite_2_def Inf_finite_2_def Sup_finite_2_def) instance finite_2 :: complete_linorder .. @@ -797,11 +797,11 @@ then have "\x. x \ B \ x = a\<^sub>3" by(case_tac x)(auto simp add: Inf_finite_3_def split: split_if_asm) then show ?thesis using a\<^sub>2_a\<^sub>3 - by(auto simp add: INF_def Inf_finite_3_def max_def less_eq_finite_3_def less_finite_3_def split: split_if_asm) - qed(auto simp add: INF_def Inf_finite_3_def max_def less_finite_3_def less_eq_finite_3_def split: split_if_asm) + by(auto simp add: Inf_finite_3_def max_def less_eq_finite_3_def less_finite_3_def split: split_if_asm) + qed (auto simp add: Inf_finite_3_def max_def less_finite_3_def less_eq_finite_3_def split: split_if_asm) show "a \ \B = (\b\B. a \ b)" - by(cases a "\B" rule: finite_3.exhaust[case_product finite_3.exhaust]) - (auto simp add: SUP_def Sup_finite_3_def min_def less_finite_3_def less_eq_finite_3_def split: split_if_asm) + by (cases a "\B" rule: finite_3.exhaust[case_product finite_3.exhaust]) + (auto simp add: Sup_finite_3_def min_def less_finite_3_def less_eq_finite_3_def split: split_if_asm) qed instance finite_3 :: complete_linorder .. @@ -920,10 +920,10 @@ fix a :: finite_4 and B show "a \ \B = (\b\B. a \ b)" by(cases a "\B" rule: finite_4.exhaust[case_product finite_4.exhaust]) - (auto simp add: sup_finite_4_def Inf_finite_4_def INF_def split: finite_4.splits split_if_asm) + (auto simp add: sup_finite_4_def Inf_finite_4_def split: finite_4.splits split_if_asm) show "a \ \B = (\b\B. a \ b)" by(cases a "\B" rule: finite_4.exhaust[case_product finite_4.exhaust]) - (auto simp add: inf_finite_4_def Sup_finite_4_def SUP_def split: finite_4.splits split_if_asm) + (auto simp add: inf_finite_4_def Sup_finite_4_def split: finite_4.splits split_if_asm) qed instantiation finite_4 :: complete_boolean_algebra begin diff -r e307a410f46c -r ab76bd43c14a src/HOL/Fields.thy --- a/src/HOL/Fields.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Fields.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1152,6 +1152,10 @@ lemma divide_le_0_abs_iff [simp]: "(a / \b\ \ 0) = (a \ 0 | b = 0)" by (auto simp: divide_le_0_iff) +lemma inverse_sgn: + "sgn (inverse a) = inverse (sgn a)" + by (simp add: sgn_if) + lemma field_le_mult_one_interval: assumes *: "\z. \ 0 < z ; z < 1 \ \ z * x \ y" shows "x \ y" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Filter.thy --- a/src/HOL/Filter.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Filter.thy Wed Feb 17 23:29:35 2016 +0100 @@ -437,8 +437,8 @@ qed lemma eventually_INF: "eventually P (INF b:B. F b) \ (\X\B. finite X \ eventually P (INF b:X. F b))" - unfolding INF_def[of B] eventually_Inf[of P "F`B"] - by (metis Inf_image_eq finite_imageI image_mono finite_subset_image) + unfolding eventually_Inf [of P "F`B"] + by (metis finite_imageI image_mono finite_subset_image) lemma Inf_filter_not_bot: fixes B :: "'a filter set" @@ -449,7 +449,7 @@ lemma INF_filter_not_bot: fixes F :: "'i \ 'a filter" shows "(\X. X \ B \ finite X \ (INF b:X. F b) \ bot) \ (INF b:B. F b) \ bot" - unfolding trivial_limit_def eventually_INF[of _ B] + unfolding trivial_limit_def eventually_INF [of _ _ B] bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp lemma eventually_Inf_base: @@ -477,7 +477,7 @@ lemma eventually_INF_base: "B \ {} \ (\a b. a \ B \ b \ B \ \x\B. F x \ inf (F a) (F b)) \ eventually P (INF b:B. F b) \ (\b\B. eventually P (F b))" - unfolding INF_def by (subst eventually_Inf_base) auto + by (subst eventually_Inf_base) auto subsubsection \Map function for filters\ @@ -573,7 +573,7 @@ by (auto intro: exI[of _ "\x. x \ A"] exI[of _ "\x. x \ B"]) lemma SUP_principal[simp]: "(SUP i : I. principal (A i)) = principal (\i\I. A i)" - unfolding filter_eq_iff eventually_Sup SUP_def by (auto simp: eventually_principal) + unfolding filter_eq_iff eventually_Sup by (auto simp: eventually_principal) lemma INF_principal_finite: "finite X \ (INF x:X. principal (f x)) = principal (\x\X. f x)" by (induct X rule: finite_induct) auto diff -r e307a410f46c -r ab76bd43c14a src/HOL/GCD.thy --- a/src/HOL/GCD.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/GCD.thy Wed Feb 17 23:29:35 2016 +0100 @@ -31,7 +31,7 @@ imports Main begin -subsection \GCD and LCM definitions\ +subsection \Abstract GCD and LCM\ class gcd = zero + one + dvd + fixes gcd :: "'a \ 'a \ 'a" @@ -44,8 +44,40 @@ end class Gcd = gcd + - fixes Gcd :: "'a set \ 'a" - and Lcm :: "'a set \ 'a" + fixes Gcd :: "'a set \ 'a" ("Gcd") + and Lcm :: "'a set \ 'a" ("Lcm") +begin + +abbreviation GCD :: "'b set \ ('b \ 'a) \ 'a" +where + "GCD A f \ Gcd (f ` A)" + +abbreviation LCM :: "'b set \ ('b \ 'a) \ 'a" +where + "LCM A f \ Lcm (f ` A)" + +end + +syntax + "_GCD1" :: "pttrns \ 'b \ 'b" ("(3Gcd _./ _)" [0, 10] 10) + "_GCD" :: "pttrn \ 'a set \ 'b \ 'b" ("(3Gcd _\_./ _)" [0, 0, 10] 10) + "_LCM1" :: "pttrns \ 'b \ 'b" ("(3Lcm _./ _)" [0, 10] 10) + "_LCM" :: "pttrn \ 'a set \ 'b \ 'b" ("(3Lcm _\_./ _)" [0, 0, 10] 10) + +translations + "Gcd x y. B" \ "Gcd x. Gcd y. B" + "Gcd x. B" \ "CONST GCD CONST UNIV (\x. B)" + "Gcd x. B" \ "Gcd x \ CONST UNIV. B" + "Gcd x\A. B" \ "CONST GCD A (\x. B)" + "Lcm x y. B" \ "Lcm x. Lcm y. B" + "Lcm x. B" \ "CONST LCM CONST UNIV (\x. B)" + "Lcm x. B" \ "Lcm x \ CONST UNIV. B" + "Lcm x\A. B" \ "CONST LCM A (\x. B)" + +print_translation \ + [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax GCD} @{syntax_const "_GCD"}, + Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax LCM} @{syntax_const "_LCM"}] +\ \ \to avoid eta-contraction of body\ class semiring_gcd = normalization_semidom + gcd + assumes gcd_dvd1 [iff]: "gcd a b dvd a" @@ -67,6 +99,14 @@ "b dvd c \ gcd a b dvd c" by (rule dvd_trans) (rule gcd_dvd2) +lemma dvd_gcdD1: + "a dvd gcd b c \ a dvd b" + using gcd_dvd1 [of b c] by (blast intro: dvd_trans) + +lemma dvd_gcdD2: + "a dvd gcd b c \ a dvd c" + using gcd_dvd2 [of b c] by (blast intro: dvd_trans) + lemma gcd_0_left [simp]: "gcd 0 a = normalize a" by (rule associated_eqI) simp_all @@ -203,6 +243,14 @@ "a dvd c \ a dvd lcm b c" by (rule dvd_trans) (assumption, blast) +lemma lcm_dvdD1: + "lcm a b dvd c \ a dvd c" + using dvd_lcm1 [of a b] by (blast intro: dvd_trans) + +lemma lcm_dvdD2: + "lcm a b dvd c \ b dvd c" + using dvd_lcm2 [of a b] by (blast intro: dvd_trans) + lemma lcm_least: assumes "a dvd c" and "b dvd c" shows "lcm a b dvd c" @@ -350,16 +398,79 @@ end +class ring_gcd = comm_ring_1 + semiring_gcd + class semiring_Gcd = semiring_gcd + Gcd + assumes Gcd_dvd: "a \ A \ Gcd A dvd a" and Gcd_greatest: "(\b. b \ A \ a dvd b) \ a dvd Gcd A" and normalize_Gcd [simp]: "normalize (Gcd A) = Gcd A" + assumes dvd_Lcm: "a \ A \ a dvd Lcm A" + and Lcm_least: "(\b. b \ A \ b dvd a) \ Lcm A dvd a" + and normalize_Lcm [simp]: "normalize (Lcm A) = Lcm A" begin +lemma Lcm_Gcd: + "Lcm A = Gcd {b. \a\A. a dvd b}" + by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) + +lemma Gcd_Lcm: + "Gcd A = Lcm {b. \a\A. b dvd a}" + by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) + lemma Gcd_empty [simp]: "Gcd {} = 0" by (rule dvd_0_left, rule Gcd_greatest) simp +lemma Lcm_empty [simp]: + "Lcm {} = 1" + by (auto intro: associated_eqI Lcm_least) + +lemma Gcd_insert [simp]: + "Gcd (insert a A) = gcd a (Gcd A)" +proof - + have "Gcd (insert a A) dvd gcd a (Gcd A)" + by (auto intro: Gcd_dvd Gcd_greatest) + moreover have "gcd a (Gcd A) dvd Gcd (insert a A)" + proof (rule Gcd_greatest) + fix b + assume "b \ insert a A" + then show "gcd a (Gcd A) dvd b" + proof + assume "b = a" then show ?thesis by simp + next + assume "b \ A" + then have "Gcd A dvd b" by (rule Gcd_dvd) + moreover have "gcd a (Gcd A) dvd Gcd A" by simp + ultimately show ?thesis by (blast intro: dvd_trans) + qed + qed + ultimately show ?thesis + by (auto intro: associated_eqI) +qed + +lemma Lcm_insert [simp]: + "Lcm (insert a A) = lcm a (Lcm A)" +proof (rule sym) + have "lcm a (Lcm A) dvd Lcm (insert a A)" + by (auto intro: dvd_Lcm Lcm_least) + moreover have "Lcm (insert a A) dvd lcm a (Lcm A)" + proof (rule Lcm_least) + fix b + assume "b \ insert a A" + then show "b dvd lcm a (Lcm A)" + proof + assume "b = a" then show ?thesis by simp + next + assume "b \ A" + then have "b dvd Lcm A" by (rule dvd_Lcm) + moreover have "Lcm A dvd lcm a (Lcm A)" by simp + ultimately show ?thesis by (blast intro: dvd_trans) + qed + qed + ultimately show "lcm a (Lcm A) = Lcm (insert a A)" + by (rule associated_eqI) (simp_all add: lcm_eq_0_iff) +qed + lemma Gcd_0_iff [simp]: "Gcd A = 0 \ A \ {0}" (is "?P \ ?Q") proof @@ -384,147 +495,6 @@ then show ?P by simp qed -lemma unit_factor_Gcd: - "unit_factor (Gcd A) = (if \a\A. a = 0 then 0 else 1)" -proof (cases "Gcd A = 0") - case True then show ?thesis by auto -next - case False - from unit_factor_mult_normalize - have "unit_factor (Gcd A) * normalize (Gcd A) = Gcd A" . - then have "unit_factor (Gcd A) * Gcd A = Gcd A" by simp - then have "unit_factor (Gcd A) * Gcd A div Gcd A = Gcd A div Gcd A" by simp - with False have "unit_factor (Gcd A) = 1" by simp - with False show ?thesis by auto -qed - -lemma Gcd_UNIV [simp]: - "Gcd UNIV = 1" - by (rule associated_eqI) (auto intro: Gcd_dvd simp add: unit_factor_Gcd) - -lemma Gcd_eq_1_I: - assumes "is_unit a" and "a \ A" - shows "Gcd A = 1" -proof - - from assms have "is_unit (Gcd A)" - by (blast intro: Gcd_dvd dvd_unit_imp_unit) - then have "normalize (Gcd A) = 1" - by (rule is_unit_normalize) - then show ?thesis - by simp -qed - -lemma Gcd_insert [simp]: - "Gcd (insert a A) = gcd a (Gcd A)" -proof - - have "Gcd (insert a A) dvd gcd a (Gcd A)" - by (auto intro: Gcd_dvd Gcd_greatest) - moreover have "gcd a (Gcd A) dvd Gcd (insert a A)" - proof (rule Gcd_greatest) - fix b - assume "b \ insert a A" - then show "gcd a (Gcd A) dvd b" - proof - assume "b = a" then show ?thesis by simp - next - assume "b \ A" - then have "Gcd A dvd b" by (rule Gcd_dvd) - moreover have "gcd a (Gcd A) dvd Gcd A" by simp - ultimately show ?thesis by (blast intro: dvd_trans) - qed - qed - ultimately show ?thesis - by (auto intro: associated_eqI) -qed - -lemma dvd_Gcd: \ \FIXME remove\ - "\b\A. a dvd b \ a dvd Gcd A" - by (blast intro: Gcd_greatest) - -lemma Gcd_set [code_unfold]: - "Gcd (set as) = foldr gcd as 0" - by (induct as) simp_all - -lemma Gcd_image_normalize [simp]: - "Gcd (normalize ` A) = Gcd A" -proof - - have "Gcd (normalize ` A) dvd a" if "a \ A" for a - proof - - from that obtain B where "A = insert a B" by blast - moreover have " gcd (normalize a) (Gcd (normalize ` B)) dvd normalize a" - by (rule gcd_dvd1) - ultimately show "Gcd (normalize ` A) dvd a" - by simp - qed - then have "Gcd (normalize ` A) dvd Gcd A" and "Gcd A dvd Gcd (normalize ` A)" - by (auto intro!: Gcd_greatest intro: Gcd_dvd) - then show ?thesis - by (auto intro: associated_eqI) -qed - -end - -class semiring_Lcm = semiring_Gcd + - assumes Lcm_Gcd: "Lcm A = Gcd {b. \a\A. a dvd b}" -begin - -lemma dvd_Lcm: - assumes "a \ A" - shows "a dvd Lcm A" - using assms by (auto intro: Gcd_greatest simp add: Lcm_Gcd) - -lemma Lcm_least: - assumes "\b. b \ A \ b dvd a" - shows "Lcm A dvd a" - using assms by (auto intro: Gcd_dvd simp add: Lcm_Gcd) - -lemma normalize_Lcm [simp]: - "normalize (Lcm A) = Lcm A" - by (simp add: Lcm_Gcd) - -lemma unit_factor_Lcm: - "unit_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)" -proof (cases "Lcm A = 0") - case True then show ?thesis by simp -next - case False - with unit_factor_normalize have "unit_factor (normalize (Lcm A)) = 1" - by blast - with False show ?thesis - by simp -qed - -lemma Gcd_Lcm: - "Gcd A = Lcm {b. \a\A. b dvd a}" - by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) - -lemma Lcm_empty [simp]: - "Lcm {} = 1" - by (simp add: Lcm_Gcd) - -lemma Lcm_insert [simp]: - "Lcm (insert a A) = lcm a (Lcm A)" -proof (rule sym) - have "lcm a (Lcm A) dvd Lcm (insert a A)" - by (auto intro: dvd_Lcm Lcm_least) - moreover have "Lcm (insert a A) dvd lcm a (Lcm A)" - proof (rule Lcm_least) - fix b - assume "b \ insert a A" - then show "b dvd lcm a (Lcm A)" - proof - assume "b = a" then show ?thesis by simp - next - assume "b \ A" - then have "b dvd Lcm A" by (rule dvd_Lcm) - moreover have "Lcm A dvd lcm a (Lcm A)" by simp - ultimately show ?thesis by (blast intro: dvd_trans) - qed - qed - ultimately show "lcm a (Lcm A) = Lcm (insert a A)" - by (rule associated_eqI) (simp_all add: lcm_eq_0_iff) -qed - lemma Lcm_1_iff [simp]: "Lcm A = 1 \ (\a\A. is_unit a)" (is "?P \ ?Q") proof @@ -548,6 +518,44 @@ by simp qed +lemma unit_factor_Gcd: + "unit_factor (Gcd A) = (if \a\A. a = 0 then 0 else 1)" +proof (cases "Gcd A = 0") + case True then show ?thesis by auto +next + case False + from unit_factor_mult_normalize + have "unit_factor (Gcd A) * normalize (Gcd A) = Gcd A" . + then have "unit_factor (Gcd A) * Gcd A = Gcd A" by simp + then have "unit_factor (Gcd A) * Gcd A div Gcd A = Gcd A div Gcd A" by simp + with False have "unit_factor (Gcd A) = 1" by simp + with False show ?thesis by auto +qed + +lemma unit_factor_Lcm: + "unit_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)" +proof (cases "Lcm A = 0") + case True then show ?thesis by simp +next + case False + with unit_factor_normalize have "unit_factor (normalize (Lcm A)) = 1" + by blast + with False show ?thesis + by simp +qed + +lemma Gcd_eq_1_I: + assumes "is_unit a" and "a \ A" + shows "Gcd A = 1" +proof - + from assms have "is_unit (Gcd A)" + by (blast intro: Gcd_dvd dvd_unit_imp_unit) + then have "normalize (Gcd A) = 1" + by (rule is_unit_normalize) + then show ?thesis + by simp +qed + lemma Lcm_eq_0_I: assumes "0 \ A" shows "Lcm A = 0" @@ -558,6 +566,10 @@ by simp qed +lemma Gcd_UNIV [simp]: + "Gcd UNIV = 1" + using dvd_refl by (rule Gcd_eq_1_I) simp + lemma Lcm_UNIV [simp]: "Lcm UNIV = 0" by (rule Lcm_eq_0_I) simp @@ -573,25 +585,57 @@ (auto simp add: lcm_eq_0_iff) qed +lemma Gcd_set [code_unfold]: + "Gcd (set as) = foldr gcd as 0" + by (induct as) simp_all + lemma Lcm_set [code_unfold]: "Lcm (set as) = foldr lcm as 1" by (induct as) simp_all - + +lemma Gcd_image_normalize [simp]: + "Gcd (normalize ` A) = Gcd A" +proof - + have "Gcd (normalize ` A) dvd a" if "a \ A" for a + proof - + from that obtain B where "A = insert a B" by blast + moreover have "gcd (normalize a) (Gcd (normalize ` B)) dvd normalize a" + by (rule gcd_dvd1) + ultimately show "Gcd (normalize ` A) dvd a" + by simp + qed + then have "Gcd (normalize ` A) dvd Gcd A" and "Gcd A dvd Gcd (normalize ` A)" + by (auto intro!: Gcd_greatest intro: Gcd_dvd) + then show ?thesis + 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 -class ring_gcd = comm_ring_1 + semiring_gcd +subsection \GCD and LCM on @{typ nat} and @{typ int}\ instantiation nat :: gcd begin -fun - gcd_nat :: "nat \ nat \ nat" -where - "gcd_nat x y = - (if y = 0 then x else gcd y (x mod y))" +fun gcd_nat :: "nat \ nat \ nat" +where "gcd_nat x y = + (if y = 0 then x else gcd y (x mod y))" -definition - lcm_nat :: "nat \ nat \ nat" +definition lcm_nat :: "nat \ nat \ nat" where "lcm_nat x y = x * y div (gcd x y)" @@ -602,22 +646,17 @@ instantiation int :: gcd begin -definition - gcd_int :: "int \ int \ int" -where - "gcd_int x y = int (gcd (nat \x\) (nat \y\))" +definition gcd_int :: "int \ int \ int" + where "gcd_int x y = int (gcd (nat \x\) (nat \y\))" -definition - lcm_int :: "int \ int \ int" -where - "lcm_int x y = int (lcm (nat \x\) (nat \y\))" +definition lcm_int :: "int \ int \ int" + where "lcm_int x y = int (lcm (nat \x\) (nat \y\))" instance .. end - -subsection \Transfer setup\ +text \Transfer setup\ lemma transfer_nat_int_gcd: "(x::int) >= 0 \ y >= 0 \ gcd (nat x) (nat y) = nat (gcd x y)" @@ -646,10 +685,6 @@ declare transfer_morphism_int_nat[transfer add return: transfer_int_nat_gcd transfer_int_nat_gcd_closures] - -subsection \GCD properties\ - -(* was gcd_induct *) lemma gcd_nat_induct: fixes m n :: nat assumes "\m. P m 0" @@ -662,30 +697,30 @@ (* 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) lemma gcd_neg2_int [simp]: "gcd (x::int) (-y) = gcd x y" by (simp add: gcd_int_def) -lemma gcd_neg_numeral_1_int [simp]: - "gcd (- numeral n :: int) x = gcd (numeral n) x" - by (fact gcd_neg1_int) - -lemma gcd_neg_numeral_2_int [simp]: - "gcd x (- numeral n :: int) = gcd x (numeral n)" - by (fact gcd_neg2_int) - -lemma abs_gcd_int[simp]: "\gcd (x::int) y\ = gcd x y" +lemma abs_gcd_int [simp]: "\gcd (x::int) y\ = gcd x y" by(simp add: gcd_int_def) lemma gcd_abs_int: "gcd (x::int) y = gcd \x\ \y\" by (simp add: gcd_int_def) -lemma gcd_abs1_int[simp]: "gcd \x\ (y::int) = gcd x y" +lemma gcd_abs1_int [simp]: "gcd \x\ (y::int) = gcd x y" by (metis abs_idempotent gcd_abs_int) -lemma gcd_abs2_int[simp]: "gcd x \y::int\ = gcd x y" +lemma gcd_abs2_int [simp]: "gcd x \y::int\ = gcd x y" by (metis abs_idempotent gcd_abs_int) lemma gcd_cases_int: @@ -712,10 +747,10 @@ lemma abs_lcm_int [simp]: "\lcm i j::int\ = lcm i j" by (simp add:lcm_int_def) -lemma lcm_abs1_int[simp]: "lcm \x\ (y::int) = lcm x y" +lemma lcm_abs1_int [simp]: "lcm \x\ (y::int) = lcm x y" by (metis abs_idempotent lcm_int_def) -lemma lcm_abs2_int[simp]: "lcm x \y::int\ = lcm x y" +lemma lcm_abs2_int [simp]: "lcm x \y::int\ = lcm x y" by (metis abs_idempotent lcm_int_def) lemma lcm_cases_int: @@ -730,11 +765,9 @@ lemma lcm_ge_0_int [simp]: "lcm (x::int) y >= 0" by (simp add: lcm_int_def) -(* was gcd_0, etc. *) lemma gcd_0_nat: "gcd (x::nat) 0 = x" by simp -(* was igcd_0, etc. *) lemma gcd_0_int [simp]: "gcd (x::int) 0 = \x\" by (unfold gcd_int_def, auto) @@ -749,7 +782,7 @@ (* weaker, but useful for the simplifier *) -lemma gcd_non_0_nat: "y ~= (0::nat) \ gcd (x::nat) y = gcd y (x mod y)" +lemma gcd_non_0_nat: "y \ (0::nat) \ gcd (x::nat) y = gcd y (x mod y)" by simp lemma gcd_1_nat [simp]: "gcd (m::nat) 1 = 1" @@ -798,18 +831,6 @@ by standard (simp_all add: dvd_int_unfold_dvd_nat gcd_int_def lcm_int_def zdiv_int nat_abs_mult_distrib [symmetric] lcm_gcd gcd_greatest) -lemma dvd_gcd_D1_nat: "k dvd gcd m n \ (k::nat) dvd m" - by (metis gcd_dvd1 dvd_trans) - -lemma dvd_gcd_D2_nat: "k dvd gcd m n \ (k::nat) dvd n" - by (metis gcd_dvd2 dvd_trans) - -lemma dvd_gcd_D1_int: "i dvd gcd m n \ (i::int) dvd m" - by (metis gcd_dvd1 dvd_trans) - -lemma dvd_gcd_D2_int: "i dvd gcd m n \ (i::int) dvd n" - by (metis gcd_dvd2 dvd_trans) - lemma gcd_le1_nat [simp]: "a \ 0 \ gcd (a::nat) b \ a" by (rule dvd_imp_le, auto) @@ -822,27 +843,11 @@ lemma gcd_le2_int [simp]: "b > 0 \ gcd (a::int) b \ b" by (rule zdvd_imp_le, auto) -lemma gcd_greatest_iff_nat: - "(k dvd gcd (m::nat) n) = (k dvd m & k dvd n)" - by (fact gcd_greatest_iff) - -lemma gcd_greatest_iff_int: - "((k::int) dvd gcd m n) = (k dvd m & k dvd n)" - by (fact gcd_greatest_iff) - -lemma gcd_zero_nat: - "(gcd (m::nat) n = 0) = (m = 0 & n = 0)" - by (fact gcd_eq_0_iff) - -lemma gcd_zero_int [simp]: - "(gcd (m::int) n = 0) = (m = 0 & n = 0)" - by (fact gcd_eq_0_iff) - lemma gcd_pos_nat [simp]: "(gcd (m::nat) n > 0) = (m ~= 0 | n ~= 0)" - by (insert gcd_zero_nat [of m n], arith) + by (insert gcd_eq_0_iff [of m n], arith) lemma gcd_pos_int [simp]: "(gcd (m::int) n > 0) = (m ~= 0 | n ~= 0)" - by (insert gcd_zero_int [of m n], insert gcd_ge_0_int [of m n], arith) + by (insert gcd_eq_0_iff [of m n], insert gcd_ge_0_int [of m n], arith) lemma gcd_unique_nat: "(d::nat) dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" @@ -862,31 +867,14 @@ done interpretation gcd_nat: - semilattice_neutr_order gcd "0::nat" Rings.dvd "(\m n. m dvd n \ \ n dvd m)" - by standard (auto simp add: gcd_unique_nat [symmetric] intro: dvd.antisym dvd_trans) - -lemmas gcd_assoc_nat = gcd.assoc [where ?'a = nat] -lemmas gcd_commute_nat = gcd.commute [where ?'a = nat] -lemmas gcd_left_commute_nat = gcd.left_commute [where ?'a = nat] -lemmas gcd_assoc_int = gcd.assoc [where ?'a = int] -lemmas gcd_commute_int = gcd.commute [where ?'a = int] -lemmas gcd_left_commute_int = gcd.left_commute [where ?'a = int] - -lemmas gcd_ac_nat = gcd_assoc_nat gcd_commute_nat gcd_left_commute_nat - -lemmas gcd_ac_int = gcd_assoc_int gcd_commute_int gcd_left_commute_int - -lemma gcd_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \ gcd x y = x" - by (fact gcd_nat.absorb1) - -lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \ gcd x y = y" - by (fact gcd_nat.absorb2) + semilattice_neutr_order gcd "0::nat" Rings.dvd "\m n. m dvd n \ m \ n" + by standard (auto simp add: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans) lemma gcd_proj1_if_dvd_int [simp]: "x dvd y \ gcd (x::int) y = \x\" by (metis abs_dvd_iff gcd_0_left_int gcd_abs_int gcd_unique_int) lemma gcd_proj2_if_dvd_int [simp]: "y dvd x \ gcd (x::int) y = \y\" - by (metis gcd_proj1_if_dvd_int gcd_commute_int) + by (metis gcd_proj1_if_dvd_int gcd.commute) text \ \medskip Multiplication laws @@ -926,21 +914,10 @@ ultimately show ?thesis by simp qed -end - -lemmas coprime_dvd_mult_nat = coprime_dvd_mult [where ?'a = nat] -lemmas coprime_dvd_mult_int = coprime_dvd_mult [where ?'a = int] - -lemma coprime_dvd_mult_iff_nat: "coprime (k::nat) n \ - (k dvd m * n) = (k dvd m)" - by (auto intro: coprime_dvd_mult_nat) - -lemma coprime_dvd_mult_iff_int: "coprime (k::int) n \ - (k dvd m * n) = (k dvd m)" - by (auto intro: coprime_dvd_mult_int) - -context semiring_gcd -begin +lemma coprime_dvd_mult_iff: + assumes "coprime a c" + shows "a dvd b * c \ a dvd b" + using assms by (auto intro: coprime_dvd_mult) lemma gcd_mult_cancel: "coprime c b \ gcd (c * a) b = gcd a b" @@ -951,65 +928,79 @@ apply (simp_all add: ac_simps) done -end - -lemmas gcd_mult_cancel_nat = gcd_mult_cancel [where ?'a = nat] -lemmas gcd_mult_cancel_int = gcd_mult_cancel [where ?'a = int] - -lemma coprime_crossproduct_nat: - fixes a b c d :: nat +lemma coprime_crossproduct: + fixes a b c d assumes "coprime a d" and "coprime b c" - shows "a * c = b * d \ a = b \ c = d" (is "?lhs \ ?rhs") + shows "normalize a * normalize c = normalize b * normalize d + \ normalize a = normalize b \ normalize c = normalize d" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by simp next assume ?lhs - from \?lhs\ have "a dvd b * d" by (auto intro: dvdI dest: sym) - with \coprime a d\ have "a dvd b" by (simp add: coprime_dvd_mult_iff_nat) - from \?lhs\ have "b dvd a * c" by (auto intro: dvdI dest: sym) - with \coprime b c\ have "b dvd a" by (simp add: coprime_dvd_mult_iff_nat) - from \?lhs\ have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult.commute) - with \coprime b c\ have "c dvd d" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat) - from \?lhs\ have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult.commute) - with \coprime a d\ have "d dvd c" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat) - from \a dvd b\ \b dvd a\ have "a = b" by (rule Nat.dvd.antisym) - moreover from \c dvd d\ \d dvd c\ have "c = d" by (rule Nat.dvd.antisym) + from \?lhs\ have "normalize a dvd normalize b * normalize d" + by (auto intro: dvdI dest: sym) + with \coprime a d\ have "a dvd b" + by (simp add: coprime_dvd_mult_iff normalize_mult [symmetric]) + from \?lhs\ have "normalize b dvd normalize a * normalize c" + by (auto intro: dvdI dest: sym) + with \coprime b c\ have "b dvd a" + by (simp add: coprime_dvd_mult_iff normalize_mult [symmetric]) + from \?lhs\ have "normalize c dvd normalize d * normalize b" + by (auto intro: dvdI dest: sym simp add: mult.commute) + with \coprime b c\ have "c dvd d" + by (simp add: coprime_dvd_mult_iff gcd.commute normalize_mult [symmetric]) + from \?lhs\ have "normalize d dvd normalize c * normalize a" + by (auto intro: dvdI dest: sym simp add: mult.commute) + with \coprime a d\ have "d dvd c" + by (simp add: coprime_dvd_mult_iff gcd.commute normalize_mult [symmetric]) + from \a dvd b\ \b dvd a\ have "normalize a = normalize b" + by (rule associatedI) + moreover from \c dvd d\ \d dvd c\ have "normalize c = normalize d" + by (rule associatedI) ultimately show ?rhs .. qed +end + +lemma coprime_crossproduct_nat: + fixes a b c d :: nat + assumes "coprime a d" and "coprime b c" + shows "a * c = b * d \ a = b \ c = d" + using assms coprime_crossproduct [of a d b c] by simp + lemma coprime_crossproduct_int: fixes a b c d :: int assumes "coprime a d" and "coprime b c" shows "\a\ * \c\ = \b\ * \d\ \ \a\ = \b\ \ \c\ = \d\" - using assms by (intro coprime_crossproduct_nat [transferred]) auto + using assms coprime_crossproduct [of a d b c] by simp text \\medskip Addition laws\ lemma gcd_add1_nat [simp]: "gcd ((m::nat) + n) n = gcd m n" apply (case_tac "n = 0") apply (simp_all add: gcd_non_0_nat) -done + done lemma gcd_add2_nat [simp]: "gcd (m::nat) (m + n) = gcd m n" - apply (subst (1 2) gcd_commute_nat) + apply (subst (1 2) gcd.commute) apply (subst add.commute) apply simp -done + done (* to do: add the other variations? *) lemma gcd_diff1_nat: "(m::nat) >= n \ gcd (m - n) n = gcd m n" - by (subst gcd_add1_nat [symmetric], auto) + by (subst gcd_add1_nat [symmetric]) auto lemma gcd_diff2_nat: "(n::nat) >= m \ gcd (n - m) n = gcd m n" - apply (subst gcd_commute_nat) + apply (subst gcd.commute) apply (subst gcd_diff1_nat [symmetric]) apply auto - apply (subst gcd_commute_nat) + apply (subst gcd.commute) apply (subst gcd_diff1_nat) apply assumption - apply (rule gcd_commute_nat) -done + apply (rule gcd.commute) + done lemma gcd_non_0_int: "(y::int) > 0 \ gcd x y = gcd y (x mod y)" apply (frule_tac b = y and a = x in pos_mod_sign) @@ -1017,10 +1008,10 @@ apply (auto simp add: gcd_non_0_nat nat_mod_distrib [symmetric] zmod_zminus1_eq_if) apply (frule_tac a = x in pos_mod_bound) - apply (subst (1 2) gcd_commute_nat) + apply (subst (1 2) gcd.commute) apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2_nat nat_le_eq_zle) -done + done lemma gcd_red_int: "gcd (x::int) y = gcd y (x mod y)" apply (case_tac "y = 0") @@ -1035,13 +1026,13 @@ by (metis gcd_red_int mod_add_self1 add.commute) lemma gcd_add2_int [simp]: "gcd m ((m::int) + n) = gcd m n" -by (metis gcd_add1_int gcd_commute_int add.commute) +by (metis gcd_add1_int gcd.commute add.commute) lemma gcd_add_mult_nat: "gcd (m::nat) (k * m + n) = gcd m n" -by (metis mod_mult_self3 gcd_commute_nat gcd_red_nat) +by (metis mod_mult_self3 gcd.commute gcd_red_nat) lemma gcd_add_mult_int: "gcd (m::int) (k * m + n) = gcd m n" -by (metis gcd_commute_int gcd_red_int mod_mult_self1 add.commute) +by (metis gcd.commute gcd_red_int mod_mult_self1 add.commute) (* to do: differences, and all variations of addition rules @@ -1052,42 +1043,48 @@ (* to do: add the three variations of these, and for ints? *) -lemma finite_divisors_nat[simp]: - assumes "(m::nat) ~= 0" shows "finite{d. d dvd m}" +lemma finite_divisors_nat [simp]: -- \FIXME move\ + fixes m :: nat + assumes "m > 0" + shows "finite {d. d dvd m}" proof- - have "finite{d. d <= m}" - by (blast intro: bounded_nat_set_is_finite) - from finite_subset[OF _ this] show ?thesis using assms - by (metis Collect_mono dvd_imp_le neq0_conv) + from assms have "{d. d dvd m} \ {d. d \ m}" + by (auto dest: dvd_imp_le) + then show ?thesis + using finite_Collect_le_nat by (rule finite_subset) qed -lemma finite_divisors_int[simp]: - assumes "(i::int) ~= 0" shows "finite{d. d dvd i}" -proof- - have "{d. \d\ <= \i\} = {- \i\ .. \i\}" by(auto simp:abs_if) - hence "finite {d. \d\ <= \i\}" by simp - from finite_subset[OF _ this] show ?thesis using assms +lemma finite_divisors_int [simp]: + fixes i :: int + assumes "i \ 0" + shows "finite {d. d dvd i}" +proof - + have "{d. \d\ \ \i\} = {- \i\..\i\}" + by (auto simp: abs_if) + then have "finite {d. \d\ <= \i\}" + by simp + from finite_subset [OF _ this] show ?thesis using assms by (simp add: dvd_imp_le_int subset_iff) qed -lemma Max_divisors_self_nat[simp]: "n\0 \ Max{d::nat. d dvd n} = n" +lemma Max_divisors_self_nat [simp]: "n\0 \ Max{d::nat. d dvd n} = n" apply(rule antisym) apply (fastforce intro: Max_le_iff[THEN iffD2] simp: dvd_imp_le) apply simp done -lemma Max_divisors_self_int[simp]: "n\0 \ Max{d::int. d dvd n} = \n\" +lemma Max_divisors_self_int [simp]: "n\0 \ Max{d::int. d dvd n} = \n\" apply(rule antisym) apply(rule Max_le_iff [THEN iffD2]) apply (auto intro: abs_le_D1 dvd_imp_le_int) done lemma gcd_is_Max_divisors_nat: - "m ~= 0 \ n ~= 0 \ gcd (m::nat) n = (Max {d. d dvd m & d dvd n})" + "m > 0 \ n > 0 \ gcd (m::nat) n = Max {d. d dvd m \ d dvd n}" apply(rule Max_eqI[THEN sym]) apply (metis finite_Collect_conjI finite_divisors_nat) apply simp - apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le gcd_greatest_iff_nat gcd_pos_nat) + apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le gcd_greatest_iff gcd_pos_nat) apply simp done @@ -1096,7 +1093,7 @@ apply(rule Max_eqI[THEN sym]) apply (metis finite_Collect_conjI finite_divisors_int) apply simp - apply (metis gcd_greatest_iff_int gcd_pos_int zdvd_imp_le) + apply (metis gcd_greatest_iff gcd_pos_int zdvd_imp_le) apply simp done @@ -1134,20 +1131,32 @@ ultimately show ?thesis using dvd_times_left_cancel_iff [of "gcd a b" _ 1] by simp qed +lemma coprime: + "coprime a b \ (\d. d dvd a \ d dvd b \ is_unit d)" (is "?P \ ?Q") +proof + assume ?P then show ?Q by auto +next + assume ?Q + then have "is_unit (gcd a b) \ gcd a b dvd a \ gcd a b dvd b" + by blast + then have "is_unit (gcd a b)" + by simp + then show ?P + by simp +qed + end -lemmas div_gcd_coprime_nat = div_gcd_coprime [where ?'a = nat] -lemmas div_gcd_coprime_int = div_gcd_coprime [where ?'a = int] - -lemma coprime_nat: "coprime (a::nat) b \ (\d. d dvd a \ d dvd b \ d = 1)" - using gcd_unique_nat[of 1 a b, simplified] by auto +lemma coprime_nat: + "coprime (a::nat) b \ (\d. d dvd a \ d dvd b \ d = 1)" + using coprime [of a b] by simp lemma coprime_Suc_0_nat: - "coprime (a::nat) b \ (\d. d dvd a \ d dvd b \ d = Suc 0)" + "coprime (a::nat) b \ (\d. d dvd a \ d dvd b \ d = Suc 0)" using coprime_nat by simp -lemma coprime_int: "coprime (a::int) b \ - (\d. d >= 0 \ d dvd a \ d dvd b \ d = 1)" +lemma coprime_int: + "coprime (a::int) b \ (\d. d \ 0 \ d dvd a \ d dvd b \ d = 1)" using gcd_unique_int [of 1 a b] apply clarsimp apply (erule subst) @@ -1165,7 +1174,7 @@ apply (erule ssubst) apply (subgoal_tac "b' = b div gcd a b") apply (erule ssubst) - apply (rule div_gcd_coprime_nat) + apply (rule div_gcd_coprime) using z apply force apply (subst (1) b) using z apply force @@ -1177,12 +1186,11 @@ assumes z: "gcd (a::int) b \ 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b" shows "coprime a' b'" - apply (subgoal_tac "a' = a div gcd a b") apply (erule ssubst) apply (subgoal_tac "b' = b div gcd a b") apply (erule ssubst) - apply (rule div_gcd_coprime_int) + apply (rule div_gcd_coprime) using z apply force apply (subst (1) b) using z apply force @@ -1204,9 +1212,6 @@ end -lemmas coprime_mult_nat = coprime_mult [where ?'a = nat] -lemmas coprime_mult_int = coprime_mult [where ?'a = int] - lemma coprime_lmult_nat: assumes dab: "coprime (d::nat) (a * b)" shows "coprime d a" proof - @@ -1246,13 +1251,13 @@ lemma coprime_mul_eq_nat: "coprime (d::nat) (a * b) \ coprime d a \ coprime d b" using coprime_rmult_nat[of d a b] coprime_lmult_nat[of d a b] - coprime_mult_nat[of d a b] + coprime_mult [of d a b] by blast lemma coprime_mul_eq_int: "coprime (d::int) (a * b) \ coprime d a \ coprime d b" using coprime_rmult_int[of d a b] coprime_lmult_int[of d a b] - coprime_mult_int[of d a b] + coprime_mult [of d a b] by blast lemma coprime_power_int: @@ -1268,7 +1273,7 @@ shows "\a' b'. a = a' * gcd a b \ b = b' * gcd a b \ coprime a' b'" apply (rule_tac x = "a div gcd a b" in exI) apply (rule_tac x = "b div gcd a b" in exI) - using nz apply (auto simp add: div_gcd_coprime_nat dvd_div_mult) + using nz apply (auto simp add: div_gcd_coprime dvd_div_mult) done lemma gcd_coprime_exists_int: @@ -1276,14 +1281,14 @@ shows "\a' b'. a = a' * gcd a b \ b = b' * gcd a b \ coprime a' b'" apply (rule_tac x = "a div gcd a b" in exI) apply (rule_tac x = "b div gcd a b" in exI) - using nz apply (auto simp add: div_gcd_coprime_int) + using nz apply (auto simp add: div_gcd_coprime) done lemma coprime_exp_nat: "coprime (d::nat) a \ coprime d (a^n)" - by (induct n) (simp_all add: coprime_mult_nat) + by (induct n) (simp_all add: coprime_mult) lemma coprime_exp_int: "coprime (d::int) a \ coprime d (a^n)" - by (induct n) (simp_all add: coprime_mult_int) + by (induct n) (simp_all add: coprime_mult) context semiring_gcd begin @@ -1303,12 +1308,6 @@ end -lemma coprime_exp2_nat [intro]: "coprime (a::nat) b \ coprime (a^n) (b^m)" - by (fact coprime_exp2) - -lemma coprime_exp2_int [intro]: "coprime (a::int) b \ coprime (a^n) (b^m)" - by (fact coprime_exp2) - lemma gcd_exp_nat: "gcd ((a :: nat) ^ n) (b ^ n) = gcd a b ^ n" proof (cases "a = 0 \ b = 0") @@ -1352,7 +1351,7 @@ from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult.assoc) with z have th_1: "a' dvd b' * c" by auto - from coprime_dvd_mult_nat[OF ab'(3)] th_1 + from coprime_dvd_mult [OF ab'(3)] th_1 have thc: "a' dvd c" by (subst (asm) mult.commute, blast) from ab' have "a = ?g*a'" by algebra with thb thc have ?thesis by blast } @@ -1376,7 +1375,7 @@ from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto hence "?g*a' dvd ?g * (b' * c)" by (simp add: ac_simps) with z have th_1: "a' dvd b' * c" by auto - from coprime_dvd_mult_int[OF ab'(3)] th_1 + from coprime_dvd_mult [OF ab'(3)] th_1 have thc: "a' dvd c" by (subst (asm) mult.commute, blast) from ab' have "a = ?g*a'" by algebra with thb thc have ?thesis by blast } @@ -1405,7 +1404,7 @@ have "a' dvd a'^n" by (simp add: m) with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp hence th1: "a' dvd b'^m * b'" by (simp add: m mult.commute) - from coprime_dvd_mult_nat[OF coprime_exp_nat [OF ab'(3), of m]] th1 + from coprime_dvd_mult [OF coprime_exp_nat [OF ab'(3), of m]] th1 have "a' dvd b'" by (subst (asm) mult.commute, blast) hence "a'*?g dvd b'*?g" by simp with ab'(1,2) have ?thesis by simp } @@ -1434,42 +1433,40 @@ with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp hence th1: "a' dvd b'^m * b'" by (simp add: m mult.commute) - from coprime_dvd_mult_int[OF coprime_exp_int [OF ab'(3), of m]] th1 + from coprime_dvd_mult [OF coprime_exp_int [OF ab'(3), of m]] th1 have "a' dvd b'" by (subst (asm) mult.commute, blast) hence "a'*?g dvd b'*?g" by simp with ab'(1,2) have ?thesis by simp } ultimately show ?thesis by blast qed -lemma pow_divides_eq_nat [simp]: "n ~= 0 \ ((a::nat)^n dvd b^n) = (a dvd b)" +lemma pow_divides_eq_nat [simp]: + "n > 0 \ (a::nat) ^ n dvd b ^ n \ a dvd b" by (auto intro: pow_divides_pow_nat dvd_power_same) -lemma pow_divides_eq_int [simp]: "n ~= 0 \ ((a::int)^n dvd b^n) = (a dvd b)" +lemma pow_divides_eq_int [simp]: + "n ~= 0 \ (a::int) ^ n dvd b ^ n \ a dvd b" by (auto intro: pow_divides_pow_int dvd_power_same) -lemma divides_mult_nat: - assumes mr: "(m::nat) dvd r" and nr: "n dvd r" and mn:"coprime m n" - shows "m * n dvd r" +context semiring_gcd +begin + +lemma divides_mult: + assumes "a dvd c" and nr: "b dvd c" and "coprime a b" + shows "a * b dvd c" proof- - from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'" - unfolding dvd_def by blast - from mr n' have "m dvd n'*n" by (simp add: mult.commute) - hence "m dvd n'" using coprime_dvd_mult_iff_nat[OF mn] by simp - then obtain k where k: "n' = m*k" unfolding dvd_def by blast - from n' k show ?thesis unfolding dvd_def by auto + from \b dvd c\ obtain b' where"c = b * b'" .. + with \a dvd c\ have "a dvd b' * b" + by (simp add: ac_simps) + with \coprime a b\ have "a dvd b'" + by (simp add: coprime_dvd_mult_iff) + then obtain a' where "b' = a * a'" .. + with \c = b * b'\ have "c = (a * b) * a'" + by (simp add: ac_simps) + then show ?thesis .. qed -lemma divides_mult_int: - assumes mr: "(m::int) dvd r" and nr: "n dvd r" and mn:"coprime m n" - shows "m * n dvd r" -proof- - from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'" - unfolding dvd_def by blast - from mr n' have "m dvd n'*n" by (simp add: mult.commute) - hence "m dvd n'" using coprime_dvd_mult_iff_int[OF mn] by simp - then obtain k where k: "n' = m*k" unfolding dvd_def by blast - from n' k show ?thesis unfolding dvd_def by auto -qed +end lemma coprime_plus_one_nat [simp]: "coprime ((n::nat) + 1) n" by (simp add: gcd.commute del: One_nat_def) @@ -1482,29 +1479,27 @@ lemma coprime_minus_one_nat: "(n::nat) \ 0 \ coprime (n - 1) n" using coprime_plus_one_nat [of "n - 1"] - gcd_commute_nat [of "n - 1" n] by auto + gcd.commute [of "n - 1" n] by auto lemma coprime_minus_one_int: "coprime ((n::int) - 1) n" using coprime_plus_one_int [of "n - 1"] - gcd_commute_int [of "n - 1" n] by auto + gcd.commute [of "n - 1" n] by auto -lemma setprod_coprime_nat [rule_format]: - "(ALL i: A. coprime (f i) (x::nat)) --> coprime (\i\A. f i) x" - apply (case_tac "finite A") - apply (induct set: finite) - apply (auto simp add: gcd_mult_cancel_nat) -done +lemma setprod_coprime_nat: + fixes x :: nat + shows "(\i. i \ A \ coprime (f i) x) \ coprime (\i\A. f i) x" + by (induct A rule: infinite_finite_induct) + (auto simp add: gcd_mult_cancel One_nat_def [symmetric] simp del: One_nat_def) -lemma setprod_coprime_int [rule_format]: - "(ALL i: A. coprime (f i) (x::int)) --> coprime (\i\A. f i) x" - apply (case_tac "finite A") - apply (induct set: finite) - apply (auto simp add: gcd_mult_cancel_int) -done +lemma setprod_coprime_int: + fixes x :: int + shows "(\i. i \ A \ coprime (f i) x) \ coprime (\i\A. f i) x" + by (induct A rule: infinite_finite_induct) + (auto simp add: gcd_mult_cancel) lemma coprime_common_divisor_nat: "coprime (a::nat) b \ x dvd a \ x dvd b \ x = 1" - by (metis gcd_greatest_iff_nat nat_dvd_1_iff_1) + by (metis gcd_greatest_iff nat_dvd_1_iff_1) lemma coprime_common_divisor_int: "coprime (a::int) b \ x dvd a \ x dvd b \ \x\ = 1" @@ -1515,10 +1510,10 @@ by (meson coprime_int dvd_trans gcd_dvd1 gcd_dvd2 gcd_ge_0_int) lemma invertible_coprime_nat: "(x::nat) * y mod m = 1 \ coprime x m" -by (metis coprime_lmult_nat gcd_1_nat gcd_commute_nat gcd_red_nat) +by (metis coprime_lmult_nat gcd_1_nat gcd.commute gcd_red_nat) lemma invertible_coprime_int: "(x::int) * y mod m = 1 \ coprime x m" -by (metis coprime_lmult_int gcd_1_int gcd_commute_int gcd_red_int) +by (metis coprime_lmult_int gcd_1_int gcd.commute gcd_red_int) subsection \Bezout's theorem\ @@ -1761,11 +1756,10 @@ qed -subsection \LCM properties\ +subsection \LCM properties on @{typ nat} and @{typ int}\ lemma lcm_altdef_int [code]: "lcm (a::int) b = \a\ * \b\ div gcd a b" - by (simp add: lcm_int_def lcm_nat_def zdiv_int - of_nat_mult gcd_int_def) + by (simp add: lcm_int_def lcm_nat_def zdiv_int gcd_int_def) lemma prod_gcd_lcm_nat: "(m::nat) * n = gcd m n * lcm m n" unfolding lcm_nat_def @@ -1773,24 +1767,12 @@ lemma prod_gcd_lcm_int: "\m::int\ * \n\ = gcd m n * lcm m n" unfolding lcm_int_def gcd_int_def - apply (subst int_mult [symmetric]) + apply (subst of_nat_mult [symmetric]) apply (subst prod_gcd_lcm_nat [symmetric]) apply (subst nat_abs_mult_distrib [symmetric]) apply (simp, simp add: abs_mult) done -lemma lcm_0_nat [simp]: "lcm (m::nat) 0 = 0" - unfolding lcm_nat_def by simp - -lemma lcm_0_int [simp]: "lcm (m::int) 0 = 0" - unfolding lcm_int_def by simp - -lemma lcm_0_left_nat [simp]: "lcm (0::nat) n = 0" - unfolding lcm_nat_def by simp - -lemma lcm_0_left_int [simp]: "lcm (0::int) n = 0" - unfolding lcm_int_def by simp - lemma lcm_pos_nat: "(m::nat) > 0 \ n>0 \ lcm m n > 0" by (metis gr0I mult_is_0 prod_gcd_lcm_nat) @@ -1800,70 +1782,21 @@ apply (subst lcm_abs_int) apply (rule lcm_pos_nat [transferred]) apply auto -done + done -lemma dvd_pos_nat: +lemma dvd_pos_nat: -- \FIXME move\ fixes n m :: nat assumes "n > 0" and "m dvd n" shows "m > 0" -using assms by (cases m) auto - -lemma lcm_least_nat: - assumes "(m::nat) dvd k" and "n dvd k" - shows "lcm m n dvd k" - using assms by (rule lcm_least) - -lemma lcm_least_int: - "(m::int) dvd k \ n dvd k \ lcm m n dvd k" - by (rule lcm_least) - -lemma lcm_dvd1_nat: "(m::nat) dvd lcm m n" - by (fact dvd_lcm1) - -lemma lcm_dvd1_int: "(m::int) dvd lcm m n" - by (fact dvd_lcm1) - -lemma lcm_dvd2_nat: "(n::nat) dvd lcm m n" - by (fact dvd_lcm2) - -lemma lcm_dvd2_int: "(n::int) dvd lcm m n" - by (fact dvd_lcm2) - -lemma dvd_lcm_I1_nat[simp]: "(k::nat) dvd m \ k dvd lcm m n" -by(metis lcm_dvd1_nat dvd_trans) - -lemma dvd_lcm_I2_nat[simp]: "(k::nat) dvd n \ k dvd lcm m n" -by(metis lcm_dvd2_nat dvd_trans) - -lemma dvd_lcm_I1_int[simp]: "(i::int) dvd m \ i dvd lcm m n" -by(metis lcm_dvd1_int dvd_trans) - -lemma dvd_lcm_I2_int[simp]: "(i::int) dvd n \ i dvd lcm m n" -by(metis lcm_dvd2_int dvd_trans) + using assms by (cases m) auto lemma lcm_unique_nat: "(a::nat) dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" - by (auto intro: dvd_antisym lcm_least_nat lcm_dvd1_nat lcm_dvd2_nat) + by (auto intro: dvd_antisym lcm_least) lemma lcm_unique_int: "d >= 0 \ (a::int) dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" - using lcm_least_int zdvd_antisym_nonneg by auto - -interpretation lcm_nat: abel_semigroup "lcm :: nat \ nat \ nat" - + lcm_nat: semilattice_neutr "lcm :: nat \ nat \ nat" 1 - by standard (simp_all del: One_nat_def) - -interpretation lcm_int: abel_semigroup "lcm :: int \ int \ int" .. - -lemmas lcm_assoc_nat = lcm.assoc [where ?'a = nat] -lemmas lcm_commute_nat = lcm.commute [where ?'a = nat] -lemmas lcm_left_commute_nat = lcm.left_commute [where ?'a = nat] -lemmas lcm_assoc_int = lcm.assoc [where ?'a = int] -lemmas lcm_commute_int = lcm.commute [where ?'a = int] -lemmas lcm_left_commute_int = lcm.left_commute [where ?'a = int] - -lemmas lcm_ac_nat = lcm_assoc_nat lcm_commute_nat lcm_left_commute_nat -lemmas lcm_ac_int = lcm_assoc_int lcm_commute_int lcm_left_commute_int + using lcm_least zdvd_antisym_nonneg by auto lemma lcm_proj2_if_dvd_nat [simp]: "(x::nat) dvd y \ lcm x y = y" apply (rule sym) @@ -1878,21 +1811,21 @@ done lemma lcm_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \ lcm y x = y" -by (subst lcm_commute_nat, erule lcm_proj2_if_dvd_nat) +by (subst lcm.commute, erule lcm_proj2_if_dvd_nat) lemma lcm_proj1_if_dvd_int [simp]: "(x::int) dvd y \ lcm y x = \y\" -by (subst lcm_commute_int, erule lcm_proj2_if_dvd_int) +by (subst lcm.commute, erule lcm_proj2_if_dvd_int) -lemma lcm_proj1_iff_nat[simp]: "lcm m n = (m::nat) \ n dvd m" +lemma lcm_proj1_iff_nat [simp]: "lcm m n = (m::nat) \ n dvd m" by (metis lcm_proj1_if_dvd_nat lcm_unique_nat) -lemma lcm_proj2_iff_nat[simp]: "lcm m n = (n::nat) \ m dvd n" +lemma lcm_proj2_iff_nat [simp]: "lcm m n = (n::nat) \ m dvd n" by (metis lcm_proj2_if_dvd_nat lcm_unique_nat) -lemma lcm_proj1_iff_int[simp]: "lcm m n = \m::int\ \ n dvd m" +lemma lcm_proj1_iff_int [simp]: "lcm m n = \m::int\ \ n dvd m" by (metis dvd_abs_iff lcm_proj1_if_dvd_int lcm_unique_int) -lemma lcm_proj2_iff_int[simp]: "lcm m n = \n::int\ \ m dvd n" +lemma lcm_proj2_iff_int [simp]: "lcm m n = \n::int\ \ m dvd n" by (metis dvd_abs_iff lcm_proj2_if_dvd_int lcm_unique_int) lemma (in semiring_gcd) comp_fun_idem_gcd: @@ -1903,49 +1836,26 @@ "comp_fun_idem lcm" by standard (simp_all add: fun_eq_iff ac_simps) -lemma comp_fun_idem_gcd_nat: "comp_fun_idem (gcd :: nat\nat\nat)" - by (fact comp_fun_idem_gcd) - -lemma comp_fun_idem_gcd_int: "comp_fun_idem (gcd :: int\int\int)" - by (fact comp_fun_idem_gcd) - -lemma comp_fun_idem_lcm_nat: "comp_fun_idem (lcm :: nat\nat\nat)" - by (fact comp_fun_idem_lcm) - -lemma comp_fun_idem_lcm_int: "comp_fun_idem (lcm :: int\int\int)" - by (fact comp_fun_idem_lcm) - -lemma lcm_0_iff_nat [simp]: "lcm (m::nat) n = 0 \ m=0 \ n=0" - by (fact lcm_eq_0_iff) - -lemma lcm_0_iff_int [simp]: "lcm (m::int) n = 0 \ m=0 \ n=0" - by (fact lcm_eq_0_iff) - -lemma lcm_1_iff_nat [simp]: "lcm (m::nat) n = 1 \ m=1 \ n=1" - by (simp only: lcm_eq_1_iff) simp +lemma lcm_1_iff_nat [simp]: + "lcm (m::nat) n = Suc 0 \ m = Suc 0 \ n = Suc 0" + using lcm_eq_1_iff [of m n] by simp -lemma lcm_1_iff_int [simp]: "lcm (m::int) n = 1 \ (m=1 \ m = -1) \ (n=1 \ n = -1)" +lemma lcm_1_iff_int [simp]: + "lcm (m::int) n = 1 \ (m=1 \ m = -1) \ (n=1 \ n = -1)" by auto -subsection \The complete divisibility lattice\ - -interpretation gcd_semilattice_nat: semilattice_inf gcd Rings.dvd "(\m n::nat. m dvd n \ \ n dvd m)" - by standard simp_all - -interpretation lcm_semilattice_nat: semilattice_sup lcm Rings.dvd "(\m n::nat. m dvd n \ \ n dvd m)" - by standard simp_all - -interpretation gcd_lcm_lattice_nat: lattice gcd Rings.dvd "(\m n::nat. m dvd n & ~ n dvd m)" lcm .. +subsection \The complete divisibility lattice on @{typ nat} and @{typ int}\ text\Lifting gcd and lcm to sets (Gcd/Lcm). Gcd is defined via Lcm to facilitate the proof that we have a complete lattice. \ -instantiation nat :: Gcd +instantiation nat :: semiring_Gcd begin -interpretation semilattice_neutr_set lcm "1::nat" .. +interpretation semilattice_neutr_set lcm "1::nat" + by standard simp_all definition "Lcm (M::nat set) = (if finite M then F M else 0)" @@ -1977,86 +1887,62 @@ fixes M :: "nat set" assumes "\m\M. m dvd n" shows "Lcm M dvd n" -proof (cases "n = 0") - case True then show ?thesis by simp +proof (cases "n > 0") + case False then show ?thesis by simp next - case False + case True then have "finite {d. d dvd n}" by (rule finite_divisors_nat) moreover have "M \ {d. d dvd n}" using assms by fast ultimately have "finite M" by (rule rev_finite_subset) - then show ?thesis using assms by (induct M) (simp_all add: Lcm_nat_empty Lcm_nat_insert) + then show ?thesis using assms + by (induct M) (simp_all add: Lcm_nat_empty Lcm_nat_insert) qed definition "Gcd (M::nat set) = Lcm {d. \m\M. d dvd m}" -instance .. - -end - -instance nat :: semiring_Gcd -proof +instance proof show "Gcd N dvd n" if "n \ N" for N and n :: nat using that by (induct N rule: infinite_finite_induct) (auto simp add: Gcd_nat_def) show "n dvd Gcd N" if "\m. m \ N \ n dvd m" for N and n :: nat using that by (induct N rule: infinite_finite_induct) (auto simp add: Gcd_nat_def) -qed simp - -instance nat :: semiring_Lcm -proof - show "Lcm N = Gcd {m. \n\N. n dvd m}" for N :: "nat set" - by (rule associated_eqI) (auto intro!: Gcd_dvd Gcd_greatest) -qed - -interpretation gcd_lcm_complete_lattice_nat: - complete_lattice Gcd Lcm gcd Rings.dvd "\m n. m dvd n \ \ n dvd m" lcm 1 "0::nat" -rewrites "Inf.INFIMUM Gcd A f = Gcd (f ` A :: nat set)" - and "Sup.SUPREMUM Lcm A f = Lcm (f ` A)" -proof - - show "class.complete_lattice Gcd Lcm gcd Rings.dvd (\m n. m dvd n \ \ n dvd m) lcm 1 (0::nat)" - by standard (auto simp add: Gcd_nat_def Lcm_nat_empty Lcm_nat_infinite) - then interpret gcd_lcm_complete_lattice_nat: - complete_lattice Gcd Lcm gcd Rings.dvd "\m n. m dvd n \ \ n dvd m" lcm 1 "0::nat" . - from gcd_lcm_complete_lattice_nat.INF_def show "Inf.INFIMUM Gcd A f = Gcd (f ` A)" . - from gcd_lcm_complete_lattice_nat.SUP_def show "Sup.SUPREMUM Lcm A f = Lcm (f ` A)" . -qed + show "n dvd Lcm N" if "n \ N" for N and n ::nat + using that by (induct N rule: infinite_finite_induct) + auto + show "Lcm N dvd n" if "\m. m \ N \ m dvd n" for N and n ::nat + using that by (induct N rule: infinite_finite_induct) + auto +qed simp_all -declare gcd_lcm_complete_lattice_nat.Inf_image_eq [simp del] -declare gcd_lcm_complete_lattice_nat.Sup_image_eq [simp del] - -lemma Lcm_empty_nat: - "Lcm {} = (1::nat)" - by (fact Lcm_empty) +end -lemma Lcm_insert_nat [simp]: - "Lcm (insert (n::nat) N) = lcm n (Lcm N)" - by (fact Lcm_insert) - -lemma Lcm_eq_0 [simp]: - "finite (M::nat set) \ 0 \ M \ Lcm M = 0" - by (rule Lcm_eq_0_I) - -lemma Lcm0_iff [simp]: - fixes M :: "nat set" - assumes "finite M" and "M \ {}" - shows "Lcm M = 0 \ 0 \ M" - using assms by (simp add: Lcm_0_iff) +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})" -apply(rule antisym) - apply(rule Max_ge) - apply (metis all_not_in_conv finite_divisors_nat finite_INT) - apply (simp add: Gcd_dvd) -apply (rule Max_le_iff[THEN iffD2]) - apply (metis all_not_in_conv finite_divisors_nat finite_INT) - apply fastforce -apply clarsimp -apply (metis Gcd_dvd Max_in dvd_0_left dvd_Gcd dvd_imp_le linorder_antisym_conv3 not_less0) -done +lemma Gcd_eq_Max: + fixes M :: "nat set" + assumes "finite (M::nat set)" and "M \ {}" and "0 \ M" + shows "Gcd M = Max (\m\M. {d. d dvd m})" +proof (rule antisym) + from assms obtain m where "m \ M" and "m > 0" + by auto + from \m > 0\ have "finite {d. d dvd m}" + by (blast intro: finite_divisors_nat) + with \m \ M\ have fin: "finite (\m\M. {d. d dvd m})" + by blast + from fin show "Gcd M \ Max (\m\M. {d. d dvd m})" + by (auto intro: Max_ge Gcd_dvd) + from fin show "Max (\m\M. {d. d dvd m}) \ Gcd M" + apply (rule Max.boundedI) + apply auto + apply (meson Gcd_dvd Gcd_greatest \0 < m\ \m \ M\ dvd_imp_le dvd_pos_nat) + done +qed lemma Gcd_remove0_nat: "finite M \ Gcd M = Gcd (M - {0::nat})" apply(induct pred:finite) @@ -2086,18 +1972,9 @@ apply(rule antisym) apply(rule Max_ge, assumption) apply(erule (2) Lcm_in_lcm_closed_set_nat) -apply clarsimp -apply (metis Lcm0_iff dvd_Lcm_nat dvd_imp_le neq0_conv) +apply (auto simp add: not_le Lcm_0_iff dvd_imp_le leD le_neq_trans) done -lemma Lcm_set_nat [code, code_unfold]: - "Lcm (set ns) = fold lcm ns (1::nat)" - by (fact gcd_lcm_complete_lattice_nat.Sup_set_fold) - -lemma Gcd_set_nat [code]: - "Gcd (set ns) = fold gcd ns (0::nat)" - by (fact gcd_lcm_complete_lattice_nat.Inf_set_fold) - lemma mult_inj_if_coprime_nat: "inj_on f A \ inj_on g B \ ALL a:A. ALL b:B. coprime (f a) (g b) \ inj_on (%(a,b). f a * g b::nat) (A \ B)" @@ -2115,116 +1992,72 @@ subsubsection \Setwise gcd and lcm for integers\ -instantiation int :: Gcd +instantiation int :: semiring_Gcd 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 + Lcm_int_def int_dvd_iff dvd_int_iff dvd_int_unfold_dvd_nat [symmetric]) + +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}\ + +instantiation integer :: gcd +begin + +context + includes integer.lifting +begin + +lift_definition gcd_integer :: "integer \ integer \ integer" + is gcd . +lift_definition lcm_integer :: "integer \ integer \ integer" + is lcm . + +end instance .. end -instance int :: semiring_Gcd - by standard (auto intro!: Gcd_dvd Gcd_greatest simp add: Gcd_int_def Lcm_int_def int_dvd_iff dvd_int_iff - dvd_int_unfold_dvd_nat [symmetric]) - -instance int :: semiring_Lcm -proof - fix K :: "int set" - have "{n. \k\K. nat \k\ dvd n} = ((\k. nat \k\) ` {l. \k\K. k dvd l})" - proof (rule set_eqI) - fix n - have "(\k\K. nat \k\ dvd n) \ (\l. (\k\K. k dvd l) \ n = nat \l\)" (is "?P \ ?Q") - proof - assume ?P - then have "(\k\K. k dvd int n) \ n = nat \int n\" - by (auto simp add: dvd_int_unfold_dvd_nat) - then show ?Q by blast - next - assume ?Q then show ?P - by (auto simp add: dvd_int_unfold_dvd_nat) - qed - then show "n \ {n. \k\K. nat \k\ dvd n} \ n \ (\k. nat \k\) ` {l. \k\K. k dvd l}" - by auto - qed - then show "Lcm K = Gcd {l. \k\K. k dvd l}" - by (simp add: Gcd_int_def Lcm_int_def Lcm_Gcd) -qed - -lemma Lcm_empty_int [simp]: "Lcm {} = (1::int)" - by (fact Lcm_empty) - -lemma Lcm_insert_int [simp]: - "Lcm (insert (n::int) N) = lcm n (Lcm N)" - by (fact Lcm_insert) - -lemma dvd_int_iff: "x dvd y \ nat \x\ dvd nat \y\" - by (fact dvd_int_unfold_dvd_nat) - -lemma dvd_Lcm_int [simp]: - fixes M :: "int set" assumes "m \ M" shows "m dvd Lcm M" - using assms by (fact dvd_Lcm) - -lemma Lcm_dvd_int [simp]: - fixes M :: "int set" - assumes "\m\M. m dvd n" shows "Lcm M dvd n" - using assms by (simp add: Lcm_int_def dvd_int_iff) - -lemma Lcm_set_int [code, code_unfold]: - "Lcm (set xs) = fold lcm xs (1::int)" - by (induct xs rule: rev_induct) (simp_all add: lcm_commute_int) - -lemma Gcd_set_int [code]: - "Gcd (set xs) = fold gcd xs (0::int)" - by (induct xs rule: rev_induct) (simp_all add: gcd_commute_int) - - -text \Fact aliasses\ - -lemmas gcd_dvd1_nat = gcd_dvd1 [where ?'a = nat] - and gcd_dvd2_nat = gcd_dvd2 [where ?'a = nat] - and gcd_greatest_nat = gcd_greatest [where ?'a = nat] - -lemmas gcd_dvd1_int = gcd_dvd1 [where ?'a = int] - and gcd_dvd2_int = gcd_dvd2 [where ?'a = int] - and gcd_greatest_int = gcd_greatest [where ?'a = int] - -lemmas Gcd_dvd_nat [simp] = Gcd_dvd [where ?'a = nat] - and dvd_Gcd_nat [simp] = dvd_Gcd [where ?'a = nat] - -lemmas Gcd_dvd_int [simp] = Gcd_dvd [where ?'a = int] - and dvd_Gcd_int [simp] = dvd_Gcd [where ?'a = int] - -lemmas Gcd_empty_nat = Gcd_empty [where ?'a = nat] - and Gcd_insert_nat = Gcd_insert [where ?'a = nat] - -lemmas Gcd_empty_int = Gcd_empty [where ?'a = int] - and Gcd_insert_int = Gcd_insert [where ?'a = int] - -subsection \gcd and lcm instances for @{typ integer}\ - -instantiation integer :: gcd begin -context includes integer.lifting begin -lift_definition gcd_integer :: "integer \ integer \ integer" is gcd . -lift_definition lcm_integer :: "integer \ integer \ integer" is lcm . -end -instance .. -end lifting_update integer.lifting lifting_forget integer.lifting -context includes integer.lifting begin +context + includes integer.lifting +begin lemma gcd_code_integer [code]: "gcd k l = \if l = (0::integer) then k else gcd l (\k\ mod \l\)\" -by transfer(fact gcd_code_int) + by transfer (fact gcd_code_int) lemma lcm_code_integer [code]: "lcm (a::integer) b = \a\ * \b\ div gcd a b" -by transfer(fact lcm_altdef_int) + by transfer (fact lcm_altdef_int) end @@ -2234,4 +2067,75 @@ and (Scala) "_.gcd'((_)')" \ \There is no gcd operation in the SML standard library, so no code setup for SML\ +text \Some code equations\ + +lemma Lcm_set_nat [code, code_unfold]: + "Lcm (set ns) = fold lcm ns (1::nat)" + using Lcm_set [of ns] by (simp_all add: fun_eq_iff ac_simps foldr_fold [symmetric]) + +lemma Gcd_set_nat [code]: + "Gcd (set ns) = fold gcd ns (0::nat)" + using Gcd_set [of ns] by (simp_all add: fun_eq_iff ac_simps foldr_fold [symmetric]) + +lemma Lcm_set_int [code, code_unfold]: + "Lcm (set xs) = fold lcm xs (1::int)" + using Lcm_set [of xs] by (simp_all add: fun_eq_iff ac_simps foldr_fold [symmetric]) + +lemma Gcd_set_int [code]: + "Gcd (set xs) = fold gcd xs (0::int)" + using Gcd_set [of xs] by (simp_all add: fun_eq_iff ac_simps foldr_fold [symmetric]) + +text \Fact aliasses\ + +lemma lcm_0_iff_nat [simp]: "lcm (m::nat) n = 0 \ m = 0 \ n = 0" + by (fact lcm_eq_0_iff) + +lemma lcm_0_iff_int [simp]: "lcm (m::int) n = 0 \ m = 0 \ n = 0" + by (fact lcm_eq_0_iff) + +lemma dvd_lcm_I1_nat [simp]: "(k::nat) dvd m \ k dvd lcm m n" + by (fact dvd_lcmI1) + +lemma dvd_lcm_I2_nat [simp]: "(k::nat) dvd n \ k dvd lcm m n" + by (fact dvd_lcmI2) + +lemma dvd_lcm_I1_int [simp]: "(i::int) dvd m \ i dvd lcm m n" + by (fact dvd_lcmI1) + +lemma dvd_lcm_I2_int [simp]: "(i::int) dvd n \ i dvd lcm m n" + by (fact dvd_lcmI2) + +lemma coprime_exp2_nat [intro]: "coprime (a::nat) b \ coprime (a^n) (b^m)" + by (fact coprime_exp2) + +lemma coprime_exp2_int [intro]: "coprime (a::int) b \ coprime (a^n) (b^m)" + by (fact coprime_exp2) + +lemmas Gcd_dvd_nat [simp] = Gcd_dvd [where ?'a = nat] +lemmas Gcd_dvd_int [simp] = Gcd_dvd [where ?'a = int] +lemmas Gcd_greatest_nat [simp] = Gcd_greatest [where ?'a = nat] +lemmas Gcd_greatest_int [simp] = Gcd_greatest [where ?'a = int] + +lemma dvd_Lcm_int [simp]: + fixes M :: "int set" assumes "m \ M" shows "m dvd Lcm M" + using assms by (fact dvd_Lcm) + +lemma gcd_neg_numeral_1_int [simp]: + "gcd (- numeral n :: int) x = gcd (numeral n) x" + by (fact gcd_neg1_int) + +lemma gcd_neg_numeral_2_int [simp]: + "gcd x (- numeral n :: int) = gcd x (numeral n)" + by (fact gcd_neg2_int) + +lemma gcd_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \ gcd x y = x" + by (fact gcd_nat.absorb1) + +lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \ gcd x y = y" + by (fact gcd_nat.absorb2) + +lemmas Lcm_eq_0_I_nat [simp] = Lcm_eq_0_I [where ?'a = nat] +lemmas Lcm_0_iff_nat [simp] = Lcm_0_iff [where ?'a = nat] +lemmas Lcm_least_int [simp] = Lcm_least [where ?'a = int] + end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Groups.thy --- a/src/HOL/Groups.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Groups.thy Wed Feb 17 23:29:35 2016 +0100 @@ -999,14 +999,17 @@ apply (simp add: algebra_simps) done -lemma diff_gt_0_iff_gt [simp]: "a - b > 0 \ a > b" -by (simp add: less_diff_eq) +lemma diff_gt_0_iff_gt [simp]: + "a - b > 0 \ a > b" + by (simp add: less_diff_eq) -lemma diff_le_eq[algebra_simps, field_simps]: "a - b \ c \ a \ c + b" -by (auto simp add: le_less diff_less_eq ) +lemma diff_le_eq [algebra_simps, field_simps]: + "a - b \ c \ a \ c + b" + by (auto simp add: le_less diff_less_eq ) -lemma le_diff_eq[algebra_simps, field_simps]: "a \ c - b \ a + b \ c" -by (auto simp add: le_less less_diff_eq) +lemma le_diff_eq [algebra_simps, field_simps]: + "a \ c - b \ a + b \ c" + by (auto simp add: le_less less_diff_eq) lemma diff_le_0_iff_le [simp]: "a - b \ 0 \ a \ b" @@ -1014,6 +1017,10 @@ lemmas le_iff_diff_le_0 = diff_le_0_iff_le [symmetric] +lemma diff_ge_0_iff_ge [simp]: + "a - b \ 0 \ a \ b" + by (simp add: le_diff_eq) + lemma diff_eq_diff_less: "a - b = c - d \ a < b \ c < d" by (auto simp only: less_iff_diff_less_0 [of a b] less_iff_diff_less_0 [of c d]) @@ -1382,7 +1389,6 @@ lemmas ab_left_minus = left_minus \ \FIXME duplicate\ lemmas diff_diff_eq = diff_diff_add \ \FIXME duplicate\ - subsection \Tools setup\ lemma add_mono_thms_linordered_semiring: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Hilbert_Choice.thy Wed Feb 17 23:29:35 2016 +0100 @@ -253,10 +253,10 @@ done lemma image_surj_f_inv_f: "surj f ==> f ` (inv f ` A) = A" -by (simp add: image_eq_UN surj_f_inv_f) + by (simp add: surj_f_inv_f image_comp comp_def) lemma image_inv_f_f: "inj f ==> inv f ` (f ` A) = A" - by (simp add: image_eq_UN) + by simp lemma inv_image_comp: "inj f ==> inv f ` (f ` X) = X" by (fact image_inv_f_f) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Hoare_Parallel/OG_Hoare.thy --- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Wed Feb 17 23:29:35 2016 +0100 @@ -111,7 +111,8 @@ apply(rule conjI) apply (blast dest: L3_5i) apply(simp add: SEM_def sem_def id_def) -apply (blast dest: Basic_ntran rtrancl_imp_UN_relpow) +apply (auto dest: Basic_ntran rtrancl_imp_UN_relpow) +apply blast done lemma atom_hoare_sound [rule_format]: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Hoare_Parallel/RG_Hoare.thy --- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1266,23 +1266,23 @@ apply simp apply clarify apply(subgoal_tac "\iassum(Pre(xs!i), Rely(xs!i))") - apply(erule_tac x=i and P="\i. H i \ \ (J i) sat [I i,K i,M i,N i]" for H J I K M N in allE,erule impE,assumption) + apply(erule_tac x=xa and P="\i. H i \ \ (J i) sat [I i,K i,M i,N i]" for H J I K M N in allE,erule impE,assumption) apply(simp add:com_validity_def) apply(erule_tac x=s in allE) - apply(erule_tac x=i and P="\j. H j \ (I j) \ cp (J j) s" for H I J in allE,simp) - apply(drule_tac c="clist!i" in subsetD) + apply(erule_tac x=xa and P="\j. H j \ (I j) \ cp (J j) s" for H I J in allE,simp) + apply(drule_tac c="clist!xa" in subsetD) apply (force simp add:Com_def) apply(simp add:comm_def conjoin_def same_program_def del:last.simps) apply clarify apply(erule_tac x="length x - 1" and P="\j. H j \ fst(I j)=(J j)" for H I J in allE) apply (simp add:All_None_def same_length_def) - apply(erule_tac x=i and P="\j. H j \ length(J j)=(K j)" for H J K in allE) + apply(erule_tac x=xa and P="\j. H j \ length(J j)=(K j)" for H J K in allE) apply(subgoal_tac "length x - 1 < length x",simp) apply(case_tac "x\[]") apply(simp add: last_conv_nth) - apply(erule_tac x="clist!i" in ballE) + apply(erule_tac x="clist!xa" in ballE) apply(simp add:same_state_def) - apply(subgoal_tac "clist!i\[]") + apply(subgoal_tac "clist!xa\[]") apply(simp add: last_conv_nth) apply(case_tac x) apply (force simp add:par_cp_def) @@ -1297,19 +1297,19 @@ apply(rule conjI) apply(simp add:conjoin_def same_state_def par_cp_def) apply clarify - apply(erule_tac x=ia and P="\j. T j \ (\i. H j i \ (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp) + apply(erule_tac x=i and P="\j. T j \ (\i. H j i \ (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp) apply(erule_tac x=0 and P="\j. H j \ (snd (d j))=(snd (e j))" for H d e in allE) apply(case_tac x,simp+) apply (simp add:par_assum_def) apply clarify - apply(drule_tac c="snd (clist ! ia ! 0)" in subsetD) + apply(drule_tac c="snd (clist ! i ! 0)" in subsetD) apply assumption apply simp apply clarify -apply(erule_tac x=ia in all_dupE) +apply(erule_tac x=i in all_dupE) apply(rule subsetD, erule mp, assumption) apply(erule_tac pre=pre and rely=rely and x=x and s=s in three) - apply(erule_tac x=ic in allE,erule mp) + apply(erule_tac x=ib in allE,erule mp) apply simp_all apply(simp add:ParallelCom_def) apply(force simp add:Com_def) diff -r e307a410f46c -r ab76bd43c14a src/HOL/IMP/Denotational.thy --- a/src/HOL/IMP/Denotational.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/IMP/Denotational.thy Wed Feb 17 23:29:35 2016 +0100 @@ -98,6 +98,9 @@ theorem lfp_if_cont: assumes "cont f" shows "lfp f = (UN n. (f^^n) {})" (is "_ = ?U") proof + from assms mono_if_cont + have mono: "(f ^^ n) {} \ (f ^^ Suc n) {}" for n + using funpow_decreasing [of n "Suc n"] by auto show "lfp f \ ?U" proof (rule lfp_lowerbound) have "f ?U = (UN n. (f^^Suc n){})" @@ -105,7 +108,7 @@ by(simp add: cont_def) also have "\ = (f^^0){} \ \" by simp also have "\ = ?U" - using assms funpow_decreasing le_SucI mono_if_cont by blast + using mono by auto (metis funpow_simps_right(2) funpow_swap1 o_apply) finally show "f ?U \ ?U" by simp qed next diff -r e307a410f46c -r ab76bd43c14a src/HOL/Inequalities.thy --- a/src/HOL/Inequalities.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Inequalities.thy Wed Feb 17 23:29:35 2016 +0100 @@ -31,7 +31,7 @@ have "m*(m-1) \ n*(n + 1)" using assms by (meson diff_le_self order_trans le_add1 mult_le_mono) hence "int(\ {m..n}) = int((n*(n+1) - m*(m-1)) div 2)" using assms - by (auto simp: Setsum_Icc_int[transferred, OF assms] zdiv_int int_mult simp del: of_nat_setsum + by (auto simp: Setsum_Icc_int[transferred, OF assms] zdiv_int of_nat_mult simp del: of_nat_setsum split: zdiff_int_split) thus ?thesis using of_nat_eq_iff by blast @@ -64,7 +64,7 @@ { fix i j::nat assume "i 0 \ b i - b j \ 0 \ a i - a j \ 0 \ b i - b j \ 0" using assms by (cases "i \ j") (auto simp: algebra_simps) - } hence "?S \ 0" + } then have "?S \ 0" by (auto intro!: setsum_nonpos simp: mult_le_0_iff) finally show ?thesis by (simp add: algebra_simps) qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Int.thy --- a/src/HOL/Int.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Int.thy Wed Feb 17 23:29:35 2016 +0100 @@ -190,6 +190,21 @@ apply arith done +lemma zabs_less_one_iff [simp]: + fixes z :: int + shows "\z\ < 1 \ z = 0" (is "?P \ ?Q") +proof + assume ?Q then show ?P by simp +next + assume ?P + with zless_imp_add1_zle [of "\z\" 1] have "\z\ + 1 \ 1" + by simp + then have "\z\ \ 0" + by simp + then show ?Q + by simp +qed + lemmas int_distrib = distrib_right [of z1 z2 w] distrib_left [of w z1 z2] @@ -320,6 +335,45 @@ lemma of_int_nonneg: "z \ 0 \ of_int z \ 0" by simp +lemma of_int_abs [simp]: + "of_int \x\ = \of_int x\" + by (auto simp add: abs_if) + +lemma of_int_lessD: + assumes "\of_int n\ < x" + shows "n = 0 \ x > 1" +proof (cases "n = 0") + case True then show ?thesis by simp +next + case False + then have "\n\ \ 0" by simp + then have "\n\ > 0" by simp + then have "\n\ \ 1" + using zless_imp_add1_zle [of 0 "\n\"] by simp + then have "\of_int n\ \ 1" + unfolding of_int_1_le_iff [of "\n\", symmetric] by simp + then have "1 < x" using assms by (rule le_less_trans) + then show ?thesis .. +qed + +lemma of_int_leD: + assumes "\of_int n\ \ x" + shows "n = 0 \ 1 \ x" +proof (cases "n = 0") + case True then show ?thesis by simp +next + case False + then have "\n\ \ 0" by simp + then have "\n\ > 0" by simp + then have "\n\ \ 1" + using zless_imp_add1_zle [of 0 "\n\"] by simp + then have "\of_int n\ \ 1" + unfolding of_int_1_le_iff [of "\n\", symmetric] by simp + then have "1 \ x" using assms by (rule order_trans) + then show ?thesis .. +qed + + end text \Comparisons involving @{term of_int}.\ @@ -516,15 +570,19 @@ lemma negative_eq_positive [simp]: "(- int n = of_nat m) = (n = 0 & m = 0)" by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) -lemma zle_iff_zadd: "w \ z \ (\n. z = w + int n)" -proof - - have "(w \ z) = (0 \ z - w)" - by (simp only: le_diff_eq add_0_left) - also have "\ = (\n. z - w = of_nat n)" - by (auto elim: zero_le_imp_eq_int) - also have "\ = (\n. z = w + of_nat n)" - by (simp only: algebra_simps) - finally show ?thesis . +lemma zle_iff_zadd: + "w \ z \ (\n. z = w + int n)" (is "?P \ ?Q") +proof + assume ?Q + then show ?P by auto +next + assume ?P + then have "0 \ z - w" by simp + then obtain n where "z - w = int n" + using zero_le_imp_eq_int [of "z - w"] by blast + then have "z = w + int n" + by simp + then show ?Q .. qed lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z" @@ -1152,9 +1210,6 @@ subsection\Products and 1, by T. M. Rasmussen\ -lemma zabs_less_one_iff [simp]: "(\z\ < 1) = (z = (0::int))" -by arith - lemma abs_zmult_eq_1: assumes mn: "\m * n\ = 1" shows "\m\ = (1::int)" @@ -1674,34 +1729,9 @@ hide_const (open) Pos Neg sub dup -subsection \Legacy theorems\ - -lemmas inj_int = inj_of_nat [where 'a=int] -lemmas zadd_int = of_nat_add [where 'a=int, symmetric] -lemmas int_mult = of_nat_mult [where 'a=int] -lemmas int_eq_0_conv = of_nat_eq_0_iff [where 'a=int and m="n"] for n -lemmas zless_int = of_nat_less_iff [where 'a=int] -lemmas int_less_0_conv = of_nat_less_0_iff [where 'a=int and m="k"] for k -lemmas zero_less_int_conv = of_nat_0_less_iff [where 'a=int] -lemmas zero_zle_int = of_nat_0_le_iff [where 'a=int] -lemmas int_le_0_conv = of_nat_le_0_iff [where 'a=int and m="n"] for n -lemmas int_0 = of_nat_0 [where 'a=int] -lemmas int_1 = of_nat_1 [where 'a=int] -lemmas int_Suc = of_nat_Suc [where 'a=int] -lemmas int_numeral = of_nat_numeral [where 'a=int] -lemmas abs_int_eq = abs_of_nat [where 'a=int and n="m"] for m -lemmas of_int_int_eq = of_int_of_nat_eq [where 'a=int] -lemmas zdiff_int = of_nat_diff [where 'a=int, symmetric] -lemmas zpower_numeral_even = power_numeral_even [where 'a=int] -lemmas zpower_numeral_odd = power_numeral_odd [where 'a=int] - text \De-register \int\ as a quotient type:\ lifting_update int.lifting lifting_forget int.lifting -text\Also the class for fields with characteristic zero.\ -class field_char_0 = field + ring_char_0 -subclass (in linordered_field) field_char_0 .. - end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Isar_Examples/Fibonacci.thy --- a/src/HOL/Isar_Examples/Fibonacci.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Isar_Examples/Fibonacci.thy Wed Feb 17 23:29:35 2016 +0100 @@ -82,7 +82,7 @@ also have "gcd (fib (n + 2)) \ = gcd (fib (n + 2)) (fib (n + 1))" by (rule gcd_add2_nat) also have "\ = gcd (fib (n + 1)) (fib (n + 1 + 1))" - by (simp add: gcd_commute_nat) + by (simp add: gcd.commute) also assume "\ = 1" finally show "?P (n + 2)" . qed @@ -104,16 +104,16 @@ next case (Suc k) then have "gcd (fib m) (fib (n + m)) = gcd (fib (n + k + 1)) (fib (k + 1))" - by (simp add: gcd_commute_nat) + by (simp add: gcd.commute) also have "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n" by (rule fib_add) also have "gcd \ (fib (k + 1)) = gcd (fib k * fib n) (fib (k + 1))" by (simp add: gcd_mult_add) also have "\ = gcd (fib n) (fib (k + 1))" - by (simp only: gcd_fib_Suc_eq_1 gcd_mult_cancel_nat) + by (simp only: gcd_fib_Suc_eq_1 gcd_mult_cancel) also have "\ = gcd (fib m) (fib n)" - using Suc by (simp add: gcd_commute_nat) + using Suc by (simp add: gcd.commute) finally show ?thesis . qed @@ -171,7 +171,7 @@ also from n have "\ = gcd (fib n) (fib m)" by (rule gcd_fib_mod) also have "\ = gcd (fib m) (fib n)" - by (rule gcd_commute_nat) + by (rule gcd.commute) finally show "fib (gcd m n) = gcd (fib m) (fib n)" . qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Countable_Set_Type.thy --- a/src/HOL/Library/Countable_Set_Type.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Countable_Set_Type.thy Wed Feb 17 23:29:35 2016 +0100 @@ -123,7 +123,11 @@ lemma cUnion_transfer [transfer_rule]: "rel_fun (pcr_cset (pcr_cset A)) (pcr_cset A) Union cUnion" -unfolding cUnion_def[abs_def] Union_conv_UNION[abs_def] by transfer_prover +proof - + have "rel_fun (pcr_cset (pcr_cset A)) (pcr_cset A) (\A. UNION A id) (\A. cUNION A id)" + by transfer_prover + then show ?thesis by (simp add: cUnion_def [symmetric]) +qed lemmas cset_eqI = set_eqI[Transfer.transferred] lemmas cset_eq_iff[no_atp] = set_eq_iff[Transfer.transferred] @@ -342,11 +346,9 @@ lemmas cin_mono = in_mono[Transfer.transferred] lemmas cLeast_mono = Least_mono[Transfer.transferred] lemmas cequalityI = equalityI[Transfer.transferred] -lemmas cUnion_cimage_eq [simp] = Union_image_eq[Transfer.transferred] lemmas cUN_iff [simp] = UN_iff[Transfer.transferred] lemmas cUN_I [intro] = UN_I[Transfer.transferred] lemmas cUN_E [elim!] = UN_E[Transfer.transferred] -lemmas cimage_eq_cUN = image_eq_UN[Transfer.transferred] lemmas cUN_upper = UN_upper[Transfer.transferred] lemmas cUN_least = UN_least[Transfer.transferred] lemmas cUN_cinsert_distrib = UN_insert_distrib[Transfer.transferred] diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Extended_Real.thy Wed Feb 17 23:29:35 2016 +0100 @@ -2002,12 +2002,11 @@ proof cases assume "(SUP i:I. f i) = - \" moreover then have "\i. i \ I \ f i = -\" - unfolding Sup_eq_MInfty Sup_image_eq[symmetric] by auto + unfolding Sup_eq_MInfty by auto ultimately show ?thesis by (cases c) (auto simp: \I \ {}\) next assume "(SUP i:I. f i) \ - \" then show ?thesis - unfolding Sup_image_eq[symmetric] by (subst continuous_at_Sup_mono[where f="\x. x + c"]) (auto simp: continuous_at_imp_continuous_at_within continuous_at mono_def ereal_add_mono \I \ {}\ \c \ -\\) qed @@ -2130,7 +2129,6 @@ by simp next assume "(SUP i:I. f i) \ 0" then show ?thesis - unfolding SUP_def by (subst continuous_at_Sup_mono[where f="\x. c * x"]) (auto simp: mono_def continuous_at continuous_at_imp_continuous_at_within \I \ {}\ intro!: ereal_mult_left_mono c) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/FSet.thy --- a/src/HOL/Library/FSet.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/FSet.thy Wed Feb 17 23:29:35 2016 +0100 @@ -147,7 +147,7 @@ parametric right_total_Compl_transfer Compl_transfer by simp instance - by (standard, simp_all only: INF_def SUP_def) (transfer, simp add: Compl_partition Diff_eq)+ + by (standard; transfer) (simp_all add: Diff_eq) end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Finite_Lattice.thy --- a/src/HOL/Library/Finite_Lattice.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Finite_Lattice.thy Wed Feb 17 23:29:35 2016 +0100 @@ -116,12 +116,12 @@ lemma finite_Inf_prod: "Inf(A :: ('a::finite_lattice_complete \ 'b::finite_lattice_complete) set) = Finite_Set.fold inf top A" - by (metis Inf_fold_inf finite_code) + by (metis Inf_fold_inf finite) lemma finite_Sup_prod: "Sup (A :: ('a::finite_lattice_complete \ 'b::finite_lattice_complete) set) = Finite_Set.fold sup bot A" - by (metis Sup_fold_sup finite_code) + by (metis Sup_fold_sup finite) instance prod :: (finite_lattice_complete, finite_lattice_complete) finite_lattice_complete by standard (auto simp: finite_bot_prod finite_top_prod finite_Inf_prod finite_Sup_prod) @@ -130,20 +130,20 @@ already form a finite lattice.\ lemma finite_bot_fun: "(bot :: ('a::finite \ 'b::finite_lattice_complete)) = Inf_fin UNIV" - by (metis Inf_UNIV Inf_fin_Inf empty_not_UNIV finite_code) + by (metis Inf_UNIV Inf_fin_Inf empty_not_UNIV finite) lemma finite_top_fun: "(top :: ('a::finite \ 'b::finite_lattice_complete)) = Sup_fin UNIV" - by (metis Sup_UNIV Sup_fin_Sup empty_not_UNIV finite_code) + by (metis Sup_UNIV Sup_fin_Sup empty_not_UNIV finite) lemma finite_Inf_fun: "Inf (A::('a::finite \ 'b::finite_lattice_complete) set) = Finite_Set.fold inf top A" - by (metis Inf_fold_inf finite_code) + by (metis Inf_fold_inf finite) lemma finite_Sup_fun: "Sup (A::('a::finite \ 'b::finite_lattice_complete) set) = Finite_Set.fold sup bot A" - by (metis Sup_fold_sup finite_code) + by (metis Sup_fold_sup finite) instance "fun" :: (finite, finite_lattice_complete) finite_lattice_complete by standard (auto simp: finite_bot_fun finite_top_fun finite_Inf_fun finite_Sup_fun) @@ -165,11 +165,7 @@ lemma finite_distrib_lattice_complete_inf_Sup: "inf (x::'a::finite_distrib_lattice_complete) (Sup A) = (SUP y:A. inf x y)" - apply (rule finite_induct) - apply (metis finite_code) - apply (metis SUP_empty Sup_empty inf_bot_right) - apply (metis SUP_insert Sup_insert inf_sup_distrib1) - done + using finite [of A] by induct (simp_all add: inf_sup_distrib1) instance finite_distrib_lattice_complete \ complete_distrib_lattice proof diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Float.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1261,7 +1261,7 @@ case True def x' \ "x * 2 ^ nat l" have "int x * 2 ^ nat l = x'" - by (simp add: x'_def int_mult of_nat_power) + by (simp add: x'_def of_nat_mult of_nat_power) moreover have "real x * 2 powr l = real x'" by (simp add: powr_realpow[symmetric] \0 \ l\ x'_def) ultimately show ?thesis @@ -1274,7 +1274,7 @@ case False def y' \ "y * 2 ^ nat (- l)" from \y \ 0\ have "y' \ 0" by (simp add: y'_def) - have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult of_nat_power) + have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def of_nat_mult of_nat_power) moreover have "real x * real_of_int (2::int) powr real_of_int l / real y = x / real y'" using \\ 0 \ l\ by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Formal_Power_Series.thy Wed Feb 17 23:29:35 2016 +0100 @@ -3181,7 +3181,7 @@ let ?KM = "{(k,m). k + m \ n}" let ?f = "\s. UNION {(0::nat)..s} (\i. {(i,s - i)})" have th0: "?KM = UNION {0..n} ?f" - by (auto simp add: set_eq_iff Bex_def) + by auto show "?l = ?r " unfolding th0 apply (subst setsum.UNION_disjoint) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Library.thy Wed Feb 17 23:29:35 2016 +0100 @@ -53,7 +53,6 @@ Parallel Permutation Permutations - Poly_Deriv Polynomial Preorder Product_Vector diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Liminf_Limsup.thy --- a/src/HOL/Library/Liminf_Limsup.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Liminf_Limsup.thy Wed Feb 17 23:29:35 2016 +0100 @@ -380,7 +380,7 @@ note * = this have "f (Liminf F g) = (SUP P : {P. eventually P F}. f (Inf (g ` Collect P)))" - unfolding Liminf_def SUP_def + unfolding Liminf_def by (subst continuous_at_Sup_mono[OF am continuous_on_imp_continuous_within[OF c]]) (auto intro: eventually_True) also have "\ = (SUP P : {P. eventually P F}. INFIMUM (g ` Collect P) f)" @@ -405,7 +405,7 @@ note * = this have "f (Limsup F g) = (INF P : {P. eventually P F}. f (Sup (g ` Collect P)))" - unfolding Limsup_def INF_def + unfolding Limsup_def by (subst continuous_at_Inf_mono[OF am continuous_on_imp_continuous_within[OF c]]) (auto intro: eventually_True) also have "\ = (INF P : {P. eventually P F}. SUPREMUM (g ` Collect P) f)" @@ -429,7 +429,7 @@ by auto qed have "f (Limsup F g) = (SUP P : {P. eventually P F}. f (Sup (g ` Collect P)))" - unfolding Limsup_def INF_def + unfolding Limsup_def by (subst continuous_at_Inf_antimono[OF am continuous_on_imp_continuous_within[OF c]]) (auto intro: eventually_True) also have "\ = (SUP P : {P. eventually P F}. INFIMUM (g ` Collect P) f)" @@ -455,7 +455,7 @@ note * = this have "f (Liminf F g) = (INF P : {P. eventually P F}. f (Inf (g ` Collect P)))" - unfolding Liminf_def SUP_def + unfolding Liminf_def by (subst continuous_at_Sup_antimono[OF am continuous_on_imp_continuous_within[OF c]]) (auto intro: eventually_True) also have "\ = (INF P : {P. eventually P F}. SUPREMUM (g ` Collect P) f)" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Lub_Glb.thy --- a/src/HOL/Library/Lub_Glb.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Lub_Glb.thy Wed Feb 17 23:29:35 2016 +0100 @@ -230,7 +230,7 @@ by (intro LIMSEQ_incseq_SUP) (auto simp: incseq_def image_def eq_commute bdd_above_setle) also have "(SUP i. X i) = u" using isLub_cSup[of "range X"] u[THEN isLubD1] - by (intro isLub_unique[OF _ u]) (auto simp add: SUP_def image_def eq_commute) + by (intro isLub_unique[OF _ u]) (auto simp add: image_def eq_commute) finally show ?thesis . qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Numeral_Type.thy Wed Feb 17 23:29:35 2016 +0100 @@ -230,7 +230,7 @@ "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" apply (rule mod_type.intro) -apply (simp add: int_mult type_definition_bit0) +apply (simp add: of_nat_mult type_definition_bit0) apply (rule one_less_int_card) apply (rule zero_bit0_def) apply (rule one_bit0_def) @@ -245,7 +245,7 @@ "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" apply (rule mod_type.intro) -apply (simp add: int_mult type_definition_bit1) +apply (simp add: of_nat_mult type_definition_bit1) apply (rule one_less_int_card) apply (rule zero_bit1_def) apply (rule one_bit1_def) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Old_SMT/old_smt_normalize.ML --- a/src/HOL/Library/Old_SMT/old_smt_normalize.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Old_SMT/old_smt_normalize.ML Wed Feb 17 23:29:35 2016 +0100 @@ -469,7 +469,7 @@ "int (n * m) = int n * int m" "int (n div m) = int n div int m" "int (n mod m) = int n mod int m" - by (auto simp add: int_mult zdiv_int zmod_int)} + by (auto simp add: of_nat_mult zdiv_int zmod_int)} val int_if = mk_meta_eq @{lemma "int (if P then n else m) = (if P then int n else int m)" @@ -479,7 +479,7 @@ let val eq = Old_SMT_Utils.mk_cequals lhs (Numeral.mk_cnumber @{ctyp int} i) val tac = - Simplifier.simp_tac (put_simpset HOL_ss ctxt addsimps [@{thm Int.int_numeral}]) 1 + Simplifier.simp_tac (put_simpset HOL_ss ctxt addsimps [@{thm of_nat_numeral}]) 1 in Goal.norm_result ctxt (Goal.prove_internal ctxt [] eq (K tac)) end fun ite_conv cv1 cv2 = diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Option_ord.thy --- a/src/HOL/Library/Option_ord.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Option_ord.thy Wed Feb 17 23:29:35 2016 +0100 @@ -289,7 +289,7 @@ show "a \ \B = (\b\B. a \ b)" proof (cases a) case None - then show ?thesis by (simp add: INF_def) + then show ?thesis by simp next case (Some c) show ?thesis @@ -303,13 +303,13 @@ from sup_Inf have "Some c \ Some (\Option.these B) = Some (\b\Option.these B. c \ b)" by simp then have "Some c \ \(Some ` Option.these B) = (\x\Some ` Option.these B. Some c \ x)" by (simp add: Some_INF Some_Inf comp_def) - with Some B show ?thesis by (simp add: Some_image_these_eq) + with Some B show ?thesis by (simp add: Some_image_these_eq cong del: strong_INF_cong) qed qed show "a \ \B = (\b\B. a \ b)" proof (cases a) case None - then show ?thesis by (simp add: SUP_def image_constant_conv bot_option_def) + then show ?thesis by (simp add: image_constant_conv bot_option_def cong del: strong_SUP_cong) next case (Some c) show ?thesis @@ -332,7 +332,7 @@ ultimately have "Some c \ \(Some ` Option.these B) = (\x\Some ` Option.these B. Some c \ x)" by (simp add: Some_SUP Some_Sup comp_def) with Some show ?thesis - by (simp add: Some_image_these_eq Sup_B SUP_B Sup_None SUP_None SUP_union Sup_union_distrib) + by (simp add: Some_image_these_eq Sup_B SUP_B Sup_None SUP_None SUP_union Sup_union_distrib cong del: strong_SUP_cong) qed qed qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Poly_Deriv.thy --- a/src/HOL/Library/Poly_Deriv.thy Wed Feb 17 23:28:58 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,634 +0,0 @@ -(* Title: HOL/Library/Poly_Deriv.thy - Author: Amine Chaieb - Author: Brian Huffman -*) - -section\Polynomials and Differentiation\ - -theory Poly_Deriv -imports Deriv Polynomial -begin - -subsection \Derivatives of univariate polynomials\ - -function pderiv :: "('a :: semidom) poly \ 'a poly" -where - [simp del]: "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" - by (auto intro: pCons_cases) - -termination pderiv - by (relation "measure degree") simp_all - -lemma pderiv_0 [simp]: - "pderiv 0 = 0" - using pderiv.simps [of 0 0] by simp - -lemma pderiv_pCons: - "pderiv (pCons a p) = p + pCons 0 (pderiv p)" - by (simp add: pderiv.simps) - -lemma pderiv_1 [simp]: "pderiv 1 = 0" - unfolding one_poly_def by (simp add: pderiv_pCons) - -lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" - and pderiv_numeral [simp]: "pderiv (numeral m) = 0" - by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) - -lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" - by (induct p arbitrary: n) - (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) - -fun pderiv_coeffs_code :: "('a :: semidom) \ 'a list \ 'a list" where - "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" -| "pderiv_coeffs_code f [] = []" - -definition pderiv_coeffs :: "('a :: semidom) list \ 'a list" where - "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" - -(* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) -lemma pderiv_coeffs_code: - "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * (nth_default 0 xs n)" -proof (induct xs arbitrary: f n) - case (Cons x xs f n) - show ?case - proof (cases n) - case 0 - thus ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0", auto simp: cCons_def) - next - case (Suc m) note n = this - show ?thesis - proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") - case False - hence "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = - nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" - by (auto simp: cCons_def n) - also have "\ = (f + of_nat n) * (nth_default 0 xs m)" - unfolding Cons by (simp add: n add_ac) - finally show ?thesis by (simp add: n) - next - case True - { - fix g - have "pderiv_coeffs_code g xs = [] \ g + of_nat m = 0 \ nth_default 0 xs m = 0" - proof (induct xs arbitrary: g m) - case (Cons x xs g) - from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" - and g: "(g = 0 \ x = 0)" - by (auto simp: cCons_def split: if_splits) - note IH = Cons(1)[OF empty] - from IH[of m] IH[of "m - 1"] g - show ?case by (cases m, auto simp: field_simps) - qed simp - } note empty = this - from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" - by (auto simp: cCons_def n) - moreover have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" using True - by (simp add: n, insert empty[of "f+1"], auto simp: field_simps) - ultimately show ?thesis by simp - qed - qed -qed simp - -lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\ i. f (Suc i)) [0 ..< n]" - by (induct n arbitrary: f, auto) - -lemma coeffs_pderiv_code [code abstract]: - "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def -proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) - case (1 n) - have id: "coeff p (Suc n) = nth_default 0 (map (\i. coeff p (Suc i)) [0.. degree p = 0" - apply (rule iffI) - apply (cases p, simp) - apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) - apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) - done - -lemma degree_pderiv: "degree (pderiv (p :: 'a poly)) = degree p - 1" - apply (rule order_antisym [OF degree_le]) - apply (simp add: coeff_pderiv coeff_eq_0) - apply (cases "degree p", simp) - apply (rule le_degree) - apply (simp add: coeff_pderiv del: of_nat_Suc) - apply (metis degree_0 leading_coeff_0_iff nat.distinct(1)) - done - -lemma not_dvd_pderiv: - assumes "degree (p :: 'a poly) \ 0" - shows "\ p dvd pderiv p" -proof - assume dvd: "p dvd pderiv p" - then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto - from dvd have le: "degree p \ degree (pderiv p)" - by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) - from this[unfolded degree_pderiv] assms show False by auto -qed - -lemma dvd_pderiv_iff [simp]: "(p :: 'a poly) dvd pderiv p \ degree p = 0" - using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) - -end - -lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" -by (simp add: pderiv_pCons) - -lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" -by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) - -lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" -by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) - -lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q" -by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) - -lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" -by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) - -lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" -by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) - -lemma pderiv_power_Suc: - "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" -apply (induct n) -apply simp -apply (subst power_Suc) -apply (subst pderiv_mult) -apply (erule ssubst) -apply (simp only: of_nat_Suc smult_add_left smult_1_left) -apply (simp add: algebra_simps) -done - -lemma pderiv_setprod: "pderiv (setprod f (as)) = - (\a \ as. setprod f (as - {a}) * pderiv (f a))" -proof (induct as rule: infinite_finite_induct) - case (insert a as) - hence id: "setprod f (insert a as) = f a * setprod f as" - "\ g. setsum g (insert a as) = g a + setsum g as" - "insert a as - {a} = as" - by auto - { - fix b - assume "b \ as" - hence id2: "insert a as - {b} = insert a (as - {b})" using \a \ as\ by auto - have "setprod f (insert a as - {b}) = f a * setprod f (as - {b})" - unfolding id2 - by (subst setprod.insert, insert insert, auto) - } note id2 = this - show ?case - unfolding id pderiv_mult insert(3) setsum_right_distrib - by (auto simp add: ac_simps id2 intro!: setsum.cong) -qed auto - -lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" -by (rule DERIV_cong, rule DERIV_pow, simp) -declare DERIV_pow2 [simp] DERIV_pow [simp] - -lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D" -by (rule DERIV_cong, rule DERIV_add, auto) - -lemma poly_DERIV [simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" - by (induct p, auto intro!: derivative_eq_intros simp add: pderiv_pCons) - -lemma continuous_on_poly [continuous_intros]: - fixes p :: "'a :: {real_normed_field} poly" - assumes "continuous_on A f" - shows "continuous_on A (\x. poly p (f x))" -proof - - have "continuous_on A (\x. (\i\degree p. (f x) ^ i * coeff p i))" - by (intro continuous_intros assms) - also have "\ = (\x. poly p (f x))" by (intro ext) (simp add: poly_altdef mult_ac) - finally show ?thesis . -qed - -text\Consequences of the derivative theorem above\ - -lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (at x::real filter)" -apply (simp add: real_differentiable_def) -apply (blast intro: poly_DERIV) -done - -lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)" -by (rule poly_DERIV [THEN DERIV_isCont]) - -lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |] - ==> \x. a < x & x < b & (poly p x = 0)" -using IVT_objl [of "poly p" a 0 b] -by (auto simp add: order_le_less) - -lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |] - ==> \x. a < x & x < b & (poly p x = 0)" -by (insert poly_IVT_pos [where p = "- p" ]) simp - -lemma poly_IVT: - fixes p::"real poly" - assumes "ax>a. x < b \ poly p x = 0" -by (metis assms(1) assms(2) less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) - -lemma poly_MVT: "(a::real) < b ==> - \x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" -using MVT [of a b "poly p"] -apply auto -apply (rule_tac x = z in exI) -apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique]) -done - -lemma poly_MVT': - assumes "{min a b..max a b} \ A" - shows "\x\A. poly p b - poly p a = (b - a) * poly (pderiv p) (x::real)" -proof (cases a b rule: linorder_cases) - case less - from poly_MVT[OF less, of p] guess x by (elim exE conjE) - thus ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) - -next - case greater - from poly_MVT[OF greater, of p] guess x by (elim exE conjE) - thus ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) -qed (insert assms, auto) - -lemma poly_pinfty_gt_lc: - fixes p:: "real poly" - assumes "lead_coeff p > 0" - shows "\ n. \ x \ n. poly p x \ lead_coeff p" using assms -proof (induct p) - case 0 - thus ?case by auto -next - case (pCons a p) - have "\a\0;p=0\ \ ?case" by auto - moreover have "p\0 \ ?case" - proof - - assume "p\0" - then obtain n1 where gte_lcoeff:"\x\n1. lead_coeff p \ poly p x" using that pCons by auto - have gt_0:"lead_coeff p >0" using pCons(3) \p\0\ by auto - def n\"max n1 (1+ \a\/(lead_coeff p))" - show ?thesis - proof (rule_tac x=n in exI,rule,rule) - fix x assume "n \ x" - hence "lead_coeff p \ poly p x" - using gte_lcoeff unfolding n_def by auto - hence " \a\/(lead_coeff p) \ \a\/(poly p x)" and "poly p x>0" using gt_0 - by (intro frac_le,auto) - hence "x\1+ \a\/(poly p x)" using \n\x\[unfolded n_def] by auto - thus "lead_coeff (pCons a p) \ poly (pCons a p) x" - using \lead_coeff p \ poly p x\ \poly p x>0\ \p\0\ - by (auto simp add:field_simps) - qed - qed - ultimately show ?case by fastforce -qed - - -subsection \Algebraic numbers\ - -text \ - Algebraic numbers can be defined in two equivalent ways: all real numbers that are - roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry - uses the rational definition, but we need the integer definition. - - The equivalence is obvious since any rational polynomial can be multiplied with the - LCM of its coefficients, yielding an integer polynomial with the same roots. -\ -subsection \Algebraic numbers\ - -definition algebraic :: "'a :: field_char_0 \ bool" where - "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" - -lemma algebraicI: - assumes "\i. coeff p i \ \" "p \ 0" "poly p x = 0" - shows "algebraic x" - using assms unfolding algebraic_def by blast - -lemma algebraicE: - assumes "algebraic x" - obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" - using assms unfolding algebraic_def by blast - -lemma quotient_of_denom_pos': "snd (quotient_of x) > 0" - using quotient_of_denom_pos[OF surjective_pairing] . - -lemma of_int_div_in_Ints: - "b dvd a \ of_int a div of_int b \ (\ :: 'a :: ring_div set)" -proof (cases "of_int b = (0 :: 'a)") - assume "b dvd a" "of_int b \ (0::'a)" - then obtain c where "a = b * c" by (elim dvdE) - with \of_int b \ (0::'a)\ show ?thesis by simp -qed auto - -lemma of_int_divide_in_Ints: - "b dvd a \ of_int a / of_int b \ (\ :: 'a :: field set)" -proof (cases "of_int b = (0 :: 'a)") - assume "b dvd a" "of_int b \ (0::'a)" - then obtain c where "a = b * c" by (elim dvdE) - with \of_int b \ (0::'a)\ show ?thesis by simp -qed auto - -lemma algebraic_altdef: - fixes p :: "'a :: field_char_0 poly" - shows "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" -proof safe - fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" - def cs \ "coeffs p" - from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast - then obtain f where f: "\i. coeff p i = of_rat (f (coeff p i))" - by (subst (asm) bchoice_iff) blast - def cs' \ "map (quotient_of \ f) (coeffs p)" - def d \ "Lcm (set (map snd cs'))" - def p' \ "smult (of_int d) p" - - have "\n. coeff p' n \ \" - proof - fix n :: nat - show "coeff p' n \ \" - proof (cases "n \ degree p") - case True - def c \ "coeff p n" - def a \ "fst (quotient_of (f (coeff p n)))" and b \ "snd (quotient_of (f (coeff p n)))" - have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp - have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) - also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def - by (subst quotient_of_div [of "f (coeff p n)", symmetric]) - (simp_all add: f [symmetric]) - also have "of_int d * \ = of_rat (of_int (a*d) / of_int b)" - by (simp add: of_rat_mult of_rat_divide) - also from nz True have "b \ snd ` set cs'" unfolding cs'_def - by (force simp: o_def b_def coeffs_def simp del: upt_Suc) - hence "b dvd (a * d)" unfolding d_def by simp - hence "of_int (a * d) / of_int b \ (\ :: rat set)" - by (rule of_int_divide_in_Ints) - hence "of_rat (of_int (a * d) / of_int b) \ \" by (elim Ints_cases) auto - finally show ?thesis . - qed (auto simp: p'_def not_le coeff_eq_0) - qed - - moreover have "set (map snd cs') \ {0<..}" - unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) - hence "d \ 0" unfolding d_def by (induction cs') simp_all - with nz have "p' \ 0" by (simp add: p'_def) - moreover from root have "poly p' x = 0" by (simp add: p'_def) - ultimately show "algebraic x" unfolding algebraic_def by blast -next - - assume "algebraic x" - then obtain p where p: "\i. coeff p i \ \" "poly p x = 0" "p \ 0" - by (force simp: algebraic_def) - moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp - ultimately show "(\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" by auto -qed - - -text\Lemmas for Derivatives\ - -lemma order_unique_lemma: - fixes p :: "'a::idom poly" - assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" - shows "n = order a p" -unfolding Polynomial.order_def -apply (rule Least_equality [symmetric]) -apply (fact assms) -apply (rule classical) -apply (erule notE) -unfolding not_less_eq_eq -using assms(1) apply (rule power_le_dvd) -apply assumption -done - -lemma lemma_order_pderiv1: - "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + - smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" -apply (simp only: pderiv_mult pderiv_power_Suc) -apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons) -done - -lemma lemma_order_pderiv: - fixes p :: "'a :: field_char_0 poly" - assumes n: "0 < n" - and pd: "pderiv p \ 0" - and pe: "p = [:- a, 1:] ^ n * q" - and nd: "~ [:- a, 1:] dvd q" - shows "n = Suc (order a (pderiv p))" -using n -proof - - have "pderiv ([:- a, 1:] ^ n * q) \ 0" - using assms by auto - obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" - using assms by (cases n) auto - have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" - by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) - have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" - proof (rule order_unique_lemma) - show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" - apply (subst lemma_order_pderiv1) - apply (rule dvd_add) - apply (metis dvdI dvd_mult2 power_Suc2) - apply (metis dvd_smult dvd_triv_right) - done - next - show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" - apply (subst lemma_order_pderiv1) - by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) - qed - then show ?thesis - by (metis \n = Suc n'\ pe) -qed - -lemma order_decomp: - assumes "p \ 0" - shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" -proof - - from assms have A: "[:- a, 1:] ^ order a p dvd p" - and B: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) - from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" .. - with B have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" - by simp - then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" - by simp - then have D: "\ [:- a, 1:] dvd q" - using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] - by auto - from C D show ?thesis by blast -qed - -lemma order_pderiv: - "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ \ - (order a p = Suc (order a (pderiv p)))" -apply (case_tac "p = 0", simp) -apply (drule_tac a = a and p = p in order_decomp) -using neq0_conv -apply (blast intro: lemma_order_pderiv) -done - -lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" -proof - - def i \ "order a p" - def j \ "order a q" - def t \ "[:-a, 1:]" - have t_dvd_iff: "\u. t dvd u \ poly u a = 0" - unfolding t_def by (simp add: dvd_iff_poly_eq_0) - assume "p * q \ 0" - then show "order a (p * q) = i + j" - apply clarsimp - apply (drule order [where a=a and p=p, folded i_def t_def]) - apply (drule order [where a=a and p=q, folded j_def t_def]) - apply clarify - apply (erule dvdE)+ - apply (rule order_unique_lemma [symmetric], fold t_def) - apply (simp_all add: power_add t_dvd_iff) - done -qed - -lemma order_smult: - assumes "c \ 0" - shows "order x (smult c p) = order x p" -proof (cases "p = 0") - case False - have "smult c p = [:c:] * p" by simp - also from assms False have "order x \ = order x [:c:] + order x p" - by (subst order_mult) simp_all - also from assms have "order x [:c:] = 0" by (intro order_0I) auto - finally show ?thesis by simp -qed simp - -(* Next two lemmas contributed by Wenda Li *) -lemma order_1_eq_0 [simp]:"order x 1 = 0" - by (metis order_root poly_1 zero_neq_one) - -lemma order_power_n_n: "order a ([:-a,1:]^n)=n" -proof (induct n) (*might be proved more concisely using nat_less_induct*) - case 0 - thus ?case by (metis order_root poly_1 power_0 zero_neq_one) -next - case (Suc n) - have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" - by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral - one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) - moreover have "order a [:-a,1:]=1" unfolding order_def - proof (rule Least_equality,rule ccontr) - assume "\ \ [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" - hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp - hence "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:] )" - by (rule dvd_imp_degree_le,auto) - thus False by auto - next - fix y assume asm:"\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" - show "1 \ y" - proof (rule ccontr) - assume "\ 1 \ y" - hence "y=0" by auto - hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto - thus False using asm by auto - qed - qed - ultimately show ?case using Suc by auto -qed - -text\Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\ - -lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" -apply (cases "p = 0", auto) -apply (drule order_2 [where a=a and p=p]) -apply (metis not_less_eq_eq power_le_dvd) -apply (erule power_le_dvd [OF order_1]) -done - -lemma poly_squarefree_decomp_order: - assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" - and p: "p = q * d" - and p': "pderiv p = e * d" - and d: "d = r * p + s * pderiv p" - shows "order a q = (if order a p = 0 then 0 else 1)" -proof (rule classical) - assume 1: "order a q \ (if order a p = 0 then 0 else 1)" - from \pderiv p \ 0\ have "p \ 0" by auto - with p have "order a p = order a q + order a d" - by (simp add: order_mult) - with 1 have "order a p \ 0" by (auto split: if_splits) - have "order a (pderiv p) = order a e + order a d" - using \pderiv p \ 0\ \pderiv p = e * d\ by (simp add: order_mult) - have "order a p = Suc (order a (pderiv p))" - using \pderiv p \ 0\ \order a p \ 0\ by (rule order_pderiv) - have "d \ 0" using \p \ 0\ \p = q * d\ by simp - have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" - apply (simp add: d) - apply (rule dvd_add) - apply (rule dvd_mult) - apply (simp add: order_divides \p \ 0\ - \order a p = Suc (order a (pderiv p))\) - apply (rule dvd_mult) - apply (simp add: order_divides) - done - then have "order a (pderiv p) \ order a d" - using \d \ 0\ by (simp add: order_divides) - show ?thesis - using \order a p = order a q + order a d\ - using \order a (pderiv p) = order a e + order a d\ - using \order a p = Suc (order a (pderiv p))\ - using \order a (pderiv p) \ order a d\ - by auto -qed - -lemma poly_squarefree_decomp_order2: - "\pderiv p \ (0 :: 'a :: field_char_0 poly); - p = q * d; - pderiv p = e * d; - d = r * p + s * pderiv p - \ \ \a. order a q = (if order a p = 0 then 0 else 1)" -by (blast intro: poly_squarefree_decomp_order) - -lemma order_pderiv2: - "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ - \ (order a (pderiv p) = n) = (order a p = Suc n)" -by (auto dest: order_pderiv) - -definition - rsquarefree :: "'a::idom poly => bool" where - "rsquarefree p = (p \ 0 & (\a. (order a p = 0) | (order a p = 1)))" - -lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h :: 'a :: {semidom,semiring_char_0}:]" -apply (simp add: pderiv_eq_0_iff) -apply (case_tac p, auto split: if_splits) -done - -lemma rsquarefree_roots: - fixes p :: "'a :: field_char_0 poly" - shows "rsquarefree p = (\a. \(poly p a = 0 \ poly (pderiv p) a = 0))" -apply (simp add: rsquarefree_def) -apply (case_tac "p = 0", simp, simp) -apply (case_tac "pderiv p = 0") -apply simp -apply (drule pderiv_iszero, clarsimp) -apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) -apply (force simp add: order_root order_pderiv2) -done - -lemma poly_squarefree_decomp: - assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" - and "p = q * d" - and "pderiv p = e * d" - and "d = r * p + s * pderiv p" - shows "rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" -proof - - from \pderiv p \ 0\ have "p \ 0" by auto - with \p = q * d\ have "q \ 0" by simp - have "\a. order a q = (if order a p = 0 then 0 else 1)" - using assms by (rule poly_squarefree_decomp_order2) - with \p \ 0\ \q \ 0\ show ?thesis - by (simp add: rsquarefree_def order_root) -qed - -end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Polynomial.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1,13 +1,15 @@ (* Title: HOL/Library/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin + Author: Amine Chaieb Author: Florian Haftmann *) section \Polynomials as type over a ring structure\ theory Polynomial -imports Main "~~/src/HOL/GCD" "~~/src/HOL/Library/More_List" "~~/src/HOL/Library/Infinite_Set" +imports Main "~~/src/HOL/Deriv" "~~/src/HOL/Library/More_List" + "~~/src/HOL/Library/Infinite_Set" begin subsection \Auxiliary: operations for lists (later) representing coefficients\ @@ -1551,7 +1553,7 @@ assumes "a \ 0" shows "is_unit (monom a 0)" proof - from assms show "1 = monom a 0 * monom (1 / a) 0" + from assms show "1 = monom a 0 * monom (inverse a) 0" by (simp add: mult_monom) qed @@ -1602,7 +1604,7 @@ begin definition normalize_poly :: "'a poly \ 'a poly" - where "normalize_poly p = smult (1 / coeff p (degree p)) p" + where "normalize_poly p = smult (inverse (coeff p (degree p))) p" definition unit_factor_poly :: "'a poly \ 'a poly" where "unit_factor_poly p = monom (coeff p (degree p)) 0" @@ -1611,8 +1613,9 @@ proof fix p :: "'a poly" show "unit_factor p * normalize p = p" - by (simp add: normalize_poly_def unit_factor_poly_def) - (simp only: mult_smult_left [symmetric] smult_monom, simp) + by (cases "p = 0") + (simp_all add: normalize_poly_def unit_factor_poly_def, + simp only: mult_smult_left [symmetric] smult_monom, simp) next show "normalize 0 = (0::'a poly)" by (simp add: normalize_poly_def) @@ -1645,6 +1648,21 @@ end +lemma unit_factor_monom [simp]: + "unit_factor (monom a n) = + (if a = 0 then 0 else monom a 0)" + by (simp add: unit_factor_poly_def degree_monom_eq) + +lemma unit_factor_pCons [simp]: + "unit_factor (pCons a p) = + (if p = 0 then monom a 0 else unit_factor p)" + by (simp add: unit_factor_poly_def) + +lemma normalize_monom [simp]: + "normalize (monom a n) = + (if a = 0 then 0 else monom 1 n)" + by (simp add: normalize_poly_def degree_monom_eq smult_monom) + lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" using pdivmod_rel [of x y] @@ -1884,137 +1902,6 @@ by (subst (asm) order_root) auto -subsection \GCD of polynomials\ - -instantiation poly :: (field) gcd -begin - -function gcd_poly :: "'a::field poly \ 'a poly \ 'a poly" -where - "gcd (x::'a poly) 0 = smult (inverse (coeff x (degree x))) x" -| "y \ 0 \ gcd (x::'a poly) y = gcd y (x mod y)" -by auto - -termination "gcd :: _ poly \ _" -by (relation "measure (\(x, y). if y = 0 then 0 else Suc (degree y))") - (auto dest: degree_mod_less) - -declare gcd_poly.simps [simp del] - -definition lcm_poly :: "'a::field poly \ 'a poly \ 'a poly" -where - "lcm_poly a b = a * b div smult (coeff a (degree a) * coeff b (degree b)) (gcd a b)" - -instance .. - -end - -lemma - fixes x y :: "_ poly" - shows poly_gcd_dvd1 [iff]: "gcd x y dvd x" - and poly_gcd_dvd2 [iff]: "gcd x y dvd y" - apply (induct x y rule: gcd_poly.induct) - apply (simp_all add: gcd_poly.simps) - apply (fastforce simp add: smult_dvd_iff dest: inverse_zero_imp_zero) - apply (blast dest: dvd_mod_imp_dvd) - done - -lemma poly_gcd_greatest: - fixes k x y :: "_ poly" - shows "k dvd x \ k dvd y \ k dvd gcd x y" - by (induct x y rule: gcd_poly.induct) - (simp_all add: gcd_poly.simps dvd_mod dvd_smult) - -lemma dvd_poly_gcd_iff [iff]: - fixes k x y :: "_ poly" - shows "k dvd gcd x y \ k dvd x \ k dvd y" - by (auto intro!: poly_gcd_greatest intro: dvd_trans [of _ "gcd x y"]) - -lemma poly_gcd_monic: - fixes x y :: "_ poly" - shows "coeff (gcd x y) (degree (gcd x y)) = - (if x = 0 \ y = 0 then 0 else 1)" - by (induct x y rule: gcd_poly.induct) - (simp_all add: gcd_poly.simps nonzero_imp_inverse_nonzero) - -lemma poly_gcd_zero_iff [simp]: - fixes x y :: "_ poly" - shows "gcd x y = 0 \ x = 0 \ y = 0" - by (simp only: dvd_0_left_iff [symmetric] dvd_poly_gcd_iff) - -lemma poly_gcd_0_0 [simp]: - "gcd (0::_ poly) 0 = 0" - by simp - -lemma poly_dvd_antisym: - fixes p q :: "'a::idom poly" - assumes coeff: "coeff p (degree p) = coeff q (degree q)" - assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" -proof (cases "p = 0") - case True with coeff show "p = q" by simp -next - case False with coeff have "q \ 0" by auto - have degree: "degree p = degree q" - using \p dvd q\ \q dvd p\ \p \ 0\ \q \ 0\ - by (intro order_antisym dvd_imp_degree_le) - - from \p dvd q\ obtain a where a: "q = p * a" .. - with \q \ 0\ have "a \ 0" by auto - with degree a \p \ 0\ have "degree a = 0" - by (simp add: degree_mult_eq) - with coeff a show "p = q" - by (cases a, auto split: if_splits) -qed - -lemma poly_gcd_unique: - fixes d x y :: "_ poly" - assumes dvd1: "d dvd x" and dvd2: "d dvd y" - and greatest: "\k. k dvd x \ k dvd y \ k dvd d" - and monic: "coeff d (degree d) = (if x = 0 \ y = 0 then 0 else 1)" - shows "gcd x y = d" -proof - - have "coeff (gcd x y) (degree (gcd x y)) = coeff d (degree d)" - by (simp_all add: poly_gcd_monic monic) - moreover have "gcd x y dvd d" - using poly_gcd_dvd1 poly_gcd_dvd2 by (rule greatest) - moreover have "d dvd gcd x y" - using dvd1 dvd2 by (rule poly_gcd_greatest) - ultimately show ?thesis - by (rule poly_dvd_antisym) -qed - -interpretation gcd_poly: abel_semigroup "gcd :: _ poly \ _" -proof - fix x y z :: "'a poly" - show "gcd (gcd x y) z = gcd x (gcd y z)" - by (rule poly_gcd_unique) (auto intro: dvd_trans simp add: poly_gcd_monic) - show "gcd x y = gcd y x" - by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) -qed - -lemmas poly_gcd_assoc = gcd_poly.assoc -lemmas poly_gcd_commute = gcd_poly.commute -lemmas poly_gcd_left_commute = gcd_poly.left_commute - -lemmas poly_gcd_ac = poly_gcd_assoc poly_gcd_commute poly_gcd_left_commute - -lemma poly_gcd_1_left [simp]: "gcd 1 y = (1 :: _ poly)" -by (rule poly_gcd_unique) simp_all - -lemma poly_gcd_1_right [simp]: "gcd x 1 = (1 :: _ poly)" -by (rule poly_gcd_unique) simp_all - -lemma poly_gcd_minus_left [simp]: "gcd (- x) y = gcd x (y :: _ poly)" -by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) - -lemma poly_gcd_minus_right [simp]: "gcd x (- y) = gcd x (y :: _ poly)" -by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) - -lemma poly_gcd_code [code]: - "gcd x y = (if y = 0 then smult (inverse (coeff x (degree x))) x else gcd y (x mod (y :: _ poly)))" - by (simp add: gcd_poly.simps) - - subsection \Additional induction rules on polynomials\ text \ @@ -2326,8 +2213,630 @@ lemma lead_coeff_nonzero: "p \ 0 \ lead_coeff p \ 0" by (simp add: lead_coeff_def) + + +subsection \Derivatives of univariate polynomials\ + +function pderiv :: "('a :: semidom) poly \ 'a poly" +where + [simp del]: "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" + by (auto intro: pCons_cases) + +termination pderiv + by (relation "measure degree") simp_all + +lemma pderiv_0 [simp]: + "pderiv 0 = 0" + using pderiv.simps [of 0 0] by simp + +lemma pderiv_pCons: + "pderiv (pCons a p) = p + pCons 0 (pderiv p)" + by (simp add: pderiv.simps) + +lemma pderiv_1 [simp]: "pderiv 1 = 0" + unfolding one_poly_def by (simp add: pderiv_pCons) + +lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" + and pderiv_numeral [simp]: "pderiv (numeral m) = 0" + by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) + +lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" + by (induct p arbitrary: n) + (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) + +fun pderiv_coeffs_code :: "('a :: semidom) \ 'a list \ 'a list" where + "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" +| "pderiv_coeffs_code f [] = []" + +definition pderiv_coeffs :: "('a :: semidom) list \ 'a list" where + "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" + +(* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) +lemma pderiv_coeffs_code: + "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * (nth_default 0 xs n)" +proof (induct xs arbitrary: f n) + case (Cons x xs f n) + show ?case + proof (cases n) + case 0 + thus ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0", auto simp: cCons_def) + next + case (Suc m) note n = this + show ?thesis + proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") + case False + hence "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = + nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" + by (auto simp: cCons_def n) + also have "\ = (f + of_nat n) * (nth_default 0 xs m)" + unfolding Cons by (simp add: n add_ac) + finally show ?thesis by (simp add: n) + next + case True + { + fix g + have "pderiv_coeffs_code g xs = [] \ g + of_nat m = 0 \ nth_default 0 xs m = 0" + proof (induct xs arbitrary: g m) + case (Cons x xs g) + from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" + and g: "(g = 0 \ x = 0)" + by (auto simp: cCons_def split: if_splits) + note IH = Cons(1)[OF empty] + from IH[of m] IH[of "m - 1"] g + show ?case by (cases m, auto simp: field_simps) + qed simp + } note empty = this + from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" + by (auto simp: cCons_def n) + moreover have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" using True + by (simp add: n, insert empty[of "f+1"], auto simp: field_simps) + ultimately show ?thesis by simp + qed + qed +qed simp + +lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\ i. f (Suc i)) [0 ..< n]" + by (induct n arbitrary: f, auto) + +lemma coeffs_pderiv_code [code abstract]: + "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def +proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) + case (1 n) + have id: "coeff p (Suc n) = nth_default 0 (map (\i. coeff p (Suc i)) [0.. degree p = 0" + apply (rule iffI) + apply (cases p, simp) + apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) + apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) + done + +lemma degree_pderiv: "degree (pderiv (p :: 'a poly)) = degree p - 1" + apply (rule order_antisym [OF degree_le]) + apply (simp add: coeff_pderiv coeff_eq_0) + apply (cases "degree p", simp) + apply (rule le_degree) + apply (simp add: coeff_pderiv del: of_nat_Suc) + apply (metis degree_0 leading_coeff_0_iff nat.distinct(1)) + done + +lemma not_dvd_pderiv: + assumes "degree (p :: 'a poly) \ 0" + shows "\ p dvd pderiv p" +proof + assume dvd: "p dvd pderiv p" + then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto + from dvd have le: "degree p \ degree (pderiv p)" + by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) + from this[unfolded degree_pderiv] assms show False by auto +qed + +lemma dvd_pderiv_iff [simp]: "(p :: 'a poly) dvd pderiv p \ degree p = 0" + using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) + +end + +lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" +by (simp add: pderiv_pCons) + +lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q" +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" +by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) + +lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" +by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) + +lemma pderiv_power_Suc: + "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" +apply (induct n) +apply simp +apply (subst power_Suc) +apply (subst pderiv_mult) +apply (erule ssubst) +apply (simp only: of_nat_Suc smult_add_left smult_1_left) +apply (simp add: algebra_simps) +done + +lemma pderiv_setprod: "pderiv (setprod f (as)) = + (\a \ as. setprod f (as - {a}) * pderiv (f a))" +proof (induct as rule: infinite_finite_induct) + case (insert a as) + hence id: "setprod f (insert a as) = f a * setprod f as" + "\ g. setsum g (insert a as) = g a + setsum g as" + "insert a as - {a} = as" + by auto + { + fix b + assume "b \ as" + hence id2: "insert a as - {b} = insert a (as - {b})" using \a \ as\ by auto + have "setprod f (insert a as - {b}) = f a * setprod f (as - {b})" + unfolding id2 + by (subst setprod.insert, insert insert, auto) + } note id2 = this + show ?case + unfolding id pderiv_mult insert(3) setsum_right_distrib + by (auto simp add: ac_simps id2 intro!: setsum.cong) +qed auto + +lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" +by (rule DERIV_cong, rule DERIV_pow, simp) +declare DERIV_pow2 [simp] DERIV_pow [simp] + +lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D" +by (rule DERIV_cong, rule DERIV_add, auto) + +lemma poly_DERIV [simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" + by (induct p, auto intro!: derivative_eq_intros simp add: pderiv_pCons) + +lemma continuous_on_poly [continuous_intros]: + fixes p :: "'a :: {real_normed_field} poly" + assumes "continuous_on A f" + shows "continuous_on A (\x. poly p (f x))" +proof - + have "continuous_on A (\x. (\i\degree p. (f x) ^ i * coeff p i))" + by (intro continuous_intros assms) + also have "\ = (\x. poly p (f x))" by (intro ext) (simp add: poly_altdef mult_ac) + finally show ?thesis . +qed + +text\Consequences of the derivative theorem above\ + +lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (at x::real filter)" +apply (simp add: real_differentiable_def) +apply (blast intro: poly_DERIV) +done + +lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)" +by (rule poly_DERIV [THEN DERIV_isCont]) + +lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |] + ==> \x. a < x & x < b & (poly p x = 0)" +using IVT_objl [of "poly p" a 0 b] +by (auto simp add: order_le_less) + +lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |] + ==> \x. a < x & x < b & (poly p x = 0)" +by (insert poly_IVT_pos [where p = "- p" ]) simp + +lemma poly_IVT: + fixes p::"real poly" + assumes "ax>a. x < b \ poly p x = 0" +by (metis assms(1) assms(2) less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) + +lemma poly_MVT: "(a::real) < b ==> + \x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" +using MVT [of a b "poly p"] +apply auto +apply (rule_tac x = z in exI) +apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique]) +done + +lemma poly_MVT': + assumes "{min a b..max a b} \ A" + shows "\x\A. poly p b - poly p a = (b - a) * poly (pderiv p) (x::real)" +proof (cases a b rule: linorder_cases) + case less + from poly_MVT[OF less, of p] guess x by (elim exE conjE) + thus ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) + +next + case greater + from poly_MVT[OF greater, of p] guess x by (elim exE conjE) + thus ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) +qed (insert assms, auto) + +lemma poly_pinfty_gt_lc: + fixes p:: "real poly" + assumes "lead_coeff p > 0" + shows "\ n. \ x \ n. poly p x \ lead_coeff p" using assms +proof (induct p) + case 0 + thus ?case by auto +next + case (pCons a p) + have "\a\0;p=0\ \ ?case" by auto + moreover have "p\0 \ ?case" + proof - + assume "p\0" + then obtain n1 where gte_lcoeff:"\x\n1. lead_coeff p \ poly p x" using that pCons by auto + have gt_0:"lead_coeff p >0" using pCons(3) \p\0\ by auto + def n\"max n1 (1+ \a\/(lead_coeff p))" + show ?thesis + proof (rule_tac x=n in exI,rule,rule) + fix x assume "n \ x" + hence "lead_coeff p \ poly p x" + using gte_lcoeff unfolding n_def by auto + hence " \a\/(lead_coeff p) \ \a\/(poly p x)" and "poly p x>0" using gt_0 + by (intro frac_le,auto) + hence "x\1+ \a\/(poly p x)" using \n\x\[unfolded n_def] by auto + thus "lead_coeff (pCons a p) \ poly (pCons a p) x" + using \lead_coeff p \ poly p x\ \poly p x>0\ \p\0\ + by (auto simp add:field_simps) + qed + qed + ultimately show ?case by fastforce +qed + + +subsection \Algebraic numbers\ + +text \ + Algebraic numbers can be defined in two equivalent ways: all real numbers that are + roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry + uses the rational definition, but we need the integer definition. + + The equivalence is obvious since any rational polynomial can be multiplied with the + LCM of its coefficients, yielding an integer polynomial with the same roots. +\ +subsection \Algebraic numbers\ + +definition algebraic :: "'a :: field_char_0 \ bool" where + "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" + +lemma algebraicI: + assumes "\i. coeff p i \ \" "p \ 0" "poly p x = 0" + shows "algebraic x" + using assms unfolding algebraic_def by blast +lemma algebraicE: + assumes "algebraic x" + obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" + using assms unfolding algebraic_def by blast + +lemma quotient_of_denom_pos': "snd (quotient_of x) > 0" + using quotient_of_denom_pos[OF surjective_pairing] . +lemma of_int_div_in_Ints: + "b dvd a \ of_int a div of_int b \ (\ :: 'a :: ring_div set)" +proof (cases "of_int b = (0 :: 'a)") + assume "b dvd a" "of_int b \ (0::'a)" + then obtain c where "a = b * c" by (elim dvdE) + with \of_int b \ (0::'a)\ show ?thesis by simp +qed auto + +lemma of_int_divide_in_Ints: + "b dvd a \ of_int a / of_int b \ (\ :: 'a :: field set)" +proof (cases "of_int b = (0 :: 'a)") + assume "b dvd a" "of_int b \ (0::'a)" + then obtain c where "a = b * c" by (elim dvdE) + with \of_int b \ (0::'a)\ show ?thesis by simp +qed auto + +lemma algebraic_altdef: + fixes p :: "'a :: field_char_0 poly" + shows "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" +proof safe + fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" + def cs \ "coeffs p" + from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast + then obtain f where f: "\i. coeff p i = of_rat (f (coeff p i))" + by (subst (asm) bchoice_iff) blast + def cs' \ "map (quotient_of \ f) (coeffs p)" + def d \ "Lcm (set (map snd cs'))" + def p' \ "smult (of_int d) p" + + have "\n. coeff p' n \ \" + proof + fix n :: nat + show "coeff p' n \ \" + proof (cases "n \ degree p") + case True + def c \ "coeff p n" + def a \ "fst (quotient_of (f (coeff p n)))" and b \ "snd (quotient_of (f (coeff p n)))" + have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp + have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) + also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def + by (subst quotient_of_div [of "f (coeff p n)", symmetric]) + (simp_all add: f [symmetric]) + also have "of_int d * \ = of_rat (of_int (a*d) / of_int b)" + by (simp add: of_rat_mult of_rat_divide) + also from nz True have "b \ snd ` set cs'" unfolding cs'_def + by (force simp: o_def b_def coeffs_def simp del: upt_Suc) + hence "b dvd (a * d)" unfolding d_def by simp + hence "of_int (a * d) / of_int b \ (\ :: rat set)" + by (rule of_int_divide_in_Ints) + hence "of_rat (of_int (a * d) / of_int b) \ \" by (elim Ints_cases) auto + finally show ?thesis . + qed (auto simp: p'_def not_le coeff_eq_0) + qed + + moreover have "set (map snd cs') \ {0<..}" + unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) + hence "d \ 0" unfolding d_def by (induction cs') simp_all + with nz have "p' \ 0" by (simp add: p'_def) + moreover from root have "poly p' x = 0" by (simp add: p'_def) + ultimately show "algebraic x" unfolding algebraic_def by blast +next + + assume "algebraic x" + then obtain p where p: "\i. coeff p i \ \" "poly p x = 0" "p \ 0" + by (force simp: algebraic_def) + moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp + ultimately show "(\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" by auto +qed + + +text\Lemmas for Derivatives\ + +lemma order_unique_lemma: + fixes p :: "'a::idom poly" + assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" + shows "n = order a p" +unfolding Polynomial.order_def +apply (rule Least_equality [symmetric]) +apply (fact assms) +apply (rule classical) +apply (erule notE) +unfolding not_less_eq_eq +using assms(1) apply (rule power_le_dvd) +apply assumption +done + +lemma lemma_order_pderiv1: + "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" +apply (simp only: pderiv_mult pderiv_power_Suc) +apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons) +done + +lemma lemma_order_pderiv: + fixes p :: "'a :: field_char_0 poly" + assumes n: "0 < n" + and pd: "pderiv p \ 0" + and pe: "p = [:- a, 1:] ^ n * q" + and nd: "~ [:- a, 1:] dvd q" + shows "n = Suc (order a (pderiv p))" +using n +proof - + have "pderiv ([:- a, 1:] ^ n * q) \ 0" + using assms by auto + obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" + using assms by (cases n) auto + have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" + by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) + have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" + proof (rule order_unique_lemma) + show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" + apply (subst lemma_order_pderiv1) + apply (rule dvd_add) + apply (metis dvdI dvd_mult2 power_Suc2) + apply (metis dvd_smult dvd_triv_right) + done + next + show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" + apply (subst lemma_order_pderiv1) + by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) + qed + then show ?thesis + by (metis \n = Suc n'\ pe) +qed + +lemma order_decomp: + assumes "p \ 0" + shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" +proof - + from assms have A: "[:- a, 1:] ^ order a p dvd p" + and B: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) + from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" .. + with B have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" + by simp + then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" + by simp + then have D: "\ [:- a, 1:] dvd q" + using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] + by auto + from C D show ?thesis by blast +qed + +lemma order_pderiv: + "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ \ + (order a p = Suc (order a (pderiv p)))" +apply (case_tac "p = 0", simp) +apply (drule_tac a = a and p = p in order_decomp) +using neq0_conv +apply (blast intro: lemma_order_pderiv) +done + +lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" +proof - + def i \ "order a p" + def j \ "order a q" + def t \ "[:-a, 1:]" + have t_dvd_iff: "\u. t dvd u \ poly u a = 0" + unfolding t_def by (simp add: dvd_iff_poly_eq_0) + assume "p * q \ 0" + then show "order a (p * q) = i + j" + apply clarsimp + apply (drule order [where a=a and p=p, folded i_def t_def]) + apply (drule order [where a=a and p=q, folded j_def t_def]) + apply clarify + apply (erule dvdE)+ + apply (rule order_unique_lemma [symmetric], fold t_def) + apply (simp_all add: power_add t_dvd_iff) + done +qed + +lemma order_smult: + assumes "c \ 0" + shows "order x (smult c p) = order x p" +proof (cases "p = 0") + case False + have "smult c p = [:c:] * p" by simp + also from assms False have "order x \ = order x [:c:] + order x p" + by (subst order_mult) simp_all + also from assms have "order x [:c:] = 0" by (intro order_0I) auto + finally show ?thesis by simp +qed simp + +(* Next two lemmas contributed by Wenda Li *) +lemma order_1_eq_0 [simp]:"order x 1 = 0" + by (metis order_root poly_1 zero_neq_one) + +lemma order_power_n_n: "order a ([:-a,1:]^n)=n" +proof (induct n) (*might be proved more concisely using nat_less_induct*) + case 0 + thus ?case by (metis order_root poly_1 power_0 zero_neq_one) +next + case (Suc n) + have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" + by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral + one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) + moreover have "order a [:-a,1:]=1" unfolding order_def + proof (rule Least_equality,rule ccontr) + assume "\ \ [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" + hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp + hence "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:] )" + by (rule dvd_imp_degree_le,auto) + thus False by auto + next + fix y assume asm:"\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" + show "1 \ y" + proof (rule ccontr) + assume "\ 1 \ y" + hence "y=0" by auto + hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto + thus False using asm by auto + qed + qed + ultimately show ?case using Suc by auto +qed + +text\Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\ + +lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" +apply (cases "p = 0", auto) +apply (drule order_2 [where a=a and p=p]) +apply (metis not_less_eq_eq power_le_dvd) +apply (erule power_le_dvd [OF order_1]) +done + +lemma poly_squarefree_decomp_order: + assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" + and p: "p = q * d" + and p': "pderiv p = e * d" + and d: "d = r * p + s * pderiv p" + shows "order a q = (if order a p = 0 then 0 else 1)" +proof (rule classical) + assume 1: "order a q \ (if order a p = 0 then 0 else 1)" + from \pderiv p \ 0\ have "p \ 0" by auto + with p have "order a p = order a q + order a d" + by (simp add: order_mult) + with 1 have "order a p \ 0" by (auto split: if_splits) + have "order a (pderiv p) = order a e + order a d" + using \pderiv p \ 0\ \pderiv p = e * d\ by (simp add: order_mult) + have "order a p = Suc (order a (pderiv p))" + using \pderiv p \ 0\ \order a p \ 0\ by (rule order_pderiv) + have "d \ 0" using \p \ 0\ \p = q * d\ by simp + have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" + apply (simp add: d) + apply (rule dvd_add) + apply (rule dvd_mult) + apply (simp add: order_divides \p \ 0\ + \order a p = Suc (order a (pderiv p))\) + apply (rule dvd_mult) + apply (simp add: order_divides) + done + then have "order a (pderiv p) \ order a d" + using \d \ 0\ by (simp add: order_divides) + show ?thesis + using \order a p = order a q + order a d\ + using \order a (pderiv p) = order a e + order a d\ + using \order a p = Suc (order a (pderiv p))\ + using \order a (pderiv p) \ order a d\ + by auto +qed + +lemma poly_squarefree_decomp_order2: + "\pderiv p \ (0 :: 'a :: field_char_0 poly); + p = q * d; + pderiv p = e * d; + d = r * p + s * pderiv p + \ \ \a. order a q = (if order a p = 0 then 0 else 1)" +by (blast intro: poly_squarefree_decomp_order) + +lemma order_pderiv2: + "\pderiv p \ 0; order a (p :: 'a :: field_char_0 poly) \ 0\ + \ (order a (pderiv p) = n) = (order a p = Suc n)" +by (auto dest: order_pderiv) + +definition + rsquarefree :: "'a::idom poly => bool" where + "rsquarefree p = (p \ 0 & (\a. (order a p = 0) | (order a p = 1)))" + +lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h :: 'a :: {semidom,semiring_char_0}:]" +apply (simp add: pderiv_eq_0_iff) +apply (case_tac p, auto split: if_splits) +done + +lemma rsquarefree_roots: + fixes p :: "'a :: field_char_0 poly" + shows "rsquarefree p = (\a. \(poly p a = 0 \ poly (pderiv p) a = 0))" +apply (simp add: rsquarefree_def) +apply (case_tac "p = 0", simp, simp) +apply (case_tac "pderiv p = 0") +apply simp +apply (drule pderiv_iszero, clarsimp) +apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) +apply (force simp add: order_root order_pderiv2) +done + +lemma poly_squarefree_decomp: + assumes "pderiv (p :: 'a :: field_char_0 poly) \ 0" + and "p = q * d" + and "pderiv p = e * d" + and "d = r * p + s * pderiv p" + shows "rsquarefree q & (\a. (poly q a = 0) = (poly p a = 0))" +proof - + from \pderiv p \ 0\ have "p \ 0" by auto + with \p = q * d\ have "q \ 0" by simp + have "\a. order a q = (if order a p = 0 then 0 else 1)" + using assms by (rule poly_squarefree_decomp_order2) + with \p \ 0\ \q \ 0\ show ?thesis + by (simp add: rsquarefree_def order_root) +qed + no_notation cCons (infixr "##" 65) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Polynomial_GCD_euclidean.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Polynomial_GCD_euclidean.thy Wed Feb 17 23:29:35 2016 +0100 @@ -0,0 +1,140 @@ +(* Title: HOL/Library/Polynomial_GCD_euclidean.thy + Author: Brian Huffman + Author: Clemens Ballarin + Author: Amine Chaieb + Author: Florian Haftmann +*) + +section \GCD and LCM on polynomials over a field\ + +theory Polynomial_GCD_euclidean +imports Main "~~/src/HOL/GCD" "~~/src/HOL/Library/Polynomial" +begin + +subsection \GCD of polynomials\ + +instantiation poly :: (field) gcd +begin + +function gcd_poly :: "'a::field poly \ 'a poly \ 'a poly" +where + "gcd (x::'a poly) 0 = smult (inverse (coeff x (degree x))) x" +| "y \ 0 \ gcd (x::'a poly) y = gcd y (x mod y)" +by auto + +termination "gcd :: _ poly \ _" +by (relation "measure (\(x, y). if y = 0 then 0 else Suc (degree y))") + (auto dest: degree_mod_less) + +declare gcd_poly.simps [simp del] + +definition lcm_poly :: "'a::field poly \ 'a poly \ 'a poly" +where + "lcm_poly a b = a * b div smult (coeff a (degree a) * coeff b (degree b)) (gcd a b)" + +instance .. + +end + +lemma + fixes x y :: "_ poly" + shows poly_gcd_dvd1 [iff]: "gcd x y dvd x" + and poly_gcd_dvd2 [iff]: "gcd x y dvd y" + apply (induct x y rule: gcd_poly.induct) + apply (simp_all add: gcd_poly.simps) + apply (fastforce simp add: smult_dvd_iff dest: inverse_zero_imp_zero) + apply (blast dest: dvd_mod_imp_dvd) + done + +lemma poly_gcd_greatest: + fixes k x y :: "_ poly" + shows "k dvd x \ k dvd y \ k dvd gcd x y" + by (induct x y rule: gcd_poly.induct) + (simp_all add: gcd_poly.simps dvd_mod dvd_smult) + +lemma dvd_poly_gcd_iff [iff]: + fixes k x y :: "_ poly" + shows "k dvd gcd x y \ k dvd x \ k dvd y" + by (auto intro!: poly_gcd_greatest intro: dvd_trans [of _ "gcd x y"]) + +lemma poly_gcd_monic: + fixes x y :: "_ poly" + shows "coeff (gcd x y) (degree (gcd x y)) = + (if x = 0 \ y = 0 then 0 else 1)" + by (induct x y rule: gcd_poly.induct) + (simp_all add: gcd_poly.simps nonzero_imp_inverse_nonzero) + +lemma poly_gcd_zero_iff [simp]: + fixes x y :: "_ poly" + shows "gcd x y = 0 \ x = 0 \ y = 0" + by (simp only: dvd_0_left_iff [symmetric] dvd_poly_gcd_iff) + +lemma poly_gcd_0_0 [simp]: + "gcd (0::_ poly) 0 = 0" + by simp + +lemma poly_dvd_antisym: + fixes p q :: "'a::idom poly" + assumes coeff: "coeff p (degree p) = coeff q (degree q)" + assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" +proof (cases "p = 0") + case True with coeff show "p = q" by simp +next + case False with coeff have "q \ 0" by auto + have degree: "degree p = degree q" + using \p dvd q\ \q dvd p\ \p \ 0\ \q \ 0\ + by (intro order_antisym dvd_imp_degree_le) + + from \p dvd q\ obtain a where a: "q = p * a" .. + with \q \ 0\ have "a \ 0" by auto + with degree a \p \ 0\ have "degree a = 0" + by (simp add: degree_mult_eq) + with coeff a show "p = q" + by (cases a, auto split: if_splits) +qed + +lemma poly_gcd_unique: + fixes d x y :: "_ poly" + assumes dvd1: "d dvd x" and dvd2: "d dvd y" + and greatest: "\k. k dvd x \ k dvd y \ k dvd d" + and monic: "coeff d (degree d) = (if x = 0 \ y = 0 then 0 else 1)" + shows "gcd x y = d" +proof - + have "coeff (gcd x y) (degree (gcd x y)) = coeff d (degree d)" + by (simp_all add: poly_gcd_monic monic) + moreover have "gcd x y dvd d" + using poly_gcd_dvd1 poly_gcd_dvd2 by (rule greatest) + moreover have "d dvd gcd x y" + using dvd1 dvd2 by (rule poly_gcd_greatest) + ultimately show ?thesis + by (rule poly_dvd_antisym) +qed + +instance poly :: (field) semiring_gcd +proof + fix p q :: "'a::field poly" + show "normalize (gcd p q) = gcd p q" + by (induct p q rule: gcd_poly.induct) + (simp_all add: gcd_poly.simps normalize_poly_def) + show "lcm p q = normalize (p * q) div gcd p q" + by (simp add: coeff_degree_mult div_smult_left div_smult_right lcm_poly_def normalize_poly_def) + (metis (no_types, lifting) div_smult_right inverse_mult_distrib inverse_zero mult.commute pdivmod_rel pdivmod_rel_def smult_eq_0_iff) +qed simp_all + +lemma poly_gcd_1_left [simp]: "gcd 1 y = (1 :: _ poly)" +by (rule poly_gcd_unique) simp_all + +lemma poly_gcd_1_right [simp]: "gcd x 1 = (1 :: _ poly)" +by (rule poly_gcd_unique) simp_all + +lemma poly_gcd_minus_left [simp]: "gcd (- x) y = gcd x (y :: _ poly)" +by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) + +lemma poly_gcd_minus_right [simp]: "gcd x (- y) = gcd x (y :: _ poly)" +by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) + +lemma poly_gcd_code [code]: + "gcd x y = (if y = 0 then normalize x else gcd y (x mod (y :: _ poly)))" + by (simp add: gcd_poly.simps) + +end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Library/Product_Order.thy --- a/src/HOL/Library/Product_Order.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Library/Product_Order.thy Wed Feb 17 23:29:35 2016 +0100 @@ -179,7 +179,6 @@ instance prod :: (conditionally_complete_lattice, conditionally_complete_lattice) conditionally_complete_lattice by standard (force simp: less_eq_prod_def Inf_prod_def Sup_prod_def bdd_below_def bdd_above_def - INF_def SUP_def simp del: Inf_image_eq Sup_image_eq intro!: cInf_lower cSup_upper cInf_greatest cSup_least)+ instance prod :: (complete_lattice, complete_lattice) complete_lattice @@ -211,10 +210,10 @@ using snd_Inf [of "f ` A", symmetric] by (simp add: comp_def) lemma SUP_Pair: "(SUP x:A. (f x, g x)) = (SUP x:A. f x, SUP x:A. g x)" - unfolding SUP_def Sup_prod_def by (simp add: comp_def) + unfolding Sup_prod_def by (simp add: comp_def) lemma INF_Pair: "(INF x:A. (f x, g x)) = (INF x:A. f x, INF x:A. g x)" - unfolding INF_def Inf_prod_def by (simp add: comp_def) + unfolding Inf_prod_def by (simp add: comp_def) text \Alternative formulations for set infima and suprema over the product @@ -222,11 +221,11 @@ lemma INF_prod_alt_def: "INFIMUM A f = (INFIMUM A (fst o f), INFIMUM A (snd o f))" - unfolding INF_def Inf_prod_def by simp + unfolding Inf_prod_def by simp lemma SUP_prod_alt_def: "SUPREMUM A f = (SUPREMUM A (fst o f), SUPREMUM A (snd o f))" - unfolding SUP_def Sup_prod_def by simp + unfolding Sup_prod_def by simp subsection \Complete distributive lattices\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Lifting_Set.thy --- a/src/HOL/Lifting_Set.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Lifting_Set.thy Wed Feb 17 23:29:35 2016 +0100 @@ -138,7 +138,7 @@ lemma UNION_transfer [transfer_rule]: "(rel_set A ===> (A ===> rel_set B) ===> rel_set B) UNION UNION" - unfolding Union_image_eq [symmetric, abs_def] by transfer_prover + by transfer_prover lemma Ball_transfer [transfer_rule]: "(rel_set A ===> (A ===> op =) ===> op =) Ball Ball" @@ -176,11 +176,11 @@ lemma INF_parametric [transfer_rule]: "(rel_set A ===> (A ===> HOL.eq) ===> HOL.eq) INFIMUM INFIMUM" - unfolding INF_def [abs_def] by transfer_prover + by transfer_prover lemma SUP_parametric [transfer_rule]: "(rel_set R ===> (R ===> HOL.eq) ===> HOL.eq) SUPREMUM SUPREMUM" - unfolding SUP_def [abs_def] by transfer_prover + by transfer_prover subsubsection \Rules requiring bi-unique, bi-total or right-total relations\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/List.thy --- a/src/HOL/List.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/List.thy Wed Feb 17 23:29:35 2016 +0100 @@ -6774,7 +6774,7 @@ lemma set_relcomp [code]: "set xys O set yzs = set ([(fst xy, snd yz). xy \ xys, yz \ yzs, snd xy = fst yz])" - by (auto simp add: Bex_def) + by auto (auto simp add: Bex_def image_def) lemma wf_set [code]: "wf (set xs) = acyclic (set xs)" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Matrix_LP/ComputeFloat.thy --- a/src/HOL/Matrix_LP/ComputeFloat.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Matrix_LP/ComputeFloat.thy Wed Feb 17 23:29:35 2016 +0100 @@ -155,8 +155,8 @@ lemma not_true_eq_false: "(~ True) = False" by simp -lemmas powerarith = nat_numeral zpower_numeral_even - zpower_numeral_odd zpower_Pls +lemmas powerarith = nat_numeral power_numeral_even + power_numeral_odd zpower_Pls definition float :: "(int \ int) \ real" where "float = (\(a, b). real_of_int a * 2 powr real_of_int b)" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Matrix_LP/ComputeNumeral.thy --- a/src/HOL/Matrix_LP/ComputeNumeral.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Matrix_LP/ComputeNumeral.thy Wed Feb 17 23:29:35 2016 +0100 @@ -42,7 +42,7 @@ nat_numeral nat_0 nat_neg_numeral of_int_numeral of_int_neg_numeral of_int_0 -lemmas zpowerarith = zpower_numeral_even zpower_numeral_odd zpower_Pls int_pow_1 +lemmas zpowerarith = power_numeral_even power_numeral_odd zpower_Pls int_pow_1 (* div, mod *) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Wed Feb 17 23:29:35 2016 +0100 @@ -116,7 +116,7 @@ also assume "dist f g = 0" finally show "f = g" by (auto simp: Rep_bcontfun_inject[symmetric] Abs_bcontfun_inverse) - qed (auto simp: dist_bcontfun_def SUP_def simp del: Sup_image_eq intro!: cSup_eq) + qed (auto simp: dist_bcontfun_def intro!: cSup_eq) show "dist f g \ dist f h + dist g h" proof (subst dist_bcontfun_def, safe intro!: cSUP_least) fix x @@ -374,7 +374,7 @@ ultimately show "norm (a *\<^sub>R f) = \a\ * norm f" by (simp add: norm_bcontfun_def dist_bcontfun_def norm_conv_dist Abs_bcontfun_inverse - zero_bcontfun_def const_bcontfun SUP_def del: Sup_image_eq) + zero_bcontfun_def const_bcontfun image_image) presburger qed qed (auto simp: norm_bcontfun_def sgn_bcontfun_def) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Wed Feb 17 23:29:35 2016 +0100 @@ -6690,7 +6690,7 @@ using \open s\ apply (simp add: open_contains_ball Ball_def) apply (erule all_forward) - using "*" by blast + using "*" by auto blast+ have 2: "closedin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" using assms by (auto intro: continuous_closed_in_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Wed Feb 17 23:29:35 2016 +0100 @@ -421,13 +421,13 @@ lemma interval_upperbound[simp]: "\i\Basis. a\i \ b\i \ interval_upperbound (cbox a b) = (b::'a::euclidean_space)" - unfolding interval_upperbound_def euclidean_representation_setsum cbox_def SUP_def + unfolding interval_upperbound_def euclidean_representation_setsum cbox_def by (safe intro!: cSup_eq) auto lemma interval_lowerbound[simp]: "\i\Basis. a\i \ b\i \ interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" - unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def INF_def + unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def by (safe intro!: cInf_eq) auto lemmas interval_bounds = interval_upperbound interval_lowerbound @@ -436,7 +436,7 @@ fixes X::"real set" shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X" and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X" - by (auto simp: interval_upperbound_def interval_lowerbound_def SUP_def INF_def) + by (auto simp: interval_upperbound_def interval_lowerbound_def) lemma interval_bounds'[simp]: assumes "cbox a b \ {}" @@ -496,7 +496,7 @@ using assms box_ne_empty(1) content_cbox by blast lemma content_real: "a \ b \ content {a..b} = b - a" - by (auto simp: interval_upperbound_def interval_lowerbound_def SUP_def INF_def content_def) + by (auto simp: interval_upperbound_def interval_lowerbound_def content_def) lemma abs_eq_content: "\y - x\ = (if x\y then content {x .. y} else content {y..x})" by (auto simp: content_real) @@ -856,7 +856,7 @@ by auto show "\?A = s1 \ s2" apply (rule set_eqI) - unfolding * and Union_image_eq UN_iff + unfolding * and UN_iff using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] apply auto done @@ -9394,7 +9394,6 @@ prefer 3 apply assumption apply rule - apply (rule finite_imageI) apply (rule r) apply safe apply (drule qq) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Wed Feb 17 23:29:35 2016 +0100 @@ -2752,7 +2752,7 @@ lemma infnorm_Max: fixes x :: "'a::euclidean_space" shows "infnorm x = Max ((\i. \x \ i\) ` Basis)" - by (simp add: infnorm_def infnorm_set_image cSup_eq_Max del: Sup_image_eq) + by (simp add: infnorm_def infnorm_set_image cSup_eq_Max) lemma infnorm_set_lemma: fixes x :: "'a::euclidean_space" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Wed Feb 17 23:29:35 2016 +0100 @@ -38,13 +38,11 @@ assume "X \ {}" hence "\i. (\x. x \ i) ` X \ {}" by simp thus "(\x. x \ X \ z \ x) \ z \ Inf X" "(\x. x \ X \ x \ z) \ Sup X \ z" - by (auto simp: eucl_Inf eucl_Sup eucl_le Inf_class.INF_def Sup_class.SUP_def - simp del: Inf_class.Inf_image_eq Sup_class.Sup_image_eq + by (auto simp: eucl_Inf eucl_Sup eucl_le intro!: cInf_greatest cSup_least) qed (force intro!: cInf_lower cSup_upper simp: bdd_below_def bdd_above_def preorder_class.bdd_below_def preorder_class.bdd_above_def - eucl_Inf eucl_Sup eucl_le Inf_class.INF_def Sup_class.SUP_def - simp del: Inf_class.Inf_image_eq Sup_class.Sup_image_eq)+ + eucl_Inf eucl_Sup eucl_le)+ lemma inner_Basis_inf_left: "i \ Basis \ inf x y \ i = inf (x \ i) (y \ i)" and inner_Basis_sup_left: "i \ Basis \ sup x y \ i = sup (x \ i) (y \ i)" @@ -89,7 +87,7 @@ shows "Sup s = X" using assms unfolding eucl_Sup euclidean_representation_setsum - by (auto simp: Sup_class.SUP_def simp del: Sup_class.Sup_image_eq intro!: conditionally_complete_lattice_class.cSup_eq_maximum) + by (auto intro!: conditionally_complete_lattice_class.cSup_eq_maximum) lemma Inf_eq_minimum_componentwise: assumes i: "\b. b \ Basis \ X \ b = i b \ b" @@ -98,7 +96,7 @@ shows "Inf s = X" using assms unfolding eucl_Inf euclidean_representation_setsum - by (auto simp: Inf_class.INF_def simp del: Inf_class.Inf_image_eq intro!: conditionally_complete_lattice_class.cInf_eq_minimum) + by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum) end @@ -117,10 +115,9 @@ and x: "x \ X" "s = x \ b" "\y. y \ X \ x \ b \ y \ b" by auto hence "Inf ?proj = x \ b" - by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum simp del: Inf_class.Inf_image_eq) + by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum) hence "x \ b = Inf X \ b" - by (auto simp: eucl_Inf Inf_class.INF_def inner_setsum_left inner_Basis if_distrib \b \ Basis\ setsum.delta - simp del: Inf_class.Inf_image_eq + by (auto simp: eucl_Inf inner_setsum_left inner_Basis if_distrib \b \ Basis\ setsum.delta cong: if_cong) with x show "\x. x \ X \ x \ b = Inf X \ b \ (\y. y \ X \ x \ b \ y \ b)" by blast qed @@ -140,10 +137,10 @@ and x: "x \ X" "s = x \ b" "\y. y \ X \ y \ b \ x \ b" by auto hence "Sup ?proj = x \ b" - by (auto intro!: cSup_eq_maximum simp del: Sup_image_eq) + by (auto intro!: cSup_eq_maximum) hence "x \ b = Sup X \ b" - by (auto simp: eucl_Sup[where 'a='a] SUP_def inner_setsum_left inner_Basis if_distrib \b \ Basis\ setsum.delta - simp del: Sup_image_eq cong: if_cong) + by (auto simp: eucl_Sup[where 'a='a] inner_setsum_left inner_Basis if_distrib \b \ Basis\ setsum.delta + cong: if_cong) with x show "\x. x \ X \ x \ b = Sup X \ b \ (\y. y \ X \ y \ b \ x \ b)" by blast qed @@ -152,7 +149,7 @@ by (auto) instance real :: ordered_euclidean_space - by standard (auto simp: INF_def SUP_def) + by standard auto lemma in_Basis_prod_iff: fixes i::"'a::euclidean_space*'b::euclidean_space" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Feb 17 23:29:35 2016 +0100 @@ -532,7 +532,6 @@ lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" - unfolding Inter_image_eq [symmetric] apply (rule closedin_Inter) using assms apply auto @@ -605,7 +604,7 @@ ultimately have "?L (\K)" by blast } ultimately show ?thesis - unfolding subset_eq mem_Collect_eq istopology_def by blast + unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" @@ -2426,7 +2425,7 @@ fix y assume "y \ {x<..} \ I" with False bnd have "Inf (f ` ({x<..} \ I)) \ f y" - by (auto intro!: cInf_lower bdd_belowI2 simp del: Inf_image_eq) + by (auto intro!: cInf_lower bdd_belowI2) with a have "a < f y" by (blast intro: less_le_trans) } @@ -3802,7 +3801,7 @@ lemma compact_UN [intro]: "finite A \ (\x. x \ A \ compact (B x)) \ compact (\x\A. B x)" - unfolding SUP_def by (rule compact_Union) auto + by (rule compact_Union) auto lemma closed_inter_compact [intro]: assumes "closed s" @@ -4090,7 +4089,7 @@ by metis def X \ "\n. X' (from_nat_into A ` {.. n})" have X: "\n. X n \ U - (\i\n. from_nat_into A i)" - using \A \ {}\ unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into) + using \A \ {}\ unfolding X_def by (intro T) (auto intro: from_nat_into) then have "range X \ U" by auto with subseq[of X] obtain r x where "x \ U" and r: "subseq r" "(X \ r) \ x" @@ -4198,7 +4197,7 @@ then have s: "\x. x \ s \ \U. x\U \ open U \ finite (U \ t)" by metis have "s \ \C" using \t \ s\ - unfolding C_def Union_image_eq + unfolding C_def apply (safe dest!: s) apply (rule_tac a="U \ t" in UN_I) apply (auto intro!: interiorI simp add: finite_subset) @@ -4211,7 +4210,7 @@ by (rule countably_compactE) then obtain E where E: "E \ {F. finite F \ F \ t }" "finite E" and s: "s \ (\F\E. interior (F \ (- t)))" - by (metis (lifting) Union_image_eq finite_subset_image C_def) + by (metis (lifting) finite_subset_image C_def) from s \t \ s\ have "t \ \E" using interior_subset by blast moreover have "finite (\E)" @@ -4348,7 +4347,7 @@ from Rats_dense_in_real[OF \0 < e / 2\] obtain r where "r \ \" "0 < r" "r < e / 2" by auto from f[rule_format, of r] \0 < r\ \x \ s\ obtain k where "k \ f r" "x \ ball k r" - unfolding Union_image_eq by auto + by auto from \r \ \\ \0 < r\ \k \ f r\ have "ball k r \ K" by (auto simp: K_def) then show "\b\K. x \ b \ b \ s \ T" @@ -7412,7 +7411,7 @@ lemma diameter_cbox: fixes a b::"'a::euclidean_space" shows "(\i \ Basis. a \ i \ b \ i) \ diameter (cbox a b) = dist a b" - by (force simp add: diameter_def SUP_def simp del: Sup_image_eq intro!: cSup_eq_maximum setL2_mono + by (force simp add: diameter_def intro!: cSup_eq_maximum setL2_mono simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) lemma eucl_less_eq_halfspaces: diff -r e307a410f46c -r ab76bd43c14a src/HOL/NSA/Examples/NSPrimes.thy --- a/src/HOL/NSA/Examples/NSPrimes.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/NSA/Examples/NSPrimes.thy Wed Feb 17 23:29:35 2016 +0100 @@ -27,13 +27,19 @@ | injf_max_Suc: "injf_max (Suc n) E = choicefun({e. e:E & injf_max n E < e})" -lemma dvd_by_all: "\M. \N. 0 < N & (\m. 0 < m & (m::nat) <= M --> m dvd N)" -apply (rule allI) -apply (induct_tac "M", auto) -apply (rule_tac x = "N * (Suc n) " in exI) -by (metis dvd.order_refl dvd_mult dvd_mult2 le_Suc_eq nat_0_less_mult_iff zero_less_Suc) +lemma dvd_by_all2: + fixes M :: nat + shows "\N>0. \m. 0 < m \ m \ M \ m dvd N" +apply (induct M) +apply auto +apply (rule_tac x = "N * (Suc M) " in exI) +apply auto +apply (metis dvdI dvd_add_times_triv_left_iff dvd_add_triv_right_iff dvd_refl dvd_trans le_Suc_eq mult_Suc_right) +done -lemmas dvd_by_all2 = dvd_by_all [THEN spec] +lemma dvd_by_all: + "\M::nat. \N>0. \m. 0 < m \ m \ M \ m dvd N" + using dvd_by_all2 by blast lemma hypnat_of_nat_le_zero_iff [simp]: "(hypnat_of_nat n <= 0) = (n = 0)" by (transfer, simp) @@ -256,8 +262,9 @@ lemma hdvd_diff: "!!k m n :: hypnat. [| k dvd m; k dvd n |] ==> k dvd (m - n)" by (transfer, rule dvd_diff_nat) -lemma hdvd_one_eq_one: "!!x. x dvd (1::hypnat) ==> x = 1" -by (transfer, rule gcd_lcm_complete_lattice_nat.le_bot) +lemma hdvd_one_eq_one: + "\x::hypnat. is_unit x \ x = 1" + by transfer simp text\Already proved as \primes_infinite\, but now using non-standard naturals.\ theorem not_finite_prime: "~ finite {p::nat. prime p}" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Nat.thy --- a/src/HOL/Nat.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Nat.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1934,11 +1934,6 @@ unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) -text \@{term "op dvd"} is a partial order\ - -interpretation dvd: order "op dvd" "\n m :: nat. n dvd m \ \ m dvd n" - proof qed (auto intro: dvd_refl dvd_trans dvd_antisym) - lemma dvd_diff_nat[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)" unfolding dvd_def by (blast intro: diff_mult_distrib2 [symmetric]) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Nat_Transfer.thy --- a/src/HOL/Nat_Transfer.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Nat_Transfer.thy Wed Feb 17 23:29:35 2016 +0100 @@ -286,7 +286,7 @@ "(int x) * (int y) = int (x * y)" "tsub (int x) (int y) = int (x - y)" "(int x)^n = int (x^n)" - by (auto simp add: int_mult tsub_def of_nat_power) + by (auto simp add: of_nat_mult tsub_def of_nat_power) lemma transfer_int_nat_function_closures: "is_nat x \ is_nat y \ is_nat (x + y)" diff -r e307a410f46c -r ab76bd43c14a src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/NthRoot.thy Wed Feb 17 23:29:35 2016 +0100 @@ -10,17 +10,6 @@ imports Deriv Binomial begin -lemma abs_sgn_eq: "\sgn x :: real\ = (if x = 0 then 0 else 1)" - by (simp add: sgn_real_def) - -lemma inverse_sgn: "sgn (inverse a) = inverse (sgn a :: real)" - by (simp add: sgn_real_def) - -lemma power_eq_iff_eq_base: - fixes a b :: "_ :: linordered_semidom" - shows "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" - using power_eq_imp_eq_base[of a n b] by auto - subsection \Existence of Nth Root\ text \Existence follows from the Intermediate Value Theorem\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Num.thy --- a/src/HOL/Num.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Num.thy Wed Feb 17 23:29:35 2016 +0100 @@ -833,6 +833,11 @@ end +text\Also the class for fields with characteristic zero.\ + +class field_char_0 = field + ring_char_0 + + subsubsection \ Structures with negation and order: class \linordered_idom\ \ @@ -1099,6 +1104,8 @@ context linordered_field begin +subclass field_char_0 .. + lemma half_gt_zero_iff: "0 < a / 2 \ 0 < a" (is "?P \ ?Q") by (auto simp add: field_simps) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Cong.thy Wed Feb 17 23:29:35 2016 +0100 @@ -267,7 +267,7 @@ lemma cong_mult_rcancel_int: "coprime k (m::int) \ [a * k = b * k] (mod m) = [a = b] (mod m)" - by (metis cong_altdef_int left_diff_distrib coprime_dvd_mult_iff_int gcd.commute) + by (metis cong_altdef_int left_diff_distrib coprime_dvd_mult_iff gcd.commute) lemma cong_mult_rcancel_nat: "coprime k (m::nat) \ [a * k = b * k] (mod m) = [a = b] (mod m)" @@ -285,7 +285,7 @@ lemma coprime_cong_mult_int: "[(a::int) = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" -by (metis divides_mult_int cong_altdef_int) +by (metis divides_mult cong_altdef_int) lemma coprime_cong_mult_nat: assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n" @@ -527,13 +527,10 @@ apply (metis cong_solve_int) done -lemma coprime_iff_invertible_nat: "m > 0 \ coprime a m = (EX x. [a * x = Suc 0] (mod m))" - apply (auto intro: cong_solve_coprime_nat) - apply (metis cong_Suc_0_nat cong_solve_nat gcd_nat.left_neutral) - apply (metis One_nat_def cong_gcd_eq_nat coprime_lmult_nat - gcd_lcm_complete_lattice_nat.inf_bot_right gcd.commute) - done - +lemma coprime_iff_invertible_nat: + "m > 0 \ coprime a m = (EX x. [a * x = Suc 0] (mod m))" + by (metis One_nat_def cong_gcd_eq_nat cong_solve_coprime_nat coprime_lmult_nat gcd.commute gcd_Suc_0) + lemma coprime_iff_invertible_int: "m > (0::int) \ coprime a m = (EX x. [a * x = 1] (mod m))" apply (auto intro: cong_solve_coprime_int) apply (metis cong_int_def coprime_mul_eq_int gcd_1_int gcd.commute gcd_red_int) @@ -557,13 +554,13 @@ lemma cong_cong_lcm_nat: "[(x::nat) = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" apply (cases "y \ x") - apply (metis cong_altdef_nat lcm_least_nat) - apply (metis cong_altdef_nat cong_diff_cong_0'_nat lcm_semilattice_nat.sup.bounded_iff le0 minus_nat.diff_0) + apply (metis cong_altdef_nat lcm_least) + apply (meson cong_altdef_nat cong_sym_nat lcm_least_iff nat_le_linear) done lemma cong_cong_lcm_int: "[(x::int) = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" - by (auto simp add: cong_altdef_int lcm_least_int) [1] + by (auto simp add: cong_altdef_int lcm_least) [1] lemma cong_cong_setprod_coprime_nat [rule_format]: "finite A \ (\i\A. (\j\A. i \ j \ coprime (m i) (m j))) \ @@ -591,7 +588,7 @@ from cong_solve_coprime_nat [OF a] obtain x1 where one: "[m1 * x1 = 1] (mod m2)" by auto from a have b: "coprime m2 m1" - by (subst gcd_commute_nat) + by (subst gcd.commute) from cong_solve_coprime_nat [OF b] obtain x2 where two: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" @@ -610,7 +607,7 @@ from cong_solve_coprime_int [OF a] obtain x1 where one: "[m1 * x1 = 1] (mod m2)" by auto from a have b: "coprime m2 m1" - by (subst gcd_commute_int) + by (subst gcd.commute) from cong_solve_coprime_int [OF b] obtain x2 where two: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Eratosthenes.thy --- a/src/HOL/Number_Theory/Eratosthenes.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Eratosthenes.thy Wed Feb 17 23:29:35 2016 +0100 @@ -283,7 +283,7 @@ next case False with \m \ 1\ have "Suc (Suc 0) \ m" by arith with \m < Suc n\ * \m dvd q\ have "q dvd m" by simp - with \m dvd q\ show ?thesis by (simp add: dvd.eq_iff) + with \m dvd q\ show ?thesis by (simp add: dvd_antisym) qed } then have aux: "\m q. Suc (Suc 0) \ q \ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Euclidean_Algorithm.thy --- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1064,11 +1064,11 @@ "Lcm {} = 1" by (simp add: Lcm_1_iff) -lemma Lcm_eq_0 [simp]: +lemma Lcm_eq_0_I [simp]: "0 \ A \ Lcm A = 0" by (drule dvd_Lcm) simp -lemma Lcm0_iff': +lemma Lcm_0_iff': "Lcm A = 0 \ \(\l. l \ 0 \ (\a\A. a dvd l))" proof assume "Lcm A = 0" @@ -1092,7 +1092,7 @@ qed qed (simp only: Lcm_Lcm_eucl Lcm_eucl_def if_False) -lemma Lcm0_iff [simp]: +lemma Lcm_0_iff [simp]: "finite A \ Lcm A = 0 \ 0 \ A" proof - assume "finite A" @@ -1108,7 +1108,7 @@ done moreover from \finite A\ have "\a\A. a dvd \A" by blast ultimately have "\l. l \ 0 \ (\a\A. a dvd l)" by blast - with Lcm0_iff' have "Lcm A \ 0" by simp + with Lcm_0_iff' have "Lcm A \ 0" by simp } ultimately show "Lcm A = 0 \ 0 \ A" by blast qed @@ -1200,7 +1200,7 @@ qed subclass semiring_Gcd - by standard (simp_all add: Gcd_greatest) + by standard (auto intro: Gcd_greatest Lcm_least) lemma GcdI: assumes "\a. a \ A \ b dvd a" and "\c. (\a. a \ A \ c dvd a) \ c dvd b" @@ -1210,10 +1210,7 @@ lemma Lcm_Gcd: "Lcm A = Gcd {m. \a\A. a dvd m}" - by (rule LcmI[symmetric]) (auto intro: dvd_Gcd Gcd_greatest) - -subclass semiring_Lcm - by standard (simp add: Lcm_Gcd) + by (rule LcmI[symmetric]) (auto intro: Gcd_greatest Gcd_greatest) lemma Gcd_1: "1 \ A \ Gcd A = 1" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Factorial_Ring.thy --- a/src/HOL/Number_Theory/Factorial_Ring.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Factorial_Ring.thy Wed Feb 17 23:29:35 2016 +0100 @@ -344,13 +344,13 @@ then have *: "is_prime (nat \p\)" by simp assume "p dvd k * l" then have "nat \p\ dvd nat \k * l\" - by (simp add: dvd_int_iff) + by (simp add: dvd_int_unfold_dvd_nat) then have "nat \p\ dvd nat \k\ * nat \l\" by (simp add: abs_mult nat_mult_distrib) with * have "nat \p\ dvd nat \k\ \ nat \p\ dvd nat \l\" using is_primeD [of "nat \p\"] by auto then show "p dvd k \ p dvd l" - by (simp add: dvd_int_iff) + by (simp add: dvd_int_unfold_dvd_nat) next fix k :: int assume P: "\l. l dvd k \ \ is_prime l" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Fib.thy --- a/src/HOL/Number_Theory/Fib.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Fib.thy Wed Feb 17 23:29:35 2016 +0100 @@ -63,11 +63,11 @@ done lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)" - apply (simp add: gcd_commute_nat [of "fib m"]) + apply (simp add: gcd.commute [of "fib m"]) apply (cases m) apply (auto simp add: fib_add) - apply (metis gcd_commute_nat mult.commute coprime_fib_Suc_nat - gcd_add_mult_nat gcd_mult_cancel_nat gcd.commute) + apply (metis gcd.commute mult.commute coprime_fib_Suc_nat + gcd_add_mult_nat gcd_mult_cancel gcd.commute) done lemma gcd_fib_diff: "m \ n \ gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)" @@ -98,7 +98,7 @@ lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)" -- \Law 6.111\ - by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat gcd_commute_nat gcd_fib_mod) + by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat gcd.commute gcd_fib_mod) theorem fib_mult_eq_setsum_nat: "fib (Suc n) * fib n = (\k \ {..n}. fib k * fib k)" by (induct n rule: nat.induct) (auto simp add: field_simps) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Gauss.thy --- a/src/HOL/Number_Theory/Gauss.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Gauss.thy Wed Feb 17 23:29:35 2016 +0100 @@ -117,7 +117,7 @@ order_le_less_trans [of x "(int p - 1) div 2" p] order_le_less_trans [of y "(int p - 1) div 2" p] have "x = y" - by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int zero_zle_int) + by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff) } note xy = this show ?thesis apply (insert p_ge_2 p_a_relprime p_minus_one_l) @@ -141,12 +141,12 @@ by (metis cong_int_def) with p_a_relprime p_prime cong_mult_rcancel_int [of a p x y] have "[x = y](mod p)" - by (metis cong_mult_self_int dvd_div_mult_self gcd_commute_int prime_imp_coprime_int) + by (metis cong_mult_self_int dvd_div_mult_self gcd.commute prime_imp_coprime_int) with cong_less_imp_eq_int [of x y p] p_minus_one_l order_le_less_trans [of x "(int p - 1) div 2" p] order_le_less_trans [of y "(int p - 1) div 2" p] have "x = y" - by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int zero_zle_int) + by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff) then have False by (simp add: f)} then show ?thesis diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/MiscAlgebra.thy --- a/src/HOL/Number_Theory/MiscAlgebra.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/MiscAlgebra.thy Wed Feb 17 23:29:35 2016 +0100 @@ -259,7 +259,7 @@ (ALL A:C. ALL B:C. A ~= B --> A Int B = {}) |] ==> finprod G f (\C) = finprod G (finprod G f) C" apply (frule finprod_UN_disjoint [of C id f]) - apply (auto simp add: SUP_def) + apply auto done lemma (in comm_monoid) finprod_one: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Pocklington.thy --- a/src/HOL/Number_Theory/Pocklington.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Pocklington.thy Wed Feb 17 23:29:35 2016 +0100 @@ -11,7 +11,7 @@ subsection\Lemmas about previously defined terms\ lemma prime: - "prime p \ p \ 0 \ p\1 \ (\m. 0 < m \ m < p \ coprime p m)" + "prime p \ p \ 0 \ p \ 1 \ (\m. 0 < m \ m < p \ coprime p m)" (is "?lhs \ ?rhs") proof- {assume "p=0 \ p=1" hence ?thesis @@ -39,7 +39,7 @@ {assume "q\p" with qp have qplt: "q < p" by arith from H qplt q0 have "coprime p q" by arith hence ?lhs using q - by (metis gcd_semilattice_nat.inf_absorb2 one_not_prime_nat)} + by (auto dest: gcd_nat.absorb2)} ultimately have ?lhs by blast} ultimately have ?thesis by blast} ultimately show ?thesis by (cases"p=0 \ p=1", auto) @@ -105,16 +105,16 @@ by (metis cong_solve_unique neq0_conv p prime_gt_0_nat px) {assume y0: "y = 0" with y(2) have th: "p dvd a" - by (metis cong_dvd_eq_nat gcd_lcm_complete_lattice_nat.top_greatest mult_0_right) + by (auto dest: cong_dvd_eq_nat) have False by (metis gcd_nat.absorb1 one_not_prime_nat p pa th)} with y show ?thesis unfolding Ex1_def using neq0_conv by blast qed lemma cong_unique_inverse_prime: - assumes p: "prime p" and x0: "0 < x" and xp: "x < p" + assumes "prime p" and "0 < x" and "x < p" shows "\!y. 0 < y \ y < p \ [x * y = 1] (mod p)" -by (metis cong_solve_unique_nontrivial gcd_lcm_complete_lattice_nat.inf_bot_left gcd.commute assms) + by (rule cong_solve_unique_nontrivial) (insert assms, simp_all) lemma chinese_remainder_coprime_unique: fixes a::nat @@ -469,7 +469,7 @@ "prime n \ n \ 1 \ (\d. d dvd n \ d\<^sup>2 \ n \ d = 1)" proof - {assume "n=0 \ n=1" hence ?thesis - by (metis dvd.order_refl le_refl one_not_prime_nat power_zero_numeral zero_not_prime_nat)} + by auto} moreover {assume n: "n\0" "n\1" hence np: "n > 1" by arith @@ -583,9 +583,8 @@ {fix d assume d: "d dvd p" "d dvd a" "d \ 1" from pp[unfolded prime_def] d have dp: "d = p" by blast from n have "n \ 0" by simp - then have False using d - by (metis coprime_minus_one_nat dp lucas_coprime_lemma an coprime_nat - gcd_lcm_complete_lattice_nat.top_greatest pn)} + then have False using d dp pn + by auto (metis One_nat_def Suc_pred an dvd_1_iff_1 gcd_greatest_iff lucas_coprime_lemma)} hence cpa: "coprime p a" by auto have arp: "coprime (a^r) p" by (metis coprime_exp_nat cpa gcd.commute) @@ -617,9 +616,8 @@ have th: "q dvd p - 1" by (metis cong_to_1_nat) have "p - 1 \ 0" using prime_ge_2_nat [OF p(1)] by arith - with pq p have False - by (metis Suc_diff_1 gcd_le2_nat gcd_semilattice_nat.inf_absorb1 not_less_eq_eq - prime_gt_0_nat th) } + with pq th have False + by (simp add: nat_dvd_not_less)} with n01 show ?ths by blast qed @@ -649,8 +647,7 @@ have "p - 1 \ 0" using prime_ge_2_nat [OF p(1)] by arith then have pS: "Suc (p - 1) = p" by arith from b have d: "n dvd a^((n - 1) div p)" unfolding b0 - by (metis b0 diff_0_eq_0 gcd_dvd2_nat gcd_lcm_complete_lattice_nat.inf_bot_left - gcd_lcm_complete_lattice_nat.inf_top_left) + by auto from divides_rexp[OF d, of "p - 1"] pS eq cong_dvd_eq_nat [OF an] n have False by simp} diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Primes.thy --- a/src/HOL/Number_Theory/Primes.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Primes.thy Wed Feb 17 23:29:35 2016 +0100 @@ -64,7 +64,7 @@ lemma prime_imp_coprime_nat: "prime p \ \ p dvd n \ coprime p n" apply (unfold prime_def) - apply (metis gcd_dvd1_nat gcd_dvd2_nat) + apply (metis gcd_dvd1 gcd_dvd2) done lemma prime_int_altdef: @@ -72,22 +72,22 @@ m = 1 \ m = p))" apply (simp add: prime_def) apply (auto simp add: ) - apply (metis One_nat_def int_1 nat_0_le nat_dvd_iff) + apply (metis One_nat_def of_nat_1 nat_0_le nat_dvd_iff) apply (metis zdvd_int One_nat_def le0 of_nat_0 of_nat_1 of_nat_eq_iff of_nat_le_iff) done lemma prime_imp_coprime_int: fixes n::int shows "prime p \ \ p dvd n \ coprime p n" apply (unfold prime_int_altdef) - apply (metis gcd_dvd1_int gcd_dvd2_int gcd_ge_0_int) + apply (metis gcd_dvd1 gcd_dvd2 gcd_ge_0_int) done lemma prime_dvd_mult_nat: "prime p \ p dvd m * n \ p dvd m \ p dvd n" - by (blast intro: coprime_dvd_mult_nat prime_imp_coprime_nat) + by (blast intro: coprime_dvd_mult prime_imp_coprime_nat) lemma prime_dvd_mult_int: fixes n::int shows "prime p \ p dvd m * n \ p dvd m \ p dvd n" - by (blast intro: coprime_dvd_mult_int prime_imp_coprime_int) + by (blast intro: coprime_dvd_mult prime_imp_coprime_int) lemma prime_dvd_mult_eq_nat [simp]: "prime p \ p dvd m * n = (p dvd m \ p dvd n)" @@ -167,14 +167,19 @@ "prime p \ prime q \ p ~= q \ coprime (p^m) (q^n)" by (rule coprime_exp2_nat, rule primes_coprime_nat) -lemma prime_factor_nat: "n \ (1::nat) \ \ p. prime p \ p dvd n" - apply (induct n rule: nat_less_induct) - apply (case_tac "n = 0") - using two_is_prime_nat - apply blast - apply (metis One_nat_def dvd.order_trans dvd_refl less_Suc0 linorder_neqE_nat - nat_dvd_not_less neq0_conv prime_def) - done +lemma prime_factor_nat: + "n \ (1::nat) \ \p. prime p \ p dvd n" +proof (induct n rule: nat_less_induct) + case (1 n) show ?case + proof (cases "n = 0") + case True then show ?thesis + by (auto intro: two_is_prime_nat) + next + case False with "1.prems" have "n > 1" by simp + with "1.hyps" show ?thesis + by (metis One_nat_def dvd_mult dvd_refl not_prime_eq_prod_nat order_less_irrefl) + qed +qed text \One property of coprimality is easier to prove via prime factors.\ @@ -198,20 +203,20 @@ { assume pa: "p dvd a" from coprime_common_divisor_nat [OF ab, OF pa] p have "\ p dvd b" by auto with p have "coprime b p" - by (subst gcd_commute_nat, intro prime_imp_coprime_nat) + by (subst gcd.commute, intro prime_imp_coprime_nat) then have pnb: "coprime (p^n) b" - by (subst gcd_commute_nat, rule coprime_exp_nat) - from coprime_dvd_mult_nat[OF pnb pab] have ?thesis by blast } + by (subst gcd.commute, rule coprime_exp_nat) + from coprime_dvd_mult[OF pnb pab] have ?thesis by blast } moreover { assume pb: "p dvd b" have pnba: "p^n dvd b*a" using pab by (simp add: mult.commute) from coprime_common_divisor_nat [OF ab, of p] pb p have "\ p dvd a" by auto with p have "coprime a p" - by (subst gcd_commute_nat, intro prime_imp_coprime_nat) + by (subst gcd.commute, intro prime_imp_coprime_nat) then have pna: "coprime (p^n) a" - by (subst gcd_commute_nat, rule coprime_exp_nat) - from coprime_dvd_mult_nat[OF pna pnba] have ?thesis by blast } + by (subst gcd.commute, rule coprime_exp_nat) + from coprime_dvd_mult[OF pna pnba] have ?thesis by blast } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed @@ -417,11 +422,13 @@ lemma bezout_prime: assumes p: "prime p" and pa: "\ p dvd a" shows "\x y. a*x = Suc (p*y)" -proof- +proof - have ap: "coprime a p" by (metis gcd.commute p pa prime_imp_coprime_nat) - from coprime_bezout_strong[OF ap] show ?thesis - by (metis Suc_eq_plus1 gcd_lcm_complete_lattice_nat.bot.extremum pa) + moreover from p have "p \ 1" by auto + ultimately have "\x y. a * x = p * y + 1" + by (rule coprime_bezout_strong) + then show ?thesis by simp qed end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/Residues.thy --- a/src/HOL/Number_Theory/Residues.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/Residues.thy Wed Feb 17 23:29:35 2016 +0100 @@ -209,7 +209,7 @@ sublocale residues_prime < residues p apply (unfold R_def residues_def) using p_prime apply auto - apply (metis (full_types) int_1 of_nat_less_iff prime_gt_1_nat) + apply (metis (full_types) of_nat_1 of_nat_less_iff prime_gt_1_nat) done context residues_prime @@ -221,7 +221,7 @@ apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq res_units_eq) apply (rule classical) apply (erule notE) - apply (subst gcd_commute_int) + apply (subst gcd.commute) apply (rule prime_imp_coprime_int) apply (rule p_prime) apply (rule notI) @@ -232,7 +232,7 @@ lemma res_prime_units_eq: "Units R = {1..p - 1}" apply (subst res_units_eq) apply auto - apply (subst gcd_commute_int) + apply (subst gcd.commute) apply (auto simp add: p_prime prime_imp_coprime_int zdvd_not_zless) done @@ -256,8 +256,8 @@ apply (rule bij_betw_same_card [of nat]) apply (auto simp add: inj_on_def bij_betw_def image_def) apply (metis dual_order.irrefl dual_order.strict_trans leI nat_1 transfer_nat_int_gcd(1)) - apply (metis One_nat_def int_0 int_1 int_less_0_conv int_nat_eq nat_int - transfer_int_nat_gcd(1) zless_int) + apply (metis One_nat_def of_nat_0 of_nat_1 of_nat_less_0_iff int_nat_eq nat_int + transfer_int_nat_gcd(1) of_nat_less_iff) done lemma prime_phi: @@ -370,7 +370,7 @@ assumes "prime p" and "\ p dvd a" shows "[a ^ (p - 1) = 1] (mod p)" using fermat_theorem [of p a] assms - by (metis int_1 of_nat_power transfer_int_nat_cong zdvd_int) + by (metis of_nat_1 of_nat_power transfer_int_nat_cong zdvd_int) subsection \Wilson's theorem\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Number_Theory/UniqueFactorization.thy --- a/src/HOL/Number_Theory/UniqueFactorization.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Number_Theory/UniqueFactorization.thy Wed Feb 17 23:29:35 2016 +0100 @@ -98,14 +98,14 @@ finally have "a ^ count M a dvd a ^ count N a * (\i \ (set_mset N - {a}). i ^ count N i)" . moreover have "coprime (a ^ count M a) (\i \ (set_mset N - {a}). i ^ count N i)" - apply (subst gcd_commute_nat) + apply (subst gcd.commute) apply (rule setprod_coprime_nat) apply (rule primes_imp_powers_coprime_nat) using assms True apply auto done ultimately have "a ^ count M a dvd a ^ count N a" - by (elim coprime_dvd_mult_nat) + by (elim coprime_dvd_mult) with a show ?thesis using power_dvd_imp_le prime_def by blast next @@ -653,15 +653,16 @@ lemma multiplicity_dvd'_nat: fixes x y :: nat - shows "0 < x \ \p. prime p \ multiplicity p x \ multiplicity p y \ x dvd y" - by (metis gcd_lcm_complete_lattice_nat.top_greatest le_refl multiplicity_dvd_nat - multiplicity_nonprime_nat neq0_conv) + assumes "0 < x" + assumes "\p. prime p \ multiplicity p x \ multiplicity p y" + shows "x dvd y" + using dvd_0_right assms by (metis (no_types) le0 multiplicity_dvd_nat multiplicity_nonprime_nat not_gr0) lemma multiplicity_dvd'_int: fixes x y :: int shows "0 < x \ 0 \ y \ \p. prime p \ multiplicity p x \ multiplicity p y \ x dvd y" - by (metis GCD.dvd_int_iff abs_int_eq multiplicity_dvd'_nat multiplicity_int_def nat_int + by (metis dvd_int_iff abs_of_nat multiplicity_dvd'_nat multiplicity_int_def nat_int zero_le_imp_eq_int zero_less_imp_eq_int) lemma dvd_multiplicity_eq_nat: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Old_Number_Theory/Legacy_GCD.thy --- a/src/HOL/Old_Number_Theory/Legacy_GCD.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy Wed Feb 17 23:29:35 2016 +0100 @@ -602,7 +602,7 @@ from relprime_dvd_mult [OF g th] obtain h' where h': "nat \k\ = nat \i\ * h'" unfolding dvd_def by blast from h' have "int (nat \k\) = int (nat \i\ * h')" by simp - then have "\k\ = \i\ * int h'" by (simp add: int_mult) + then have "\k\ = \i\ * int h'" by (simp add: of_nat_mult) then show ?thesis apply (subst abs_dvd_iff [symmetric]) apply (subst dvd_abs_iff [symmetric]) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Old_Number_Theory/Primes.thy --- a/src/HOL/Old_Number_Theory/Primes.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Old_Number_Theory/Primes.thy Wed Feb 17 23:29:35 2016 +0100 @@ -819,11 +819,11 @@ by (auto simp add: dvd_def coprime) lemma mult_inj_if_coprime_nat: - "inj_on f A \ inj_on g B \ ALL a:A. ALL b:B. coprime (f a) (g b) - \ inj_on (%(a,b). f a * g b::nat) (A \ B)" -apply(auto simp add:inj_on_def) -apply(metis coprime_def dvd_triv_left gcd_proj2_if_dvd_nat gcd_semilattice_nat.inf_commute relprime_dvd_mult) -apply(metis coprime_commute coprime_divprod dvd.neq_le_trans dvd_triv_right) + "inj_on f A \ inj_on g B \ \a\A. \b\B. Primes.coprime (f a) (g b) \ + inj_on (\(a, b). f a * g b) (A \ B)" +apply (auto simp add: inj_on_def) +apply (metis coprime_def dvd_antisym dvd_triv_left relprime_dvd_mult_iff) +apply (metis coprime_commute coprime_divprod dvd_antisym dvd_triv_right) done declare power_Suc0[simp del] diff -r e307a410f46c -r ab76bd43c14a src/HOL/Power.thy --- a/src/HOL/Power.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Power.thy Wed Feb 17 23:29:35 2016 +0100 @@ -9,19 +9,6 @@ imports Num Equiv_Relations begin -context linordered_ring (* TODO: move *) -begin - -lemma sum_squares_ge_zero: - "0 \ x * x + y * y" - by (intro add_nonneg_nonneg zero_le_square) - -lemma not_sum_squares_lt_zero: - "\ x * x + y * y < 0" - by (simp add: not_less sum_squares_ge_zero) - -end - subsection \Powers for Arbitrary Monoids\ class power = one + times @@ -123,6 +110,10 @@ finally show ?case . qed simp +lemma power_minus_mult: + "0 < n \ a ^ (n - 1) * a = a ^ n" + by (simp add: power_commutes split add: nat_diff_split) + end context comm_monoid_mult @@ -584,6 +575,10 @@ "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" by (cases n) (simp_all del: power_Suc, rule power_inject_base) +lemma power_eq_iff_eq_base: + "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" + using power_eq_imp_eq_base [of a n b] by auto + lemma power2_le_imp_le: "x\<^sup>2 \ y\<^sup>2 \ 0 \ y \ x \ y" unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) @@ -596,6 +591,10 @@ "x\<^sup>2 = y\<^sup>2 \ 0 \ x \ 0 \ y \ x = y" unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp +lemma power_Suc_le_self: + shows "0 \ a \ a \ 1 \ a ^ Suc n \ a" + using power_decreasing [of 1 "Suc n" a] by simp + end context linordered_ring_strict diff -r e307a410f46c -r ab76bd43c14a src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Presburger.thy Wed Feb 17 23:29:35 2016 +0100 @@ -376,7 +376,7 @@ lemma zdiff_int_split: "P (int (x - y)) = ((y \ x \ P (int x - int y)) \ (x < y \ P 0))" - by (cases "y \ x") (simp_all add: zdiff_int) + by (cases "y \ x") (simp_all add: of_nat_diff) text \ \medskip Specific instances of congruence rules, to prevent diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Caratheodory.thy Wed Feb 17 23:29:35 2016 +0100 @@ -396,8 +396,8 @@ from \outer_measure M f X \ \\ have fin: "\outer_measure M f X\ \ \" using outer_measure_nonneg[OF posf, of X] by auto show ?thesis - using Inf_ereal_close[OF fin[unfolded outer_measure_def INF_def], OF \0 < e\] - unfolding INF_def[symmetric] outer_measure_def[symmetric] by (auto intro: less_imp_le) + using Inf_ereal_close [OF fin [unfolded outer_measure_def], OF \0 < e\] + by (auto intro: less_imp_le simp add: outer_measure_def) qed lemma (in ring_of_sets) countably_subadditive_outer_measure: @@ -574,7 +574,7 @@ shows "f (UNION I A) = (\i\I. f (A i))" proof - have "A`I \ M" "disjoint (A`I)" "finite (A`I)" "\(A`I) \ M" - using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image) + using A by (auto simp: disjoint_family_on_disjoint_image) with \volume M f\ have "f (\(A`I)) = (\a\A`I. f a)" unfolding volume_def by blast also have "\ = (\i\I. f (A i))" @@ -713,7 +713,7 @@ then have "disjoint_family F" using F'_disj by (auto simp: disjoint_family_on_def) moreover from F' have "(\i. F i) = \C" - by (auto simp: F_def set_eq_iff split: split_if_asm) + by (auto simp add: F_def split: split_if_asm) blast moreover have sets_F: "\i. F i \ M" using F' sets_C by (auto simp: F_def) moreover note sets_C @@ -783,9 +783,10 @@ by (auto simp: disjoint_family_on_def f_def) moreover have Ai_eq: "A i = (\xrange f = A i" - using f C Ai unfolding bij_betw_def by (auto simp: f_def) + using f C Ai unfolding bij_betw_def + by (auto simp add: f_def cong del: strong_SUP_cong) moreover { have "(\j. \_r (f j)) = (\j. if j \ {..< card C} then \_r (f j) else 0)" using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def) @@ -841,7 +842,7 @@ have "(\n. \_r (A' n)) = (\n. \c\C'. \_r (A' n \ c))" using C' A' by (subst volume_finite_additive[symmetric, OF V(1)]) - (auto simp: disjoint_def disjoint_family_on_def Union_image_eq[symmetric] simp del: Sup_image_eq Union_image_eq + (auto simp: disjoint_def disjoint_family_on_def intro!: G.Int G.finite_Union arg_cong[where f="\X. suminf (\i. \_r (X i))"] ext intro: generated_ringI_Basic) also have "\ = (\c\C'. \n. \_r (A' n \ c))" @@ -853,7 +854,7 @@ also have "\ = \_r (\C')" using C' Un_A by (subst volume_finite_additive[symmetric, OF V(1)]) - (auto simp: disjoint_family_on_def disjoint_def Union_image_eq[symmetric] simp del: Sup_image_eq Union_image_eq + (auto simp: disjoint_family_on_def disjoint_def intro: generated_ringI_Basic) finally show "(\n. \_r (A' n)) = \_r (\i. A' i)" using C' by simp diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Complete_Measure.thy Wed Feb 17 23:29:35 2016 +0100 @@ -205,7 +205,7 @@ unfolding binary_def by (auto split: split_if_asm) show "emeasure M (main_part M (S \ T)) = emeasure M (main_part M S \ main_part M T)" using emeasure_main_part_UN[of "binary S T" M] assms - unfolding range_binary_eq Un_range_binary UN by auto + by (simp add: range_binary_eq, simp add: Un_range_binary UN) qed (auto intro: S T) lemma sets_completionI_sub: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Embed_Measure.thy --- a/src/HOL/Probability/Embed_Measure.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Embed_Measure.thy Wed Feb 17 23:29:35 2016 +0100 @@ -36,7 +36,8 @@ then obtain A' where "\i. A i = f ` A' i" "\i. A' i \ sets M" by (auto simp: subset_eq choice_iff) moreover from this have "(\x. f ` A' x) = f ` (\x. A' x)" by blast - ultimately show "(\i. A i) \ {f ` A |A. A \ sets M}" by blast + ultimately show "(\i. A i) \ {f ` A |A. A \ sets M}" + by simp blast qed (auto dest: sets.sets_into_space) lemma the_inv_into_vimage: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Fin_Map.thy Wed Feb 17 23:29:35 2016 +0100 @@ -696,7 +696,7 @@ have "A -` y \ space (PiF I M) = (\J\I. A -` y \ space (PiF {J} M))" by (auto simp: space_PiF) also have "\ \ sets (PiF I M)" - proof + proof (rule sets.finite_UN) show "finite I" by fact fix J assume "J \ I" with assms have "finite J" by simp @@ -1055,8 +1055,8 @@ by (intro Pi'_cong) (simp_all add: S_union) also have "\ = (\xs\{xs. length xs = card I}. \' j\I. if i = j then A else S j (xs ! T j))" using T - apply auto - apply (simp_all add: Pi'_iff bchoice_iff) + apply (auto simp del: Union_iff) + apply (simp_all add: Pi'_iff bchoice_iff del: Union_iff) apply (erule conjE exE)+ apply (rule_tac x="map (\n. f (the_inv_into I T n)) [0..k\Ka. Ea k)" "finite Ka" "Ka \ {}" "Ka \ I j" "\k. k\Ka \ Ea k \ E k" by auto fix b assume "b \ ?E j" then obtain Kb Eb where b: "b = (\k\Kb. Eb k)" "finite Kb" "Kb \ {}" "Kb \ I j" "\k. k\Kb \ Eb k \ E k" by auto - let ?A = "\k. (if k \ Ka \ Kb then Ea k \ Eb k else if k \ Kb then Eb k else if k \ Ka then Ea k else {})" - have "a \ b = INTER (Ka \ Kb) ?A" - by (simp add: a b set_eq_iff) auto + let ?f = "\k. (if k \ Ka \ Kb then Ea k \ Eb k else if k \ Kb then Eb k else if k \ Ka then Ea k else {})" + have "Ka \ Kb = (Ka \ Kb) \ (Kb - Ka) \ (Ka - Kb)" + by blast + moreover have "(\x\Ka \ Kb. Ea x \ Eb x) \ + (\x\Kb - Ka. Eb x) \ (\x\Ka - Kb. Ea x) = (\k\Ka. Ea k) \ (\k\Kb. Eb k)" + by auto + ultimately have "(\k\Ka \ Kb. ?f k) = (\k\Ka. Ea k) \ (\k\Kb. Eb k)" (is "?lhs = ?rhs") + by (simp only: image_Un Inter_Un_distrib) simp + then have "a \ b = (\k\Ka \ Kb. ?f k)" + by (simp only: a(1) b(1)) with a b \j \ J\ Int_stableD[OF Int_stable] show "a \ b \ ?E j" - by (intro CollectI exI[of _ "Ka \ Kb"] exI[of _ ?A]) auto + by (intro CollectI exI[of _ "Ka \ Kb"] exI[of _ ?f]) auto qed qed ultimately show ?thesis @@ -804,14 +811,18 @@ proof - have [measurable]: "\m. {x\space M. P m x} \ sets M" using assms by (auto simp: indep_events_def) - have "prob (\n. \m\{n..}. ?P m) = 0 \ prob (\n. \m\{n..}. ?P m) = 1" - by (rule borel_0_1_law) fact + have *: "(\n. \m\{n..}. {x \ space M. P m x}) \ events" + by simp + from assms have "prob (\n. \m\{n..}. ?P m) = 0 \ prob (\n. \m\{n..}. ?P m) = 1" + by (rule borel_0_1_law) + also have "prob (\n. \m\{n..}. ?P m) = 1 \ (AE x in M. infinite {m. P m x})" + using * by (simp add: prob_eq_1) + (simp add: Bex_def infinite_nat_iff_unbounded_le) also have "prob (\n. \m\{n..}. ?P m) = 0 \ (AE x in M. finite {m. P m x})" - by (subst prob_eq_0) (auto simp add: finite_nat_iff_bounded Ball_def not_less[symmetric]) - also have "prob (\n. \m\{n..}. ?P m) = 1 \ (AE x in M. infinite {m. P m x})" - by (subst prob_eq_1) (simp_all add: Bex_def infinite_nat_iff_unbounded_le) + using * by (simp add: prob_eq_0) + (auto simp add: Ball_def finite_nat_iff_bounded not_less [symmetric]) finally show ?thesis - by metis + by blast qed lemma (in prob_space) indep_sets_finite: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Infinite_Product_Measure.thy --- a/src/HOL/Probability/Infinite_Product_Measure.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy Wed Feb 17 23:29:35 2016 +0100 @@ -181,7 +181,7 @@ proof - have "Pi UNIV E = (\i. emb UNIV {..i} (\\<^sub>E j\{..i}. E j))" using E E[THEN sets.sets_into_space] - by (auto simp: prod_emb_def Pi_iff extensional_def) blast + by (auto simp: prod_emb_def Pi_iff extensional_def) with E show ?thesis by auto qed @@ -194,7 +194,7 @@ using E by (simp add: measure_PiM_emb) moreover have "Pi UNIV E = (\n. ?E n)" using E E[THEN sets.sets_into_space] - by (auto simp: prod_emb_def extensional_def Pi_iff) blast + by (auto simp: prod_emb_def extensional_def Pi_iff) moreover have "range ?E \ sets S" using E by auto moreover have "decseq ?E" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Wed Feb 17 23:29:35 2016 +0100 @@ -536,7 +536,7 @@ moreover have "emeasure lborel (\i. {from_nat_into A i}) = 0" by (rule emeasure_UN_eq_0) auto ultimately have "emeasure lborel A \ 0" using emeasure_mono - by (metis assms bot.extremum_unique emeasure_empty image_eq_UN range_from_nat_into sets.empty_sets) + by (smt UN_E emeasure_empty equalityI from_nat_into order_refl singletonD subsetI) thus ?thesis by (auto simp add: emeasure_le_0_iff) qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Measurable.thy --- a/src/HOL/Probability/Measurable.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Measurable.thy Wed Feb 17 23:29:35 2016 +0100 @@ -152,7 +152,7 @@ "(\i. i \ X \ pred M (\x. x \ N x i)) \ pred M (\x. x \ (\i\X. N x i))" "(\i. i \ X \ pred M (\x. P x i)) \ pred M (\x. \i\X. P x i)" "(\i. i \ X \ pred M (\x. P x i)) \ pred M (\x. \i\X. P x i)" - by (auto simp: Bex_def Ball_def) + by simp_all (auto simp: Bex_def Ball_def) lemma pred_intros_finite[measurable (raw)]: "finite I \ (\i. i \ I \ pred M (\x. x \ N x i)) \ pred M (\x. x \ (\i\I. N x i))" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Measure_Space.thy --- a/src/HOL/Probability/Measure_Space.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Measure_Space.thy Wed Feb 17 23:29:35 2016 +0100 @@ -884,7 +884,7 @@ have "(\i\I. N i) = \(N ` range (from_nat_into I))" using I by simp also have "\ = (\i. (N \ from_nat_into I) i)" - by (simp only: SUP_def image_comp) + by simp finally show ?thesis by simp qed @@ -1228,7 +1228,7 @@ range: "range A \ sets M" and space: "(\i. A i) = space M" and measure: "\i. emeasure M (A i) \ \" - using sigma_finite by auto + using sigma_finite by blast show thesis proof (rule that[of "disjointed A"]) show "range (disjointed A) \ sets M" @@ -1252,7 +1252,7 @@ proof - obtain F :: "nat \ 'a set" where F: "range F \ sets M" "(\i. F i) = space M" "\i. emeasure M (F i) \ \" - using sigma_finite by auto + using sigma_finite by blast show thesis proof (rule that[of "\n. \i\n. F i"]) show "range (\n. \i\n. F i) \ sets M" @@ -1731,8 +1731,10 @@ assumes upper: "space M \ (\i \ s. f i)" shows "measure M e = (\ x \ s. measure M (e \ f x))" proof - - have e: "e = (\i \ s. e \ f i)" + have "e \ (\i\s. f i)" using \e \ sets M\ sets.sets_into_space upper by blast + then have e: "e = (\i \ s. e \ f i)" + by auto hence "measure M e = measure M (\i \ s. e \ f i)" by simp also have "\ = (\ x \ s. measure M (e \ f x))" proof (rule finite_measure_finite_Union) @@ -2074,7 +2076,7 @@ using E \ by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) next show "range (from_nat_into A) \ E" "(\i. from_nat_into A i) = \" - unfolding Sup_image_eq[symmetric, where f="from_nat_into A"] using A by auto + using A by (auto cong del: strong_SUP_cong) next fix i have "emeasure (restrict_space M \) (from_nat_into A i) = emeasure M (from_nat_into A i)" @@ -2333,12 +2335,12 @@ assumes A: "Complete_Partial_Order.chain op \ (f ` A)" and a: "a \ A" "f a \ bot" shows space_SUP: "space (SUP M:A. f M) = space (f a)" and sets_SUP: "sets (SUP M:A. f M) = sets (f a)" -unfolding SUP_def by(rule space_Sup[OF A a(2) imageI[OF a(1)]] sets_Sup[OF A a(2) imageI[OF a(1)]])+ +by(rule space_Sup[OF A a(2) imageI[OF a(1)]] sets_Sup[OF A a(2) imageI[OF a(1)]])+ lemma emeasure_SUP: assumes A: "Complete_Partial_Order.chain op \ (f ` A)" "A \ {}" assumes "X \ sets (SUP M:A. f M)" shows "emeasure (SUP M:A. f M) X = (SUP M:A. emeasure (f M)) X" -using \X \ _\ unfolding SUP_def by(subst emeasure_Sup[OF A(1)]; simp add: A) +using \X \ _\ by(subst emeasure_Sup[OF A(1)]; simp add: A) end diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Wed Feb 17 23:29:35 2016 +0100 @@ -932,7 +932,7 @@ have "a * u x < 1 * u x" by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) - finally obtain i where "a * u x < f i x" unfolding SUP_def + finally obtain i where "a * u x < f i x" by (auto simp add: less_SUP_iff) hence "a * u x \ f i x" by auto thus ?thesis using \x \ space M\ by auto @@ -2042,7 +2042,7 @@ by (auto intro!: SUP_least SUP_upper nn_integral_mono) next have "\x. \g. incseq g \ range g \ (\i. f i x) ` Y \ (SUP i:Y. f i x) = (SUP i. g i)" - unfolding Sup_class.SUP_def by(rule Sup_countable_SUP[unfolded Sup_class.SUP_def])(simp add: nonempty) + by (rule Sup_countable_SUP) (simp add: nonempty) then obtain g where incseq: "\x. incseq (g x)" and range: "\x. range (g x) \ (\i. f i x) ` Y" and sup: "\x. (SUP i:Y. f i x) = (SUP i. g x i)" by moura @@ -2088,8 +2088,8 @@ assume "x \ {.. \ (SUP i:?Y. f i) x" unfolding SUP_def Sup_fun_def image_image - using \x \ {.. by(rule Sup_upper[OF imageI]) + also have "\ \ (SUP i:?Y. f i) x" unfolding Sup_fun_def image_image + using \x \ {.. by (rule Sup_upper [OF imageI]) also have "\ = f (I m') x" unfolding m' by simp finally show "g x n \ f (I m') x" . qed @@ -2264,7 +2264,7 @@ by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) qed then show ?thesis - unfolding nn_integral_def_finite SUP_def by simp + unfolding nn_integral_def_finite by (simp cong del: strong_SUP_cong) qed lemma nn_integral_count_space_indicator: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Probability_Measure.thy Wed Feb 17 23:29:35 2016 +0100 @@ -247,7 +247,7 @@ moreover { fix y assume y: "y \ I" with q(2) \open I\ have "Sup ((\x. q x + ?F x * (y - x)) ` I) = q y" - by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI [OF _ y] simp: interior_open simp del: Sup_image_eq Inf_image_eq) } + by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI [OF _ y] simp: interior_open) } ultimately have "q (expectation X) = Sup ((\x. q x + ?F x * (expectation X - x)) ` I)" by simp also have "\ \ expectation (\w. q (X w))" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Radon_Nikodym.thy Wed Feb 17 23:29:35 2016 +0100 @@ -51,7 +51,7 @@ space: "(\i. A i) = space M" and measure: "\i. emeasure M (A i) \ \" and disjoint: "disjoint_family A" - using sigma_finite_disjoint by auto + using sigma_finite_disjoint by blast let ?B = "\i. 2^Suc i * emeasure M (A i)" have "\i. \x. 0 < x \ x < inverse (?B i)" proof @@ -560,7 +560,7 @@ proof (rule antisym) show "?a \ (SUP i. emeasure M (?O i))" unfolding a_Lim using Q' by (auto intro!: SUP_mono emeasure_mono) - show "(SUP i. emeasure M (?O i)) \ ?a" unfolding SUP_def + show "(SUP i. emeasure M (?O i)) \ ?a" proof (safe intro!: Sup_mono, unfold bex_simps) fix i have *: "(\(Q' ` {..i})) = ?O i" by auto diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Regularity.thy --- a/src/HOL/Probability/Regularity.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Regularity.thy Wed Feb 17 23:29:35 2016 +0100 @@ -138,7 +138,8 @@ proof safe fix x from X(2)[OF open_ball[of x r]] \r > 0\ obtain d where d: "d\X" "d \ ball x r" by auto show "x \ ?U" - using X(1) d by (auto intro!: exI[where x="to_nat_on X d"] simp: dist_commute Bex_def) + using X(1) d + by simp (auto intro!: exI [where x = "to_nat_on X d"] simp: dist_commute Bex_def) qed (simp add: sU) finally have "(\k. M (\n\{0..k}. cball (from_nat_into X n) r)) \ M (space M)" . } note M_space = this @@ -319,8 +320,8 @@ by (rule INF_superset_mono) (auto simp add: compact_imp_closed) also have "(INF U:{U. U \ B \ closed U}. M (space M - U)) = (INF U:{U. space M - B \ U \ open U}. emeasure M U)" - by (subst INF_image [of "\u. space M - u", symmetric, unfolded comp_def]) - (rule INF_cong, auto simp add: sU intro!: INF_cong) + unfolding INF_image [of _ "\u. space M - u" _, symmetric, unfolded comp_def] + by (rule INF_cong) (auto simp add: sU open_Compl Compl_eq_Diff_UNIV [symmetric, simp]) finally have "(INF U:{U. space M - B \ U \ open U}. emeasure M U) \ emeasure M (space M - B)" . moreover have @@ -335,8 +336,8 @@ also have "\ = (SUP U:{U. B \ U \ open U}. M (space M - U))" by (rule SUP_cong) (auto simp add: emeasure_compl sb compact_imp_closed) also have "\ = (SUP K:{K. K \ space M - B \ closed K}. emeasure M K)" - by (subst SUP_image [of "\u. space M - u", symmetric, simplified comp_def]) - (rule SUP_cong, auto simp: sU) + unfolding SUP_image [of _ "\u. space M - u" _, symmetric, unfolded comp_def] + by (rule SUP_cong) (auto simp add: sU) also have "\ = (SUP K:{K. K \ space M - B \ compact K}. emeasure M K)" proof (safe intro!: antisym SUP_least) fix K assume "closed K" "K \ space M - B" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Wed Feb 17 23:29:35 2016 +0100 @@ -432,10 +432,10 @@ by (auto simp add: binary_def) lemma Un_range_binary: "a \ b = (\i::nat. binary a b i)" - by (simp add: SUP_def range_binary_eq) + by (simp add: range_binary_eq cong del: strong_SUP_cong) lemma Int_range_binary: "a \ b = (\i::nat. binary a b i)" - by (simp add: INF_def range_binary_eq) + by (simp add: range_binary_eq cong del: strong_INF_cong) lemma sigma_algebra_iff2: "sigma_algebra \ M \ @@ -535,11 +535,14 @@ finally show ?thesis . qed -lemma sigma_sets_UNION: "countable B \ (\b. b \ B \ b \ sigma_sets X A) \ (\B) \ sigma_sets X A" - using from_nat_into[of B] range_from_nat_into[of B] sigma_sets.Union[of "from_nat_into B" X A] +lemma sigma_sets_UNION: + "countable B \ (\b. b \ B \ b \ sigma_sets X A) \ (\B) \ sigma_sets X A" apply (cases "B = {}") apply (simp add: sigma_sets.Empty) - apply (simp del: Union_image_eq add: Union_image_eq[symmetric]) + using from_nat_into [of B] range_from_nat_into [of B] sigma_sets.Union [of "from_nat_into B" X A] + apply simp + apply auto + apply (metis Sup_bot_conv(1) Union_empty `\B \ {}; countable B\ \ range (from_nat_into B) = B`) done lemma (in sigma_algebra) sigma_sets_eq: @@ -835,7 +838,7 @@ lemma (in semiring_of_sets) generated_ring_disjoint_UNION: "finite I \ disjoint (A ` I) \ (\i. i \ I \ A i \ generated_ring) \ UNION I A \ generated_ring" - unfolding SUP_def by (intro generated_ring_disjoint_Union) auto + by (intro generated_ring_disjoint_Union) auto lemma (in semiring_of_sets) generated_ring_Int: assumes a: "a \ generated_ring" and b: "b \ generated_ring" @@ -873,7 +876,7 @@ lemma (in semiring_of_sets) generated_ring_INTER: "finite I \ I \ {} \ (\i. i \ I \ A i \ generated_ring) \ INTER I A \ generated_ring" - unfolding INF_def by (intro generated_ring_Inter) auto + by (intro generated_ring_Inter) auto lemma (in semiring_of_sets) generating_ring: "ring_of_sets \ generated_ring" @@ -898,6 +901,7 @@ fix a b assume "a \ Ca" "b \ Cb" with Ca Cb Diff_cover[of a b] show "a - b \ ?R" by (auto simp add: generated_ring_def) + (metis DiffI Diff_eq_empty_iff empty_iff) next show "disjoint ((\a'. \b'\Cb. a' - b')`Ca)" using Ca by (auto simp add: disjoint_def \Cb \ {}\) @@ -935,7 +939,7 @@ done lemma UN_binaryset_eq: "(\i. binaryset A B i) = A \ B" - by (simp add: SUP_def range_binaryset_eq) + by (simp add: range_binaryset_eq cong del: strong_SUP_cong) subsubsection \Closed CDI\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/ROOT --- a/src/HOL/ROOT Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/ROOT Wed Feb 17 23:29:35 2016 +0100 @@ -38,6 +38,7 @@ Product_Lexorder Product_Order Finite_Lattice + Polynomial_GCD_euclidean (*data refinements and dependent applications*) AList_Mapping Code_Binary_Nat diff -r e307a410f46c -r ab76bd43c14a src/HOL/Rat.thy --- a/src/HOL/Rat.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Rat.thy Wed Feb 17 23:29:35 2016 +0100 @@ -78,7 +78,7 @@ from \q = Fract a b\ \b \ 0\ \?b \ 0\ have q: "q = Fract ?a ?b" by (simp add: eq_rat dvd_div_mult mult.commute [of a]) from \b \ 0\ have coprime: "coprime ?a ?b" - by (auto intro: div_gcd_coprime_int) + by (auto intro: div_gcd_coprime) show C proof (cases "b > 0") case True note assms @@ -287,7 +287,7 @@ split:split_if_asm) lemma normalize_coprime: "normalize r = (p, q) \ coprime p q" - by (auto simp add: normalize_def Let_def dvd_div_neg div_gcd_coprime_int + by (auto simp add: normalize_def Let_def dvd_div_neg div_gcd_coprime split:split_if_asm) lemma normalize_stable [simp]: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Real.thy --- a/src/HOL/Real.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Real.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1037,9 +1037,6 @@ declare of_int_1_less_iff [algebra, presburger] declare of_int_1_le_iff [algebra, presburger] -lemma of_int_abs [simp]: "of_int \x\ = (\of_int x\ :: 'a::linordered_idom)" - by (auto simp add: abs_if) - lemma int_less_real_le: "(n < m) = (real_of_int n + 1 <= real_of_int m)" proof - have "(0::real) \ 1" @@ -1285,7 +1282,7 @@ @{thm of_nat_mult}, @{thm of_int_0}, @{thm of_int_1}, @{thm of_int_add}, @{thm of_int_minus}, @{thm of_int_diff}, @{thm of_int_mult}, @{thm of_int_of_nat_eq}, - @{thm of_nat_numeral}, @{thm int_numeral}, @{thm of_int_neg_numeral}] + @{thm of_nat_numeral}, @{thm of_nat_numeral}, @{thm of_int_neg_numeral}] #> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat \ real"}) #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int \ real"})) \ @@ -1309,22 +1306,9 @@ subsection \Lemmas about powers\ -(* used by Import/HOL/real.imp *) lemma two_realpow_ge_one: "(1::real) \ 2 ^ n" by simp -text \FIXME: no longer real-specific; rename and move elsewhere\ -lemma realpow_Suc_le_self: - fixes r :: "'a::linordered_semidom" - shows "[| 0 \ r; r \ 1 |] ==> r ^ Suc n \ r" -by (insert power_decreasing [of 1 "Suc n" r], simp) - -text \FIXME: no longer real-specific; rename and move elsewhere\ -lemma realpow_minus_mult: - fixes x :: "'a::monoid_mult" - shows "0 < n \ x ^ (n - 1) * x = x ^ n" -by (simp add: power_Suc power_commutes split add: nat_diff_split) - text \FIXME: declare this [simp] for all types, or not at all\ declare sum_squares_eq_zero_iff [simp] sum_power2_eq_zero_iff [simp] diff -r e307a410f46c -r ab76bd43c14a src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Wed Feb 17 23:29:35 2016 +0100 @@ -9,9 +9,6 @@ imports Real Topological_Spaces begin -lemma (in ordered_ab_group_add) diff_ge_0_iff_ge [simp]: "a - b \ 0 \ a \ b" - by (simp add: le_diff_eq) - subsection \Locale for additive functions\ locale additive = diff -r e307a410f46c -r ab76bd43c14a src/HOL/Relation.thy --- a/src/HOL/Relation.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Relation.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1060,10 +1060,11 @@ text\Converse inclusion requires some assumptions\ lemma Image_INT_eq: - "[|single_valued (r\); A\{}|] ==> r `` INTER A B = (\x\A. r `` B x)" + "single_valued (r\) \ A \ {} \ r `` INTER A B = (\x\A. r `` B x)" apply (rule equalityI) apply (rule Image_INT_subset) -apply (simp add: single_valued_def, blast) +apply (auto simp add: single_valued_def) +apply blast done lemma Image_subset_eq: "(r``A \ B) = (A \ - ((r^-1) `` (-B)))" diff -r e307a410f46c -r ab76bd43c14a src/HOL/Rings.thy --- a/src/HOL/Rings.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Rings.thy Wed Feb 17 23:29:35 2016 +0100 @@ -145,7 +145,7 @@ show "a = a * 1" by simp qed -lemma dvd_trans: +lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - @@ -1562,6 +1562,14 @@ proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: split_if_asm) +lemma sum_squares_ge_zero: + "0 \ x * x + y * y" + by (intro add_nonneg_nonneg zero_le_square) + +lemma not_sum_squares_lt_zero: + "\ x * x + y * y < 0" + by (simp add: not_less sum_squares_ge_zero) + end class linordered_ring_strict = ring + linordered_semiring_strict @@ -1867,6 +1875,10 @@ "sgn a < 0 \ a < 0" unfolding sgn_if by auto +lemma abs_sgn_eq: + "\sgn a\ = (if a = 0 then 0 else 1)" + by (simp add: sgn_if) + lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) diff -r e307a410f46c -r ab76bd43c14a src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Set_Interval.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1085,26 +1085,33 @@ qed qed -lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" +lemma UN_UN_finite_eq: + "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) -lemma UN_finite_subset: "(!!n::nat. (\i\{0.. C) \ (\n. A n) \ C" +lemma UN_finite_subset: + "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: - "(!!n::nat. (\i\{0.. (\i\{0.. (\n. A n) \ (\n. B n)" - apply (rule UN_finite_subset) - apply (subst UN_UN_finite_eq [symmetric, of B]) - apply blast - done + assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" +proof (rule UN_finite_subset, rule) + fix n and a + from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) +qed lemma UN_finite2_eq: - "(!!n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" + "(\n::nat. (\i\{0..i\{0.. + (\n. A n) = (\n. B n)" apply (rule subset_antisym) apply (rule UN_finite2_subset, blast) - apply (rule UN_finite2_subset [where k=k]) - apply (force simp add: atLeastLessThan_add_Un [of 0]) - done + apply (rule UN_finite2_subset [where k=k]) + apply (force simp add: atLeastLessThan_add_Un [of 0]) + done subsubsection \Cardinality\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/BNF/bnf_comp_tactics.ML --- a/src/HOL/Tools/BNF/bnf_comp_tactics.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/BNF/bnf_comp_tactics.ML Wed Feb 17 23:29:35 2016 +0100 @@ -251,7 +251,7 @@ val simplified_set_simps = @{thms collect_def[abs_def] UN_insert UN_empty Un_empty_right Un_empty_left - o_def Union_Un_distrib Union_image_eq UN_empty2 UN_singleton id_bnf_def}; + o_def Union_Un_distrib UN_empty2 UN_singleton id_bnf_def}; fun mk_simplified_set_tac ctxt collect_set_map = unfold_thms_tac ctxt (collect_set_map :: @{thms comp_assoc}) THEN diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/BNF/bnf_lfp_tactics.ML --- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML Wed Feb 17 23:29:35 2016 +0100 @@ -540,8 +540,8 @@ fun mk_ctor_set_tac ctxt set set_map set_maps = let val n = length set_maps; - fun mk_UN thm = rtac ctxt (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN' - rtac ctxt @{thm Union_image_eq}; + fun mk_UN thm = rtac ctxt (thm RS @{thm arg_cong[of _ _ Union]} RS trans) + THEN' rtac ctxt @{thm refl}; in EVERY' [rtac ctxt (set RS @{thm comp_eq_dest} RS trans), rtac ctxt Un_cong, rtac ctxt (trans OF [set_map, trans_fun_cong_image_id_id_apply]), diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/BNF/bnf_util.ML --- a/src/HOL/Tools/BNF/bnf_util.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/BNF/bnf_util.ML Wed Feb 17 23:29:35 2016 +0100 @@ -257,8 +257,12 @@ Const (@{const_name Bex}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f; fun mk_UNION X f = - let val (T, U) = dest_funT (fastype_of f); - in Const (@{const_name SUPREMUM}, fastype_of X --> (T --> U) --> U) $ X $ f end; + let + val (T, U) = dest_funT (fastype_of f); + in + Const (@{const_name Sup}, HOLogic.mk_setT U --> U) + $ (Const (@{const_name image}, (T --> U) --> fastype_of X --> HOLogic.mk_setT U) $ f $ X) + end; fun mk_Union T = Const (@{const_name Sup}, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T); diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Feb 17 23:29:35 2016 +0100 @@ -1672,7 +1672,7 @@ fun do_numeral depth Ts mult T some_t0 t1 t2 = (if is_number_type ctxt T then let - val j = mult * HOLogic.dest_num t2 + val j = mult * HOLogic.dest_numeral t2 in if j = 1 then raise SAME () diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/Qelim/cooper.ML Wed Feb 17 23:29:35 2016 +0100 @@ -812,14 +812,14 @@ simpset_of (put_simpset comp_ss @{context} addsimps @{thms simp_thms} @ [@{thm "nat_numeral"} RS sym, @{thm "zdvd_int"}, @{thm "of_nat_add"}, @{thm "of_nat_mult"}] - @ map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}] + @ map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "of_nat_less_iff" [where ?'a = int]}] |> Splitter.add_split @{thm "zdiff_int_split"}) val ss2 = simpset_of (put_simpset HOL_basic_ss @{context} - addsimps [@{thm "nat_0_le"}, @{thm "int_numeral"}, + addsimps [@{thm "nat_0_le"}, @{thm "of_nat_numeral"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "zero_le_numeral"}, - @{thm "le_numeral_extra"(3)}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}] + @{thm "le_numeral_extra"(3)}, @{thm "of_nat_0"}, @{thm "of_nat_1"}, @{thm "Suc_eq_plus1"}] |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]) val div_mod_ss = simpset_of (put_simpset HOL_basic_ss @{context} diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/hologic.ML --- a/src/HOL/Tools/hologic.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/hologic.ML Wed Feb 17 23:29:35 2016 +0100 @@ -100,7 +100,7 @@ val mk_bit: int -> term val dest_bit: term -> int val mk_numeral: int -> term - val dest_num: term -> int + val dest_numeral: term -> int val numeral_const: typ -> term val add_numerals: term -> (term * typ) list -> (term * typ) list val mk_number: typ -> int -> term @@ -533,10 +533,10 @@ in if i > 0 then mk i else raise TERM ("mk_numeral: " ^ string_of_int i, []) end -fun dest_num (Const ("Num.num.One", _)) = 1 - | dest_num (Const ("Num.num.Bit0", _) $ bs) = 2 * dest_num bs - | dest_num (Const ("Num.num.Bit1", _) $ bs) = 2 * dest_num bs + 1 - | dest_num t = raise TERM ("dest_num", [t]); +fun dest_numeral (Const ("Num.num.One", _)) = 1 + | dest_numeral (Const ("Num.num.Bit0", _) $ bs) = 2 * dest_numeral bs + | dest_numeral (Const ("Num.num.Bit1", _) $ bs) = 2 * dest_numeral bs + 1 + | dest_numeral t = raise TERM ("dest_num", [t]); fun numeral_const T = Const ("Num.numeral_class.numeral", numT --> T); @@ -554,7 +554,7 @@ fun dest_number (Const ("Groups.zero_class.zero", T)) = (T, 0) | dest_number (Const ("Groups.one_class.one", T)) = (T, 1) | dest_number (Const ("Num.numeral_class.numeral", Type ("fun", [_, T])) $ t) = - (T, dest_num t) + (T, dest_numeral t) | dest_number (Const ("Groups.uminus_class.uminus", Type ("fun", [_, T])) $ t) = apsnd (op ~) (dest_number t) | dest_number t = raise TERM ("dest_number", [t]); diff -r e307a410f46c -r ab76bd43c14a src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Tools/lin_arith.ML Wed Feb 17 23:29:35 2016 +0100 @@ -175,11 +175,11 @@ | demult (Const (@{const_name Groups.zero}, _), _) = (NONE, Rat.zero) | demult (Const (@{const_name Groups.one}, _), m) = (NONE, m) (*Warning: in rare cases (neg_)numeral encloses a non-numeral, - in which case dest_num raises TERM; hence all the handles below. + in which case dest_numeral raises TERM; hence all the handles below. Same for Suc-terms that turn out not to be numerals - although the simplifier should eliminate those anyway ...*) | demult (t as Const ("Num.numeral_class.numeral", _) $ n, m) = - ((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_num n))) + ((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_numeral n))) handle TERM _ => (SOME t, m)) | demult (t as Const (@{const_name Suc}, _) $ _, m) = ((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_nat t))) @@ -208,7 +208,7 @@ | poly (Const (@{const_name Groups.one}, _), m, (p, i)) = (p, Rat.add i m) | poly (all as Const ("Num.numeral_class.numeral", Type(_,[_,_])) $ t, m, pi as (p, i)) = - (let val k = HOLogic.dest_num t + (let val k = HOLogic.dest_numeral t in (p, Rat.add i (Rat.mult m (Rat.rat_of_int k))) end handle TERM _ => add_atom all m pi) | poly (Const (@{const_name Suc}, _) $ t, m, (p, i)) = diff -r e307a410f46c -r ab76bd43c14a src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Topological_Spaces.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1133,10 +1133,11 @@ using A by auto show "nhds x = (INF n. principal (\i\n. A i))" using A unfolding nhds_def - apply (intro INF_eq) + apply - + apply (rule INF_eq) apply simp_all - apply force - apply (intro exI[of _ "\i\n. A i" for n] conjI open_INT) + apply fastforce + apply (intro exI [of _ "\i\n. A i" for n] conjI open_INT) apply auto done qed @@ -1464,7 +1465,7 @@ lemma continuous_on_open_UN: "(\s. s \ S \ open (A s)) \ (\s. s \ S \ continuous_on (A s) f) \ continuous_on (\s\S. A s) f" - unfolding Union_image_eq[symmetric] by (rule continuous_on_open_Union) auto + by (rule continuous_on_open_Union) auto lemma continuous_on_open_Un: "open s \ open t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" @@ -1689,7 +1690,7 @@ lemma compactE_image: assumes "compact s" and "\t\C. open (f t)" and "s \ (\c\C. f c)" obtains C' where "C' \ C" and "finite C'" and "s \ (\c\C'. f c)" - using assms unfolding ball_simps[symmetric] SUP_def + using assms unfolding ball_simps [symmetric] by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s]) lemma compact_inter_closed [intro]: diff -r e307a410f46c -r ab76bd43c14a src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Transcendental.thy Wed Feb 17 23:29:35 2016 +0100 @@ -29,27 +29,15 @@ qed -lemma of_int_leD: - fixes x :: "'a :: linordered_idom" - shows "\of_int n\ \ x \ n=0 \ x\1" - by (metis (no_types) le_less_trans not_less of_int_abs of_int_less_1_iff zabs_less_one_iff) - -lemma of_int_lessD: - fixes x :: "'a :: linordered_idom" - shows "\of_int n\ < x \ n=0 \ x>1" - by (metis less_le_trans not_less of_int_abs of_int_less_1_iff zabs_less_one_iff) - -lemma fact_in_Reals: "fact n \ \" by (induction n) auto +lemma fact_in_Reals: "fact n \ \" + by (induction n) auto + +lemma of_real_fact [simp]: "of_real (fact n) = fact n" + by (metis of_nat_fact of_real_of_nat_eq) lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" by (simp add: pochhammer_def) -lemma of_real_fact [simp]: "of_real (fact n) = fact n" - by (metis of_nat_fact of_real_of_nat_eq) - -lemma of_int_fact [simp]: "of_int (fact n :: int) = fact n" - by (metis of_int_of_nat_eq of_nat_fact) - lemma norm_fact [simp]: "norm (fact n :: 'a :: {real_normed_algebra_1}) = fact n" proof - diff -r e307a410f46c -r ab76bd43c14a src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Transitive_Closure.thy Wed Feb 17 23:29:35 2016 +0100 @@ -946,7 +946,7 @@ apply (rule iffI) apply (drule tranclD2) apply (clarsimp simp: rtrancl_is_UN_relpow) - apply (rule_tac x="Suc n" in exI) + apply (rule_tac x="Suc x" in exI) apply (clarsimp simp: relcomp_unfold) apply fastforce apply clarsimp @@ -1093,9 +1093,9 @@ assumes "finite R" and "finite S" shows "finite(R O S)" proof- - have "R O S = (UN (x,y) : R. \((%(u,v). if u=y then {(x,v)} else {}) ` S))" - by(force simp add: split_def) - thus ?thesis using assms by(clarsimp) + have "R O S = (\(x, y)\R. \(u, v)\S. if u = y then {(x, v)} else {})" + by (force simp add: split_def image_constant_conv split: if_splits) + then show ?thesis using assms by clarsimp qed lemma finite_relpow[simp,intro]: diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/Comp/Alloc.thy --- a/src/HOL/UNITY/Comp/Alloc.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/Comp/Alloc.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1035,7 +1035,7 @@ apply (simp only: o_apply sub_def) apply (insert Alloc_Progress [THEN rename_guarantees_sysOfAlloc_I]) apply (simp add: o_def del: INT_iff) - apply (erule component_guaranteesD) + apply (drule component_guaranteesD) apply (auto simp add: System_Increasing_allocRel [simplified sub_apply o_def] System_Increasing_allocAsk [simplified sub_apply o_def] @@ -1047,6 +1047,7 @@ lemma System_Progress: "System : system_progress" apply (unfold system_progress_def) apply (cut_tac System_Alloc_Progress) + apply auto apply (blast intro: LeadsTo_Trans System_Follows_allocGiv [THEN Follows_LeadsTo_pfixLe] System_Follows_ask [THEN Follows_LeadsTo]) diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/ELT.thy --- a/src/HOL/UNITY/ELT.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/ELT.thy Wed Feb 17 23:29:35 2016 +0100 @@ -141,7 +141,6 @@ lemma leadsETo_UN: "(!!i. i : I ==> F : (A i) leadsTo[CC] B) ==> F : (UN i:I. A i) leadsTo[CC] B" -apply (subst Union_image_eq [symmetric]) apply (blast intro: leadsETo_Union) done @@ -397,7 +396,6 @@ lemma LeadsETo_UN: "(!!i. i : I ==> F : (A i) LeadsTo[CC] B) ==> F : (UN i:I. A i) LeadsTo[CC] B" -apply (simp only: Union_image_eq [symmetric]) apply (blast intro: LeadsETo_Union) done diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/Extend.thy --- a/src/HOL/UNITY/Extend.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/Extend.thy Wed Feb 17 23:29:35 2016 +0100 @@ -382,21 +382,25 @@ (project_act h -` AllowedActs F)})" apply (rule program_equalityI) apply simp - apply (simp add: image_eq_UN) + apply (simp add: image_image) apply (simp add: project_def) done lemma extend_inverse [simp]: "project h UNIV (extend h F) = F" -apply (simp (no_asm_simp) add: project_extend_eq image_eq_UN +apply (simp (no_asm_simp) add: project_extend_eq subset_UNIV [THEN subset_trans, THEN Restrict_triv]) apply (rule program_equalityI) apply (simp_all (no_asm)) apply (subst insert_absorb) apply (simp (no_asm) add: bexI [of _ Id]) apply auto +apply (simp add: image_def) +using project_act_Id apply blast +apply (simp add: image_def) apply (rename_tac "act") -apply (rule_tac x = "extend_act h act" in bexI, auto) +apply (rule_tac x = "extend_act h act" in exI) +apply simp done lemma inj_extend: "inj (extend h)" @@ -641,11 +645,12 @@ subsection{*Guarantees*} lemma project_extend_Join: "project h UNIV ((extend h F)\G) = F\(project h UNIV G)" -apply (rule program_equalityI) - apply (simp add: project_set_extend_set_Int) - apply (auto simp add: image_eq_UN) -done - + apply (rule program_equalityI) + apply (auto simp add: project_set_extend_set_Int image_iff) + apply (metis Un_iff extend_act_inverse image_iff) + apply (metis Un_iff extend_act_inverse image_iff) + done + lemma extend_Join_eq_extend_D: "(extend h F)\G = extend h H ==> H = F\(project h UNIV G)" apply (drule_tac f = "project h UNIV" in arg_cong) diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/ProgressSets.thy --- a/src/HOL/UNITY/ProgressSets.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/ProgressSets.thy Wed Feb 17 23:29:35 2016 +0100 @@ -42,13 +42,11 @@ lemma UN_in_lattice: "[|lattice L; !!i. i\I ==> r i \ L|] ==> (\i\I. r i) \ L" -apply (unfold SUP_def) apply (blast intro: Union_in_lattice) done lemma INT_in_lattice: "[|lattice L; !!i. i\I ==> r i \ L|] ==> (\i\I. r i) \ L" -apply (unfold INF_def) apply (blast intro: Inter_in_lattice) done diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/Rename.thy --- a/src/HOL/UNITY/Rename.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/Rename.thy Wed Feb 17 23:29:35 2016 +0100 @@ -272,7 +272,7 @@ (F \ (rename (inv h) ` X) guarantees (rename (inv h) ` Y))" apply (subst rename_rename_guarantees_eq [symmetric], assumption) -apply (simp add: image_eq_UN o_def bij_is_surj [THEN surj_f_inv_f]) +apply (simp add: o_def bij_is_surj [THEN surj_f_inv_f] image_comp) done lemma rename_preserves: diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/SubstAx.thy --- a/src/HOL/UNITY/SubstAx.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/SubstAx.thy Wed Feb 17 23:29:35 2016 +0100 @@ -84,7 +84,6 @@ lemma LeadsTo_UN: "(!!i. i \ I ==> F \ (A i) LeadsTo B) ==> F \ (\i \ I. A i) LeadsTo B" -apply (unfold SUP_def) apply (blast intro: LeadsTo_Union) done @@ -188,7 +187,6 @@ lemma LeadsTo_UN_UN: "(!! i. i \ I ==> F \ (A i) LeadsTo (A' i)) ==> F \ (\i \ I. A i) LeadsTo (\i \ I. A' i)" -apply (simp only: Union_image_eq [symmetric]) apply (blast intro: LeadsTo_Union LeadsTo_weaken_R) done diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/Transformers.thy --- a/src/HOL/UNITY/Transformers.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/Transformers.thy Wed Feb 17 23:29:35 2016 +0100 @@ -467,7 +467,7 @@ "single_valued act ==> insert (wens_single act B) (range (wens_single_finite act B)) \ wens_set (mk_program (init, {act}, allowed)) B" -apply (simp add: SUP_def image_def wens_single_eq_Union) +apply (simp add: image_def wens_single_eq_Union) apply (blast intro: wens_set.Union wens_single_finite_in_wens_set) done diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/Union.thy --- a/src/HOL/UNITY/Union.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/Union.thy Wed Feb 17 23:29:35 2016 +0100 @@ -396,11 +396,30 @@ lemma safety_prop_Int [simp]: "safety_prop X \ safety_prop Y \ safety_prop (X \ Y)" - by (simp add: safety_prop_def) blast +proof (clarsimp simp add: safety_prop_def) + fix G + assume "\G. Acts G \ (\x\X. Acts x) \ G \ X" + then have X: "Acts G \ (\x\X. Acts x) \ G \ X" by blast + assume "\G. Acts G \ (\x\Y. Acts x) \ G \ Y" + then have Y: "Acts G \ (\x\Y. Acts x) \ G \ Y" by blast + assume Acts: "Acts G \ (\x\X \ Y. Acts x)" + with X and Y show "G \ X \ G \ Y" by auto +qed lemma safety_prop_INTER [simp]: "(\i. i \ I \ safety_prop (X i)) \ safety_prop (\i\I. X i)" - by (simp add: safety_prop_def) blast +proof (clarsimp simp add: safety_prop_def) + fix G and i + assume "\i. i \ I \ \ \ X i \ + (\G. Acts G \ (\x\X i. Acts x) \ G \ X i)" + then have *: "i \ I \ Acts G \ (\x\X i. Acts x) \ G \ X i" + by blast + assume "i \ I" + moreover assume "Acts G \ (\j\\i\I. X i. Acts j)" + ultimately have "Acts G \ (\i\X i. Acts i)" + by auto + with * `i \ I` show "G \ X i" by blast +qed lemma safety_prop_INTER1 [simp]: "(\i. safety_prop (X i)) \ safety_prop (\i. X i)" diff -r e307a410f46c -r ab76bd43c14a src/HOL/UNITY/WFair.thy --- a/src/HOL/UNITY/WFair.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/UNITY/WFair.thy Wed Feb 17 23:29:35 2016 +0100 @@ -203,7 +203,6 @@ lemma leadsTo_UN: "(!!i. i \ I ==> F \ (A i) leadsTo B) ==> F \ (\i \ I. A i) leadsTo B" -apply (subst Union_image_eq [symmetric]) apply (blast intro: leadsTo_Union) done @@ -310,7 +309,6 @@ lemma leadsTo_UN_UN: "(!! i. i \ I ==> F \ (A i) leadsTo (A' i)) ==> F \ (\i \ I. A i) leadsTo (\i \ I. A' i)" -apply (simp only: Union_image_eq [symmetric]) apply (blast intro: leadsTo_Union leadsTo_weaken_R) done diff -r e307a410f46c -r ab76bd43c14a src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/Word/Word.thy Wed Feb 17 23:29:35 2016 +0100 @@ -1416,7 +1416,7 @@ apply (unfold udvd_def) apply safe apply (simp add: unat_def nat_mult_distrib) - apply (simp add: uint_nat int_mult) + apply (simp add: uint_nat of_nat_mult) apply (rule exI) apply safe prefer 2 @@ -3295,7 +3295,7 @@ apply (unfold dvd_def) apply auto apply (drule uint_ge_0 [THEN nat_int.Abs_inverse' [simplified], symmetric]) - apply (simp add : int_mult of_nat_power) + apply (simp add : of_nat_mult of_nat_power) apply (simp add : nat_mult_distrib nat_power_eq) done @@ -4240,9 +4240,10 @@ apply (simp add: rotater_add [symmetric] rotate_gal [symmetric]) apply (rule rotater_eqs) apply (simp add: word_size nat_mod_distrib) - apply (rule int_eq_0_conv [THEN iffD1]) - apply (simp only: zmod_int of_nat_add) - apply (simp add: rdmods) + apply (rule of_nat_eq_0_iff [THEN iffD1]) + apply (auto simp add: not_le mod_eq_0_iff_dvd zdvd_int nat_add_distrib [symmetric]) + using mod_mod_trivial zmod_eq_dvd_iff + apply blast done lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] diff -r e307a410f46c -r ab76bd43c14a src/HOL/ex/LocaleTest2.thy --- a/src/HOL/ex/LocaleTest2.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/ex/LocaleTest2.thy Wed Feb 17 23:29:35 2016 +0100 @@ -599,7 +599,7 @@ apply (rule_tac x = "gcd x y" in exI) apply auto [1] apply (rule_tac x = "lcm x y" in exI) - apply (auto intro: lcm_dvd1_nat lcm_dvd2_nat lcm_least_nat) + apply (auto intro: dvd_lcm1 dvd_lcm2 lcm_least) done then interpret nat_dvd: dlat "op dvd :: [nat, nat] => bool" . txt \Interpretation to ease use of definitions, which are @@ -613,7 +613,7 @@ apply (unfold nat_dvd.join_def) apply (rule the_equality) apply (unfold nat_dvd.is_sup_def) - by (auto intro: lcm_dvd1_nat lcm_dvd2_nat lcm_least_nat) + by (auto intro: dvd_lcm1 dvd_lcm2 lcm_least) qed text \Interpreted theorems from the locales, involving defined terms.\ diff -r e307a410f46c -r ab76bd43c14a src/HOL/ex/NatSum.thy --- a/src/HOL/ex/NatSum.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/ex/NatSum.thy Wed Feb 17 23:29:35 2016 +0100 @@ -98,7 +98,7 @@ "30 * int (\i=0..i=0..2 = p * k\<^sup>2" by (simp add: power2_eq_square) then have "p dvd n\<^sup>2" .. with \prime p\ have "p dvd n" by (rule prime_dvd_power_nat) - with dvd_m have "p dvd gcd m n" by (rule gcd_greatest_nat) + with dvd_m have "p dvd gcd m n" by (rule gcd_greatest) with \coprime m n\ have "p = 1" by simp with p show False by simp qed diff -r e307a410f46c -r ab76bd43c14a src/HOL/ex/Transfer_Int_Nat.thy --- a/src/HOL/ex/Transfer_Int_Nat.thy Wed Feb 17 23:28:58 2016 +0100 +++ b/src/HOL/ex/Transfer_Int_Nat.thy Wed Feb 17 23:29:35 2016 +0100 @@ -40,10 +40,10 @@ unfolding rel_fun_def ZN_def by simp lemma ZN_mult [transfer_rule]: "(ZN ===> ZN ===> ZN) (op *) (op *)" - unfolding rel_fun_def ZN_def by (simp add: int_mult) + unfolding rel_fun_def ZN_def by (simp add: of_nat_mult) lemma ZN_diff [transfer_rule]: "(ZN ===> ZN ===> ZN) tsub (op -)" - unfolding rel_fun_def ZN_def tsub_def by (simp add: zdiff_int) + unfolding rel_fun_def ZN_def tsub_def by (simp add: of_nat_diff) lemma ZN_power [transfer_rule]: "(ZN ===> op = ===> ZN) (op ^) (op ^)" unfolding rel_fun_def ZN_def by (simp add: of_nat_power) @@ -65,7 +65,7 @@ lemma ZN_Ex [transfer_rule]: "((ZN ===> op =) ===> op =) (Bex {0..}) Ex" unfolding rel_fun_def ZN_def Bex_def atLeast_iff - by (metis zero_le_imp_eq_int zero_zle_int) + by (metis zero_le_imp_eq_int of_nat_0_le_iff) lemma ZN_le [transfer_rule]: "(ZN ===> ZN ===> op =) (op \) (op \)" unfolding rel_fun_def ZN_def by simp diff -r e307a410f46c -r ab76bd43c14a src/Pure/ML/exn_output.ML diff -r e307a410f46c -r ab76bd43c14a src/Pure/ML/ml_compiler.ML