--- 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
-----------------------------
--- 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 ***
--- 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 \<subseteq> carrier G"
unfolding subgroup_def by simp
from xcarr Hcarr show "(\<Union>h\<in>H. {h \<oplus>\<^bsub>G\<^esub> x}) = (\<Union>h\<in>H. {x \<oplus>\<^bsub>G\<^esub> h})"
- using m_comm [simplified] by fast
+ using m_comm [simplified] by fastforce
qed
qed
--- 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 \<Longrightarrow> 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 _ "\<one>"])
apply (auto simp add: subgroup.one_closed subgroup.subset [THEN subsetD])
@@ -932,7 +932,7 @@
obtain g where g: "g \<in> 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 \<subseteq> carrier G" and X'sub: "X' \<subseteq> carrier G"
by (force simp add: kernel_def r_coset_def image_def)+
hence "h ` (X <#> X') = {h g \<otimes>\<^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) \<otimes>\<^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) \<otimes>\<^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:
"\<lbrakk>g \<in> carrier G; g' \<in> carrier G; h g = h g'\<rbrakk>
\<Longrightarrow> kernel G H h #> g \<subseteq> 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 \<otimes> g \<otimes> 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 "(\<Union>x\<in>kernel G H h #> g. {h x}) = {y}"
by (auto simp add: y kernel_def r_coset_def)
with g show "y \<in> (\<lambda>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
--- 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\<open>Prime Theorems\<close>
lemma prime_iff:
- "(prime p) = (Suc 0 < p & (\<forall>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 \<longleftrightarrow> Suc 0 < p \<and> (\<forall>a b. p dvd a * b \<longrightarrow> p dvd a \<or> 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 |]
--- 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 "\<bar>int p * x\<bar> = 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
--- 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 \<longleftrightarrow> x - a \<in> \<int> \<and> 0 \<le> a \<and> 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 \<longleftrightarrow> x - a \<in> \<int> \<and> 0 \<le> a \<and> a < 1"
+ apply (auto simp: Ints_def frac_def algebra_simps floor_unique)
+ apply linarith+
+ done
lemma frac_eq: "(frac x) = x \<longleftrightarrow> 0 \<le> x \<and> x < 1"
by (simp add: frac_unique_iff)
--- 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"
--- 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(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
-by (intro equalityI parts_UN_subset1 parts_UN_subset2)
+lemma parts_UN [simp]:
+ "parts (\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts (H x))"
+ by (intro equalityI parts_UN_subset1 parts_UN_subset2)
+
+lemma parts_image [simp]:
+ "parts (f ` A) = (\<Union>x\<in>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\<open>Added to simplify arguments to parts, analz and synth.
NOTE: the UN versions are no longer used!\<close>
@@ -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\<open>In any message, there is an upper bound N on its greatest nonce.\<close>
lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
--- 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 \<in> responses evs |]
==> (Key K \<in> parts{RB} | Key K \<in> 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\<open>Simplification using two distinct treatments of "image"\<close>
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\<open>Base case of respond\<close>
--- 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 \<union> Nonce`N) = Key`K \<union> Nonce`N"
-apply auto
-done
+lemma analz_image_Key_Un_Nonce:
+ "analz (Key ` K \<union> Nonce ` N) = Key ` K \<union> Nonce ` N"
+ by (auto simp del: parts_image)
method_setup sc_analz_freshK = \<open>
Scan.succeed (fn ctxt =>
--- 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 \<union> Nonce`N) = Key`K \<union> Nonce`N"
-apply auto
-done
+lemma analz_image_Key_Un_Nonce:
+ "analz (Key ` K \<union> Nonce ` N) = Key ` K \<union> Nonce ` N"
+ by (auto simp del: parts_image)
method_setup sc_analz_freshK = \<open>
Scan.succeed (fn ctxt =>
--- 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\<open>Added to extend initstate with set of nonces\<close>
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)
--- 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:
"|\<Union>i \<in> I. A i| \<le>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 \<noteq> a2"
--- 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 \<Longrightarrow> 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)
--- 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 \<Longrightarrow> A = \<Union>{under a| a. a \<in> A}"
-using ofilter_under_UNION[of A]
-by(unfold Union_eq, auto)
+using ofilter_under_UNION [of A] by auto
subsubsection {* Other properties *}
--- 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"
--- 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 \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
begin
-definition INFIMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
- INF_def: "INFIMUM A f = \<Sqinter>(f ` A)"
-
-lemma Inf_image_eq [simp]:
- "\<Sqinter>(f ` A) = INFIMUM A f"
- by (simp add: INF_def)
+abbreviation INFIMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+ "INFIMUM A f \<equiv> \<Sqinter>(f ` A)"
lemma INF_image [simp]:
"INFIMUM (f ` A) g = INFIMUM A (g \<circ> f)"
- by (simp only: INF_def image_comp)
+ by (simp add: image_comp)
lemma INF_identity_eq [simp]:
"INFIMUM A (\<lambda>x. x) = \<Sqinter>A"
- by (simp add: INF_def)
+ by simp
lemma INF_id_eq [simp]:
"INFIMUM A id = \<Sqinter>A"
- by (simp add: id_def)
+ by simp
lemma INF_cong:
"A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> 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 \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> INFIMUM A C = INFIMUM B D"
@@ -50,20 +47,17 @@
fixes Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
begin
-definition SUPREMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
- SUP_def: "SUPREMUM A f = \<Squnion>(f ` A)"
-
-lemma Sup_image_eq [simp]:
- "\<Squnion>(f ` A) = SUPREMUM A f"
- by (simp add: SUP_def)
+abbreviation SUPREMUM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+ "SUPREMUM A f \<equiv> \<Squnion>(f ` A)"
lemma SUP_image [simp]:
"SUPREMUM (f ` A) g = SUPREMUM A (g \<circ> f)"
- by (simp only: SUP_def image_comp)
+ by (simp add: image_comp)
lemma SUP_identity_eq [simp]:
"SUPREMUM A (\<lambda>x. x) = \<Squnion>A"
- by (simp add: SUP_def)
+ by simp
lemma SUP_id_eq [simp]:
"SUPREMUM A id = \<Squnion>A"
@@ -71,7 +65,7 @@
lemma SUP_cong:
"A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> 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 \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> 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:
"(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> (\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> \<Squnion>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]: "(\<Sqinter>x\<in>insert a A. f x) = f a \<sqinter> INFIMUM A f"
- unfolding INF_def Inf_insert by simp
+ by (simp cong del: strong_INF_cong)
lemma Sup_insert [simp]: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper)
lemma SUP_insert [simp]: "(\<Squnion>x\<in>insert a A. f x) = f a \<squnion> SUPREMUM A f"
- unfolding SUP_def Sup_insert by simp
+ by (simp cong del: strong_SUP_cong)
lemma INF_empty [simp]: "(\<Sqinter>x\<in>{}. f x) = \<top>"
- by (simp add: INF_def)
+ by (simp cong del: strong_INF_cong)
lemma SUP_empty [simp]: "(\<Squnion>x\<in>{}. f x) = \<bottom>"
- by (simp add: SUP_def)
+ by (simp cong del: strong_SUP_cong)
lemma Inf_UNIV [simp]:
"\<Sqinter>UNIV = \<bottom>"
@@ -308,18 +294,18 @@
ultimately show ?thesis by (rule Sup_upper2)
qed
+lemma INF_eq:
+ assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<ge> g j"
+ assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<ge> f i"
+ shows "INFIMUM A f = INFIMUM B g"
+ by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+
+
lemma SUP_eq:
assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<le> g j"
assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<le> f i"
- shows "(\<Squnion>i\<in>A. f i) = (\<Squnion>j\<in>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 "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<ge> g j"
- assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<ge> f i"
- shows "(\<Sqinter>i\<in>A. f i) = (\<Sqinter>j\<in>B. g j)"
- by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+
-
lemma less_eq_Inf_inter: "\<Sqinter>A \<squnion> \<Sqinter>B \<sqsubseteq> \<Sqinter>(A \<inter> B)"
by (auto intro: Inf_greatest Inf_lower)
@@ -498,24 +484,24 @@
lemma sup_INF:
"a \<squnion> (\<Sqinter>b\<in>B. f b) = (\<Sqinter>b\<in>B. a \<squnion> f b)"
- by (simp only: INF_def sup_Inf image_image)
+ unfolding sup_Inf by simp
lemma inf_SUP:
"a \<sqinter> (\<Squnion>b\<in>B. f b) = (\<Squnion>b\<in>B. a \<sqinter> 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 \<ge>) (op >) inf \<top> \<bottom>"
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 \<squnion> \<Sqinter>{b, c} = (\<Sqinter>d\<in>{b, c}. a \<squnion> d)" .
- then show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)" by (simp add: INF_def)
+ then show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)" by simp
qed
lemma Inf_sup:
@@ -592,7 +578,7 @@
qed
lemma uminus_INF: "- (\<Sqinter>x\<in>A. B x) = (\<Squnion>x\<in>A. - B x)"
- by (simp only: INF_def SUP_def uminus_Inf image_image)
+ by (simp add: uminus_Inf image_image)
lemma uminus_Sup:
"- (\<Squnion>A) = \<Sqinter>(uminus ` A)"
@@ -602,7 +588,7 @@
qed
lemma uminus_SUP: "- (\<Squnion>x\<in>A. B x) = (\<Sqinter>x\<in>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 \<open>Inter\<close>
@@ -1011,22 +996,18 @@
"(\<Inter>x\<in>A. B x) = {y. \<forall>x\<in>A. y \<in> B x}"
by (auto intro!: INF_eqI)
-lemma Inter_image_eq:
- "\<Inter>(B ` A) = (\<Inter>x\<in>A. B x)"
- by (fact Inf_image_eq)
-
lemma INT_iff [simp]: "b \<in> (\<Inter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. b \<in> B x)"
using Inter_iff [of _ "B ` A"] by simp
lemma INT_I [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> b \<in> B x) \<Longrightarrow> b \<in> (\<Inter>x\<in>A. B x)"
- by (auto simp add: INF_def image_def)
+ by auto
lemma INT_D [elim, Pure.elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> B a"
by auto
lemma INT_E [elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> (b \<in> B a \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R"
\<comment> \<open>"Classical" elimination -- by the Excluded Middle on @{prop "a\<in>A"}.\<close>
- by (auto simp add: INF_def image_def)
+ by auto
lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
by blast
@@ -1198,10 +1179,6 @@
"x \<in> Set.bind P f \<longleftrightarrow> x \<in> UNION P f "
by (simp add: bind_UNION)
-lemma Union_image_eq:
- "\<Union>(B ` A) = (\<Union>x\<in>A. B x)"
- by (fact Sup_image_eq)
-
lemma Union_SetCompr_eq: "\<Union>{f x| x. P x} = {a. \<exists>x. P x \<and> a \<in> f x}"
by blast
@@ -1214,10 +1191,7 @@
by auto
lemma UN_E [elim!]: "b \<in> (\<Union>x\<in>A. B x) \<Longrightarrow> (\<And>x. x\<in>A \<Longrightarrow> b \<in> B x \<Longrightarrow> R) \<Longrightarrow> R"
- by (auto simp add: SUP_def image_def)
-
-lemma image_eq_UN: "f ` A = (\<Union>x\<in>A. {f x})"
- by blast
+ by auto
lemma UN_upper: "a \<in> A \<Longrightarrow> B a \<subseteq> (\<Union>x\<in>A. B x)"
by (fact SUP_upper)
@@ -1273,7 +1247,7 @@
by blast
lemma Un_eq_UN: "A \<union> B = (\<Union>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: "(\<Union>b. A b) = (A True \<union> A False)"
by (fact SUP_UNIV_bool_expand)
@@ -1355,7 +1329,7 @@
lemma inj_on_INTER:
"I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> inj_on f (A i)) \<Longrightarrow> inj_on f (\<Inter>i \<in> I. A i)"
- unfolding inj_on_def by blast
+ unfolding inj_on_def by safe simp
lemma inj_on_UNION_chain:
assumes CH: "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> 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 \<union> (if i\<in>J then B else {}))"
-by (auto split: if_splits)
-
+ "UNION J (A(i := B)) = UNION (J - {i}) A \<union> (if i \<in> J then B else {})"
+ by (auto simp add: set_eq_iff)
+
subsubsection \<open>Complement\<close>
--- 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 \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> 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 \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> 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 \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFIMUM A f \<le> INFIMUM B g"
using cInf_mono [of "g ` B" "f ` A"] by auto
--- 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 :: _ \<Rightarrow> 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 :: _ \<Rightarrow> 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 \<Rightarrow> int"} $ t') = @{code Neg} (num_of_term vs t')
| num_of_term vs (@{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
--- 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 \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t')) =
- mk_C (HOLogic.dest_num t')
+ mk_C (HOLogic.dest_numeral t')
| num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t')) =
- mk_C (~ (HOLogic.dest_num t'))
+ mk_C (~ (HOLogic.dest_numeral t'))
| num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) =
@{code Floor} (num_of_term vs t')
| num_of_term vs (@{term "of_int :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ t')) =
@{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t')))
| num_of_term vs (@{term "numeral :: _ \<Rightarrow> real"} $ t') =
- mk_C (HOLogic.dest_num t')
+ mk_C (HOLogic.dest_numeral t')
| num_of_term vs (@{term "- numeral :: _ \<Rightarrow> 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 \<Rightarrow> real \<Rightarrow> 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 \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> 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 \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> 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 \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
@{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
| fm_of_term vs (@{term HOL.conj} $ t1 $ t2) =
--- 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' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+
from ab have stupid: "a \<noteq> 0 \<or> b \<noteq> 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 \<Longrightarrow> 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 \<noteq> 0 \<Longrightarrow> 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 \<open>a \<noteq> 0\<close> \<open>a' \<noteq> 0\<close> 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 \<Longrightarrow> isnormNum y \<Longrightarrow> 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})"
--- 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}
--- 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)
--- 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 "\<And>x. x \<in> B \<Longrightarrow> 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 \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
- by(cases a "\<Squnion>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 "\<Squnion>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 \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
by(cases a "\<Sqinter>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 \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
by(cases a "\<Squnion>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
--- 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 / \<bar>b\<bar> \<le> 0) = (a \<le> 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 *: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
shows "x \<le> y"
--- 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) \<longleftrightarrow> (\<exists>X\<subseteq>B. finite X \<and> 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 \<Rightarrow> 'a filter"
shows "(\<And>X. X \<subseteq> B \<Longrightarrow> finite X \<Longrightarrow> (INF b:X. F b) \<noteq> bot) \<Longrightarrow> (INF b:B. F b) \<noteq> 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 \<noteq> {} \<Longrightarrow> (\<And>a b. a \<in> B \<Longrightarrow> b \<in> B \<Longrightarrow> \<exists>x\<in>B. F x \<le> inf (F a) (F b)) \<Longrightarrow>
eventually P (INF b:B. F b) \<longleftrightarrow> (\<exists>b\<in>B. eventually P (F b))"
- unfolding INF_def by (subst eventually_Inf_base) auto
+ by (subst eventually_Inf_base) auto
subsubsection \<open>Map function for filters\<close>
@@ -573,7 +573,7 @@
by (auto intro: exI[of _ "\<lambda>x. x \<in> A"] exI[of _ "\<lambda>x. x \<in> B"])
lemma SUP_principal[simp]: "(SUP i : I. principal (A i)) = principal (\<Union>i\<in>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 \<Longrightarrow> (INF x:X. principal (f x)) = principal (\<Inter>x\<in>X. f x)"
by (induct X rule: finite_induct) auto
--- 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 \<open>GCD and LCM definitions\<close>
+subsection \<open>Abstract GCD and LCM\<close>
class gcd = zero + one + dvd +
fixes gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
@@ -44,8 +44,40 @@
end
class Gcd = gcd +
- fixes Gcd :: "'a set \<Rightarrow> 'a"
- and Lcm :: "'a set \<Rightarrow> 'a"
+ fixes Gcd :: "'a set \<Rightarrow> 'a" ("Gcd")
+ and Lcm :: "'a set \<Rightarrow> 'a" ("Lcm")
+begin
+
+abbreviation GCD :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+ "GCD A f \<equiv> Gcd (f ` A)"
+
+abbreviation LCM :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+ "LCM A f \<equiv> Lcm (f ` A)"
+
+end
+
+syntax
+ "_GCD1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3Gcd _./ _)" [0, 10] 10)
+ "_GCD" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3Gcd _\<in>_./ _)" [0, 0, 10] 10)
+ "_LCM1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3Lcm _./ _)" [0, 10] 10)
+ "_LCM" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3Lcm _\<in>_./ _)" [0, 0, 10] 10)
+
+translations
+ "Gcd x y. B" \<rightleftharpoons> "Gcd x. Gcd y. B"
+ "Gcd x. B" \<rightleftharpoons> "CONST GCD CONST UNIV (\<lambda>x. B)"
+ "Gcd x. B" \<rightleftharpoons> "Gcd x \<in> CONST UNIV. B"
+ "Gcd x\<in>A. B" \<rightleftharpoons> "CONST GCD A (\<lambda>x. B)"
+ "Lcm x y. B" \<rightleftharpoons> "Lcm x. Lcm y. B"
+ "Lcm x. B" \<rightleftharpoons> "CONST LCM CONST UNIV (\<lambda>x. B)"
+ "Lcm x. B" \<rightleftharpoons> "Lcm x \<in> CONST UNIV. B"
+ "Lcm x\<in>A. B" \<rightleftharpoons> "CONST LCM A (\<lambda>x. B)"
+
+print_translation \<open>
+ [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax GCD} @{syntax_const "_GCD"},
+ Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax LCM} @{syntax_const "_LCM"}]
+\<close> \<comment> \<open>to avoid eta-contraction of body\<close>
class semiring_gcd = normalization_semidom + gcd +
assumes gcd_dvd1 [iff]: "gcd a b dvd a"
@@ -67,6 +99,14 @@
"b dvd c \<Longrightarrow> gcd a b dvd c"
by (rule dvd_trans) (rule gcd_dvd2)
+lemma dvd_gcdD1:
+ "a dvd gcd b c \<Longrightarrow> a dvd b"
+ using gcd_dvd1 [of b c] by (blast intro: dvd_trans)
+
+lemma dvd_gcdD2:
+ "a dvd gcd b c \<Longrightarrow> 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 \<Longrightarrow> a dvd lcm b c"
by (rule dvd_trans) (assumption, blast)
+lemma lcm_dvdD1:
+ "lcm a b dvd c \<Longrightarrow> a dvd c"
+ using dvd_lcm1 [of a b] by (blast intro: dvd_trans)
+
+lemma lcm_dvdD2:
+ "lcm a b dvd c \<Longrightarrow> 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 \<in> A \<Longrightarrow> Gcd A dvd a"
and Gcd_greatest: "(\<And>b. b \<in> A \<Longrightarrow> a dvd b) \<Longrightarrow> a dvd Gcd A"
and normalize_Gcd [simp]: "normalize (Gcd A) = Gcd A"
+ assumes dvd_Lcm: "a \<in> A \<Longrightarrow> a dvd Lcm A"
+ and Lcm_least: "(\<And>b. b \<in> A \<Longrightarrow> b dvd a) \<Longrightarrow> Lcm A dvd a"
+ and normalize_Lcm [simp]: "normalize (Lcm A) = Lcm A"
begin
+lemma Lcm_Gcd:
+ "Lcm A = Gcd {b. \<forall>a\<in>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. \<forall>a\<in>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 \<in> insert a A"
+ then show "gcd a (Gcd A) dvd b"
+ proof
+ assume "b = a" then show ?thesis by simp
+ next
+ assume "b \<in> 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 \<in> insert a A"
+ then show "b dvd lcm a (Lcm A)"
+ proof
+ assume "b = a" then show ?thesis by simp
+ next
+ assume "b \<in> 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 \<longleftrightarrow> A \<subseteq> {0}" (is "?P \<longleftrightarrow> ?Q")
proof
@@ -384,147 +495,6 @@
then show ?P by simp
qed
-lemma unit_factor_Gcd:
- "unit_factor (Gcd A) = (if \<forall>a\<in>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 \<in> 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 \<in> insert a A"
- then show "gcd a (Gcd A) dvd b"
- proof
- assume "b = a" then show ?thesis by simp
- next
- assume "b \<in> 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: \<comment> \<open>FIXME remove\<close>
- "\<forall>b\<in>A. a dvd b \<Longrightarrow> 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 \<in> 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. \<forall>a\<in>A. a dvd b}"
-begin
-
-lemma dvd_Lcm:
- assumes "a \<in> A"
- shows "a dvd Lcm A"
- using assms by (auto intro: Gcd_greatest simp add: Lcm_Gcd)
-
-lemma Lcm_least:
- assumes "\<And>b. b \<in> A \<Longrightarrow> 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. \<forall>a\<in>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 \<in> insert a A"
- then show "b dvd lcm a (Lcm A)"
- proof
- assume "b = a" then show ?thesis by simp
- next
- assume "b \<in> 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 \<longleftrightarrow> (\<forall>a\<in>A. is_unit a)" (is "?P \<longleftrightarrow> ?Q")
proof
@@ -548,6 +518,44 @@
by simp
qed
+lemma unit_factor_Gcd:
+ "unit_factor (Gcd A) = (if \<forall>a\<in>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 \<in> 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 \<in> 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 \<in> 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 "\<And>b. b \<in> A \<Longrightarrow> a dvd b"
+ and "\<And>c. (\<And>b. b \<in> A \<Longrightarrow> c dvd b) \<Longrightarrow> 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 "\<And>b. b \<in> A \<Longrightarrow> b dvd a"
+ and "\<And>c. (\<And>b. b \<in> A \<Longrightarrow> b dvd c) \<Longrightarrow> 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 \<open>GCD and LCM on @{typ nat} and @{typ int}\<close>
instantiation nat :: gcd
begin
-fun
- gcd_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
- "gcd_nat x y =
- (if y = 0 then x else gcd y (x mod y))"
+fun gcd_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where "gcd_nat x y =
+ (if y = 0 then x else gcd y (x mod y))"
-definition
- lcm_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+definition lcm_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
where
"lcm_nat x y = x * y div (gcd x y)"
@@ -602,22 +646,17 @@
instantiation int :: gcd
begin
-definition
- gcd_int :: "int \<Rightarrow> int \<Rightarrow> int"
-where
- "gcd_int x y = int (gcd (nat \<bar>x\<bar>) (nat \<bar>y\<bar>))"
+definition gcd_int :: "int \<Rightarrow> int \<Rightarrow> int"
+ where "gcd_int x y = int (gcd (nat \<bar>x\<bar>) (nat \<bar>y\<bar>))"
-definition
- lcm_int :: "int \<Rightarrow> int \<Rightarrow> int"
-where
- "lcm_int x y = int (lcm (nat \<bar>x\<bar>) (nat \<bar>y\<bar>))"
+definition lcm_int :: "int \<Rightarrow> int \<Rightarrow> int"
+ where "lcm_int x y = int (lcm (nat \<bar>x\<bar>) (nat \<bar>y\<bar>))"
instance ..
end
-
-subsection \<open>Transfer setup\<close>
+text \<open>Transfer setup\<close>
lemma transfer_nat_int_gcd:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> 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 \<open>GCD properties\<close>
-
-(* was gcd_induct *)
lemma gcd_nat_induct:
fixes m n :: nat
assumes "\<And>m. P m 0"
@@ -662,30 +697,30 @@
(* specific to int *)
+lemma gcd_eq_int_iff:
+ "gcd k l = int n \<longleftrightarrow> gcd (nat \<bar>k\<bar>) (nat \<bar>l\<bar>) = n"
+ by (simp add: gcd_int_def)
+
+lemma lcm_eq_int_iff:
+ "lcm k l = int n \<longleftrightarrow> lcm (nat \<bar>k\<bar>) (nat \<bar>l\<bar>) = 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]: "\<bar>gcd (x::int) y\<bar> = gcd x y"
+lemma abs_gcd_int [simp]: "\<bar>gcd (x::int) y\<bar> = gcd x y"
by(simp add: gcd_int_def)
lemma gcd_abs_int: "gcd (x::int) y = gcd \<bar>x\<bar> \<bar>y\<bar>"
by (simp add: gcd_int_def)
-lemma gcd_abs1_int[simp]: "gcd \<bar>x\<bar> (y::int) = gcd x y"
+lemma gcd_abs1_int [simp]: "gcd \<bar>x\<bar> (y::int) = gcd x y"
by (metis abs_idempotent gcd_abs_int)
-lemma gcd_abs2_int[simp]: "gcd x \<bar>y::int\<bar> = gcd x y"
+lemma gcd_abs2_int [simp]: "gcd x \<bar>y::int\<bar> = gcd x y"
by (metis abs_idempotent gcd_abs_int)
lemma gcd_cases_int:
@@ -712,10 +747,10 @@
lemma abs_lcm_int [simp]: "\<bar>lcm i j::int\<bar> = lcm i j"
by (simp add:lcm_int_def)
-lemma lcm_abs1_int[simp]: "lcm \<bar>x\<bar> (y::int) = lcm x y"
+lemma lcm_abs1_int [simp]: "lcm \<bar>x\<bar> (y::int) = lcm x y"
by (metis abs_idempotent lcm_int_def)
-lemma lcm_abs2_int[simp]: "lcm x \<bar>y::int\<bar> = lcm x y"
+lemma lcm_abs2_int [simp]: "lcm x \<bar>y::int\<bar> = 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 = \<bar>x\<bar>"
by (unfold gcd_int_def, auto)
@@ -749,7 +782,7 @@
(* weaker, but useful for the simplifier *)
-lemma gcd_non_0_nat: "y ~= (0::nat) \<Longrightarrow> gcd (x::nat) y = gcd y (x mod y)"
+lemma gcd_non_0_nat: "y \<noteq> (0::nat) \<Longrightarrow> 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 \<Longrightarrow> (k::nat) dvd m"
- by (metis gcd_dvd1 dvd_trans)
-
-lemma dvd_gcd_D2_nat: "k dvd gcd m n \<Longrightarrow> (k::nat) dvd n"
- by (metis gcd_dvd2 dvd_trans)
-
-lemma dvd_gcd_D1_int: "i dvd gcd m n \<Longrightarrow> (i::int) dvd m"
- by (metis gcd_dvd1 dvd_trans)
-
-lemma dvd_gcd_D2_int: "i dvd gcd m n \<Longrightarrow> (i::int) dvd n"
- by (metis gcd_dvd2 dvd_trans)
-
lemma gcd_le1_nat [simp]: "a \<noteq> 0 \<Longrightarrow> gcd (a::nat) b \<le> a"
by (rule dvd_imp_le, auto)
@@ -822,27 +843,11 @@
lemma gcd_le2_int [simp]: "b > 0 \<Longrightarrow> gcd (a::int) b \<le> 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 \<and> d dvd b \<and>
(\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
@@ -862,31 +867,14 @@
done
interpretation gcd_nat:
- semilattice_neutr_order gcd "0::nat" Rings.dvd "(\<lambda>m n. m dvd n \<and> \<not> 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 \<Longrightarrow> gcd x y = x"
- by (fact gcd_nat.absorb1)
-
-lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \<Longrightarrow> gcd x y = y"
- by (fact gcd_nat.absorb2)
+ semilattice_neutr_order gcd "0::nat" Rings.dvd "\<lambda>m n. m dvd n \<and> m \<noteq> 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 \<Longrightarrow> gcd (x::int) y = \<bar>x\<bar>"
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 \<Longrightarrow> gcd (x::int) y = \<bar>y\<bar>"
- by (metis gcd_proj1_if_dvd_int gcd_commute_int)
+ by (metis gcd_proj1_if_dvd_int gcd.commute)
text \<open>
\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 \<Longrightarrow>
- (k dvd m * n) = (k dvd m)"
- by (auto intro: coprime_dvd_mult_nat)
-
-lemma coprime_dvd_mult_iff_int: "coprime (k::int) n \<Longrightarrow>
- (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 \<longleftrightarrow> a dvd b"
+ using assms by (auto intro: coprime_dvd_mult)
lemma gcd_mult_cancel:
"coprime c b \<Longrightarrow> 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 \<longleftrightarrow> a = b \<and> c = d" (is "?lhs \<longleftrightarrow> ?rhs")
+ shows "normalize a * normalize c = normalize b * normalize d
+ \<longleftrightarrow> normalize a = normalize b \<and> normalize c = normalize d" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?rhs then show ?lhs by simp
next
assume ?lhs
- from \<open>?lhs\<close> have "a dvd b * d" by (auto intro: dvdI dest: sym)
- with \<open>coprime a d\<close> have "a dvd b" by (simp add: coprime_dvd_mult_iff_nat)
- from \<open>?lhs\<close> have "b dvd a * c" by (auto intro: dvdI dest: sym)
- with \<open>coprime b c\<close> have "b dvd a" by (simp add: coprime_dvd_mult_iff_nat)
- from \<open>?lhs\<close> have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult.commute)
- with \<open>coprime b c\<close> have "c dvd d" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
- from \<open>?lhs\<close> have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult.commute)
- with \<open>coprime a d\<close> have "d dvd c" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
- from \<open>a dvd b\<close> \<open>b dvd a\<close> have "a = b" by (rule Nat.dvd.antisym)
- moreover from \<open>c dvd d\<close> \<open>d dvd c\<close> have "c = d" by (rule Nat.dvd.antisym)
+ from \<open>?lhs\<close> have "normalize a dvd normalize b * normalize d"
+ by (auto intro: dvdI dest: sym)
+ with \<open>coprime a d\<close> have "a dvd b"
+ by (simp add: coprime_dvd_mult_iff normalize_mult [symmetric])
+ from \<open>?lhs\<close> have "normalize b dvd normalize a * normalize c"
+ by (auto intro: dvdI dest: sym)
+ with \<open>coprime b c\<close> have "b dvd a"
+ by (simp add: coprime_dvd_mult_iff normalize_mult [symmetric])
+ from \<open>?lhs\<close> have "normalize c dvd normalize d * normalize b"
+ by (auto intro: dvdI dest: sym simp add: mult.commute)
+ with \<open>coprime b c\<close> have "c dvd d"
+ by (simp add: coprime_dvd_mult_iff gcd.commute normalize_mult [symmetric])
+ from \<open>?lhs\<close> have "normalize d dvd normalize c * normalize a"
+ by (auto intro: dvdI dest: sym simp add: mult.commute)
+ with \<open>coprime a d\<close> have "d dvd c"
+ by (simp add: coprime_dvd_mult_iff gcd.commute normalize_mult [symmetric])
+ from \<open>a dvd b\<close> \<open>b dvd a\<close> have "normalize a = normalize b"
+ by (rule associatedI)
+ moreover from \<open>c dvd d\<close> \<open>d dvd c\<close> 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 \<longleftrightarrow> a = b \<and> 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 "\<bar>a\<bar> * \<bar>c\<bar> = \<bar>b\<bar> * \<bar>d\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>b\<bar> \<and> \<bar>c\<bar> = \<bar>d\<bar>"
- using assms by (intro coprime_crossproduct_nat [transferred]) auto
+ using assms coprime_crossproduct [of a d b c] by simp
text \<open>\medskip Addition laws\<close>
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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> 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]: -- \<open>FIXME move\<close>
+ 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} \<subseteq> {d. d \<le> 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. \<bar>d\<bar> <= \<bar>i\<bar>} = {- \<bar>i\<bar> .. \<bar>i\<bar>}" by(auto simp:abs_if)
- hence "finite {d. \<bar>d\<bar> <= \<bar>i\<bar>}" by simp
- from finite_subset[OF _ this] show ?thesis using assms
+lemma finite_divisors_int [simp]:
+ fixes i :: int
+ assumes "i \<noteq> 0"
+ shows "finite {d. d dvd i}"
+proof -
+ have "{d. \<bar>d\<bar> \<le> \<bar>i\<bar>} = {- \<bar>i\<bar>..\<bar>i\<bar>}"
+ by (auto simp: abs_if)
+ then have "finite {d. \<bar>d\<bar> <= \<bar>i\<bar>}"
+ 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\<noteq>0 \<Longrightarrow> Max{d::nat. d dvd n} = n"
+lemma Max_divisors_self_nat [simp]: "n\<noteq>0 \<Longrightarrow> 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\<noteq>0 \<Longrightarrow> Max{d::int. d dvd n} = \<bar>n\<bar>"
+lemma Max_divisors_self_int [simp]: "n\<noteq>0 \<Longrightarrow> Max{d::int. d dvd n} = \<bar>n\<bar>"
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 \<Longrightarrow> n ~= 0 \<Longrightarrow> gcd (m::nat) n = (Max {d. d dvd m & d dvd n})"
+ "m > 0 \<Longrightarrow> n > 0 \<Longrightarrow> gcd (m::nat) n = Max {d. d dvd m \<and> 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 \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> is_unit d)" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?P then show ?Q by auto
+next
+ assume ?Q
+ then have "is_unit (gcd a b) \<longleftrightarrow> gcd a b dvd a \<and> 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 \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
- using gcd_unique_nat[of 1 a b, simplified] by auto
+lemma coprime_nat:
+ "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
+ using coprime [of a b] by simp
lemma coprime_Suc_0_nat:
- "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = Suc 0)"
+ "coprime (a::nat) b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = Suc 0)"
using coprime_nat by simp
-lemma coprime_int: "coprime (a::int) b \<longleftrightarrow>
- (\<forall>d. d >= 0 \<and> d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
+lemma coprime_int:
+ "coprime (a::int) b \<longleftrightarrow> (\<forall>d. d \<ge> 0 \<and> d dvd a \<and> d dvd b \<longleftrightarrow> 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 \<noteq> 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) \<longleftrightarrow>
coprime d a \<and> 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) \<longleftrightarrow>
coprime d a \<and> 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 "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> 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 "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> coprime (a^n) (b^m)"
- by (fact coprime_exp2)
-
-lemma coprime_exp2_int [intro]: "coprime (a::int) b \<Longrightarrow> 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 \<and> 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 \<Longrightarrow> ((a::nat)^n dvd b^n) = (a dvd b)"
+lemma pow_divides_eq_nat [simp]:
+ "n > 0 \<Longrightarrow> (a::nat) ^ n dvd b ^ n \<longleftrightarrow> a dvd b"
by (auto intro: pow_divides_pow_nat dvd_power_same)
-lemma pow_divides_eq_int [simp]: "n ~= 0 \<Longrightarrow> ((a::int)^n dvd b^n) = (a dvd b)"
+lemma pow_divides_eq_int [simp]:
+ "n ~= 0 \<Longrightarrow> (a::int) ^ n dvd b ^ n \<longleftrightarrow> 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 \<open>b dvd c\<close> obtain b' where"c = b * b'" ..
+ with \<open>a dvd c\<close> have "a dvd b' * b"
+ by (simp add: ac_simps)
+ with \<open>coprime a b\<close> have "a dvd b'"
+ by (simp add: coprime_dvd_mult_iff)
+ then obtain a' where "b' = a * a'" ..
+ with \<open>c = b * b'\<close> 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) \<noteq> 0 \<Longrightarrow> 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 (\<Prod>i\<in>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 "(\<And>i. i \<in> A \<Longrightarrow> coprime (f i) x) \<Longrightarrow> coprime (\<Prod>i\<in>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 (\<Prod>i\<in>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 "(\<And>i. i \<in> A \<Longrightarrow> coprime (f i) x) \<Longrightarrow> coprime (\<Prod>i\<in>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 \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> 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 \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> \<bar>x\<bar> = 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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<open>Bezout's theorem\<close>
@@ -1761,11 +1756,10 @@
qed
-subsection \<open>LCM properties\<close>
+subsection \<open>LCM properties on @{typ nat} and @{typ int}\<close>
lemma lcm_altdef_int [code]: "lcm (a::int) b = \<bar>a\<bar> * \<bar>b\<bar> 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: "\<bar>m::int\<bar> * \<bar>n\<bar> = 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 \<Longrightarrow> n>0 \<Longrightarrow> 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: -- \<open>FIXME move\<close>
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 \<Longrightarrow> n dvd k \<Longrightarrow> 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 \<Longrightarrow> k dvd lcm m n"
-by(metis lcm_dvd1_nat dvd_trans)
-
-lemma dvd_lcm_I2_nat[simp]: "(k::nat) dvd n \<Longrightarrow> k dvd lcm m n"
-by(metis lcm_dvd2_nat dvd_trans)
-
-lemma dvd_lcm_I1_int[simp]: "(i::int) dvd m \<Longrightarrow> i dvd lcm m n"
-by(metis lcm_dvd1_int dvd_trans)
-
-lemma dvd_lcm_I2_int[simp]: "(i::int) dvd n \<Longrightarrow> 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 \<and> b dvd d \<and>
(\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> 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 \<and> (a::int) dvd d \<and> b dvd d \<and>
(\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
- using lcm_least_int zdvd_antisym_nonneg by auto
-
-interpretation lcm_nat: abel_semigroup "lcm :: nat \<Rightarrow> nat \<Rightarrow> nat"
- + lcm_nat: semilattice_neutr "lcm :: nat \<Rightarrow> nat \<Rightarrow> nat" 1
- by standard (simp_all del: One_nat_def)
-
-interpretation lcm_int: abel_semigroup "lcm :: int \<Rightarrow> int \<Rightarrow> 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 \<Longrightarrow> lcm x y = y"
apply (rule sym)
@@ -1878,21 +1811,21 @@
done
lemma lcm_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \<Longrightarrow> 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 \<Longrightarrow> lcm y x = \<bar>y\<bar>"
-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) \<longleftrightarrow> n dvd m"
+lemma lcm_proj1_iff_nat [simp]: "lcm m n = (m::nat) \<longleftrightarrow> n dvd m"
by (metis lcm_proj1_if_dvd_nat lcm_unique_nat)
-lemma lcm_proj2_iff_nat[simp]: "lcm m n = (n::nat) \<longleftrightarrow> m dvd n"
+lemma lcm_proj2_iff_nat [simp]: "lcm m n = (n::nat) \<longleftrightarrow> m dvd n"
by (metis lcm_proj2_if_dvd_nat lcm_unique_nat)
-lemma lcm_proj1_iff_int[simp]: "lcm m n = \<bar>m::int\<bar> \<longleftrightarrow> n dvd m"
+lemma lcm_proj1_iff_int [simp]: "lcm m n = \<bar>m::int\<bar> \<longleftrightarrow> 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 = \<bar>n::int\<bar> \<longleftrightarrow> m dvd n"
+lemma lcm_proj2_iff_int [simp]: "lcm m n = \<bar>n::int\<bar> \<longleftrightarrow> 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\<Rightarrow>nat\<Rightarrow>nat)"
- by (fact comp_fun_idem_gcd)
-
-lemma comp_fun_idem_gcd_int: "comp_fun_idem (gcd :: int\<Rightarrow>int\<Rightarrow>int)"
- by (fact comp_fun_idem_gcd)
-
-lemma comp_fun_idem_lcm_nat: "comp_fun_idem (lcm :: nat\<Rightarrow>nat\<Rightarrow>nat)"
- by (fact comp_fun_idem_lcm)
-
-lemma comp_fun_idem_lcm_int: "comp_fun_idem (lcm :: int\<Rightarrow>int\<Rightarrow>int)"
- by (fact comp_fun_idem_lcm)
-
-lemma lcm_0_iff_nat [simp]: "lcm (m::nat) n = 0 \<longleftrightarrow> m=0 \<or> n=0"
- by (fact lcm_eq_0_iff)
-
-lemma lcm_0_iff_int [simp]: "lcm (m::int) n = 0 \<longleftrightarrow> m=0 \<or> n=0"
- by (fact lcm_eq_0_iff)
-
-lemma lcm_1_iff_nat [simp]: "lcm (m::nat) n = 1 \<longleftrightarrow> m=1 \<and> n=1"
- by (simp only: lcm_eq_1_iff) simp
+lemma lcm_1_iff_nat [simp]:
+ "lcm (m::nat) n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = Suc 0"
+ using lcm_eq_1_iff [of m n] by simp
-lemma lcm_1_iff_int [simp]: "lcm (m::int) n = 1 \<longleftrightarrow> (m=1 \<or> m = -1) \<and> (n=1 \<or> n = -1)"
+lemma lcm_1_iff_int [simp]:
+ "lcm (m::int) n = 1 \<longleftrightarrow> (m=1 \<or> m = -1) \<and> (n=1 \<or> n = -1)"
by auto
-subsection \<open>The complete divisibility lattice\<close>
-
-interpretation gcd_semilattice_nat: semilattice_inf gcd Rings.dvd "(\<lambda>m n::nat. m dvd n \<and> \<not> n dvd m)"
- by standard simp_all
-
-interpretation lcm_semilattice_nat: semilattice_sup lcm Rings.dvd "(\<lambda>m n::nat. m dvd n \<and> \<not> n dvd m)"
- by standard simp_all
-
-interpretation gcd_lcm_lattice_nat: lattice gcd Rings.dvd "(\<lambda>m n::nat. m dvd n & ~ n dvd m)" lcm ..
+subsection \<open>The complete divisibility lattice on @{typ nat} and @{typ int}\<close>
text\<open>Lifting gcd and lcm to sets (Gcd/Lcm).
Gcd is defined via Lcm to facilitate the proof that we have a complete lattice.
\<close>
-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 "\<forall>m\<in>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 \<subseteq> {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. \<forall>m\<in>M. d dvd m}"
-instance ..
-
-end
-
-instance nat :: semiring_Gcd
-proof
+instance proof
show "Gcd N dvd n" if "n \<in> 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 "\<And>m. m \<in> N \<Longrightarrow> 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. \<forall>n\<in>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 "\<lambda>m n. m dvd n \<and> \<not> 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 (\<lambda>m n. m dvd n \<and> \<not> 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 "\<lambda>m n. m dvd n \<and> \<not> 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 \<in> N" for N and n ::nat
+ using that by (induct N rule: infinite_finite_induct)
+ auto
+ show "Lcm N dvd n" if "\<And>m. m \<in> N \<Longrightarrow> 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) \<Longrightarrow> 0 \<in> M \<Longrightarrow> Lcm M = 0"
- by (rule Lcm_eq_0_I)
-
-lemma Lcm0_iff [simp]:
- fixes M :: "nat set"
- assumes "finite M" and "M \<noteq> {}"
- shows "Lcm M = 0 \<longleftrightarrow> 0 \<in> M"
- using assms by (simp add: Lcm_0_iff)
+lemma Gcd_nat_eq_one:
+ "1 \<in> N \<Longrightarrow> Gcd N = (1::nat)"
+ by (rule Gcd_eq_1_I) auto
text\<open>Alternative characterizations of Gcd:\<close>
-lemma Gcd_eq_Max: "finite(M::nat set) \<Longrightarrow> M \<noteq> {} \<Longrightarrow> 0 \<notin> M \<Longrightarrow> Gcd M = Max(\<Inter>m\<in>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 \<noteq> {}" and "0 \<notin> M"
+ shows "Gcd M = Max (\<Inter>m\<in>M. {d. d dvd m})"
+proof (rule antisym)
+ from assms obtain m where "m \<in> M" and "m > 0"
+ by auto
+ from \<open>m > 0\<close> have "finite {d. d dvd m}"
+ by (blast intro: finite_divisors_nat)
+ with \<open>m \<in> M\<close> have fin: "finite (\<Inter>m\<in>M. {d. d dvd m})"
+ by blast
+ from fin show "Gcd M \<le> Max (\<Inter>m\<in>M. {d. d dvd m})"
+ by (auto intro: Max_ge Gcd_dvd)
+ from fin show "Max (\<Inter>m\<in>M. {d. d dvd m}) \<le> Gcd M"
+ apply (rule Max.boundedI)
+ apply auto
+ apply (meson Gcd_dvd Gcd_greatest \<open>0 < m\<close> \<open>m \<in> M\<close> dvd_imp_le dvd_pos_nat)
+ done
+qed
lemma Gcd_remove0_nat: "finite M \<Longrightarrow> 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 \<Longrightarrow> inj_on g B \<Longrightarrow> ALL a:A. ALL b:B. coprime (f a) (g b)
\<Longrightarrow> inj_on (%(a,b). f a * g b::nat) (A \<times> B)"
@@ -2115,116 +1992,72 @@
subsubsection \<open>Setwise gcd and lcm for integers\<close>
-instantiation int :: Gcd
+instantiation int :: semiring_Gcd
begin
definition
- "Lcm M = int (Lcm (nat ` abs ` M))"
+ "Lcm M = int (Lcm m\<in>M. (nat \<circ> abs) m)"
definition
- "Gcd M = int (Gcd (nat ` abs ` M))"
+ "Gcd M = int (Gcd m\<in>M. (nat \<circ> 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 "\<bar>Gcd K\<bar> = Gcd K"
+ using normalize_Gcd [of K] by simp
+
+lemma abs_Lcm [simp]:
+ fixes K :: "int set"
+ shows "\<bar>Lcm K\<bar> = Lcm K"
+ using normalize_Lcm [of K] by simp
+lemma Gcm_eq_int_iff:
+ "Gcd K = int n \<longleftrightarrow> Gcd ((nat \<circ> abs) ` K) = n"
+ by (simp add: Gcd_int_def comp_def image_image)
+
+lemma Lcm_eq_int_iff:
+ "Lcm K = int n \<longleftrightarrow> Lcm ((nat \<circ> abs) ` K) = n"
+ by (simp add: Lcm_int_def comp_def image_image)
+
+
+subsection \<open>GCD and LCM on @{typ integer}\<close>
+
+instantiation integer :: gcd
+begin
+
+context
+ includes integer.lifting
+begin
+
+lift_definition gcd_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+ is gcd .
+lift_definition lcm_integer :: "integer \<Rightarrow> integer \<Rightarrow> 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. \<forall>k\<in>K. nat \<bar>k\<bar> dvd n} = ((\<lambda>k. nat \<bar>k\<bar>) ` {l. \<forall>k\<in>K. k dvd l})"
- proof (rule set_eqI)
- fix n
- have "(\<forall>k\<in>K. nat \<bar>k\<bar> dvd n) \<longleftrightarrow> (\<exists>l. (\<forall>k\<in>K. k dvd l) \<and> n = nat \<bar>l\<bar>)" (is "?P \<longleftrightarrow> ?Q")
- proof
- assume ?P
- then have "(\<forall>k\<in>K. k dvd int n) \<and> n = nat \<bar>int n\<bar>"
- 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 \<in> {n. \<forall>k\<in>K. nat \<bar>k\<bar> dvd n} \<longleftrightarrow> n \<in> (\<lambda>k. nat \<bar>k\<bar>) ` {l. \<forall>k\<in>K. k dvd l}"
- by auto
- qed
- then show "Lcm K = Gcd {l. \<forall>k\<in>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 \<longleftrightarrow> nat \<bar>x\<bar> dvd nat \<bar>y\<bar>"
- by (fact dvd_int_unfold_dvd_nat)
-
-lemma dvd_Lcm_int [simp]:
- fixes M :: "int set" assumes "m \<in> M" shows "m dvd Lcm M"
- using assms by (fact dvd_Lcm)
-
-lemma Lcm_dvd_int [simp]:
- fixes M :: "int set"
- assumes "\<forall>m\<in>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 \<open>Fact aliasses\<close>
-
-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 \<open>gcd and lcm instances for @{typ integer}\<close>
-
-instantiation integer :: gcd begin
-context includes integer.lifting begin
-lift_definition gcd_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer" is gcd .
-lift_definition lcm_integer :: "integer \<Rightarrow> integer \<Rightarrow> 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 = \<bar>if l = (0::integer) then k else gcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
-by transfer(fact gcd_code_int)
+ by transfer (fact gcd_code_int)
lemma lcm_code_integer [code]: "lcm (a::integer) b = \<bar>a\<bar> * \<bar>b\<bar> div gcd a b"
-by transfer(fact lcm_altdef_int)
+ by transfer (fact lcm_altdef_int)
end
@@ -2234,4 +2067,75 @@
and (Scala) "_.gcd'((_)')"
\<comment> \<open>There is no gcd operation in the SML standard library, so no code setup for SML\<close>
+text \<open>Some code equations\<close>
+
+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 \<open>Fact aliasses\<close>
+
+lemma lcm_0_iff_nat [simp]: "lcm (m::nat) n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"
+ by (fact lcm_eq_0_iff)
+
+lemma lcm_0_iff_int [simp]: "lcm (m::int) n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"
+ by (fact lcm_eq_0_iff)
+
+lemma dvd_lcm_I1_nat [simp]: "(k::nat) dvd m \<Longrightarrow> k dvd lcm m n"
+ by (fact dvd_lcmI1)
+
+lemma dvd_lcm_I2_nat [simp]: "(k::nat) dvd n \<Longrightarrow> k dvd lcm m n"
+ by (fact dvd_lcmI2)
+
+lemma dvd_lcm_I1_int [simp]: "(i::int) dvd m \<Longrightarrow> i dvd lcm m n"
+ by (fact dvd_lcmI1)
+
+lemma dvd_lcm_I2_int [simp]: "(i::int) dvd n \<Longrightarrow> i dvd lcm m n"
+ by (fact dvd_lcmI2)
+
+lemma coprime_exp2_nat [intro]: "coprime (a::nat) b \<Longrightarrow> coprime (a^n) (b^m)"
+ by (fact coprime_exp2)
+
+lemma coprime_exp2_int [intro]: "coprime (a::int) b \<Longrightarrow> 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 \<in> 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 \<Longrightarrow> gcd x y = x"
+ by (fact gcd_nat.absorb1)
+
+lemma gcd_proj2_if_dvd_nat [simp]: "(y::nat) dvd x \<Longrightarrow> 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
--- 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 \<longleftrightarrow> a > b"
-by (simp add: less_diff_eq)
+lemma diff_gt_0_iff_gt [simp]:
+ "a - b > 0 \<longleftrightarrow> a > b"
+ by (simp add: less_diff_eq)
-lemma diff_le_eq[algebra_simps, field_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
-by (auto simp add: le_less diff_less_eq )
+lemma diff_le_eq [algebra_simps, field_simps]:
+ "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
+ by (auto simp add: le_less diff_less_eq )
-lemma le_diff_eq[algebra_simps, field_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
-by (auto simp add: le_less less_diff_eq)
+lemma le_diff_eq [algebra_simps, field_simps]:
+ "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
+ by (auto simp add: le_less less_diff_eq)
lemma diff_le_0_iff_le [simp]:
"a - b \<le> 0 \<longleftrightarrow> a \<le> 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 \<ge> 0 \<longleftrightarrow> a \<ge> b"
+ by (simp add: le_diff_eq)
+
lemma diff_eq_diff_less:
"a - b = c - d \<Longrightarrow> a < b \<longleftrightarrow> 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 \<comment> \<open>FIXME duplicate\<close>
lemmas diff_diff_eq = diff_diff_add \<comment> \<open>FIXME duplicate\<close>
-
subsection \<open>Tools setup\<close>
lemma add_mono_thms_linordered_semiring:
--- 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)
--- 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]:
--- 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 "\<forall>i<length clist. clist!i\<in>assum(Pre(xs!i), Rely(xs!i))")
- apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> \<Turnstile> (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="\<lambda>i. H i \<longrightarrow> \<Turnstile> (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="\<lambda>j. H j \<longrightarrow> (I j) \<in> 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="\<lambda>j. H j \<longrightarrow> (I j) \<in> 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="\<lambda>j. H j \<longrightarrow> 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="\<lambda>j. H j \<longrightarrow> length(J j)=(K j)" for H J K in allE)
+ apply(erule_tac x=xa and P="\<lambda>j. H j \<longrightarrow> length(J j)=(K j)" for H J K in allE)
apply(subgoal_tac "length x - 1 < length x",simp)
apply(case_tac "x\<noteq>[]")
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\<noteq>[]")
+ apply(subgoal_tac "clist!xa\<noteq>[]")
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="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. T j \<longrightarrow> (\<forall>i. H j i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for T H d e in allE,simp)
apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (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)
--- 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) {} \<subseteq> (f ^^ Suc n) {}" for n
+ using funpow_decreasing [of n "Suc n"] by auto
show "lfp f \<subseteq> ?U"
proof (rule lfp_lowerbound)
have "f ?U = (UN n. (f^^Suc n){})"
@@ -105,7 +108,7 @@
by(simp add: cont_def)
also have "\<dots> = (f^^0){} \<union> \<dots>" by simp
also have "\<dots> = ?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 \<subseteq> ?U" by simp
qed
next
--- 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) \<le> n*(n + 1)"
using assms by (meson diff_le_self order_trans le_add1 mult_le_mono)
hence "int(\<Sum> {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<n" "j<n"
hence "a i - a j \<le> 0 \<and> b i - b j \<ge> 0 \<or> a i - a j \<ge> 0 \<and> b i - b j \<le> 0"
using assms by (cases "i \<le> j") (auto simp: algebra_simps)
- } hence "?S \<le> 0"
+ } then have "?S \<le> 0"
by (auto intro!: setsum_nonpos simp: mult_le_0_iff)
finally show ?thesis by (simp add: algebra_simps)
qed
--- 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 "\<bar>z\<bar> < 1 \<longleftrightarrow> z = 0" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?Q then show ?P by simp
+next
+ assume ?P
+ with zless_imp_add1_zle [of "\<bar>z\<bar>" 1] have "\<bar>z\<bar> + 1 \<le> 1"
+ by simp
+ then have "\<bar>z\<bar> \<le> 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 \<ge> 0 \<Longrightarrow> of_int z \<ge> 0"
by simp
+lemma of_int_abs [simp]:
+ "of_int \<bar>x\<bar> = \<bar>of_int x\<bar>"
+ by (auto simp add: abs_if)
+
+lemma of_int_lessD:
+ assumes "\<bar>of_int n\<bar> < x"
+ shows "n = 0 \<or> x > 1"
+proof (cases "n = 0")
+ case True then show ?thesis by simp
+next
+ case False
+ then have "\<bar>n\<bar> \<noteq> 0" by simp
+ then have "\<bar>n\<bar> > 0" by simp
+ then have "\<bar>n\<bar> \<ge> 1"
+ using zless_imp_add1_zle [of 0 "\<bar>n\<bar>"] by simp
+ then have "\<bar>of_int n\<bar> \<ge> 1"
+ unfolding of_int_1_le_iff [of "\<bar>n\<bar>", symmetric] by simp
+ then have "1 < x" using assms by (rule le_less_trans)
+ then show ?thesis ..
+qed
+
+lemma of_int_leD:
+ assumes "\<bar>of_int n\<bar> \<le> x"
+ shows "n = 0 \<or> 1 \<le> x"
+proof (cases "n = 0")
+ case True then show ?thesis by simp
+next
+ case False
+ then have "\<bar>n\<bar> \<noteq> 0" by simp
+ then have "\<bar>n\<bar> > 0" by simp
+ then have "\<bar>n\<bar> \<ge> 1"
+ using zless_imp_add1_zle [of 0 "\<bar>n\<bar>"] by simp
+ then have "\<bar>of_int n\<bar> \<ge> 1"
+ unfolding of_int_1_le_iff [of "\<bar>n\<bar>", symmetric] by simp
+ then have "1 \<le> x" using assms by (rule order_trans)
+ then show ?thesis ..
+qed
+
+
end
text \<open>Comparisons involving @{term of_int}.\<close>
@@ -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 \<le> z \<longleftrightarrow> (\<exists>n. z = w + int n)"
-proof -
- have "(w \<le> z) = (0 \<le> z - w)"
- by (simp only: le_diff_eq add_0_left)
- also have "\<dots> = (\<exists>n. z - w = of_nat n)"
- by (auto elim: zero_le_imp_eq_int)
- also have "\<dots> = (\<exists>n. z = w + of_nat n)"
- by (simp only: algebra_simps)
- finally show ?thesis .
+lemma zle_iff_zadd:
+ "w \<le> z \<longleftrightarrow> (\<exists>n. z = w + int n)" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?Q
+ then show ?P by auto
+next
+ assume ?P
+ then have "0 \<le> 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\<open>Products and 1, by T. M. Rasmussen\<close>
-lemma zabs_less_one_iff [simp]: "(\<bar>z\<bar> < 1) = (z = (0::int))"
-by arith
-
lemma abs_zmult_eq_1:
assumes mn: "\<bar>m * n\<bar> = 1"
shows "\<bar>m\<bar> = (1::int)"
@@ -1674,34 +1729,9 @@
hide_const (open) Pos Neg sub dup
-subsection \<open>Legacy theorems\<close>
-
-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 \<open>De-register \<open>int\<close> as a quotient type:\<close>
lifting_update int.lifting
lifting_forget int.lifting
-text\<open>Also the class for fields with characteristic zero.\<close>
-class field_char_0 = field + ring_char_0
-subclass (in linordered_field) field_char_0 ..
-
end
--- 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)) \<dots> = gcd (fib (n + 2)) (fib (n + 1))"
by (rule gcd_add2_nat)
also have "\<dots> = gcd (fib (n + 1)) (fib (n + 1 + 1))"
- by (simp add: gcd_commute_nat)
+ by (simp add: gcd.commute)
also assume "\<dots> = 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 \<dots> (fib (k + 1)) = gcd (fib k * fib n) (fib (k + 1))"
by (simp add: gcd_mult_add)
also have "\<dots> = 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 "\<dots> = 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 "\<dots> = gcd (fib n) (fib m)"
by (rule gcd_fib_mod)
also have "\<dots> = 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
--- 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) (\<lambda>A. UNION A id) (\<lambda>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]
--- 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) = - \<infinity>"
moreover then have "\<And>i. i \<in> I \<Longrightarrow> f i = -\<infinity>"
- unfolding Sup_eq_MInfty Sup_image_eq[symmetric] by auto
+ unfolding Sup_eq_MInfty by auto
ultimately show ?thesis
by (cases c) (auto simp: \<open>I \<noteq> {}\<close>)
next
assume "(SUP i:I. f i) \<noteq> - \<infinity>" then show ?thesis
- unfolding Sup_image_eq[symmetric]
by (subst continuous_at_Sup_mono[where f="\<lambda>x. x + c"])
(auto simp: continuous_at_imp_continuous_at_within continuous_at mono_def ereal_add_mono \<open>I \<noteq> {}\<close> \<open>c \<noteq> -\<infinity>\<close>)
qed
@@ -2130,7 +2129,6 @@
by simp
next
assume "(SUP i:I. f i) \<noteq> 0" then show ?thesis
- unfolding SUP_def
by (subst continuous_at_Sup_mono[where f="\<lambda>x. c * x"])
(auto simp: mono_def continuous_at continuous_at_imp_continuous_at_within \<open>I \<noteq> {}\<close>
intro!: ereal_mult_left_mono c)
--- 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
--- 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 \<times> '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 \<times> '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.\<close>
lemma finite_bot_fun: "(bot :: ('a::finite \<Rightarrow> '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 \<Rightarrow> '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 \<Rightarrow> '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 \<Rightarrow> '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 \<subseteq> complete_distrib_lattice
proof
--- 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' \<equiv> "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] \<open>0 \<le> l\<close> x'_def)
ultimately show ?thesis
@@ -1274,7 +1274,7 @@
case False
def y' \<equiv> "y * 2 ^ nat (- l)"
from \<open>y \<noteq> 0\<close> have "y' \<noteq> 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 \<open>\<not> 0 \<le> l\<close>
by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps)
--- 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 \<le> n}"
let ?f = "\<lambda>s. UNION {(0::nat)..s} (\<lambda>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)
--- 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
--- 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 "\<dots> = (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 "\<dots> = (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 "\<dots> = (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 "\<dots> = (INF P : {P. eventually P F}. SUPREMUM (g ` Collect P) f)"
--- 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
--- 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 \<Rightarrow> int"
"Abs_bit0 :: int \<Rightarrow> '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 \<Rightarrow> int"
"Abs_bit1 :: int \<Rightarrow> '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)
--- 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 =
--- 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 \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> 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 \<squnion> Some (\<Sqinter>Option.these B) = Some (\<Sqinter>b\<in>Option.these B. c \<squnion> b)" by simp
then have "Some c \<squnion> \<Sqinter>(Some ` Option.these B) = (\<Sqinter>x\<in>Some ` Option.these B. Some c \<squnion> 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 \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> 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 \<sqinter> \<Squnion>(Some ` Option.these B) = (\<Squnion>x\<in>Some ` Option.these B. Some c \<sqinter> 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
--- 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\<open>Polynomials and Differentiation\<close>
-
-theory Poly_Deriv
-imports Deriv Polynomial
-begin
-
-subsection \<open>Derivatives of univariate polynomials\<close>
-
-function pderiv :: "('a :: semidom) poly \<Rightarrow> '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) \<Rightarrow> 'a list \<Rightarrow> '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 \<Rightarrow> '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 = [] \<and> f * x = 0", auto simp: cCons_def)
- next
- case (Suc m) note n = this
- show ?thesis
- proof (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> 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 "\<dots> = (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 = [] \<Longrightarrow> g + of_nat m = 0 \<or> 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 \<or> 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 (\<lambda> 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 (\<lambda>i. coeff p (Suc i)) [0..<degree p]) n"
- by (cases "n < degree p", auto simp: nth_default_def coeff_eq_0)
- show ?case unfolding coeffs_def map_upt_Suc by (auto simp: id)
-next
- case 2
- obtain n xs where id: "tl (coeffs p) = xs" "(1 :: 'a) = n" by auto
- from 2 show ?case
- unfolding id by (induct xs arbitrary: n, auto simp: cCons_def)
-qed
-
-context
- assumes "SORT_CONSTRAINT('a::{semidom, semiring_char_0})"
-begin
-
-lemma pderiv_eq_0_iff:
- "pderiv (p :: 'a poly) = 0 \<longleftrightarrow> 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) \<noteq> 0"
- shows "\<not> 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 \<le> 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 \<longleftrightarrow> 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)) =
- (\<Sum>a \<in> 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"
- "\<And> g. setsum g (insert a as) = g a + setsum g as"
- "insert a as - {a} = as"
- by auto
- {
- fix b
- assume "b \<in> as"
- hence id2: "insert a as - {b} = insert a (as - {b})" using \<open>a \<notin> as\<close> 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 (\<lambda>x. poly p (f x))"
-proof -
- have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))"
- by (intro continuous_intros assms)
- also have "\<dots> = (\<lambda>x. poly p (f x))" by (intro ext) (simp add: poly_altdef mult_ac)
- finally show ?thesis .
-qed
-
-text\<open>Consequences of the derivative theorem above\<close>
-
-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 |]
- ==> \<exists>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 |]
- ==> \<exists>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 "a<b" and "poly p a * poly p b < 0"
- shows "\<exists>x>a. x < b \<and> 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 ==>
- \<exists>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} \<subseteq> A"
- shows "\<exists>x\<in>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 "\<exists> n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p" using assms
-proof (induct p)
- case 0
- thus ?case by auto
-next
- case (pCons a p)
- have "\<lbrakk>a\<noteq>0;p=0\<rbrakk> \<Longrightarrow> ?case" by auto
- moreover have "p\<noteq>0 \<Longrightarrow> ?case"
- proof -
- assume "p\<noteq>0"
- then obtain n1 where gte_lcoeff:"\<forall>x\<ge>n1. lead_coeff p \<le> poly p x" using that pCons by auto
- have gt_0:"lead_coeff p >0" using pCons(3) \<open>p\<noteq>0\<close> by auto
- def n\<equiv>"max n1 (1+ \<bar>a\<bar>/(lead_coeff p))"
- show ?thesis
- proof (rule_tac x=n in exI,rule,rule)
- fix x assume "n \<le> x"
- hence "lead_coeff p \<le> poly p x"
- using gte_lcoeff unfolding n_def by auto
- hence " \<bar>a\<bar>/(lead_coeff p) \<ge> \<bar>a\<bar>/(poly p x)" and "poly p x>0" using gt_0
- by (intro frac_le,auto)
- hence "x\<ge>1+ \<bar>a\<bar>/(poly p x)" using \<open>n\<le>x\<close>[unfolded n_def] by auto
- thus "lead_coeff (pCons a p) \<le> poly (pCons a p) x"
- using \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x>0\<close> \<open>p\<noteq>0\<close>
- by (auto simp add:field_simps)
- qed
- qed
- ultimately show ?case by fastforce
-qed
-
-
-subsection \<open>Algebraic numbers\<close>
-
-text \<open>
- 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.
-\<close>
-subsection \<open>Algebraic numbers\<close>
-
-definition algebraic :: "'a :: field_char_0 \<Rightarrow> bool" where
- "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<int>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
-
-lemma algebraicI:
- assumes "\<And>i. coeff p i \<in> \<int>" "p \<noteq> 0" "poly p x = 0"
- shows "algebraic x"
- using assms unfolding algebraic_def by blast
-
-lemma algebraicE:
- assumes "algebraic x"
- obtains p where "\<And>i. coeff p i \<in> \<int>" "p \<noteq> 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 \<Longrightarrow> of_int a div of_int b \<in> (\<int> :: 'a :: ring_div set)"
-proof (cases "of_int b = (0 :: 'a)")
- assume "b dvd a" "of_int b \<noteq> (0::'a)"
- then obtain c where "a = b * c" by (elim dvdE)
- with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
-qed auto
-
-lemma of_int_divide_in_Ints:
- "b dvd a \<Longrightarrow> of_int a / of_int b \<in> (\<int> :: 'a :: field set)"
-proof (cases "of_int b = (0 :: 'a)")
- assume "b dvd a" "of_int b \<noteq> (0::'a)"
- then obtain c where "a = b * c" by (elim dvdE)
- with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
-qed auto
-
-lemma algebraic_altdef:
- fixes p :: "'a :: field_char_0 poly"
- shows "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
-proof safe
- fix p assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0"
- def cs \<equiv> "coeffs p"
- from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'" unfolding Rats_def by blast
- then obtain f where f: "\<And>i. coeff p i = of_rat (f (coeff p i))"
- by (subst (asm) bchoice_iff) blast
- def cs' \<equiv> "map (quotient_of \<circ> f) (coeffs p)"
- def d \<equiv> "Lcm (set (map snd cs'))"
- def p' \<equiv> "smult (of_int d) p"
-
- have "\<forall>n. coeff p' n \<in> \<int>"
- proof
- fix n :: nat
- show "coeff p' n \<in> \<int>"
- proof (cases "n \<le> degree p")
- case True
- def c \<equiv> "coeff p n"
- def a \<equiv> "fst (quotient_of (f (coeff p n)))" and b \<equiv> "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 * \<dots> = of_rat (of_int (a*d) / of_int b)"
- by (simp add: of_rat_mult of_rat_divide)
- also from nz True have "b \<in> 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 \<in> (\<int> :: rat set)"
- by (rule of_int_divide_in_Ints)
- hence "of_rat (of_int (a * d) / of_int b) \<in> \<int>" 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') \<subseteq> {0<..}"
- unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc)
- hence "d \<noteq> 0" unfolding d_def by (induction cs') simp_all
- with nz have "p' \<noteq> 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: "\<And>i. coeff p i \<in> \<int>" "poly p x = 0" "p \<noteq> 0"
- by (force simp: algebraic_def)
- moreover have "coeff p i \<in> \<int> \<Longrightarrow> coeff p i \<in> \<rat>" for i by (elim Ints_cases) simp
- ultimately show "(\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)" by auto
-qed
-
-
-text\<open>Lemmas for Derivatives\<close>
-
-lemma order_unique_lemma:
- fixes p :: "'a::idom poly"
- assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-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 \<noteq> 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) \<noteq> 0"
- using assms by auto
- obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0"
- using assms by (cases n) auto
- have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> 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 "\<not> [:- 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 \<open>n = Suc n'\<close> pe)
-qed
-
-lemma order_decomp:
- assumes "p \<noteq> 0"
- shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q"
-proof -
- from assms have A: "[:- a, 1:] ^ order a p dvd p"
- and B: "\<not> [:- 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 "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q"
- by simp
- then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q"
- by simp
- then have D: "\<not> [:- 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:
- "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk> \<Longrightarrow>
- (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 \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q"
-proof -
- def i \<equiv> "order a p"
- def j \<equiv> "order a q"
- def t \<equiv> "[:-a, 1:]"
- have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0"
- unfolding t_def by (simp add: dvd_iff_poly_eq_0)
- assume "p * q \<noteq> 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 \<noteq> 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 \<dots> = 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 "\<not> \<not> [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]"
- hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp
- hence "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:] )"
- by (rule dvd_imp_degree_le,auto)
- thus False by auto
- next
- fix y assume asm:"\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]"
- show "1 \<le> y"
- proof (rule ccontr)
- assume "\<not> 1 \<le> 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\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close>
-
-lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> 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) \<noteq> 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 \<noteq> (if order a p = 0 then 0 else 1)"
- from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 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 \<noteq> 0" by (auto split: if_splits)
- have "order a (pderiv p) = order a e + order a d"
- using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult)
- have "order a p = Suc (order a (pderiv p))"
- using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv)
- have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> 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 \<open>p \<noteq> 0\<close>
- \<open>order a p = Suc (order a (pderiv p))\<close>)
- apply (rule dvd_mult)
- apply (simp add: order_divides)
- done
- then have "order a (pderiv p) \<le> order a d"
- using \<open>d \<noteq> 0\<close> by (simp add: order_divides)
- show ?thesis
- using \<open>order a p = order a q + order a d\<close>
- using \<open>order a (pderiv p) = order a e + order a d\<close>
- using \<open>order a p = Suc (order a (pderiv p))\<close>
- using \<open>order a (pderiv p) \<le> order a d\<close>
- by auto
-qed
-
-lemma poly_squarefree_decomp_order2:
- "\<lbrakk>pderiv p \<noteq> (0 :: 'a :: field_char_0 poly);
- p = q * d;
- pderiv p = e * d;
- d = r * p + s * pderiv p
- \<rbrakk> \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
-by (blast intro: poly_squarefree_decomp_order)
-
-lemma order_pderiv2:
- "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk>
- \<Longrightarrow> (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 \<noteq> 0 & (\<forall>a. (order a p = 0) | (order a p = 1)))"
-
-lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>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 = (\<forall>a. \<not>(poly p a = 0 \<and> 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) \<noteq> 0"
- and "p = q * d"
- and "pderiv p = e * d"
- and "d = r * p + s * pderiv p"
- shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"
-proof -
- from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
- with \<open>p = q * d\<close> have "q \<noteq> 0" by simp
- have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
- using assms by (rule poly_squarefree_decomp_order2)
- with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis
- by (simp add: rsquarefree_def order_root)
-qed
-
-end
--- 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 \<open>Polynomials as type over a ring structure\<close>
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 \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
@@ -1551,7 +1553,7 @@
assumes "a \<noteq> 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 \<Rightarrow> '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 \<Rightarrow> '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 \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y"
using pdivmod_rel [of x y]
@@ -1884,137 +1902,6 @@
by (subst (asm) order_root) auto
-subsection \<open>GCD of polynomials\<close>
-
-instantiation poly :: (field) gcd
-begin
-
-function gcd_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
-where
- "gcd (x::'a poly) 0 = smult (inverse (coeff x (degree x))) x"
-| "y \<noteq> 0 \<Longrightarrow> gcd (x::'a poly) y = gcd y (x mod y)"
-by auto
-
-termination "gcd :: _ poly \<Rightarrow> _"
-by (relation "measure (\<lambda>(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 \<Rightarrow> 'a poly \<Rightarrow> '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 \<Longrightarrow> k dvd y \<Longrightarrow> 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 \<longleftrightarrow> k dvd x \<and> 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 \<and> 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 \<longleftrightarrow> x = 0 \<and> 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 \<noteq> 0" by auto
- have degree: "degree p = degree q"
- using \<open>p dvd q\<close> \<open>q dvd p\<close> \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close>
- by (intro order_antisym dvd_imp_degree_le)
-
- from \<open>p dvd q\<close> obtain a where a: "q = p * a" ..
- with \<open>q \<noteq> 0\<close> have "a \<noteq> 0" by auto
- with degree a \<open>p \<noteq> 0\<close> 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: "\<And>k. k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd d"
- and monic: "coeff d (degree d) = (if x = 0 \<and> 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 \<Rightarrow> _"
-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 \<open>Additional induction rules on polynomials\<close>
text \<open>
@@ -2326,8 +2213,630 @@
lemma lead_coeff_nonzero: "p \<noteq> 0 \<Longrightarrow> lead_coeff p \<noteq> 0"
by (simp add: lead_coeff_def)
+
+
+subsection \<open>Derivatives of univariate polynomials\<close>
+
+function pderiv :: "('a :: semidom) poly \<Rightarrow> '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) \<Rightarrow> 'a list \<Rightarrow> '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 \<Rightarrow> '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 = [] \<and> f * x = 0", auto simp: cCons_def)
+ next
+ case (Suc m) note n = this
+ show ?thesis
+ proof (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> 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 "\<dots> = (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 = [] \<Longrightarrow> g + of_nat m = 0 \<or> 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 \<or> 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 (\<lambda> 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 (\<lambda>i. coeff p (Suc i)) [0..<degree p]) n"
+ by (cases "n < degree p", auto simp: nth_default_def coeff_eq_0)
+ show ?case unfolding coeffs_def map_upt_Suc by (auto simp: id)
+next
+ case 2
+ obtain n xs where id: "tl (coeffs p) = xs" "(1 :: 'a) = n" by auto
+ from 2 show ?case
+ unfolding id by (induct xs arbitrary: n, auto simp: cCons_def)
+qed
+
+context
+ assumes "SORT_CONSTRAINT('a::{semidom, semiring_char_0})"
+begin
+
+lemma pderiv_eq_0_iff:
+ "pderiv (p :: 'a poly) = 0 \<longleftrightarrow> 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) \<noteq> 0"
+ shows "\<not> 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 \<le> 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 \<longleftrightarrow> 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)) =
+ (\<Sum>a \<in> 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"
+ "\<And> g. setsum g (insert a as) = g a + setsum g as"
+ "insert a as - {a} = as"
+ by auto
+ {
+ fix b
+ assume "b \<in> as"
+ hence id2: "insert a as - {b} = insert a (as - {b})" using \<open>a \<notin> as\<close> 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 (\<lambda>x. poly p (f x))"
+proof -
+ have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))"
+ by (intro continuous_intros assms)
+ also have "\<dots> = (\<lambda>x. poly p (f x))" by (intro ext) (simp add: poly_altdef mult_ac)
+ finally show ?thesis .
+qed
+
+text\<open>Consequences of the derivative theorem above\<close>
+
+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 |]
+ ==> \<exists>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 |]
+ ==> \<exists>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 "a<b" and "poly p a * poly p b < 0"
+ shows "\<exists>x>a. x < b \<and> 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 ==>
+ \<exists>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} \<subseteq> A"
+ shows "\<exists>x\<in>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 "\<exists> n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p" using assms
+proof (induct p)
+ case 0
+ thus ?case by auto
+next
+ case (pCons a p)
+ have "\<lbrakk>a\<noteq>0;p=0\<rbrakk> \<Longrightarrow> ?case" by auto
+ moreover have "p\<noteq>0 \<Longrightarrow> ?case"
+ proof -
+ assume "p\<noteq>0"
+ then obtain n1 where gte_lcoeff:"\<forall>x\<ge>n1. lead_coeff p \<le> poly p x" using that pCons by auto
+ have gt_0:"lead_coeff p >0" using pCons(3) \<open>p\<noteq>0\<close> by auto
+ def n\<equiv>"max n1 (1+ \<bar>a\<bar>/(lead_coeff p))"
+ show ?thesis
+ proof (rule_tac x=n in exI,rule,rule)
+ fix x assume "n \<le> x"
+ hence "lead_coeff p \<le> poly p x"
+ using gte_lcoeff unfolding n_def by auto
+ hence " \<bar>a\<bar>/(lead_coeff p) \<ge> \<bar>a\<bar>/(poly p x)" and "poly p x>0" using gt_0
+ by (intro frac_le,auto)
+ hence "x\<ge>1+ \<bar>a\<bar>/(poly p x)" using \<open>n\<le>x\<close>[unfolded n_def] by auto
+ thus "lead_coeff (pCons a p) \<le> poly (pCons a p) x"
+ using \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x>0\<close> \<open>p\<noteq>0\<close>
+ by (auto simp add:field_simps)
+ qed
+ qed
+ ultimately show ?case by fastforce
+qed
+
+
+subsection \<open>Algebraic numbers\<close>
+
+text \<open>
+ 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.
+\<close>
+subsection \<open>Algebraic numbers\<close>
+
+definition algebraic :: "'a :: field_char_0 \<Rightarrow> bool" where
+ "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<int>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
+
+lemma algebraicI:
+ assumes "\<And>i. coeff p i \<in> \<int>" "p \<noteq> 0" "poly p x = 0"
+ shows "algebraic x"
+ using assms unfolding algebraic_def by blast
+lemma algebraicE:
+ assumes "algebraic x"
+ obtains p where "\<And>i. coeff p i \<in> \<int>" "p \<noteq> 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 \<Longrightarrow> of_int a div of_int b \<in> (\<int> :: 'a :: ring_div set)"
+proof (cases "of_int b = (0 :: 'a)")
+ assume "b dvd a" "of_int b \<noteq> (0::'a)"
+ then obtain c where "a = b * c" by (elim dvdE)
+ with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
+qed auto
+
+lemma of_int_divide_in_Ints:
+ "b dvd a \<Longrightarrow> of_int a / of_int b \<in> (\<int> :: 'a :: field set)"
+proof (cases "of_int b = (0 :: 'a)")
+ assume "b dvd a" "of_int b \<noteq> (0::'a)"
+ then obtain c where "a = b * c" by (elim dvdE)
+ with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
+qed auto
+
+lemma algebraic_altdef:
+ fixes p :: "'a :: field_char_0 poly"
+ shows "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
+proof safe
+ fix p assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0"
+ def cs \<equiv> "coeffs p"
+ from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'" unfolding Rats_def by blast
+ then obtain f where f: "\<And>i. coeff p i = of_rat (f (coeff p i))"
+ by (subst (asm) bchoice_iff) blast
+ def cs' \<equiv> "map (quotient_of \<circ> f) (coeffs p)"
+ def d \<equiv> "Lcm (set (map snd cs'))"
+ def p' \<equiv> "smult (of_int d) p"
+
+ have "\<forall>n. coeff p' n \<in> \<int>"
+ proof
+ fix n :: nat
+ show "coeff p' n \<in> \<int>"
+ proof (cases "n \<le> degree p")
+ case True
+ def c \<equiv> "coeff p n"
+ def a \<equiv> "fst (quotient_of (f (coeff p n)))" and b \<equiv> "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 * \<dots> = of_rat (of_int (a*d) / of_int b)"
+ by (simp add: of_rat_mult of_rat_divide)
+ also from nz True have "b \<in> 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 \<in> (\<int> :: rat set)"
+ by (rule of_int_divide_in_Ints)
+ hence "of_rat (of_int (a * d) / of_int b) \<in> \<int>" 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') \<subseteq> {0<..}"
+ unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc)
+ hence "d \<noteq> 0" unfolding d_def by (induction cs') simp_all
+ with nz have "p' \<noteq> 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: "\<And>i. coeff p i \<in> \<int>" "poly p x = 0" "p \<noteq> 0"
+ by (force simp: algebraic_def)
+ moreover have "coeff p i \<in> \<int> \<Longrightarrow> coeff p i \<in> \<rat>" for i by (elim Ints_cases) simp
+ ultimately show "(\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)" by auto
+qed
+
+
+text\<open>Lemmas for Derivatives\<close>
+
+lemma order_unique_lemma:
+ fixes p :: "'a::idom poly"
+ assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-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 \<noteq> 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) \<noteq> 0"
+ using assms by auto
+ obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0"
+ using assms by (cases n) auto
+ have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> 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 "\<not> [:- 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 \<open>n = Suc n'\<close> pe)
+qed
+
+lemma order_decomp:
+ assumes "p \<noteq> 0"
+ shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q"
+proof -
+ from assms have A: "[:- a, 1:] ^ order a p dvd p"
+ and B: "\<not> [:- 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 "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q"
+ by simp
+ then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q"
+ by simp
+ then have D: "\<not> [:- 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:
+ "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk> \<Longrightarrow>
+ (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 \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q"
+proof -
+ def i \<equiv> "order a p"
+ def j \<equiv> "order a q"
+ def t \<equiv> "[:-a, 1:]"
+ have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0"
+ unfolding t_def by (simp add: dvd_iff_poly_eq_0)
+ assume "p * q \<noteq> 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 \<noteq> 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 \<dots> = 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 "\<not> \<not> [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]"
+ hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp
+ hence "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:] )"
+ by (rule dvd_imp_degree_le,auto)
+ thus False by auto
+ next
+ fix y assume asm:"\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]"
+ show "1 \<le> y"
+ proof (rule ccontr)
+ assume "\<not> 1 \<le> 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\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close>
+
+lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> 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) \<noteq> 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 \<noteq> (if order a p = 0 then 0 else 1)"
+ from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 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 \<noteq> 0" by (auto split: if_splits)
+ have "order a (pderiv p) = order a e + order a d"
+ using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult)
+ have "order a p = Suc (order a (pderiv p))"
+ using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv)
+ have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> 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 \<open>p \<noteq> 0\<close>
+ \<open>order a p = Suc (order a (pderiv p))\<close>)
+ apply (rule dvd_mult)
+ apply (simp add: order_divides)
+ done
+ then have "order a (pderiv p) \<le> order a d"
+ using \<open>d \<noteq> 0\<close> by (simp add: order_divides)
+ show ?thesis
+ using \<open>order a p = order a q + order a d\<close>
+ using \<open>order a (pderiv p) = order a e + order a d\<close>
+ using \<open>order a p = Suc (order a (pderiv p))\<close>
+ using \<open>order a (pderiv p) \<le> order a d\<close>
+ by auto
+qed
+
+lemma poly_squarefree_decomp_order2:
+ "\<lbrakk>pderiv p \<noteq> (0 :: 'a :: field_char_0 poly);
+ p = q * d;
+ pderiv p = e * d;
+ d = r * p + s * pderiv p
+ \<rbrakk> \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
+by (blast intro: poly_squarefree_decomp_order)
+
+lemma order_pderiv2:
+ "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk>
+ \<Longrightarrow> (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 \<noteq> 0 & (\<forall>a. (order a p = 0) | (order a p = 1)))"
+
+lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>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 = (\<forall>a. \<not>(poly p a = 0 \<and> 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) \<noteq> 0"
+ and "p = q * d"
+ and "pderiv p = e * d"
+ and "d = r * p + s * pderiv p"
+ shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"
+proof -
+ from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
+ with \<open>p = q * d\<close> have "q \<noteq> 0" by simp
+ have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
+ using assms by (rule poly_squarefree_decomp_order2)
+ with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis
+ by (simp add: rsquarefree_def order_root)
+qed
+
no_notation cCons (infixr "##" 65)
--- /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 \<open>GCD and LCM on polynomials over a field\<close>
+
+theory Polynomial_GCD_euclidean
+imports Main "~~/src/HOL/GCD" "~~/src/HOL/Library/Polynomial"
+begin
+
+subsection \<open>GCD of polynomials\<close>
+
+instantiation poly :: (field) gcd
+begin
+
+function gcd_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
+where
+ "gcd (x::'a poly) 0 = smult (inverse (coeff x (degree x))) x"
+| "y \<noteq> 0 \<Longrightarrow> gcd (x::'a poly) y = gcd y (x mod y)"
+by auto
+
+termination "gcd :: _ poly \<Rightarrow> _"
+by (relation "measure (\<lambda>(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 \<Rightarrow> 'a poly \<Rightarrow> '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 \<Longrightarrow> k dvd y \<Longrightarrow> 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 \<longleftrightarrow> k dvd x \<and> 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 \<and> 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 \<longleftrightarrow> x = 0 \<and> 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 \<noteq> 0" by auto
+ have degree: "degree p = degree q"
+ using \<open>p dvd q\<close> \<open>q dvd p\<close> \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close>
+ by (intro order_antisym dvd_imp_degree_le)
+
+ from \<open>p dvd q\<close> obtain a where a: "q = p * a" ..
+ with \<open>q \<noteq> 0\<close> have "a \<noteq> 0" by auto
+ with degree a \<open>p \<noteq> 0\<close> 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: "\<And>k. k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd d"
+ and monic: "coeff d (degree d) = (if x = 0 \<and> 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
--- 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 \<open>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 \<open>Complete distributive lattices\<close>
--- 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 \<open>Rules requiring bi-unique, bi-total or right-total relations\<close>
--- 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 \<leftarrow> xys, yz \<leftarrow> 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)"
--- 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 \<times> int) \<Rightarrow> real" where
"float = (\<lambda>(a, b). real_of_int a * 2 powr real_of_int b)"
--- 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 *)
--- 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 \<le> 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) = \<bar>a\<bar> * 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)
--- 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>open s\<close>
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) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
using assms
by (auto intro: continuous_closed_in_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv)
--- 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]:
"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
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]:
"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
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 \<noteq> {}"
@@ -496,7 +496,7 @@
using assms box_ne_empty(1) content_cbox by blast
lemma content_real: "a \<le> b \<Longrightarrow> 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: "\<bar>y - x\<bar> = (if x\<le>y then content {x .. y} else content {y..x})"
by (auto simp: content_real)
@@ -856,7 +856,7 @@
by auto
show "\<Union>?A = s1 \<inter> 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)
--- 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 ((\<lambda>i. \<bar>x \<bullet> i\<bar>) ` 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"
--- 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 \<noteq> {}"
hence "\<And>i. (\<lambda>x. x \<bullet> i) ` X \<noteq> {}" by simp
thus "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X" "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> 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 \<in> Basis \<Longrightarrow> inf x y \<bullet> i = inf (x \<bullet> i) (y \<bullet> i)"
and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> 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: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> 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 \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b"
by auto
hence "Inf ?proj = x \<bullet> 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 \<bullet> b = Inf X \<bullet> b"
- by (auto simp: eucl_Inf Inf_class.INF_def inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> setsum.delta
- simp del: Inf_class.Inf_image_eq
+ by (auto simp: eucl_Inf inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> setsum.delta
cong: if_cong)
with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Inf X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> x \<bullet> b \<le> y \<bullet> b)" by blast
qed
@@ -140,10 +137,10 @@
and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b"
by auto
hence "Sup ?proj = x \<bullet> b"
- by (auto intro!: cSup_eq_maximum simp del: Sup_image_eq)
+ by (auto intro!: cSup_eq_maximum)
hence "x \<bullet> b = Sup X \<bullet> b"
- by (auto simp: eucl_Sup[where 'a='a] SUP_def inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> 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 \<open>b \<in> Basis\<close> setsum.delta
+ cong: if_cong)
with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Sup X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> y \<bullet> b \<le> x \<bullet> 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"
--- 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 \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> closedin U (B x)"
shows "closedin U (\<Inter>x\<in>A. B x)"
- unfolding Inter_image_eq [symmetric]
apply (rule closedin_Inter)
using assms
apply auto
@@ -605,7 +604,7 @@
ultimately have "?L (\<Union>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 \<longleftrightarrow> (\<exists>T. openin U T \<and> S = T \<inter> V)"
@@ -2426,7 +2425,7 @@
fix y
assume "y \<in> {x<..} \<inter> I"
with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> 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 \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> compact (B x)) \<Longrightarrow> compact (\<Union>x\<in>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 \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
have X: "\<And>n. X n \<in> U - (\<Union>i\<le>n. from_nat_into A i)"
- using \<open>A \<noteq> {}\<close> unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into)
+ using \<open>A \<noteq> {}\<close> unfolding X_def by (intro T) (auto intro: from_nat_into)
then have "range X \<subseteq> U"
by auto
with subseq[of X] obtain r x where "x \<in> U" and r: "subseq r" "(X \<circ> r) \<longlonglongrightarrow> x"
@@ -4198,7 +4197,7 @@
then have s: "\<And>x. x \<in> s \<Longrightarrow> \<exists>U. x\<in>U \<and> open U \<and> finite (U \<inter> t)" by metis
have "s \<subseteq> \<Union>C"
using \<open>t \<subseteq> s\<close>
- unfolding C_def Union_image_eq
+ unfolding C_def
apply (safe dest!: s)
apply (rule_tac a="U \<inter> 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 \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E"
and s: "s \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- t)))"
- by (metis (lifting) Union_image_eq finite_subset_image C_def)
+ by (metis (lifting) finite_subset_image C_def)
from s \<open>t \<subseteq> s\<close> have "t \<subseteq> \<Union>E"
using interior_subset by blast
moreover have "finite (\<Union>E)"
@@ -4348,7 +4347,7 @@
from Rats_dense_in_real[OF \<open>0 < e / 2\<close>] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2"
by auto
from f[rule_format, of r] \<open>0 < r\<close> \<open>x \<in> s\<close> obtain k where "k \<in> f r" "x \<in> ball k r"
- unfolding Union_image_eq by auto
+ by auto
from \<open>r \<in> \<rat>\<close> \<open>0 < r\<close> \<open>k \<in> f r\<close> have "ball k r \<in> K"
by (auto simp: K_def)
then show "\<exists>b\<in>K. x \<in> b \<and> b \<inter> s \<subseteq> T"
@@ -7412,7 +7411,7 @@
lemma diameter_cbox:
fixes a b::"'a::euclidean_space"
shows "(\<forall>i \<in> Basis. a \<bullet> i \<le> b \<bullet> i) \<Longrightarrow> 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:
--- 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: "\<forall>M. \<exists>N. 0 < N & (\<forall>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 "\<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> 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:
+ "\<forall>M::nat. \<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> 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:
+ "\<And>x::hypnat. is_unit x \<Longrightarrow> x = 1"
+ by transfer simp
text\<open>Already proved as \<open>primes_infinite\<close>, but now using non-standard naturals.\<close>
theorem not_finite_prime: "~ finite {p::nat. prime p}"
--- 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 \<open>@{term "op dvd"} is a partial order\<close>
-
-interpretation dvd: order "op dvd" "\<lambda>n m :: nat. n dvd m \<and> \<not> 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])
--- 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 \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)"
--- 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: "\<bar>sgn x :: real\<bar> = (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 \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a ^ n = b ^ n \<longleftrightarrow> a = b"
- using power_eq_imp_eq_base[of a n b] by auto
-
subsection \<open>Existence of Nth Root\<close>
text \<open>Existence follows from the Intermediate Value Theorem\<close>
--- 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\<open>Also the class for fields with characteristic zero.\<close>
+
+class field_char_0 = field + ring_char_0
+
+
subsubsection \<open>
Structures with negation and order: class \<open>linordered_idom\<close>
\<close>
@@ -1099,6 +1104,8 @@
context linordered_field
begin
+subclass field_char_0 ..
+
lemma half_gt_zero_iff:
"0 < a / 2 \<longleftrightarrow> 0 < a" (is "?P \<longleftrightarrow> ?Q")
by (auto simp add: field_simps)
--- 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) \<Longrightarrow> [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) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"
@@ -285,7 +285,7 @@
lemma coprime_cong_mult_int:
"[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n
\<Longrightarrow> [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 \<Longrightarrow> 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 \<Longrightarrow> 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) \<Longrightarrow> 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) \<Longrightarrow>
[x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
apply (cases "y \<le> 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) \<Longrightarrow>
[x = y] (mod b) \<Longrightarrow> [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 \<Longrightarrow>
(\<forall>i\<in>A. (\<forall>j\<in>A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
@@ -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)"
--- 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 \<open>m \<noteq> 1\<close> have "Suc (Suc 0) \<le> m" by arith
with \<open>m < Suc n\<close> * \<open>m dvd q\<close> have "q dvd m" by simp
- with \<open>m dvd q\<close> show ?thesis by (simp add: dvd.eq_iff)
+ with \<open>m dvd q\<close> show ?thesis by (simp add: dvd_antisym)
qed
}
then have aux: "\<And>m q. Suc (Suc 0) \<le> q \<Longrightarrow>
--- 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 \<in> A \<Longrightarrow> Lcm A = 0"
by (drule dvd_Lcm) simp
-lemma Lcm0_iff':
+lemma Lcm_0_iff':
"Lcm A = 0 \<longleftrightarrow> \<not>(\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>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 \<Longrightarrow> Lcm A = 0 \<longleftrightarrow> 0 \<in> A"
proof -
assume "finite A"
@@ -1108,7 +1108,7 @@
done
moreover from \<open>finite A\<close> have "\<forall>a\<in>A. a dvd \<Prod>A" by blast
ultimately have "\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l)" by blast
- with Lcm0_iff' have "Lcm A \<noteq> 0" by simp
+ with Lcm_0_iff' have "Lcm A \<noteq> 0" by simp
}
ultimately show "Lcm A = 0 \<longleftrightarrow> 0 \<in> 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 "\<And>a. a \<in> A \<Longrightarrow> b dvd a" and "\<And>c. (\<And>a. a \<in> A \<Longrightarrow> c dvd a) \<Longrightarrow> c dvd b"
@@ -1210,10 +1210,7 @@
lemma Lcm_Gcd:
"Lcm A = Gcd {m. \<forall>a\<in>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 \<in> A \<Longrightarrow> Gcd A = 1"
--- 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 \<bar>p\<bar>)" by simp
assume "p dvd k * l"
then have "nat \<bar>p\<bar> dvd nat \<bar>k * l\<bar>"
- by (simp add: dvd_int_iff)
+ by (simp add: dvd_int_unfold_dvd_nat)
then have "nat \<bar>p\<bar> dvd nat \<bar>k\<bar> * nat \<bar>l\<bar>"
by (simp add: abs_mult nat_mult_distrib)
with * have "nat \<bar>p\<bar> dvd nat \<bar>k\<bar> \<or> nat \<bar>p\<bar> dvd nat \<bar>l\<bar>"
using is_primeD [of "nat \<bar>p\<bar>"] by auto
then show "p dvd k \<or> p dvd l"
- by (simp add: dvd_int_iff)
+ by (simp add: dvd_int_unfold_dvd_nat)
next
fix k :: int
assume P: "\<And>l. l dvd k \<Longrightarrow> \<not> is_prime l"
--- 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 \<le> n \<Longrightarrow> 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)"
-- \<open>Law 6.111\<close>
- 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 = (\<Sum>k \<in> {..n}. fib k * fib k)"
by (induct n rule: nat.induct) (auto simp add: field_simps)
--- 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
--- 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 (\<Union>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:
--- 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\<open>Lemmas about previously defined terms\<close>
lemma prime:
- "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
+ "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p \<noteq> 1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof-
{assume "p=0 \<or> p=1" hence ?thesis
@@ -39,7 +39,7 @@
{assume "q\<noteq>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 \<or> 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 "\<exists>!y. 0 < y \<and> y < p \<and> [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 \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d\<^sup>2 \<le> n \<longrightarrow> d = 1)"
proof -
{assume "n=0 \<or> 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\<noteq>0" "n\<noteq>1"
hence np: "n > 1" by arith
@@ -583,9 +583,8 @@
{fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
from pp[unfolded prime_def] d have dp: "d = p" by blast
from n have "n \<noteq> 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 \<noteq> 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 \<noteq> 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}
--- 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 \<Longrightarrow> \<not> p dvd n \<Longrightarrow> 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 \<or> 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 \<Longrightarrow> \<not> p dvd n \<Longrightarrow> 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 \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> 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 \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> 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 \<Longrightarrow>
p dvd m * n = (p dvd m \<or> p dvd n)"
@@ -167,14 +167,19 @@
"prime p \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
by (rule coprime_exp2_nat, rule primes_coprime_nat)
-lemma prime_factor_nat: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> 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 \<noteq> (1::nat) \<Longrightarrow> \<exists>p. prime p \<and> 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 \<open>One property of coprimality is easier to prove via prime factors.\<close>
@@ -198,20 +203,20 @@
{ assume pa: "p dvd a"
from coprime_common_divisor_nat [OF ab, OF pa] p have "\<not> 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 "\<not> 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: "\<not> p dvd a"
shows "\<exists>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 \<noteq> 1" by auto
+ ultimately have "\<exists>x y. a * x = p * y + 1"
+ by (rule coprime_bezout_strong)
+ then show ?thesis by simp
qed
end
--- 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 "\<not> 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 \<open>Wilson's theorem\<close>
--- 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 * (\<Prod>i \<in> (set_mset N - {a}). i ^ count N i)" .
moreover
have "coprime (a ^ count M a) (\<Prod>i \<in> (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 \<Longrightarrow> \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> 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 "\<forall>p. prime p \<longrightarrow> multiplicity p x \<le> 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 \<Longrightarrow> 0 \<le> y \<Longrightarrow>
\<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> 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:
--- 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 \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
unfolding dvd_def by blast
from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
- then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
+ then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: of_nat_mult)
then show ?thesis
apply (subst abs_dvd_iff [symmetric])
apply (subst dvd_abs_iff [symmetric])
--- 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 \<Longrightarrow> inj_on g B \<Longrightarrow> ALL a:A. ALL b:B. coprime (f a) (g b)
- \<Longrightarrow> inj_on (%(a,b). f a * g b::nat) (A \<times> 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 \<Longrightarrow> inj_on g B \<Longrightarrow> \<forall>a\<in>A. \<forall>b\<in>B. Primes.coprime (f a) (g b) \<Longrightarrow>
+ inj_on (\<lambda>(a, b). f a * g b) (A \<times> 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]
--- 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 \<le> x * x + y * y"
- by (intro add_nonneg_nonneg zero_le_square)
-
-lemma not_sum_squares_lt_zero:
- "\<not> x * x + y * y < 0"
- by (simp add: not_less sum_squares_ge_zero)
-
-end
-
subsection \<open>Powers for Arbitrary Monoids\<close>
class power = one + times
@@ -123,6 +110,10 @@
finally show ?case .
qed simp
+lemma power_minus_mult:
+ "0 < n \<Longrightarrow> 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 \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 < n \<Longrightarrow> a = b"
by (cases n) (simp_all del: power_Suc, rule power_inject_base)
+lemma power_eq_iff_eq_base:
+ "0 < n \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a ^ n = b ^ n \<longleftrightarrow> a = b"
+ using power_eq_imp_eq_base [of a n b] by auto
+
lemma power2_le_imp_le:
"x\<^sup>2 \<le> y\<^sup>2 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x \<le> y"
unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)
@@ -596,6 +591,10 @@
"x\<^sup>2 = y\<^sup>2 \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x = y"
unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp
+lemma power_Suc_le_self:
+ shows "0 \<le> a \<Longrightarrow> a \<le> 1 \<Longrightarrow> a ^ Suc n \<le> a"
+ using power_decreasing [of 1 "Suc n" a] by simp
+
end
context linordered_ring_strict
--- 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 \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))"
- by (cases "y \<le> x") (simp_all add: zdiff_int)
+ by (cases "y \<le> x") (simp_all add: of_nat_diff)
text \<open>
\medskip Specific instances of congruence rules, to prevent
--- 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 \<open>outer_measure M f X \<noteq> \<infinity>\<close> have fin: "\<bar>outer_measure M f X\<bar> \<noteq> \<infinity>"
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 \<open>0 < e\<close>]
- 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 \<open>0 < e\<close>]
+ 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) = (\<Sum>i\<in>I. f (A i))"
proof -
have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>(A`I) \<in> 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 \<open>volume M f\<close> have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
unfolding volume_def by blast
also have "\<dots> = (\<Sum>i\<in>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 "(\<Union>i. F i) = \<Union>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: "\<And>i. F i \<in> 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 = (\<Union>x<card C. f x)"
- using f C Ai unfolding bij_betw_def by (simp add: Union_image_eq[symmetric])
+ using f C Ai unfolding bij_betw_def by auto
then have "\<Union>range 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 "(\<Sum>j. \<mu>_r (f j)) = (\<Sum>j. if j \<in> {..< card C} then \<mu>_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 "(\<Sum>n. \<mu>_r (A' n)) = (\<Sum>n. \<Sum>c\<in>C'. \<mu>_r (A' n \<inter> 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="\<lambda>X. suminf (\<lambda>i. \<mu>_r (X i))"] ext
intro: generated_ringI_Basic)
also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))"
@@ -853,7 +854,7 @@
also have "\<dots> = \<mu>_r (\<Union>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 "(\<Sum>n. \<mu>_r (A' n)) = \<mu>_r (\<Union>i. A' i)"
using C' by simp
--- 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 \<union> T)) = emeasure M (main_part M S \<union> 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:
--- 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 "\<And>i. A i = f ` A' i" "\<And>i. A' i \<in> sets M"
by (auto simp: subset_eq choice_iff)
moreover from this have "(\<Union>x. f ` A' x) = f ` (\<Union>x. A' x)" by blast
- ultimately show "(\<Union>i. A i) \<in> {f ` A |A. A \<in> sets M}" by blast
+ ultimately show "(\<Union>i. A i) \<in> {f ` A |A. A \<in> sets M}"
+ by simp blast
qed (auto dest: sets.sets_into_space)
lemma the_inv_into_vimage:
--- 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 \<inter> space (PiF I M) = (\<Union>J\<in>I. A -` y \<inter> space (PiF {J} M))"
by (auto simp: space_PiF)
also have "\<dots> \<in> sets (PiF I M)"
- proof
+ proof (rule sets.finite_UN)
show "finite I" by fact
fix J assume "J \<in> I"
with assms have "finite J" by simp
@@ -1055,8 +1055,8 @@
by (intro Pi'_cong) (simp_all add: S_union)
also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>' j\<in>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 (\<lambda>n. f (the_inv_into I T n)) [0..<card I]" in exI)
apply (auto simp: bij_betw_def)
--- a/src/HOL/Probability/Independent_Family.thy Wed Feb 17 23:28:58 2016 +0100
+++ b/src/HOL/Probability/Independent_Family.thy Wed Feb 17 23:29:35 2016 +0100
@@ -491,11 +491,18 @@
where a: "a = (\<Inter>k\<in>Ka. Ea k)" "finite Ka" "Ka \<noteq> {}" "Ka \<subseteq> I j" "\<And>k. k\<in>Ka \<Longrightarrow> Ea k \<in> E k" by auto
fix b assume "b \<in> ?E j" then obtain Kb Eb
where b: "b = (\<Inter>k\<in>Kb. Eb k)" "finite Kb" "Kb \<noteq> {}" "Kb \<subseteq> I j" "\<And>k. k\<in>Kb \<Longrightarrow> Eb k \<in> E k" by auto
- let ?A = "\<lambda>k. (if k \<in> Ka \<inter> Kb then Ea k \<inter> Eb k else if k \<in> Kb then Eb k else if k \<in> Ka then Ea k else {})"
- have "a \<inter> b = INTER (Ka \<union> Kb) ?A"
- by (simp add: a b set_eq_iff) auto
+ let ?f = "\<lambda>k. (if k \<in> Ka \<inter> Kb then Ea k \<inter> Eb k else if k \<in> Kb then Eb k else if k \<in> Ka then Ea k else {})"
+ have "Ka \<union> Kb = (Ka \<inter> Kb) \<union> (Kb - Ka) \<union> (Ka - Kb)"
+ by blast
+ moreover have "(\<Inter>x\<in>Ka \<inter> Kb. Ea x \<inter> Eb x) \<inter>
+ (\<Inter>x\<in>Kb - Ka. Eb x) \<inter> (\<Inter>x\<in>Ka - Kb. Ea x) = (\<Inter>k\<in>Ka. Ea k) \<inter> (\<Inter>k\<in>Kb. Eb k)"
+ by auto
+ ultimately have "(\<Inter>k\<in>Ka \<union> Kb. ?f k) = (\<Inter>k\<in>Ka. Ea k) \<inter> (\<Inter>k\<in>Kb. Eb k)" (is "?lhs = ?rhs")
+ by (simp only: image_Un Inter_Un_distrib) simp
+ then have "a \<inter> b = (\<Inter>k\<in>Ka \<union> Kb. ?f k)"
+ by (simp only: a(1) b(1))
with a b \<open>j \<in> J\<close> Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j"
- by (intro CollectI exI[of _ "Ka \<union> Kb"] exI[of _ ?A]) auto
+ by (intro CollectI exI[of _ "Ka \<union> Kb"] exI[of _ ?f]) auto
qed
qed
ultimately show ?thesis
@@ -804,14 +811,18 @@
proof -
have [measurable]: "\<And>m. {x\<in>space M. P m x} \<in> sets M"
using assms by (auto simp: indep_events_def)
- have "prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 1"
- by (rule borel_0_1_law) fact
+ have *: "(\<Inter>n. \<Union>m\<in>{n..}. {x \<in> space M. P m x}) \<in> events"
+ by simp
+ from assms have "prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 1"
+ by (rule borel_0_1_law)
+ also have "prob (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 1 \<longleftrightarrow> (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 (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 0 \<longleftrightarrow> (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 (\<Inter>n. \<Union>m\<in>{n..}. ?P m) = 1 \<longleftrightarrow> (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:
--- 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 = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^sub>E j\<in>{..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 = (\<Inter>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 \<subseteq> sets S"
using E by auto
moreover have "decseq ?E"
--- 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 (\<Union>i. {from_nat_into A i}) = 0"
by (rule emeasure_UN_eq_0) auto
ultimately have "emeasure lborel A \<le> 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
--- 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 @@
"(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>X. N x i))"
"(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
"(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>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 \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>I. N x i))"
--- 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 "(\<Union>i\<in>I. N i) = \<Union>(N ` range (from_nat_into I))"
using I by simp
also have "\<dots> = (\<Union>i. (N \<circ> 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 \<subseteq> sets M" and
space: "(\<Union>i. A i) = space M" and
measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
- using sigma_finite by auto
+ using sigma_finite by blast
show thesis
proof (rule that[of "disjointed A"])
show "range (disjointed A) \<subseteq> sets M"
@@ -1252,7 +1252,7 @@
proof -
obtain F :: "nat \<Rightarrow> 'a set" where
F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
- using sigma_finite by auto
+ using sigma_finite by blast
show thesis
proof (rule that[of "\<lambda>n. \<Union>i\<le>n. F i"])
show "range (\<lambda>n. \<Union>i\<le>n. F i) \<subseteq> sets M"
@@ -1731,8 +1731,10 @@
assumes upper: "space M \<subseteq> (\<Union>i \<in> s. f i)"
shows "measure M e = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
proof -
- have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
+ have "e \<subseteq> (\<Union>i\<in>s. f i)"
using \<open>e \<in> sets M\<close> sets.sets_into_space upper by blast
+ then have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
+ by auto
hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp
also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
proof (rule finite_measure_finite_Union)
@@ -2074,7 +2076,7 @@
using E \<Omega> 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) \<subseteq> E" "(\<Union>i. from_nat_into A i) = \<Omega>"
- 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 \<Omega>) (from_nat_into A i) = emeasure M (from_nat_into A i)"
@@ -2333,12 +2335,12 @@
assumes A: "Complete_Partial_Order.chain op \<le> (f ` A)" and a: "a \<in> A" "f a \<noteq> 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 \<le> (f ` A)" "A \<noteq> {}"
assumes "X \<in> sets (SUP M:A. f M)"
shows "emeasure (SUP M:A. f M) X = (SUP M:A. emeasure (f M)) X"
-using \<open>X \<in> _\<close> unfolding SUP_def by(subst emeasure_Sup[OF A(1)]; simp add: A)
+using \<open>X \<in> _\<close> by(subst emeasure_Sup[OF A(1)]; simp add: A)
end
--- 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 "\<dots> \<le> (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 \<le> f i x" by auto
thus ?thesis using \<open>x \<in> space M\<close> by auto
@@ -2042,7 +2042,7 @@
by (auto intro!: SUP_least SUP_upper nn_integral_mono)
next
have "\<And>x. \<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (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: "\<And>x. incseq (g x)"
and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y"
and sup: "\<And>x. (SUP i:Y. f i x) = (SUP i. g x i)" by moura
@@ -2088,8 +2088,8 @@
assume "x \<in> {..<m}"
hence "x < m" by simp
have "g x n = f (I x) x" by(simp add: I)
- also have "\<dots> \<le> (SUP i:?Y. f i) x" unfolding SUP_def Sup_fun_def image_image
- using \<open>x \<in> {..<m}\<close> by(rule Sup_upper[OF imageI])
+ also have "\<dots> \<le> (SUP i:?Y. f i) x" unfolding Sup_fun_def image_image
+ using \<open>x \<in> {..<m}\<close> by (rule Sup_upper [OF imageI])
also have "\<dots> = f (I m') x" unfolding m' by simp
finally show "g x n \<le> 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:
--- 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 \<in> I"
with q(2) \<open>open I\<close> have "Sup ((\<lambda>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 ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)"
by simp
also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
--- 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: "(\<Union>i. A i) = space M" and
measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and
disjoint: "disjoint_family A"
- using sigma_finite_disjoint by auto
+ using sigma_finite_disjoint by blast
let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)"
have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
proof
@@ -560,7 +560,7 @@
proof (rule antisym)
show "?a \<le> (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)) \<le> ?a" unfolding SUP_def
+ show "(SUP i. emeasure M (?O i)) \<le> ?a"
proof (safe intro!: Sup_mono, unfold bex_simps)
fix i
have *: "(\<Union>(Q' ` {..i})) = ?O i" by auto
--- 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]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto
show "x \<in> ?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 "(\<lambda>k. M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> 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 \<subseteq> B \<and> closed U}. M (space M - U)) =
(INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U)"
- by (subst INF_image [of "\<lambda>u. space M - u", symmetric, unfolded comp_def])
- (rule INF_cong, auto simp add: sU intro!: INF_cong)
+ unfolding INF_image [of _ "\<lambda>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 \<subseteq> U \<and> open U}. emeasure M U) \<le> emeasure M (space M - B)" .
moreover have
@@ -335,8 +336,8 @@
also have "\<dots> = (SUP U:{U. B \<subseteq> U \<and> open U}. M (space M - U))"
by (rule SUP_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> closed K}. emeasure M K)"
- by (subst SUP_image [of "\<lambda>u. space M - u", symmetric, simplified comp_def])
- (rule SUP_cong, auto simp: sU)
+ unfolding SUP_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
+ by (rule SUP_cong) (auto simp add: sU)
also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
proof (safe intro!: antisym SUP_least)
fix K assume "closed K" "K \<subseteq> space M - B"
--- 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 \<union> b = (\<Union>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 \<inter> b = (\<Inter>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 \<Omega> M \<longleftrightarrow>
@@ -535,11 +535,14 @@
finally show ?thesis .
qed
-lemma sigma_sets_UNION: "countable B \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets X A) \<Longrightarrow> (\<Union>B) \<in> 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 \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets X A) \<Longrightarrow> (\<Union>B) \<in> 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 `\<lbrakk>B \<noteq> {}; countable B\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> disjoint (A ` I) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> UNION I A \<in> 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 \<in> generated_ring" and b: "b \<in> generated_ring"
@@ -873,7 +876,7 @@
lemma (in semiring_of_sets) generated_ring_INTER:
"finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> INTER I A \<in> 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 \<Omega> generated_ring"
@@ -898,6 +901,7 @@
fix a b assume "a \<in> Ca" "b \<in> Cb"
with Ca Cb Diff_cover[of a b] show "a - b \<in> ?R"
by (auto simp add: generated_ring_def)
+ (metis DiffI Diff_eq_empty_iff empty_iff)
next
show "disjoint ((\<lambda>a'. \<Inter>b'\<in>Cb. a' - b')`Ca)"
using Ca by (auto simp add: disjoint_def \<open>Cb \<noteq> {}\<close>)
@@ -935,7 +939,7 @@
done
lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
- by (simp add: SUP_def range_binaryset_eq)
+ by (simp add: range_binaryset_eq cong del: strong_SUP_cong)
subsubsection \<open>Closed CDI\<close>
--- 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
--- 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 \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> \<open>?b \<noteq> 0\<close> have q: "q = Fract ?a ?b"
by (simp add: eq_rat dvd_div_mult mult.commute [of a])
from \<open>b \<noteq> 0\<close> 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) \<Longrightarrow> 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]:
--- 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 \<bar>x\<bar> = (\<bar>of_int x\<bar> :: '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) \<le> 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 \<Rightarrow> real"})
#> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int \<Rightarrow> real"}))
\<close>
@@ -1309,22 +1306,9 @@
subsection \<open>Lemmas about powers\<close>
-(* used by Import/HOL/real.imp *)
lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
by simp
-text \<open>FIXME: no longer real-specific; rename and move elsewhere\<close>
-lemma realpow_Suc_le_self:
- fixes r :: "'a::linordered_semidom"
- shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
-by (insert power_decreasing [of 1 "Suc n" r], simp)
-
-text \<open>FIXME: no longer real-specific; rename and move elsewhere\<close>
-lemma realpow_minus_mult:
- fixes x :: "'a::monoid_mult"
- shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
-by (simp add: power_Suc power_commutes split add: nat_diff_split)
-
text \<open>FIXME: declare this [simp] for all types, or not at all\<close>
declare sum_squares_eq_zero_iff [simp] sum_power2_eq_zero_iff [simp]
--- 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 \<ge> 0 \<longleftrightarrow> a \<ge> b"
- by (simp add: le_diff_eq)
-
subsection \<open>Locale for additive functions\<close>
locale additive =
--- 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\<open>Converse inclusion requires some assumptions\<close>
lemma Image_INT_eq:
- "[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
+ "single_valued (r\<inverse>) \<Longrightarrow> A \<noteq> {} \<Longrightarrow> r `` INTER A B = (\<Inter>x\<in>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 \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
--- 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: "\<bar>x\<bar> = \<bar>y\<bar> \<longleftrightarrow> x = y \<or> x = -y"
by (auto simp add: abs_if split: split_if_asm)
+lemma sum_squares_ge_zero:
+ "0 \<le> x * x + y * y"
+ by (intro add_nonneg_nonneg zero_le_square)
+
+lemma not_sum_squares_lt_zero:
+ "\<not> 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 \<longleftrightarrow> a < 0"
unfolding sgn_if by auto
+lemma abs_sgn_eq:
+ "\<bar>sgn a\<bar> = (if a = 0 then 0 else 1)"
+ by (simp add: sgn_if)
+
lemma abs_dvd_iff [simp]: "\<bar>m\<bar> dvd k \<longleftrightarrow> m dvd k"
by (simp add: abs_if)
--- 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: "(\<Union>n::nat. \<Union>i\<in>{0..<n}. A i) = (\<Union>n. A n)"
+lemma UN_UN_finite_eq:
+ "(\<Union>n::nat. \<Union>i\<in>{0..<n}. A i) = (\<Union>n. A n)"
by (auto simp add: atLeast0LessThan)
-lemma UN_finite_subset: "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> C) \<Longrightarrow> (\<Union>n. A n) \<subseteq> C"
+lemma UN_finite_subset:
+ "(\<And>n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> C) \<Longrightarrow> (\<Union>n. A n) \<subseteq> C"
by (subst UN_UN_finite_eq [symmetric]) blast
lemma UN_finite2_subset:
- "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> (\<Union>i\<in>{0..<n+k}. B i)) \<Longrightarrow> (\<Union>n. A n) \<subseteq> (\<Union>n. B n)"
- apply (rule UN_finite_subset)
- apply (subst UN_UN_finite_eq [symmetric, of B])
- apply blast
- done
+ assumes "\<And>n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> (\<Union>i\<in>{0..<n + k}. B i)"
+ shows "(\<Union>n. A n) \<subseteq> (\<Union>n. B n)"
+proof (rule UN_finite_subset, rule)
+ fix n and a
+ from assms have "(\<Union>i\<in>{0..<n}. A i) \<subseteq> (\<Union>i\<in>{0..<n + k}. B i)" .
+ moreover assume "a \<in> (\<Union>i\<in>{0..<n}. A i)"
+ ultimately have "a \<in> (\<Union>i\<in>{0..<n + k}. B i)" by blast
+ then show "a \<in> (\<Union>i. B i)" by (auto simp add: UN_UN_finite_eq)
+qed
lemma UN_finite2_eq:
- "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) = (\<Union>i\<in>{0..<n+k}. B i)) \<Longrightarrow> (\<Union>n. A n) = (\<Union>n. B n)"
+ "(\<And>n::nat. (\<Union>i\<in>{0..<n}. A i) = (\<Union>i\<in>{0..<n + k}. B i)) \<Longrightarrow>
+ (\<Union>n. A n) = (\<Union>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 \<open>Cardinality\<close>
--- 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
--- 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]),
--- 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);
--- 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 ()
--- 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}
--- 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]);
--- 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)) =
--- 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 (\<Inter>i\<le>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 _ "\<Inter>i\<le>n. A i" for n] conjI open_INT)
+ apply fastforce
+ apply (intro exI [of _ "\<Inter>i\<le>n. A i" for n] conjI open_INT)
apply auto
done
qed
@@ -1464,7 +1465,7 @@
lemma continuous_on_open_UN:
"(\<And>s. s \<in> S \<Longrightarrow> open (A s)) \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> continuous_on (A s) f) \<Longrightarrow> continuous_on (\<Union>s\<in>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 \<Longrightarrow> open t \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on t f \<Longrightarrow> continuous_on (s \<union> t) f"
@@ -1689,7 +1690,7 @@
lemma compactE_image:
assumes "compact s" and "\<forall>t\<in>C. open (f t)" and "s \<subseteq> (\<Union>c\<in>C. f c)"
obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> (\<Union>c\<in>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]:
--- 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 "\<bar>of_int n\<bar> \<le> x \<Longrightarrow> n=0 \<or> x\<ge>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 "\<bar>of_int n\<bar> < x \<Longrightarrow> n=0 \<or> 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 \<in> \<real>" by (induction n) auto
+lemma fact_in_Reals: "fact n \<in> \<real>"
+ 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 -
--- 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. \<Union>((%(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 = (\<Union>(x, y)\<in>R. \<Union>(u, v)\<in>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]:
--- 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])
--- 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
--- 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)\<squnion>G) = F\<squnion>(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)\<squnion>G = extend h H ==> H = F\<squnion>(project h UNIV G)"
apply (drule_tac f = "project h UNIV" in arg_cong)
--- 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\<in>I ==> r i \<in> L|] ==> (\<Union>i\<in>I. r i) \<in> L"
-apply (unfold SUP_def)
apply (blast intro: Union_in_lattice)
done
lemma INT_in_lattice:
"[|lattice L; !!i. i\<in>I ==> r i \<in> L|] ==> (\<Inter>i\<in>I. r i) \<in> L"
-apply (unfold INF_def)
apply (blast intro: Inter_in_lattice)
done
--- 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 \<in> (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:
--- 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 \<in> I ==> F \<in> (A i) LeadsTo B) ==> F \<in> (\<Union>i \<in> 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 \<in> I ==> F \<in> (A i) LeadsTo (A' i))
==> F \<in> (\<Union>i \<in> I. A i) LeadsTo (\<Union>i \<in> I. A' i)"
-apply (simp only: Union_image_eq [symmetric])
apply (blast intro: LeadsTo_Union LeadsTo_weaken_R)
done
--- 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)) \<subseteq>
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
--- 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 \<Longrightarrow> safety_prop Y \<Longrightarrow> safety_prop (X \<inter> Y)"
- by (simp add: safety_prop_def) blast
+proof (clarsimp simp add: safety_prop_def)
+ fix G
+ assume "\<forall>G. Acts G \<subseteq> (\<Union>x\<in>X. Acts x) \<longrightarrow> G \<in> X"
+ then have X: "Acts G \<subseteq> (\<Union>x\<in>X. Acts x) \<Longrightarrow> G \<in> X" by blast
+ assume "\<forall>G. Acts G \<subseteq> (\<Union>x\<in>Y. Acts x) \<longrightarrow> G \<in> Y"
+ then have Y: "Acts G \<subseteq> (\<Union>x\<in>Y. Acts x) \<Longrightarrow> G \<in> Y" by blast
+ assume Acts: "Acts G \<subseteq> (\<Union>x\<in>X \<inter> Y. Acts x)"
+ with X and Y show "G \<in> X \<and> G \<in> Y" by auto
+qed
lemma safety_prop_INTER [simp]:
"(\<And>i. i \<in> I \<Longrightarrow> safety_prop (X i)) \<Longrightarrow> safety_prop (\<Inter>i\<in>I. X i)"
- by (simp add: safety_prop_def) blast
+proof (clarsimp simp add: safety_prop_def)
+ fix G and i
+ assume "\<And>i. i \<in> I \<Longrightarrow> \<bottom> \<in> X i \<and>
+ (\<forall>G. Acts G \<subseteq> (\<Union>x\<in>X i. Acts x) \<longrightarrow> G \<in> X i)"
+ then have *: "i \<in> I \<Longrightarrow> Acts G \<subseteq> (\<Union>x\<in>X i. Acts x) \<Longrightarrow> G \<in> X i"
+ by blast
+ assume "i \<in> I"
+ moreover assume "Acts G \<subseteq> (\<Union>j\<in>\<Inter>i\<in>I. X i. Acts j)"
+ ultimately have "Acts G \<subseteq> (\<Union>i\<in>X i. Acts i)"
+ by auto
+ with * `i \<in> I` show "G \<in> X i" by blast
+qed
lemma safety_prop_INTER1 [simp]:
"(\<And>i. safety_prop (X i)) \<Longrightarrow> safety_prop (\<Inter>i. X i)"
--- 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 \<in> I ==> F \<in> (A i) leadsTo B) ==> F \<in> (\<Union>i \<in> 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 \<in> I ==> F \<in> (A i) leadsTo (A' i))
==> F \<in> (\<Union>i \<in> I. A i) leadsTo (\<Union>i \<in> I. A' i)"
-apply (simp only: Union_image_eq [symmetric])
apply (blast intro: leadsTo_Union leadsTo_weaken_R)
done
--- 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]
--- 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 \<open>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 \<open>Interpreted theorems from the locales, involving defined terms.\<close>
--- 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 (\<Sum>i=0..<m. i * i * i * i) =
int m * (int m - 1) * (int(2 * m) - 1) *
(int(3 * m * m) - int(3 * m) - 1)"
- by (induct m) (simp_all add: int_mult)
+ by (induct m) (simp_all add: of_nat_mult)
lemma of_nat_sum_of_fourth_powers:
"30 * of_nat (\<Sum>i=0..<m. i * i * i * i) =
--- a/src/HOL/ex/Sqrt.thy Wed Feb 17 23:28:58 2016 +0100
+++ b/src/HOL/ex/Sqrt.thy Wed Feb 17 23:29:35 2016 +0100
@@ -77,7 +77,7 @@
with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
then have "p dvd n\<^sup>2" ..
with \<open>prime p\<close> 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 \<open>coprime m n\<close> have "p = 1" by simp
with p show False by simp
qed
--- 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 \<le>) (op \<le>)"
unfolding rel_fun_def ZN_def by simp