diff --git a/src/HOL/BNF_Wellorder_Constructions.thy b/src/HOL/BNF_Wellorder_Constructions.thy --- a/src/HOL/BNF_Wellorder_Constructions.thy +++ b/src/HOL/BNF_Wellorder_Constructions.thy @@ -1,1667 +1,1681 @@ (* Title: HOL/BNF_Wellorder_Constructions.thy Author: Andrei Popescu, TU Muenchen Copyright 2012 Constructions on wellorders as needed by bounded natural functors. *) section \Constructions on Wellorders as Needed by Bounded Natural Functors\ theory BNF_Wellorder_Constructions imports BNF_Wellorder_Embedding begin text \In this section, we study basic constructions on well-orders, such as restriction to a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders, and bounded square. We also define between well-orders the relations \ordLeq\, of being embedded (abbreviated \\o\), \ordLess\, of being strictly embedded (abbreviated \), and \ordIso\, of being isomorphic (abbreviated \=o\). We study the connections between these relations, order filters, and the aforementioned constructions. A main result of this section is that \ is well-founded.\ subsection \Restriction to a set\ abbreviation Restr :: "'a rel \ 'a set \ 'a rel" where "Restr r A \ r Int (A \ A)" lemma Restr_subset: "A \ B \ Restr (Restr r B) A = Restr r A" by blast lemma Restr_Field: "Restr r (Field r) = r" unfolding Field_def by auto lemma Refl_Restr: "Refl r \ Refl(Restr r A)" unfolding refl_on_def Field_def by auto lemma linear_order_on_Restr: "linear_order_on A r \ linear_order_on (A \ above r x) (Restr r (above r x))" by(simp add: order_on_defs refl_on_def trans_def antisym_def total_on_def)(safe; blast) lemma antisym_Restr: "antisym r \ antisym(Restr r A)" unfolding antisym_def Field_def by auto lemma Total_Restr: "Total r \ Total(Restr r A)" unfolding total_on_def Field_def by auto lemma total_on_imp_Total_Restr: "total_on A r \ Total (Restr r A)" by (auto simp: Field_def total_on_def) lemma trans_Restr: "trans r \ trans(Restr r A)" unfolding trans_def Field_def by blast lemma Preorder_Restr: "Preorder r \ Preorder(Restr r A)" unfolding preorder_on_def by (simp add: Refl_Restr trans_Restr) lemma Partial_order_Restr: "Partial_order r \ Partial_order(Restr r A)" unfolding partial_order_on_def by (simp add: Preorder_Restr antisym_Restr) lemma Linear_order_Restr: "Linear_order r \ Linear_order(Restr r A)" unfolding linear_order_on_def by (simp add: Partial_order_Restr Total_Restr) lemma Well_order_Restr: assumes "Well_order r" shows "Well_order(Restr r A)" proof- have "Restr r A - Id \ r - Id" using Restr_subset by blast hence "wf(Restr r A - Id)" using assms using well_order_on_def wf_subset by blast thus ?thesis using assms unfolding well_order_on_def by (simp add: Linear_order_Restr) qed lemma Field_Restr_subset: "Field(Restr r A) \ A" by (auto simp add: Field_def) lemma Refl_Field_Restr: "Refl r \ Field(Restr r A) = (Field r) Int A" unfolding refl_on_def Field_def by blast lemma Refl_Field_Restr2: "\Refl r; A \ Field r\ \ Field(Restr r A) = A" by (auto simp add: Refl_Field_Restr) lemma well_order_on_Restr: assumes WELL: "Well_order r" and SUB: "A \ Field r" shows "well_order_on A (Restr r A)" using assms using Well_order_Restr[of r A] Refl_Field_Restr2[of r A] order_on_defs[of "Field r" r] by auto subsection \Order filters versus restrictions and embeddings\ lemma Field_Restr_ofilter: "\Well_order r; wo_rel.ofilter r A\ \ Field(Restr r A) = A" by (auto simp add: wo_rel_def wo_rel.ofilter_def wo_rel.REFL Refl_Field_Restr2) lemma ofilter_Restr_under: assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \ A" shows "under (Restr r A) a = under r a" -using assms wo_rel_def -proof(auto simp add: wo_rel.ofilter_def under_def) - fix b assume *: "a \ A" and "(b,a) \ r" - hence "b \ under r a \ a \ Field r" - unfolding under_def using Field_def by fastforce - thus "b \ A" using * assms by (auto simp add: wo_rel_def wo_rel.ofilter_def) + unfolding wo_rel.ofilter_def under_def +proof + show "{b. (b, a) \ Restr r A} \ {b. (b, a) \ r}" + by auto +next + have "under r a \ A" + proof + fix x + assume *: "x \ under r a" + then have "a \ Field r" + unfolding under_def using Field_def by fastforce + then show "x \ A" using IN assms * + by (auto simp add: wo_rel_def wo_rel.ofilter_def) + qed + then show "{b. (b, a) \ r} \ {b. (b, a) \ Restr r A}" + unfolding under_def using assms by auto qed lemma ofilter_embed: assumes "Well_order r" shows "wo_rel.ofilter r A = (A \ Field r \ embed (Restr r A) r id)" proof assume *: "wo_rel.ofilter r A" show "A \ Field r \ embed (Restr r A) r id" - proof(unfold embed_def, auto) + unfolding embed_def + proof safe fix a assume "a \ A" thus "a \ Field r" using assms * by (auto simp add: wo_rel_def wo_rel.ofilter_def) next fix a assume "a \ Field (Restr r A)" - thus "bij_betw id (under (Restr r A) a) (under r a)" using assms * + thus "bij_betw id (under (Restr r A) a) (under r (id a))" using assms * by (simp add: ofilter_Restr_under Field_Restr_ofilter) qed next assume *: "A \ Field r \ embed (Restr r A) r id" hence "Field(Restr r A) \ Field r" using assms embed_Field[of "Restr r A" r id] id_def Well_order_Restr[of r] by auto {fix a assume "a \ A" hence "a \ Field(Restr r A)" using * assms by (simp add: order_on_defs Refl_Field_Restr2) hence "bij_betw id (under (Restr r A) a) (under r a)" using * unfolding embed_def by auto hence "under r a \ under (Restr r A) a" unfolding bij_betw_def by auto also have "\ \ Field(Restr r A)" by (simp add: under_Field) also have "\ \ A" by (simp add: Field_Restr_subset) finally have "under r a \ A" . } thus "wo_rel.ofilter r A" using assms * by (simp add: wo_rel_def wo_rel.ofilter_def) qed lemma ofilter_Restr_Int: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" shows "wo_rel.ofilter (Restr r B) (A Int B)" proof- let ?rB = "Restr r B" have Well: "wo_rel r" unfolding wo_rel_def using WELL . hence Refl: "Refl r" by (simp add: wo_rel.REFL) hence Field: "Field ?rB = Field r Int B" using Refl_Field_Restr by blast have WellB: "wo_rel ?rB \ Well_order ?rB" using WELL by (simp add: Well_order_Restr wo_rel_def) (* Main proof *) show ?thesis using WellB assms - proof(auto simp add: wo_rel.ofilter_def under_def) + unfolding wo_rel.ofilter_def under_def ofilter_def + proof safe fix a assume "a \ A" and *: "a \ B" hence "a \ Field r" using OFA Well by (auto simp add: wo_rel.ofilter_def) with * show "a \ Field ?rB" using Field by auto next fix a b assume "a \ A" and "(b,a) \ r" thus "b \ A" using Well OFA by (auto simp add: wo_rel.ofilter_def under_def) qed qed lemma ofilter_Restr_subset: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and SUB: "A \ B" shows "wo_rel.ofilter (Restr r B) A" proof- have "A Int B = A" using SUB by blast thus ?thesis using assms ofilter_Restr_Int[of r A B] by auto qed lemma ofilter_subset_embed: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" shows "(A \ B) = (embed (Restr r A) (Restr r B) id)" proof- let ?rA = "Restr r A" let ?rB = "Restr r B" have Well: "wo_rel r" unfolding wo_rel_def using WELL . hence Refl: "Refl r" by (simp add: wo_rel.REFL) hence FieldA: "Field ?rA = Field r Int A" using Refl_Field_Restr by blast have FieldB: "Field ?rB = Field r Int B" using Refl Refl_Field_Restr by blast have WellA: "wo_rel ?rA \ Well_order ?rA" using WELL by (simp add: Well_order_Restr wo_rel_def) have WellB: "wo_rel ?rB \ Well_order ?rB" using WELL by (simp add: Well_order_Restr wo_rel_def) (* Main proof *) show ?thesis proof assume *: "A \ B" hence "wo_rel.ofilter (Restr r B) A" using assms by (simp add: ofilter_Restr_subset) hence "embed (Restr ?rB A) (Restr r B) id" using WellB ofilter_embed[of "?rB" A] by auto thus "embed (Restr r A) (Restr r B) id" using * by (simp add: Restr_subset) next assume *: "embed (Restr r A) (Restr r B) id" {fix a assume **: "a \ A" hence "a \ Field r" using Well OFA by (auto simp add: wo_rel.ofilter_def) with ** FieldA have "a \ Field ?rA" by auto hence "a \ Field ?rB" using * WellA embed_Field[of ?rA ?rB id] by auto hence "a \ B" using FieldB by auto } thus "A \ B" by blast qed qed lemma ofilter_subset_embedS_iso: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" shows "((A < B) = (embedS (Restr r A) (Restr r B) id)) \ ((A = B) = (iso (Restr r A) (Restr r B) id))" proof- let ?rA = "Restr r A" let ?rB = "Restr r B" have Well: "wo_rel r" unfolding wo_rel_def using WELL . hence Refl: "Refl r" by (simp add: wo_rel.REFL) hence "Field ?rA = Field r Int A" using Refl_Field_Restr by blast hence FieldA: "Field ?rA = A" using OFA Well by (auto simp add: wo_rel.ofilter_def) have "Field ?rB = Field r Int B" using Refl Refl_Field_Restr by blast hence FieldB: "Field ?rB = B" using OFB Well by (auto simp add: wo_rel.ofilter_def) (* Main proof *) show ?thesis unfolding embedS_def iso_def using assms ofilter_subset_embed[of r A B] FieldA FieldB bij_betw_id_iff[of A B] by auto qed lemma ofilter_subset_embedS: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" shows "(A < B) = embedS (Restr r A) (Restr r B) id" using assms by (simp add: ofilter_subset_embedS_iso) lemma embed_implies_iso_Restr: assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f" shows "iso r' (Restr r (f ` (Field r'))) f" proof- let ?A' = "Field r'" let ?r'' = "Restr r (f ` ?A')" have 0: "Well_order ?r''" using WELL Well_order_Restr by blast have 1: "wo_rel.ofilter r (f ` ?A')" using assms embed_Field_ofilter by blast hence "Field ?r'' = f ` (Field r')" using WELL Field_Restr_ofilter by blast hence "bij_betw f ?A' (Field ?r'')" using EMB embed_inj_on WELL' unfolding bij_betw_def by blast moreover {have "\a b. (a,b) \ r' \ a \ Field r' \ b \ Field r'" unfolding Field_def by auto hence "compat r' ?r'' f" using assms embed_iff_compat_inj_on_ofilter unfolding compat_def by blast } ultimately show ?thesis using WELL' 0 iso_iff3 by blast qed subsection \The strict inclusion on proper ofilters is well-founded\ definition ofilterIncl :: "'a rel \ 'a set rel" where "ofilterIncl r \ {(A,B). wo_rel.ofilter r A \ A \ Field r \ wo_rel.ofilter r B \ B \ Field r \ A < B}" lemma wf_ofilterIncl: assumes WELL: "Well_order r" shows "wf(ofilterIncl r)" proof- have Well: "wo_rel r" using WELL by (simp add: wo_rel_def) hence Lo: "Linear_order r" by (simp add: wo_rel.LIN) let ?h = "(\ A. wo_rel.suc r A)" let ?rS = "r - Id" have "wf ?rS" using WELL by (simp add: order_on_defs) moreover have "compat (ofilterIncl r) ?rS ?h" proof(unfold compat_def ofilterIncl_def, intro allI impI, simp, elim conjE) fix A B assume *: "wo_rel.ofilter r A" "A \ Field r" and **: "wo_rel.ofilter r B" "B \ Field r" and ***: "A < B" then obtain a and b where 0: "a \ Field r \ b \ Field r" and 1: "A = underS r a \ B = underS r b" using Well by (auto simp add: wo_rel.ofilter_underS_Field) hence "a \ b" using *** by auto moreover have "(a,b) \ r" using 0 1 Lo *** by (auto simp add: underS_incl_iff) moreover have "a = wo_rel.suc r A \ b = wo_rel.suc r B" using Well 0 1 by (simp add: wo_rel.suc_underS) ultimately show "(wo_rel.suc r A, wo_rel.suc r B) \ r \ wo_rel.suc r A \ wo_rel.suc r B" by simp qed ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf) qed subsection \Ordering the well-orders by existence of embeddings\ text \We define three relations between well-orders: \begin{itemize} \item \ordLeq\, of being embedded (abbreviated \\o\); \item \ordLess\, of being strictly embedded (abbreviated \); \item \ordIso\, of being isomorphic (abbreviated \=o\). \end{itemize} % The prefix "ord" and the index "o" in these names stand for "ordinal-like". These relations shall be proved to be inter-connected in a similar fashion as the trio \\\, \<\, \=\ associated to a total order on a set. \ definition ordLeq :: "('a rel * 'a' rel) set" where "ordLeq = {(r,r'). Well_order r \ Well_order r' \ (\f. embed r r' f)}" abbreviation ordLeq2 :: "'a rel \ 'a' rel \ bool" (infix "<=o" 50) where "r <=o r' \ (r,r') \ ordLeq" abbreviation ordLeq3 :: "'a rel \ 'a' rel \ bool" (infix "\o" 50) where "r \o r' \ r <=o r'" definition ordLess :: "('a rel * 'a' rel) set" where "ordLess = {(r,r'). Well_order r \ Well_order r' \ (\f. embedS r r' f)}" abbreviation ordLess2 :: "'a rel \ 'a' rel \ bool" (infix " (r,r') \ ordLess" definition ordIso :: "('a rel * 'a' rel) set" where "ordIso = {(r,r'). Well_order r \ Well_order r' \ (\f. iso r r' f)}" abbreviation ordIso2 :: "'a rel \ 'a' rel \ bool" (infix "=o" 50) where "r =o r' \ (r,r') \ ordIso" lemmas ordRels_def = ordLeq_def ordLess_def ordIso_def lemma ordLeq_Well_order_simp: assumes "r \o r'" shows "Well_order r \ Well_order r'" using assms unfolding ordLeq_def by simp text\Notice that the relations \\o\, \, \=o\ connect well-orders on potentially {\em distinct} types. However, some of the lemmas below, including the next one, restrict implicitly the type of these relations to \(('a rel) * ('a rel)) set\ , i.e., to \'a rel rel\.\ lemma ordLeq_reflexive: "Well_order r \ r \o r" unfolding ordLeq_def using id_embed[of r] by blast lemma ordLeq_transitive[trans]: assumes *: "r \o r'" and **: "r' \o r''" shows "r \o r''" proof- obtain f and f' where 1: "Well_order r \ Well_order r' \ Well_order r''" and "embed r r' f" and "embed r' r'' f'" using * ** unfolding ordLeq_def by blast hence "embed r r'' (f' \ f)" using comp_embed[of r r' f r'' f'] by auto thus "r \o r''" unfolding ordLeq_def using 1 by auto qed lemma ordLeq_total: "\Well_order r; Well_order r'\ \ r \o r' \ r' \o r" unfolding ordLeq_def using wellorders_totally_ordered by blast lemma ordIso_reflexive: "Well_order r \ r =o r" unfolding ordIso_def using id_iso[of r] by blast lemma ordIso_transitive[trans]: assumes *: "r =o r'" and **: "r' =o r''" shows "r =o r''" proof- obtain f and f' where 1: "Well_order r \ Well_order r' \ Well_order r''" and "iso r r' f" and 3: "iso r' r'' f'" using * ** unfolding ordIso_def by auto hence "iso r r'' (f' \ f)" using comp_iso[of r r' f r'' f'] by auto thus "r =o r''" unfolding ordIso_def using 1 by auto qed lemma ordIso_symmetric: assumes *: "r =o r'" shows "r' =o r" proof- obtain f where 1: "Well_order r \ Well_order r'" and 2: "embed r r' f \ bij_betw f (Field r) (Field r')" using * by (auto simp add: ordIso_def iso_def) let ?f' = "inv_into (Field r) f" have "embed r' r ?f' \ bij_betw ?f' (Field r') (Field r)" using 1 2 by (simp add: bij_betw_inv_into inv_into_Field_embed_bij_betw) thus "r' =o r" unfolding ordIso_def using 1 by (auto simp add: iso_def) qed lemma ordLeq_ordLess_trans[trans]: assumes "r \o r'" and " r' Well_order r''" using assms unfolding ordLeq_def ordLess_def by auto thus ?thesis using assms unfolding ordLeq_def ordLess_def using embed_comp_embedS by blast qed lemma ordLess_ordLeq_trans[trans]: assumes "r o r''" shows "r Well_order r''" using assms unfolding ordLeq_def ordLess_def by auto thus ?thesis using assms unfolding ordLeq_def ordLess_def using embedS_comp_embed by blast qed lemma ordLeq_ordIso_trans[trans]: assumes "r \o r'" and " r' =o r''" shows "r \o r''" proof- have "Well_order r \ Well_order r''" using assms unfolding ordLeq_def ordIso_def by auto thus ?thesis using assms unfolding ordLeq_def ordIso_def using embed_comp_iso by blast qed lemma ordIso_ordLeq_trans[trans]: assumes "r =o r'" and " r' \o r''" shows "r \o r''" proof- have "Well_order r \ Well_order r''" using assms unfolding ordLeq_def ordIso_def by auto thus ?thesis using assms unfolding ordLeq_def ordIso_def using iso_comp_embed by blast qed lemma ordLess_ordIso_trans[trans]: assumes "r Well_order r''" using assms unfolding ordLess_def ordIso_def by auto thus ?thesis using assms unfolding ordLess_def ordIso_def using embedS_comp_iso by blast qed lemma ordIso_ordLess_trans[trans]: assumes "r =o r'" and " r' Well_order r''" using assms unfolding ordLess_def ordIso_def by auto thus ?thesis using assms unfolding ordLess_def ordIso_def using iso_comp_embedS by blast qed lemma ordLess_not_embed: assumes "r (\f'. embed r' r f')" proof- obtain f where 1: "Well_order r \ Well_order r'" and 2: "embed r r' f" and 3: " \ bij_betw f (Field r) (Field r')" using assms unfolding ordLess_def by (auto simp add: embedS_def) {fix f' assume *: "embed r' r f'" hence "bij_betw f (Field r) (Field r')" using 1 2 by (simp add: embed_bothWays_Field_bij_betw) with 3 have False by contradiction } thus ?thesis by blast qed lemma ordLess_Field: assumes OL: "r1 (f`(Field r1) = Field r2)" proof- let ?A1 = "Field r1" let ?A2 = "Field r2" obtain g where 0: "Well_order r1 \ Well_order r2" and 1: "embed r1 r2 g \ \(bij_betw g ?A1 ?A2)" using OL unfolding ordLess_def by (auto simp add: embedS_def) hence "\a \ ?A1. f a = g a" using 0 EMB embed_unique[of r1] by auto hence "\(bij_betw f ?A1 ?A2)" using 1 bij_betw_cong[of ?A1] by blast moreover have "inj_on f ?A1" using EMB 0 by (simp add: embed_inj_on) ultimately show ?thesis by (simp add: bij_betw_def) qed lemma ordLess_iff: "r Well_order r' \ \(\f'. embed r' r f'))" proof assume *: "r (\f'. embed r' r f')" using ordLess_not_embed[of r r'] by simp with * show "Well_order r \ Well_order r' \ \ (\f'. embed r' r f')" unfolding ordLess_def by auto next assume *: "Well_order r \ Well_order r' \ \ (\f'. embed r' r f')" then obtain f where 1: "embed r r' f" using wellorders_totally_ordered[of r r'] by blast moreover {assume "bij_betw f (Field r) (Field r')" with * 1 have "embed r' r (inv_into (Field r) f) " using inv_into_Field_embed_bij_betw[of r r' f] by auto with * have False by blast } ultimately show "(r,r') \ ordLess" unfolding ordLess_def using * by (fastforce simp add: embedS_def) qed lemma ordLess_irreflexive: "\ r \(\f. embed r r f)" unfolding ordLess_iff .. moreover have "embed r r id" using id_embed[of r] . ultimately show False by blast qed lemma ordLeq_iff_ordLess_or_ordIso: "r \o r' = (r r =o r')" unfolding ordRels_def embedS_defs iso_defs by blast lemma ordIso_iff_ordLeq: "(r =o r') = (r \o r' \ r' \o r)" proof assume "r =o r'" then obtain f where 1: "Well_order r \ Well_order r' \ embed r r' f \ bij_betw f (Field r) (Field r')" unfolding ordIso_def iso_defs by auto hence "embed r r' f \ embed r' r (inv_into (Field r) f)" by (simp add: inv_into_Field_embed_bij_betw) thus "r \o r' \ r' \o r" unfolding ordLeq_def using 1 by auto next assume "r \o r' \ r' \o r" then obtain f and g where 1: "Well_order r \ Well_order r' \ embed r r' f \ embed r' r g" unfolding ordLeq_def by auto hence "iso r r' f" by (auto simp add: embed_bothWays_iso) thus "r =o r'" unfolding ordIso_def using 1 by auto qed lemma not_ordLess_ordLeq: "r \ r' \o r" using ordLess_ordLeq_trans ordLess_irreflexive by blast lemma not_ordLeq_ordLess: "r \o r' \ \ r' r' \o r" proof- have "r \o r' \ r' \o r" using assms by (simp add: ordLeq_total) moreover {assume "\ r r \o r'" hence "r =o r'" using ordLeq_iff_ordLess_or_ordIso by blast hence "r' \o r" using ordIso_symmetric ordIso_iff_ordLeq by blast } ultimately show ?thesis by blast qed lemma not_ordLess_ordIso: "r \ r =o r'" using ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast lemma not_ordLeq_iff_ordLess: assumes WELL: "Well_order r" and WELL': "Well_order r'" shows "(\ r' \o r) = (r r' o r')" using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast lemma ordLess_transitive[trans]: "\r \ r r \o r'" using ordIso_iff_ordLeq by blast lemma ordLess_imp_ordLeq: "r r \o r'" using ordLeq_iff_ordLess_or_ordIso by blast lemma ofilter_subset_ordLeq: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" shows "(A \ B) = (Restr r A \o Restr r B)" proof assume "A \ B" thus "Restr r A \o Restr r B" unfolding ordLeq_def using assms Well_order_Restr Well_order_Restr ofilter_subset_embed by blast next assume *: "Restr r A \o Restr r B" then obtain f where "embed (Restr r A) (Restr r B) f" unfolding ordLeq_def by blast {assume "B < A" hence "Restr r B B" using OFA OFB WELL wo_rel_def[of r] wo_rel.ofilter_linord[of r A B] by blast qed lemma ofilter_subset_ordLess: assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B" shows "(A < B) = (Restr r A Well_order ?rB" using WELL Well_order_Restr by blast have "(A < B) = (\ B \ A)" using assms wo_rel_def wo_rel.ofilter_linord[of r A B] by blast also have "\ = (\ Restr r B \o Restr r A)" using assms ofilter_subset_ordLeq by blast also have "\ = (Restr r A Well_order r; wo_rel.ofilter r A\ \ (A < Field r) = (Restr r A {}" shows "Restr r (underS r a) (ofilterIncl r3)" proof- have OL13: "r1 Well_order r2 \ Well_order r3" and 1: "embed r1 r2 f12 \ \(bij_betw f12 ?A1 ?A2)" and 2: "embed r2 r3 g23 \ \(bij_betw g23 ?A2 ?A3)" using OL12 OL23 by (auto simp add: ordLess_def embedS_def) hence "\a \ ?A2. f23 a = g23 a" using EMB23 embed_unique[of r2 r3] by blast hence 3: "\(bij_betw f23 ?A2 ?A3)" using 2 bij_betw_cong[of ?A2 f23 g23] by blast (* *) have 4: "wo_rel.ofilter r2 (f12 ` ?A1) \ f12 ` ?A1 \ ?A2" using 0 1 OL12 by (simp add: embed_Field_ofilter ordLess_Field) have 5: "wo_rel.ofilter r3 (f23 ` ?A2) \ f23 ` ?A2 \ ?A3" using 0 EMB23 OL23 by (simp add: embed_Field_ofilter ordLess_Field) have 6: "wo_rel.ofilter r3 (f13 ` ?A1) \ f13 ` ?A1 \ ?A3" using 0 EMB13 OL13 by (simp add: embed_Field_ofilter ordLess_Field) (* *) have "f12 ` ?A1 < ?A2" using 0 4 by (auto simp add: wo_rel_def wo_rel.ofilter_def) moreover have "inj_on f23 ?A2" using EMB23 0 by (simp add: wo_rel_def embed_inj_on) ultimately have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset) moreover {have "embed r1 r3 (f23 \ f12)" using 1 EMB23 0 by (auto simp add: comp_embed) hence "\a \ ?A1. f23(f12 a) = f13 a" using EMB13 0 embed_unique[of r1 r3 "f23 \ f12" f13] by auto hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force } ultimately have "f13 ` ?A1 < f23 ` ?A2" by simp (* *) with 5 6 show ?thesis unfolding ofilterIncl_def by auto qed lemma ordLess_iff_ordIso_Restr: assumes WELL: "Well_order r" and WELL': "Well_order r'" shows "(r' a \ Field r. r' =o Restr r (underS r a))" -proof(auto) +proof safe fix a assume *: "a \ Field r" and **: "r' =o Restr r (underS r a)" hence "Restr r (underS r a) Well_order r'" and 2: "embed r' r f \ f ` (Field r') \ Field r" unfolding ordLess_def embedS_def[abs_def] bij_betw_def using embed_inj_on by blast hence "wo_rel.ofilter r (f ` (Field r'))" using embed_Field_ofilter by blast then obtain a where 3: "a \ Field r" and 4: "underS r a = f ` (Field r')" using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def) have "iso r' (Restr r (f ` (Field r'))) f" using embed_implies_iso_Restr 2 assms by blast moreover have "Well_order (Restr r (f ` (Field r')))" using WELL Well_order_Restr by blast ultimately have "r' =o Restr r (f ` (Field r'))" using WELL' unfolding ordIso_def by auto hence "r' =o Restr r (underS r a)" using 4 by auto thus "\a \ Field r. r' =o Restr r (underS r a)" using 3 by auto qed lemma internalize_ordLess: "(r' p. Field p < Field r \ r' =o p \ p Well_order r'" unfolding ordLess_def by auto with * obtain a where 1: "a \ Field r" and 2: "r' =o Restr r (underS r a)" using ordLess_iff_ordIso_Restr by blast let ?p = "Restr r (underS r a)" have "wo_rel.ofilter r (underS r a)" using 0 by (simp add: wo_rel_def wo_rel.underS_ofilter) hence "Field ?p = underS r a" using 0 Field_Restr_ofilter by blast hence "Field ?p < Field r" using underS_Field2 1 by fast moreover have "?p p. Field p < Field r \ r' =o p \ p p. Field p < Field r \ r' =o p \ p o r) = (\p. Field p \ Field r \ r' =o p \ p \o r)" proof assume *: "r' \o r" moreover {assume "r' r' =o p \ p p. Field p \ Field r \ r' =o p \ p \o r" using ordLeq_iff_ordLess_or_ordIso by blast } moreover have "r \o r" using * ordLeq_def ordLeq_reflexive by blast ultimately show "\p. Field p \ Field r \ r' =o p \ p \o r" using ordLeq_iff_ordLess_or_ordIso by blast next assume "\p. Field p \ Field r \ r' =o p \ p \o r" thus "r' \o r" using ordIso_ordLeq_trans by blast qed lemma ordLeq_iff_ordLess_Restr: assumes WELL: "Well_order r" and WELL': "Well_order r'" shows "(r \o r') = (\a \ Field r. Restr r (underS r a) o r'" fix a assume "a \ Field r" hence "Restr r (underS r a) a \ Field r. Restr r (underS r a) Field r \ r' =o Restr r (underS r a)" using assms ordLess_iff_ordIso_Restr by blast hence False using * not_ordLess_ordIso ordIso_symmetric by blast } thus "r \o r'" using ordLess_or_ordLeq assms by blast qed lemma finite_ordLess_infinite: assumes WELL: "Well_order r" and WELL': "Well_order r'" and FIN: "finite(Field r)" and INF: "\finite(Field r')" shows "r o r" then obtain h where "inj_on h (Field r') \ h ` (Field r') \ Field r" unfolding ordLeq_def using assms embed_inj_on embed_Field by blast hence False using finite_imageD finite_subset FIN INF by blast } thus ?thesis using WELL WELL' ordLess_or_ordLeq by blast qed lemma finite_well_order_on_ordIso: assumes FIN: "finite A" and WELL: "well_order_on A r" and WELL': "well_order_on A r'" shows "r =o r'" proof- have 0: "Well_order r \ Well_order r' \ Field r = A \ Field r' = A" using assms well_order_on_Well_order by blast moreover have "\r r'. well_order_on A r \ well_order_on A r' \ r \o r' \ r =o r'" proof(clarify) fix r r' assume *: "well_order_on A r" and **: "well_order_on A r'" have 2: "Well_order r \ Well_order r' \ Field r = A \ Field r' = A" using * ** well_order_on_Well_order by blast assume "r \o r'" then obtain f where 1: "embed r r' f" and "inj_on f A \ f ` A \ A" unfolding ordLeq_def using 2 embed_inj_on embed_Field by blast hence "bij_betw f A A" unfolding bij_betw_def using FIN endo_inj_surj by blast thus "r =o r'" unfolding ordIso_def iso_def[abs_def] using 1 2 by auto qed ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast qed subsection\\ is well-founded\ text \Of course, it only makes sense to state that the \ is well-founded on the restricted type \'a rel rel\. We prove this by first showing that, for any set of well-orders all embedded in a fixed well-order, the function mapping each well-order in the set to an order filter of the fixed well-order is compatible w.r.t. to \ versus {\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded.\ definition ord_to_filter :: "'a rel \ 'a rel \ 'a set" where "ord_to_filter r0 r \ (SOME f. embed r r0 f) ` (Field r)" lemma ord_to_filter_compat: "compat (ordLess Int (ordLess\``{r0} \ ordLess\``{r0})) (ofilterIncl r0) (ord_to_filter r0)" proof(unfold compat_def ord_to_filter_def, clarify) fix r1::"'a rel" and r2::"'a rel" let ?A1 = "Field r1" let ?A2 ="Field r2" let ?A0 ="Field r0" let ?phi10 = "\ f10. embed r1 r0 f10" let ?f10 = "SOME f. ?phi10 f" let ?phi20 = "\ f20. embed r2 r0 f20" let ?f20 = "SOME f. ?phi20 f" assume *: "r1 f. ?phi10 f) \ (\f. ?phi20 f)" by (auto simp add: ordLess_def embedS_def) hence "?phi10 ?f10 \ ?phi20 ?f20" by (auto simp add: someI_ex) thus "(?f10 ` ?A1, ?f20 ` ?A2) \ ofilterIncl r0" using * ** by (simp add: embed_ordLess_ofilterIncl) qed theorem wf_ordLess: "wf ordLess" proof- {fix r0 :: "('a \ 'a) set" (* need to annotate here!*) let ?ordLess = "ordLess::('d rel * 'd rel) set" let ?R = "?ordLess Int (?ordLess\``{r0} \ ?ordLess\``{r0})" {assume Case1: "Well_order r0" hence "wf ?R" using wf_ofilterIncl[of r0] compat_wf[of ?R "ofilterIncl r0" "ord_to_filter r0"] ord_to_filter_compat[of r0] by auto } moreover {assume Case2: "\ Well_order r0" hence "?R = {}" unfolding ordLess_def by auto hence "wf ?R" using wf_empty by simp } ultimately have "wf ?R" by blast } thus ?thesis by (simp add: trans_wf_iff ordLess_trans) qed corollary exists_minim_Well_order: assumes NE: "R \ {}" and WELL: "\r \ R. Well_order r" shows "\r \ R. \r' \ R. r \o r'" proof- obtain r where "r \ R \ (\r' \ R. \ r' Copy via direct images\ text\The direct image operator is the dual of the inverse image operator \inv_image\ from \Relation.thy\. It is useful for transporting a well-order between different types.\ definition dir_image :: "'a rel \ ('a \ 'a') \ 'a' rel" where "dir_image r f = {(f a, f b)| a b. (a,b) \ r}" lemma dir_image_Field: "Field(dir_image r f) = f ` (Field r)" unfolding dir_image_def Field_def Range_def Domain_def by fast lemma dir_image_minus_Id: "inj_on f (Field r) \ (dir_image r f) - Id = dir_image (r - Id) f" unfolding inj_on_def Field_def dir_image_def by auto lemma Refl_dir_image: assumes "Refl r" shows "Refl(dir_image r f)" proof- {fix a' b' assume "(a',b') \ dir_image r f" then obtain a b where 1: "a' = f a \ b' = f b \ (a,b) \ r" unfolding dir_image_def by blast hence "a \ Field r \ b \ Field r" using Field_def by fastforce hence "(a,a) \ r \ (b,b) \ r" using assms by (simp add: refl_on_def) with 1 have "(a',a') \ dir_image r f \ (b',b') \ dir_image r f" unfolding dir_image_def by auto } thus ?thesis by(unfold refl_on_def Field_def Domain_def Range_def, auto) qed lemma trans_dir_image: assumes TRANS: "trans r" and INJ: "inj_on f (Field r)" shows "trans(dir_image r f)" -proof(unfold trans_def, auto) +unfolding trans_def +proof safe fix a' b' c' assume "(a',b') \ dir_image r f" "(b',c') \ dir_image r f" then obtain a b1 b2 c where 1: "a' = f a \ b' = f b1 \ b' = f b2 \ c' = f c" and 2: "(a,b1) \ r \ (b2,c) \ r" unfolding dir_image_def by blast hence "b1 \ Field r \ b2 \ Field r" unfolding Field_def by auto hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto hence "(a,c) \ r" using 2 TRANS unfolding trans_def by blast thus "(a',c') \ dir_image r f" unfolding dir_image_def using 1 by auto qed lemma Preorder_dir_image: "\Preorder r; inj_on f (Field r)\ \ Preorder (dir_image r f)" by (simp add: preorder_on_def Refl_dir_image trans_dir_image) lemma antisym_dir_image: assumes AN: "antisym r" and INJ: "inj_on f (Field r)" shows "antisym(dir_image r f)" -proof(unfold antisym_def, auto) +unfolding antisym_def +proof safe fix a' b' assume "(a',b') \ dir_image r f" "(b',a') \ dir_image r f" then obtain a1 b1 a2 b2 where 1: "a' = f a1 \ a' = f a2 \ b' = f b1 \ b' = f b2" and 2: "(a1,b1) \ r \ (b2,a2) \ r " and 3: "{a1,a2,b1,b2} \ Field r" unfolding dir_image_def Field_def by blast hence "a1 = a2 \ b1 = b2" using INJ unfolding inj_on_def by auto hence "a1 = b2" using 2 AN unfolding antisym_def by auto thus "a' = b'" using 1 by auto qed lemma Partial_order_dir_image: "\Partial_order r; inj_on f (Field r)\ \ Partial_order (dir_image r f)" by (simp add: partial_order_on_def Preorder_dir_image antisym_dir_image) lemma Total_dir_image: assumes TOT: "Total r" and INJ: "inj_on f (Field r)" shows "Total(dir_image r f)" proof(unfold total_on_def, intro ballI impI) fix a' b' assume "a' \ Field (dir_image r f)" "b' \ Field (dir_image r f)" then obtain a and b where 1: "a \ Field r \ b \ Field r \ f a = a' \ f b = b'" unfolding dir_image_Field[of r f] by blast moreover assume "a' \ b'" ultimately have "a \ b" using INJ unfolding inj_on_def by auto hence "(a,b) \ r \ (b,a) \ r" using 1 TOT unfolding total_on_def by auto thus "(a',b') \ dir_image r f \ (b',a') \ dir_image r f" using 1 unfolding dir_image_def by auto qed lemma Linear_order_dir_image: "\Linear_order r; inj_on f (Field r)\ \ Linear_order (dir_image r f)" by (simp add: linear_order_on_def Partial_order_dir_image Total_dir_image) lemma wf_dir_image: assumes WF: "wf r" and INJ: "inj_on f (Field r)" shows "wf(dir_image r f)" proof(unfold wf_eq_minimal2, intro allI impI, elim conjE) fix A'::"'b set" assume SUB: "A' \ Field(dir_image r f)" and NE: "A' \ {}" obtain A where A_def: "A = {a \ Field r. f a \ A'}" by blast have "A \ {} \ A \ Field r" using A_def SUB NE by (auto simp: dir_image_Field) then obtain a where 1: "a \ A \ (\b \ A. (b,a) \ r)" using spec[OF WF[unfolded wf_eq_minimal2], of A] by blast have "\b' \ A'. (b',f a) \ dir_image r f" proof(clarify) fix b' assume *: "b' \ A'" and **: "(b',f a) \ dir_image r f" obtain b1 a1 where 2: "b' = f b1 \ f a = f a1" and 3: "(b1,a1) \ r \ {a1,b1} \ Field r" using ** unfolding dir_image_def Field_def by blast hence "a = a1" using 1 A_def INJ unfolding inj_on_def by auto hence "b1 \ A \ (b1,a) \ r" using 2 3 A_def * by auto with 1 show False by auto qed thus "\a'\A'. \b'\A'. (b', a') \ dir_image r f" using A_def 1 by blast qed lemma Well_order_dir_image: "\Well_order r; inj_on f (Field r)\ \ Well_order (dir_image r f)" unfolding well_order_on_def using Linear_order_dir_image[of r f] wf_dir_image[of "r - Id" f] dir_image_minus_Id[of f r] subset_inj_on[of f "Field r" "Field(r - Id)"] mono_Field[of "r - Id" r] by auto lemma dir_image_bij_betw: "\inj_on f (Field r)\ \ bij_betw f (Field r) (Field (dir_image r f))" unfolding bij_betw_def by (simp add: dir_image_Field order_on_defs) lemma dir_image_compat: "compat r (dir_image r f) f" unfolding compat_def dir_image_def by auto lemma dir_image_iso: "\Well_order r; inj_on f (Field r)\ \ iso r (dir_image r f) f" using iso_iff3 dir_image_compat dir_image_bij_betw Well_order_dir_image by blast lemma dir_image_ordIso: "\Well_order r; inj_on f (Field r)\ \ r =o dir_image r f" unfolding ordIso_def using dir_image_iso Well_order_dir_image by blast lemma Well_order_iso_copy: assumes WELL: "well_order_on A r" and BIJ: "bij_betw f A A'" shows "\r'. well_order_on A' r' \ r =o r'" proof- let ?r' = "dir_image r f" have 1: "A = Field r \ Well_order r" using WELL well_order_on_Well_order by blast hence 2: "iso r ?r' f" using dir_image_iso using BIJ unfolding bij_betw_def by auto hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast hence "Field ?r' = A'" using 1 BIJ unfolding bij_betw_def by auto moreover have "Well_order ?r'" using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast ultimately show ?thesis unfolding ordIso_def using 1 2 by blast qed subsection \Bounded square\ text\This construction essentially defines, for an order relation \r\, a lexicographic order \bsqr r\ on \(Field r) \ (Field r)\, applying the following criteria (in this order): \begin{itemize} \item compare the maximums; \item compare the first components; \item compare the second components. \end{itemize} % The only application of this construction that we are aware of is at proving that the square of an infinite set has the same cardinal as that set. The essential property required there (and which is ensured by this construction) is that any proper order filter of the product order is included in a rectangle, i.e., in a product of proper filters on the original relation (assumed to be a well-order).\ definition bsqr :: "'a rel => ('a * 'a)rel" where "bsqr r = {((a1,a2),(b1,b2)). {a1,a2,b1,b2} \ Field r \ (a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id )}" lemma Field_bsqr: "Field (bsqr r) = Field r \ Field r" proof show "Field (bsqr r) \ Field r \ Field r" proof- {fix a1 a2 assume "(a1,a2) \ Field (bsqr r)" moreover have "\ b1 b2. ((a1,a2),(b1,b2)) \ bsqr r \ ((b1,b2),(a1,a2)) \ bsqr r \ a1 \ Field r \ a2 \ Field r" unfolding bsqr_def by auto ultimately have "a1 \ Field r \ a2 \ Field r" unfolding Field_def by auto } thus ?thesis unfolding Field_def by force qed next show "Field r \ Field r \ Field (bsqr r)" - proof(auto) + proof safe fix a1 a2 assume "a1 \ Field r" and "a2 \ Field r" hence "((a1,a2),(a1,a2)) \ bsqr r" unfolding bsqr_def by blast thus "(a1,a2) \ Field (bsqr r)" unfolding Field_def by auto qed qed lemma bsqr_Refl: "Refl(bsqr r)" by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def) lemma bsqr_Trans: assumes "Well_order r" shows "trans (bsqr r)" -proof(unfold trans_def, auto) +unfolding trans_def +proof safe (* Preliminary facts *) have Well: "wo_rel r" using assms wo_rel_def by auto hence Trans: "trans r" using wo_rel.TRANS by auto have Anti: "antisym r" using wo_rel.ANTISYM Well by auto hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id) (* Main proof *) fix a1 a2 b1 b2 c1 c2 assume *: "((a1,a2),(b1,b2)) \ bsqr r" and **: "((b1,b2),(c1,c2)) \ bsqr r" hence 0: "{a1,a2,b1,b2,c1,c2} \ Field r" unfolding bsqr_def by auto have 1: "a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" using * unfolding bsqr_def by auto have 2: "b1 = c1 \ b2 = c2 \ (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id \ wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id \ wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1 \ (b2,c2) \ r - Id" using ** unfolding bsqr_def by auto show "((a1,a2),(c1,c2)) \ bsqr r" proof- {assume Case1: "a1 = b1 \ a2 = b2" hence ?thesis using ** by simp } moreover {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id" {assume Case21: "b1 = c1 \ b2 = c2" hence ?thesis using * by simp } moreover {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \ r - Id" using Case2 TransS trans_def[of "r - Id"] by blast hence ?thesis using 0 unfolding bsqr_def by auto } moreover {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2" hence ?thesis using Case2 0 unfolding bsqr_def by auto } ultimately have ?thesis using 0 2 by auto } moreover {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id" {assume Case31: "b1 = c1 \ b2 = c2" hence ?thesis using * by simp } moreover {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" hence ?thesis using Case3 0 unfolding bsqr_def by auto } moreover {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id" hence "(a1,c1) \ r - Id" using Case3 TransS trans_def[of "r - Id"] by blast hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto } moreover {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1" hence ?thesis using Case3 0 unfolding bsqr_def by auto } ultimately have ?thesis using 0 2 by auto } moreover {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" {assume Case41: "b1 = c1 \ b2 = c2" hence ?thesis using * by simp } moreover {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \ r - Id" hence ?thesis using Case4 0 unfolding bsqr_def by force } moreover {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ (b1,c1) \ r - Id" hence ?thesis using Case4 0 unfolding bsqr_def by auto } moreover {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \ b1 = c1 \ (b2,c2) \ r - Id" hence "(a2,c2) \ r - Id" using Case4 TransS trans_def[of "r - Id"] by blast hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto } ultimately have ?thesis using 0 2 by auto } ultimately show ?thesis using 0 1 by auto qed qed lemma bsqr_antisym: assumes "Well_order r" shows "antisym (bsqr r)" proof(unfold antisym_def, clarify) (* Preliminary facts *) have Well: "wo_rel r" using assms wo_rel_def by auto hence Trans: "trans r" using wo_rel.TRANS by auto have Anti: "antisym r" using wo_rel.ANTISYM Well by auto hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id) hence IrrS: "\a b. \((a,b) \ r - Id \ (b,a) \ r - Id)" using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast (* Main proof *) fix a1 a2 b1 b2 assume *: "((a1,a2),(b1,b2)) \ bsqr r" and **: "((b1,b2),(a1,a2)) \ bsqr r" hence 0: "{a1,a2,b1,b2} \ Field r" unfolding bsqr_def by auto have 1: "a1 = b1 \ a2 = b2 \ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" using * unfolding bsqr_def by auto have 2: "b1 = a1 \ b2 = a2 \ (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id \ wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \ (b1,a1) \ r - Id \ wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \ b1 = a1 \ (b2,a2) \ r - Id" using ** unfolding bsqr_def by auto show "a1 = b1 \ a2 = b2" proof- {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r - Id" {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" hence False using Case1 IrrS by blast } moreover {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2" hence False using Case1 by auto } ultimately have ?thesis using 0 2 by auto } moreover {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ (a1,b1) \ r - Id" {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" hence False using Case2 by auto } moreover {assume Case22: "(b1,a1) \ r - Id" hence False using Case2 IrrS by blast } moreover {assume Case23: "b1 = a1" hence False using Case2 by auto } ultimately have ?thesis using 0 2 by auto } moreover {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \ a1 = b1 \ (a2,b2) \ r - Id" moreover {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \ r - Id" hence False using Case3 by auto } moreover {assume Case32: "(b1,a1) \ r - Id" hence False using Case3 by auto } moreover {assume Case33: "(b2,a2) \ r - Id" hence False using Case3 IrrS by blast } ultimately have ?thesis using 0 2 by auto } ultimately show ?thesis using 0 1 by blast qed qed lemma bsqr_Total: assumes "Well_order r" shows "Total(bsqr r)" proof- (* Preliminary facts *) have Well: "wo_rel r" using assms wo_rel_def by auto hence Total: "\a \ Field r. \b \ Field r. (a,b) \ r \ (b,a) \ r" using wo_rel.TOTALS by auto (* Main proof *) {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \ Field(bsqr r)" hence 0: "a1 \ Field r \ a2 \ Field r \ b1 \ Field r \ b2 \ Field r" using Field_bsqr by blast have "((a1,a2) = (b1,b2) \ ((a1,a2),(b1,b2)) \ bsqr r \ ((b1,b2),(a1,a2)) \ bsqr r)" proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0) (* Why didn't clarsimp simp add: Well 0 do the same job? *) assume Case1: "(a1,a2) \ r" hence 1: "wo_rel.max2 r a1 a2 = a2" using Well 0 by (simp add: wo_rel.max2_equals2) show ?thesis proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0) assume Case11: "(b1,b2) \ r" hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2) show ?thesis proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) assume Case111: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" thus ?thesis using 0 1 2 unfolding bsqr_def by auto next assume Case112: "a2 = b2" show ?thesis proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) assume Case1121: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto next assume Case1122: "a1 = b1" thus ?thesis using Case112 by auto qed qed next assume Case12: "(b2,b1) \ r" hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1) show ?thesis proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0) assume Case121: "(a2,b1) \ r - Id \ (b1,a2) \ r - Id" thus ?thesis using 0 1 3 unfolding bsqr_def by auto next assume Case122: "a2 = b1" show ?thesis proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) assume Case1221: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto next assume Case1222: "a1 = b1" show ?thesis proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) assume Case12221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto next assume Case12222: "a2 = b2" thus ?thesis using Case122 Case1222 by auto qed qed qed qed next assume Case2: "(a2,a1) \ r" hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1) show ?thesis proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0) assume Case21: "(b1,b2) \ r" hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2) show ?thesis proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0) assume Case211: "(a1,b2) \ r - Id \ (b2,a1) \ r - Id" thus ?thesis using 0 1 2 unfolding bsqr_def by auto next assume Case212: "a1 = b2" show ?thesis proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) assume Case2121: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto next assume Case2122: "a1 = b1" show ?thesis proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) assume Case21221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto next assume Case21222: "a2 = b2" thus ?thesis using Case2122 Case212 by auto qed qed qed next assume Case22: "(b2,b1) \ r" hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1) show ?thesis proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0) assume Case221: "(a1,b1) \ r - Id \ (b1,a1) \ r - Id" thus ?thesis using 0 1 3 unfolding bsqr_def by auto next assume Case222: "a1 = b1" show ?thesis proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0) assume Case2221: "(a2,b2) \ r - Id \ (b2,a2) \ r - Id" thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto next assume Case2222: "a2 = b2" thus ?thesis using Case222 by auto qed qed qed qed } thus ?thesis unfolding total_on_def by fast qed lemma bsqr_Linear_order: assumes "Well_order r" shows "Linear_order(bsqr r)" unfolding order_on_defs using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast lemma bsqr_Well_order: assumes "Well_order r" shows "Well_order(bsqr r)" using assms proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI) have 0: "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a,a') \ r)" using assms well_order_on_def Linear_order_Well_order_iff by blast fix D assume *: "D \ Field (bsqr r)" and **: "D \ {}" hence 1: "D \ Field r \ Field r" unfolding Field_bsqr by simp (* *) obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \ D}" by blast have "M \ {}" using 1 M_def ** by auto moreover have "M \ Field r" unfolding M_def using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce ultimately obtain m where m_min: "m \ M \ (\a \ M. (m,a) \ r)" using 0 by blast (* *) obtain A1 where A1_def: "A1 = {a1. \a2. (a1,a2) \ D \ wo_rel.max2 r a1 a2 = m}" by blast have "A1 \ Field r" unfolding A1_def using 1 by auto moreover have "A1 \ {}" unfolding A1_def using m_min unfolding M_def by blast ultimately obtain a1 where a1_min: "a1 \ A1 \ (\a \ A1. (a1,a) \ r)" using 0 by blast (* *) obtain A2 where A2_def: "A2 = {a2. (a1,a2) \ D \ wo_rel.max2 r a1 a2 = m}" by blast have "A2 \ Field r" unfolding A2_def using 1 by auto moreover have "A2 \ {}" unfolding A2_def using m_min a1_min unfolding A1_def M_def by blast ultimately obtain a2 where a2_min: "a2 \ A2 \ (\a \ A2. (a2,a) \ r)" using 0 by blast (* *) have 2: "wo_rel.max2 r a1 a2 = m" using a1_min a2_min unfolding A1_def A2_def by auto have 3: "(a1,a2) \ D" using a2_min unfolding A2_def by auto (* *) moreover {fix b1 b2 assume ***: "(b1,b2) \ D" hence 4: "{a1,a2,b1,b2} \ Field r" using 1 3 by blast have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r" using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto have "((a1,a2),(b1,b2)) \ bsqr r" proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2") assume Case1: "wo_rel.max2 r a1 a2 \ wo_rel.max2 r b1 b2" thus ?thesis unfolding bsqr_def using 4 5 by auto next assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2" hence "b1 \ A1" unfolding A1_def using 2 *** by auto hence 6: "(a1,b1) \ r" using a1_min by auto show ?thesis proof(cases "a1 = b1") assume Case21: "a1 \ b1" thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto next assume Case22: "a1 = b1" hence "b2 \ A2" unfolding A2_def using 2 *** Case2 by auto hence 7: "(a2,b2) \ r" using a2_min by auto thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto qed qed } (* *) ultimately show "\d \ D. \d' \ D. (d,d') \ bsqr r" by fastforce qed lemma bsqr_max2: assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \ bsqr r" shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r" proof- have "{(a1,a2),(b1,b2)} \ Field(bsqr r)" using LEQ unfolding Field_def by auto hence "{a1,a2,b1,b2} \ Field r" unfolding Field_bsqr by auto hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \ Field r" using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \ r \ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2" using LEQ unfolding bsqr_def by auto ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto qed lemma bsqr_ofilter: assumes WELL: "Well_order r" and OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \ Field r" and NE: "\ (\a. Field r = under r a)" shows "\A. wo_rel.ofilter r A \ A < Field r \ D \ A \ A" proof- let ?r' = "bsqr r" have Well: "wo_rel r" using WELL wo_rel_def by blast hence Trans: "trans r" using wo_rel.TRANS by blast have Well': "Well_order ?r' \ wo_rel ?r'" using WELL bsqr_Well_order wo_rel_def by blast (* *) have "D < Field ?r'" unfolding Field_bsqr using SUB . with OF obtain a1 and a2 where "(a1,a2) \ Field ?r'" and 1: "D = underS ?r' (a1,a2)" using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto hence 2: "{a1,a2} \ Field r" unfolding Field_bsqr by auto let ?m = "wo_rel.max2 r a1 a2" have "D \ (under r ?m) \ (under r ?m)" proof(unfold 1) {fix b1 b2 let ?n = "wo_rel.max2 r b1 b2" assume "(b1,b2) \ underS ?r' (a1,a2)" hence 3: "((b1,b2),(a1,a2)) \ ?r'" unfolding underS_def by blast hence "(?n,?m) \ r" using WELL by (simp add: bsqr_max2) moreover {have "(b1,b2) \ Field ?r'" using 3 unfolding Field_def by auto hence "{b1,b2} \ Field r" unfolding Field_bsqr by auto hence "(b1,?n) \ r \ (b2,?n) \ r" using Well by (simp add: wo_rel.max2_greater) } ultimately have "(b1,?m) \ r \ (b2,?m) \ r" using Trans trans_def[of r] by blast hence "(b1,b2) \ (under r ?m) \ (under r ?m)" unfolding under_def by simp} thus "underS ?r' (a1,a2) \ (under r ?m) \ (under r ?m)" by auto qed moreover have "wo_rel.ofilter r (under r ?m)" using Well by (simp add: wo_rel.under_ofilter) moreover have "under r ?m < Field r" using NE under_Field[of r ?m] by blast ultimately show ?thesis by blast qed definition Func where "Func A B = {f . (\ a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" lemma Func_empty: "Func {} B = {\x. undefined}" unfolding Func_def by auto lemma Func_elim: assumes "g \ Func A B" and "a \ A" shows "\ b. b \ B \ g a = b" using assms unfolding Func_def by (cases "g a = undefined") auto definition curr where "curr A f \ \ a. if a \ A then \b. f (a,b) else undefined" lemma curr_in: assumes f: "f \ Func (A \ B) C" shows "curr A f \ Func A (Func B C)" using assms unfolding curr_def Func_def by auto lemma curr_inj: assumes "f1 \ Func (A \ B) C" and "f2 \ Func (A \ B) C" shows "curr A f1 = curr A f2 \ f1 = f2" proof safe assume c: "curr A f1 = curr A f2" show "f1 = f2" proof (rule ext, clarify) fix a b show "f1 (a, b) = f2 (a, b)" proof (cases "(a,b) \ A \ B") case False thus ?thesis using assms unfolding Func_def by auto next case True hence a: "a \ A" and b: "b \ B" by auto thus ?thesis using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp qed qed qed lemma curr_surj: assumes "g \ Func A (Func B C)" shows "\ f \ Func (A \ B) C. curr A f = g" proof let ?f = "\ ab. if fst ab \ A \ snd ab \ B then g (fst ab) (snd ab) else undefined" show "curr A ?f = g" proof (rule ext) fix a show "curr A ?f a = g a" proof (cases "a \ A") case False hence "g a = undefined" using assms unfolding Func_def by auto thus ?thesis unfolding curr_def using False by simp next case True obtain g1 where "g1 \ Func B C" and "g a = g1" using assms using Func_elim[OF assms True] by blast thus ?thesis using True unfolding Func_def curr_def by auto qed qed show "?f \ Func (A \ B) C" using assms unfolding Func_def mem_Collect_eq by auto qed lemma bij_betw_curr: "bij_betw (curr A) (Func (A \ B) C) (Func A (Func B C))" -unfolding bij_betw_def inj_on_def image_def -apply (intro impI conjI ballI) -apply (erule curr_inj[THEN iffD1], assumption+) -apply auto -apply (erule curr_in) -using curr_surj by blast + unfolding bij_betw_def inj_on_def image_def + apply (intro impI conjI ballI) + apply (erule curr_inj[THEN iffD1], assumption+, safe) + using curr_surj curr_in apply blast+ + done definition Func_map where "Func_map B2 f1 f2 g b2 \ if b2 \ B2 then f1 (g (f2 b2)) else undefined" lemma Func_map: assumes g: "g \ Func A2 A1" and f1: "f1 ` A1 \ B1" and f2: "f2 ` B2 \ A2" shows "Func_map B2 f1 f2 g \ Func B2 B1" using assms unfolding Func_def Func_map_def mem_Collect_eq by auto lemma Func_non_emp: assumes "B \ {}" shows "Func A B \ {}" proof- obtain b where b: "b \ B" using assms by auto hence "(\ a. if a \ A then b else undefined) \ Func A B" unfolding Func_def by auto thus ?thesis by blast qed lemma Func_is_emp: "Func A B = {} \ A \ {} \ B = {}" (is "?L \ ?R") proof assume L: ?L moreover {assume "A = {}" hence False using L Func_empty by auto} moreover {assume "B \ {}" hence False using L Func_non_emp[of B A] by simp } ultimately show ?R by blast next assume R: ?R moreover {fix f assume "f \ Func A B" moreover obtain a where "a \ A" using R by blast ultimately obtain b where "b \ B" unfolding Func_def by blast with R have False by blast } thus ?L by blast qed lemma Func_map_surj: assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \ A2" and B2A2: "B2 = {} \ A2 = {}" shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1" proof(cases "B2 = {}") case True thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def) next case False note B2 = False show ?thesis proof safe fix h assume h: "h \ Func B2 B1" define j1 where "j1 = inv_into A1 f1" have "\ a2 \ f2 ` B2. \ b2. b2 \ B2 \ f2 b2 = a2" by blast then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by atomize_elim (rule bchoice) {fix b2 assume b2: "b2 \ B2" hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto moreover have "k (f2 b2) \ B2" using b2 A2(2) k by auto ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast } note kk = this obtain b22 where b22: "b22 \ B2" using B2 by auto define j2 where [abs_def]: "j2 a2 = (if a2 \ f2 ` B2 then k a2 else b22)" for a2 have j2A2: "j2 ` A2 \ B2" unfolding j2_def using k b22 by auto have j2: "\ b2. b2 \ B2 \ j2 (f2 b2) = b2" using kk unfolding j2_def by auto define g where "g = Func_map A2 j1 j2 h" have "Func_map B2 f1 f2 g = h" proof (rule ext) fix b2 show "Func_map B2 f1 f2 g b2 = h b2" proof(cases "b2 \ B2") case True show ?thesis proof (cases "h b2 = undefined") case True hence b1: "h b2 \ f1 ` A1" using h \b2 \ B2\ unfolding B1 Func_def by auto show ?thesis using A2 f_inv_into_f[OF b1] unfolding True g_def Func_map_def j1_def j2[OF \b2 \ B2\] by auto qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def, auto intro: f_inv_into_f) qed(insert h, unfold Func_def Func_map_def, auto) qed moreover have "g \ Func A2 A1" unfolding g_def apply(rule Func_map[OF h]) using j2A2 B1 A2 unfolding j1_def by (fast intro: inv_into_into)+ ultimately show "h \ Func_map B2 f1 f2 ` Func A2 A1" unfolding Func_map_def[abs_def] by auto - qed(insert B1 Func_map[OF _ _ A2(2)], auto) + qed(use B1 Func_map[OF _ _ A2(2)] in auto) qed end diff --git a/src/HOL/BNF_Wellorder_Relation.thy b/src/HOL/BNF_Wellorder_Relation.thy --- a/src/HOL/BNF_Wellorder_Relation.thy +++ b/src/HOL/BNF_Wellorder_Relation.thy @@ -1,589 +1,607 @@ (* Title: HOL/BNF_Wellorder_Relation.thy Author: Andrei Popescu, TU Muenchen Copyright 2012 Well-order relations as needed by bounded natural functors. *) section \Well-Order Relations as Needed by Bounded Natural Functors\ theory BNF_Wellorder_Relation imports Order_Relation begin text\In this section, we develop basic concepts and results pertaining to well-order relations. Note that we consider well-order relations as {\em non-strict relations}, i.e., as containing the diagonals of their fields.\ locale wo_rel = fixes r :: "'a rel" assumes WELL: "Well_order r" begin text\The following context encompasses all this section. In other words, for the whole section, we consider a fixed well-order relation \<^term>\r\.\ (* context wo_rel *) abbreviation under where "under \ Order_Relation.under r" abbreviation underS where "underS \ Order_Relation.underS r" abbreviation Under where "Under \ Order_Relation.Under r" abbreviation UnderS where "UnderS \ Order_Relation.UnderS r" abbreviation above where "above \ Order_Relation.above r" abbreviation aboveS where "aboveS \ Order_Relation.aboveS r" abbreviation Above where "Above \ Order_Relation.Above r" abbreviation AboveS where "AboveS \ Order_Relation.AboveS r" abbreviation ofilter where "ofilter \ Order_Relation.ofilter r" lemmas ofilter_def = Order_Relation.ofilter_def[of r] subsection \Auxiliaries\ lemma REFL: "Refl r" using WELL order_on_defs[of _ r] by auto lemma TRANS: "trans r" using WELL order_on_defs[of _ r] by auto lemma ANTISYM: "antisym r" using WELL order_on_defs[of _ r] by auto lemma TOTAL: "Total r" using WELL order_on_defs[of _ r] by auto lemma TOTALS: "\a \ Field r. \b \ Field r. (a,b) \ r \ (b,a) \ r" using REFL TOTAL refl_on_def[of _ r] total_on_def[of _ r] by force lemma LIN: "Linear_order r" using WELL well_order_on_def[of _ r] by auto lemma WF: "wf (r - Id)" using WELL well_order_on_def[of _ r] by auto lemma cases_Total: "\ phi a b. \{a,b} <= Field r; ((a,b) \ r \ phi a b); ((b,a) \ r \ phi a b)\ \ phi a b" using TOTALS by auto lemma cases_Total3: "\ phi a b. \{a,b} \ Field r; ((a,b) \ r - Id \ (b,a) \ r - Id \ phi a b); (a = b \ phi a b)\ \ phi a b" using TOTALS by auto subsection \Well-founded induction and recursion adapted to non-strict well-order relations\ text\Here we provide induction and recursion principles specific to {\em non-strict} well-order relations. Although minor variations of those for well-founded relations, they will be useful for doing away with the tediousness of having to take out the diagonal each time in order to switch to a well-founded relation.\ lemma well_order_induct: assumes IND: "\x. \y. y \ x \ (y, x) \ r \ P y \ P x" shows "P a" proof- have "\x. \y. (y, x) \ r - Id \ P y \ P x" using IND by blast thus "P a" using WF wf_induct[of "r - Id" P a] by blast qed definition worec :: "(('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" where "worec F \ wfrec (r - Id) F" definition adm_wo :: "(('a \ 'b) \ 'a \ 'b) \ bool" where "adm_wo H \ \f g x. (\y \ underS x. f y = g y) \ H f x = H g x" lemma worec_fixpoint: assumes ADM: "adm_wo H" shows "worec H = H (worec H)" proof- let ?rS = "r - Id" have "adm_wf (r - Id) H" unfolding adm_wf_def using ADM adm_wo_def[of H] underS_def[of r] by auto hence "wfrec ?rS H = H (wfrec ?rS H)" using WF wfrec_fixpoint[of ?rS H] by simp thus ?thesis unfolding worec_def . qed subsection \The notions of maximum, minimum, supremum, successor and order filter\ text\ We define the successor {\em of a set}, and not of an element (the latter is of course a particular case). Also, we define the maximum {\em of two elements}, \max2\, and the minimum {\em of a set}, \minim\ -- we chose these variants since we consider them the most useful for well-orders. The minimum is defined in terms of the auxiliary relational operator \isMinim\. Then, supremum and successor are defined in terms of minimum as expected. The minimum is only meaningful for non-empty sets, and the successor is only meaningful for sets for which strict upper bounds exist. Order filters for well-orders are also known as ``initial segments".\ definition max2 :: "'a \ 'a \ 'a" where "max2 a b \ if (a,b) \ r then b else a" definition isMinim :: "'a set \ 'a \ bool" where "isMinim A b \ b \ A \ (\a \ A. (b,a) \ r)" definition minim :: "'a set \ 'a" where "minim A \ THE b. isMinim A b" definition supr :: "'a set \ 'a" where "supr A \ minim (Above A)" definition suc :: "'a set \ 'a" where "suc A \ minim (AboveS A)" subsubsection \Properties of max2\ lemma max2_greater_among: assumes "a \ Field r" and "b \ Field r" shows "(a, max2 a b) \ r \ (b, max2 a b) \ r \ max2 a b \ {a,b}" proof- {assume "(a,b) \ r" hence ?thesis using max2_def assms REFL refl_on_def by (auto simp add: refl_on_def) } moreover {assume "a = b" hence "(a,b) \ r" using REFL assms by (auto simp add: refl_on_def) } moreover {assume *: "a \ b \ (b,a) \ r" hence "(a,b) \ r" using ANTISYM by (auto simp add: antisym_def) hence ?thesis using * max2_def assms REFL refl_on_def by (auto simp add: refl_on_def) } ultimately show ?thesis using assms TOTAL total_on_def[of "Field r" r] by blast qed lemma max2_greater: assumes "a \ Field r" and "b \ Field r" shows "(a, max2 a b) \ r \ (b, max2 a b) \ r" using assms by (auto simp add: max2_greater_among) lemma max2_among: assumes "a \ Field r" and "b \ Field r" shows "max2 a b \ {a, b}" using assms max2_greater_among[of a b] by simp lemma max2_equals1: assumes "a \ Field r" and "b \ Field r" shows "(max2 a b = a) = ((b,a) \ r)" using assms ANTISYM unfolding antisym_def using TOTALS by(auto simp add: max2_def max2_among) lemma max2_equals2: assumes "a \ Field r" and "b \ Field r" shows "(max2 a b = b) = ((a,b) \ r)" using assms ANTISYM unfolding antisym_def using TOTALS unfolding max2_def by auto lemma in_notinI: assumes "(j,i) \ r \ j = i" and "i \ Field r" and "j \ Field r" shows "(i,j) \ r" using assms max2_def max2_greater_among by fastforce subsubsection \Existence and uniqueness for isMinim and well-definedness of minim\ lemma isMinim_unique: assumes MINIM: "isMinim B a" and MINIM': "isMinim B a'" shows "a = a'" proof- {have "a \ B" using MINIM isMinim_def by simp hence "(a',a) \ r" using MINIM' isMinim_def by simp } moreover {have "a' \ B" using MINIM' isMinim_def by simp hence "(a,a') \ r" using MINIM isMinim_def by simp } ultimately show ?thesis using ANTISYM antisym_def[of r] by blast qed lemma Well_order_isMinim_exists: assumes SUB: "B \ Field r" and NE: "B \ {}" shows "\b. isMinim B b" proof- from spec[OF WF[unfolded wf_eq_minimal[of "r - Id"]], of B] NE obtain b where - *: "b \ B \ (\b'. b' \ b \ (b',b) \ r \ b' \ B)" by auto - show ?thesis - proof(simp add: isMinim_def, rule exI[of _ b], auto) - show "b \ B" using * by simp - next - fix b' assume As: "b' \ B" - hence **: "b \ Field r \ b' \ Field r" using As SUB * by auto - (* *) - from As * have "b' = b \ (b',b) \ r" by auto - moreover - {assume "b' = b" - hence "(b,b') \ r" - using ** REFL by (auto simp add: refl_on_def) - } - moreover - {assume "b' \ b \ (b',b) \ r" - hence "(b,b') \ r" - using ** TOTAL by (auto simp add: total_on_def) - } - ultimately show "(b,b') \ r" by blast + *: "b \ B \ (\b'. b' \ b \ (b',b) \ r \ b' \ B)" by auto + have "\b'. b' \ B \ (b, b') \ r" + proof + fix b' + show "b' \ B \ (b, b') \ r" + proof + assume As: "b' \ B" + hence **: "b \ Field r \ b' \ Field r" using As SUB * by auto + from As * have "b' = b \ (b',b) \ r" by auto + moreover have "b' = b \ (b, b') \ r" + using ** REFL by (auto simp add: refl_on_def) + moreover have "b' \ b \ (b',b) \ r \ (b,b') \ r" + using ** TOTAL by (auto simp add: total_on_def) + ultimately show "(b,b') \ r" by blast + qed qed + then have "isMinim B b" + unfolding isMinim_def using * by auto + then show ?thesis + by auto qed lemma minim_isMinim: assumes SUB: "B \ Field r" and NE: "B \ {}" shows "isMinim B (minim B)" proof- let ?phi = "(\ b. isMinim B b)" from assms Well_order_isMinim_exists obtain b where *: "?phi b" by blast moreover have "\ b'. ?phi b' \ b' = b" using isMinim_unique * by auto ultimately show ?thesis unfolding minim_def using theI[of ?phi b] by blast qed subsubsection\Properties of minim\ lemma minim_in: assumes "B \ Field r" and "B \ {}" shows "minim B \ B" proof- from minim_isMinim[of B] assms have "isMinim B (minim B)" by simp thus ?thesis by (simp add: isMinim_def) qed lemma minim_inField: assumes "B \ Field r" and "B \ {}" shows "minim B \ Field r" proof- have "minim B \ B" using assms by (simp add: minim_in) thus ?thesis using assms by blast qed lemma minim_least: assumes SUB: "B \ Field r" and IN: "b \ B" shows "(minim B, b) \ r" proof- from minim_isMinim[of B] assms have "isMinim B (minim B)" by auto thus ?thesis by (auto simp add: isMinim_def IN) qed lemma equals_minim: assumes SUB: "B \ Field r" and IN: "a \ B" and LEAST: "\ b. b \ B \ (a,b) \ r" shows "a = minim B" proof- from minim_isMinim[of B] assms have "isMinim B (minim B)" by auto moreover have "isMinim B a" using IN LEAST isMinim_def by auto ultimately show ?thesis using isMinim_unique by auto qed subsubsection\Properties of successor\ lemma suc_AboveS: assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" shows "suc B \ AboveS B" proof(unfold suc_def) have "AboveS B \ Field r" using AboveS_Field[of r] by auto thus "minim (AboveS B) \ AboveS B" using assms by (simp add: minim_in) qed lemma suc_greater: assumes SUB: "B \ Field r" and ABOVES: "AboveS B \ {}" and IN: "b \ B" shows "suc B \ b \ (b,suc B) \ r" proof- from assms suc_AboveS have "suc B \ AboveS B" by simp with IN AboveS_def[of r] show ?thesis by simp qed lemma suc_least_AboveS: assumes ABOVES: "a \ AboveS B" shows "(suc B,a) \ r" proof(unfold suc_def) have "AboveS B \ Field r" using AboveS_Field[of r] by auto thus "(minim (AboveS B),a) \ r" using assms minim_least by simp qed lemma suc_inField: assumes "B \ Field r" and "AboveS B \ {}" shows "suc B \ Field r" proof- have "suc B \ AboveS B" using suc_AboveS assms by simp thus ?thesis using assms AboveS_Field[of r] by auto qed lemma equals_suc_AboveS: assumes SUB: "B \ Field r" and ABV: "a \ AboveS B" and MINIM: "\ a'. a' \ AboveS B \ (a,a') \ r" shows "a = suc B" proof(unfold suc_def) have "AboveS B \ Field r" using AboveS_Field[of r B] by auto thus "a = minim (AboveS B)" using assms equals_minim by simp qed lemma suc_underS: assumes IN: "a \ Field r" shows "a = suc (underS a)" proof- have "underS a \ Field r" using underS_Field[of r] by auto moreover have "a \ AboveS (underS a)" using in_AboveS_underS IN by fast moreover have "\a' \ AboveS (underS a). (a,a') \ r" proof(clarify) fix a' assume *: "a' \ AboveS (underS a)" hence **: "a' \ Field r" using AboveS_Field by fast {assume "(a,a') \ r" hence "a' = a \ (a',a) \ r" using TOTAL IN ** by (auto simp add: total_on_def) moreover {assume "a' = a" hence "(a,a') \ r" using REFL IN ** by (auto simp add: refl_on_def) } moreover {assume "a' \ a \ (a',a) \ r" hence "a' \ underS a" unfolding underS_def by simp hence "a' \ AboveS (underS a)" using AboveS_disjoint by fast with * have False by simp } ultimately have "(a,a') \ r" by blast } thus "(a, a') \ r" by blast qed ultimately show ?thesis using equals_suc_AboveS by auto qed subsubsection \Properties of order filters\ lemma under_ofilter: "ofilter (under a)" -proof(unfold ofilter_def under_def, auto simp add: Field_def) - fix aa x - assume "(aa,a) \ r" "(x,aa) \ r" - thus "(x,a) \ r" - using TRANS trans_def[of r] by blast +proof - + have "\aa x. (aa, a) \ r \ (x, aa) \ r \ (x, a) \ r" + proof - + fix aa x + assume "(aa,a) \ r" "(x,aa) \ r" + then show "(x,a) \ r" + using TRANS trans_def[of r] by blast + qed + then show ?thesis unfolding ofilter_def under_def + by (auto simp add: Field_def) qed lemma underS_ofilter: "ofilter (underS a)" -proof(unfold ofilter_def underS_def under_def, auto simp add: Field_def) + unfolding ofilter_def underS_def under_def +proof safe fix aa assume "(a, aa) \ r" "(aa, a) \ r" and DIFF: "aa \ a" thus False using ANTISYM antisym_def[of r] by blast next fix aa x assume "(aa,a) \ r" "aa \ a" "(x,aa) \ r" thus "(x,a) \ r" - using TRANS trans_def[of r] by blast + using TRANS trans_def[of r] by blast +next + fix x + assume "x \ a" and "(x, a) \ r" + then show "x \ Field r" + unfolding Field_def + by auto qed lemma Field_ofilter: "ofilter (Field r)" by(unfold ofilter_def under_def, auto simp add: Field_def) lemma ofilter_underS_Field: "ofilter A = ((\a \ Field r. A = underS a) \ (A = Field r))" proof assume "(\a\Field r. A = underS a) \ A = Field r" thus "ofilter A" by (auto simp: underS_ofilter Field_ofilter) next assume *: "ofilter A" let ?One = "(\a\Field r. A = underS a)" let ?Two = "(A = Field r)" show "?One \ ?Two" - proof(cases ?Two, simp) + proof(cases ?Two) let ?B = "(Field r) - A" let ?a = "minim ?B" assume "A \ Field r" moreover have "A \ Field r" using * ofilter_def by simp ultimately have 1: "?B \ {}" by blast hence 2: "?a \ Field r" using minim_inField[of ?B] by blast have 3: "?a \ ?B" using minim_in[of ?B] 1 by blast hence 4: "?a \ A" by blast have 5: "A \ Field r" using * ofilter_def by auto (* *) moreover have "A = underS ?a" proof show "A \ underS ?a" - proof(unfold underS_def, auto simp add: 4) + proof fix x assume **: "x \ A" hence 11: "x \ Field r" using 5 by auto have 12: "x \ ?a" using 4 ** by auto have 13: "under x \ A" using * ofilter_def ** by auto {assume "(x,?a) \ r" hence "(?a,x) \ r" using TOTAL total_on_def[of "Field r" r] 2 4 11 12 by auto hence "?a \ under x" using under_def[of r] by auto hence "?a \ A" using ** 13 by blast with 4 have False by simp } - thus "(x,?a) \ r" by blast + then have "(x,?a) \ r" by blast + thus "x \ underS ?a" + unfolding underS_def by (auto simp add: 12) qed next show "underS ?a \ A" - proof(unfold underS_def, auto) + proof fix x - assume **: "x \ ?a" and ***: "(x,?a) \ r" - hence 11: "x \ Field r" using Field_def by fastforce + assume **: "x \ underS ?a" + hence 11: "x \ Field r" + using Field_def unfolding underS_def by fastforce {assume "x \ A" hence "x \ ?B" using 11 by auto hence "(?a,x) \ r" using 3 minim_least[of ?B x] by blast hence False - using ANTISYM antisym_def[of r] ** *** by auto + using ANTISYM antisym_def[of r] ** unfolding underS_def by auto } thus "x \ A" by blast qed qed ultimately have ?One using 2 by blast thus ?thesis by simp + next + assume "A = Field r" + then show ?thesis + by simp qed qed lemma ofilter_UNION: "(\ i. i \ I \ ofilter(A i)) \ ofilter (\i \ I. A i)" unfolding ofilter_def by blast lemma ofilter_under_UNION: assumes "ofilter A" shows "A = (\a \ A. under a)" proof have "\a \ A. under a \ A" using assms ofilter_def by auto thus "(\a \ A. under a) \ A" by blast next have "\a \ A. a \ under a" using REFL Refl_under_in[of r] assms ofilter_def[of A] by blast thus "A \ (\a \ A. under a)" by blast qed subsubsection\Other properties\ lemma ofilter_linord: assumes OF1: "ofilter A" and OF2: "ofilter B" shows "A \ B \ B \ A" proof(cases "A = Field r") assume Case1: "A = Field r" hence "B \ A" using OF2 ofilter_def by auto thus ?thesis by simp next assume Case2: "A \ Field r" with ofilter_underS_Field OF1 obtain a where 1: "a \ Field r \ A = underS a" by auto show ?thesis proof(cases "B = Field r") assume Case21: "B = Field r" hence "A \ B" using OF1 ofilter_def by auto thus ?thesis by simp next assume Case22: "B \ Field r" with ofilter_underS_Field OF2 obtain b where 2: "b \ Field r \ B = underS b" by auto have "a = b \ (a,b) \ r \ (b,a) \ r" using 1 2 TOTAL total_on_def[of _ r] by auto moreover {assume "a = b" with 1 2 have ?thesis by auto } moreover {assume "(a,b) \ r" with underS_incr[of r] TRANS ANTISYM 1 2 have "A \ B" by auto hence ?thesis by auto } moreover {assume "(b,a) \ r" with underS_incr[of r] TRANS ANTISYM 1 2 have "B \ A" by auto hence ?thesis by auto } ultimately show ?thesis by blast qed qed lemma ofilter_AboveS_Field: assumes "ofilter A" shows "A \ (AboveS A) = Field r" proof show "A \ (AboveS A) \ Field r" using assms ofilter_def AboveS_Field[of r] by auto next {fix x assume *: "x \ Field r" and **: "x \ A" {fix y assume ***: "y \ A" with ** have 1: "y \ x" by auto {assume "(y,x) \ r" moreover have "y \ Field r" using assms ofilter_def *** by auto ultimately have "(x,y) \ r" using 1 * TOTAL total_on_def[of _ r] by auto with *** assms ofilter_def under_def[of r] have "x \ A" by auto with ** have False by contradiction } hence "(y,x) \ r" by blast with 1 have "y \ x \ (y,x) \ r" by auto } with * have "x \ AboveS A" unfolding AboveS_def by auto } thus "Field r \ A \ (AboveS A)" by blast qed lemma suc_ofilter_in: assumes OF: "ofilter A" and ABOVE_NE: "AboveS A \ {}" and REL: "(b,suc A) \ r" and DIFF: "b \ suc A" shows "b \ A" proof- have *: "suc A \ Field r \ b \ Field r" using WELL REL well_order_on_domain[of "Field r"] by auto {assume **: "b \ A" hence "b \ AboveS A" using OF * ofilter_AboveS_Field by auto hence "(suc A, b) \ r" using suc_least_AboveS by auto hence False using REL DIFF ANTISYM * by (auto simp add: antisym_def) } thus ?thesis by blast qed end (* context wo_rel *) end diff --git a/src/HOL/Complete_Lattices.thy b/src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy +++ b/src/HOL/Complete_Lattices.thy @@ -1,1364 +1,1364 @@ (* Title: HOL/Complete_Lattices.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Florian Haftmann Author: Viorel Preoteasa (Complete Distributive Lattices) *) section \Complete lattices\ theory Complete_Lattices imports Fun begin subsection \Syntactic infimum and supremum operations\ class Inf = fixes Inf :: "'a set \ 'a" ("\ _" [900] 900) class Sup = fixes Sup :: "'a set \ 'a" ("\ _" [900] 900) syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _\_./ _)" [0, 0, 10] 10) syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) translations "\x y. f" \ "\x. \y. f" "\x. f" \ "\(CONST range (\x. f))" "\x\A. f" \ "CONST Inf ((\x. f) ` A)" "\x y. f" \ "\x. \y. f" "\x. f" \ "\(CONST range (\x. f))" "\x\A. f" \ "CONST Sup ((\x. f) ` A)" context Inf begin lemma INF_image: "\ (g ` f ` A) = \ ((g \ f) ` A)" by (simp add: image_comp) lemma INF_identity_eq [simp]: "(\x\A. x) = \A" by simp lemma INF_id_eq [simp]: "\(id ` A) = \A" by simp lemma INF_cong: "A = B \ (\x. x \ B \ C x = D x) \ \(C ` A) = \(D ` B)" by (simp add: image_def) lemma INF_cong_simp: "A = B \ (\x. x \ B =simp=> C x = D x) \ \(C ` A) = \(D ` B)" unfolding simp_implies_def by (fact INF_cong) end context Sup begin lemma SUP_image: "\ (g ` f ` A) = \ ((g \ f) ` A)" by(fact Inf.INF_image) lemma SUP_identity_eq [simp]: "(\x\A. x) = \A" by(fact Inf.INF_identity_eq) lemma SUP_id_eq [simp]: "\(id ` A) = \A" by(fact Inf.INF_id_eq) lemma SUP_cong: "A = B \ (\x. x \ B \ C x = D x) \ \(C ` A) = \(D ` B)" by (fact Inf.INF_cong) lemma SUP_cong_simp: "A = B \ (\x. x \ B =simp=> C x = D x) \ \(C ` A) = \(D ` B)" by (fact Inf.INF_cong_simp) end subsection \Abstract complete lattices\ text \A complete lattice always has a bottom and a top, so we include them into the following type class, along with assumptions that define bottom and top in terms of infimum and supremum.\ class complete_lattice = lattice + Inf + Sup + bot + top + assumes Inf_lower: "x \ A \ \A \ x" and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" and Sup_upper: "x \ A \ x \ \A" and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" and Inf_empty [simp]: "\{} = \" and Sup_empty [simp]: "\{} = \" begin subclass bounded_lattice proof fix a show "\ \ a" by (auto intro: Sup_least simp only: Sup_empty [symmetric]) show "a \ \" by (auto intro: Inf_greatest simp only: Inf_empty [symmetric]) qed lemma dual_complete_lattice: "class.complete_lattice Sup Inf sup (\) (>) inf \ \" by (auto intro!: class.complete_lattice.intro dual_lattice) (unfold_locales, (fact Inf_empty Sup_empty Sup_upper Sup_least Inf_lower Inf_greatest)+) end context complete_lattice begin lemma Sup_eqI: "(\y. y \ A \ y \ x) \ (\y. (\z. z \ A \ z \ y) \ x \ y) \ \A = x" by (blast intro: order.antisym Sup_least Sup_upper) lemma Inf_eqI: "(\i. i \ A \ x \ i) \ (\y. (\i. i \ A \ y \ i) \ y \ x) \ \A = x" by (blast intro: order.antisym Inf_greatest Inf_lower) lemma SUP_eqI: "(\i. i \ A \ f i \ x) \ (\y. (\i. i \ A \ f i \ y) \ x \ y) \ (\i\A. f i) = x" using Sup_eqI [of "f ` A" x] by auto lemma INF_eqI: "(\i. i \ A \ x \ f i) \ (\y. (\i. i \ A \ f i \ y) \ x \ y) \ (\i\A. f i) = x" using Inf_eqI [of "f ` A" x] by auto lemma INF_lower: "i \ A \ (\i\A. f i) \ f i" using Inf_lower [of _ "f ` A"] by simp lemma INF_greatest: "(\i. i \ A \ u \ f i) \ u \ (\i\A. f i)" using Inf_greatest [of "f ` A"] by auto lemma SUP_upper: "i \ A \ f i \ (\i\A. f i)" using Sup_upper [of _ "f ` A"] by simp lemma SUP_least: "(\i. i \ A \ f i \ u) \ (\i\A. f i) \ u" using Sup_least [of "f ` A"] by auto lemma Inf_lower2: "u \ A \ u \ v \ \A \ v" using Inf_lower [of u A] by auto lemma INF_lower2: "i \ A \ f i \ u \ (\i\A. f i) \ u" using INF_lower [of i A f] by auto lemma Sup_upper2: "u \ A \ v \ u \ v \ \A" using Sup_upper [of u A] by auto lemma SUP_upper2: "i \ A \ u \ f i \ u \ (\i\A. f i)" using SUP_upper [of i A f] by auto lemma le_Inf_iff: "b \ \A \ (\a\A. b \ a)" by (auto intro: Inf_greatest dest: Inf_lower) lemma le_INF_iff: "u \ (\i\A. f i) \ (\i\A. u \ f i)" using le_Inf_iff [of _ "f ` A"] by simp lemma Sup_le_iff: "\A \ b \ (\a\A. a \ b)" by (auto intro: Sup_least dest: Sup_upper) lemma SUP_le_iff: "(\i\A. f i) \ u \ (\i\A. f i \ u)" using Sup_le_iff [of "f ` A"] by simp lemma Inf_insert [simp]: "\(insert a A) = a \ \A" by (auto intro: le_infI le_infI1 le_infI2 order.antisym Inf_greatest Inf_lower) lemma INF_insert: "(\x\insert a A. f x) = f a \ \(f ` A)" by simp lemma Sup_insert [simp]: "\(insert a A) = a \ \A" by (auto intro: le_supI le_supI1 le_supI2 order.antisym Sup_least Sup_upper) lemma SUP_insert: "(\x\insert a A. f x) = f a \ \(f ` A)" by simp lemma INF_empty: "(\x\{}. f x) = \" by simp lemma SUP_empty: "(\x\{}. f x) = \" by simp lemma Inf_UNIV [simp]: "\UNIV = \" by (auto intro!: order.antisym Inf_lower) lemma Sup_UNIV [simp]: "\UNIV = \" by (auto intro!: order.antisym Sup_upper) lemma Inf_eq_Sup: "\A = \{b. \a \ A. b \ a}" by (auto intro: order.antisym Inf_lower Inf_greatest Sup_upper Sup_least) lemma Sup_eq_Inf: "\A = \{b. \a \ A. a \ b}" by (auto intro: order.antisym Inf_lower Inf_greatest Sup_upper Sup_least) lemma Inf_superset_mono: "B \ A \ \A \ \B" by (auto intro: Inf_greatest Inf_lower) lemma Sup_subset_mono: "A \ B \ \A \ \B" by (auto intro: Sup_least Sup_upper) lemma Inf_mono: assumes "\b. b \ B \ \a\A. a \ b" shows "\A \ \B" proof (rule Inf_greatest) fix b assume "b \ B" with assms obtain a where "a \ A" and "a \ b" by blast from \a \ A\ have "\A \ a" by (rule Inf_lower) with \a \ b\ show "\A \ b" by auto qed lemma INF_mono: "(\m. m \ B \ \n\A. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" using Inf_mono [of "g ` B" "f ` A"] by auto lemma INF_mono': "(\x. f x \ g x) \ (\x\A. f x) \ (\x\A. g x)" by (rule INF_mono) auto lemma Sup_mono: assumes "\a. a \ A \ \b\B. a \ b" shows "\A \ \B" proof (rule Sup_least) fix a assume "a \ A" with assms obtain b where "b \ B" and "a \ b" by blast from \b \ B\ have "b \ \B" by (rule Sup_upper) with \a \ b\ show "a \ \B" by auto qed lemma SUP_mono: "(\n. n \ A \ \m\B. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" using Sup_mono [of "f ` A" "g ` B"] by auto lemma SUP_mono': "(\x. f x \ g x) \ (\x\A. f x) \ (\x\A. g x)" by (rule SUP_mono) auto lemma INF_superset_mono: "B \ A \ (\x. x \ B \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" \ \The last inclusion is POSITIVE!\ by (blast intro: INF_mono dest: subsetD) lemma SUP_subset_mono: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" by (blast intro: SUP_mono dest: subsetD) lemma Inf_less_eq: assumes "\v. v \ A \ v \ u" and "A \ {}" shows "\A \ u" proof - from \A \ {}\ obtain v where "v \ A" by blast moreover from \v \ A\ assms(1) have "v \ u" by blast ultimately show ?thesis by (rule Inf_lower2) qed lemma less_eq_Sup: assumes "\v. v \ A \ u \ v" and "A \ {}" shows "u \ \A" proof - from \A \ {}\ obtain v where "v \ A" by blast moreover from \v \ A\ assms(1) have "u \ v" by blast ultimately show ?thesis by (rule Sup_upper2) qed lemma INF_eq: assumes "\i. i \ A \ \j\B. f i \ g j" and "\j. j \ B \ \i\A. g j \ f i" shows "\(f ` A) = \(g ` B)" by (intro order.antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+ lemma SUP_eq: assumes "\i. i \ A \ \j\B. f i \ g j" and "\j. j \ B \ \i\A. g j \ f i" shows "\(f ` A) = \(g ` B)" by (intro order.antisym SUP_least) (blast intro: SUP_upper2 dest: assms)+ lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" by (auto intro: Inf_greatest Inf_lower) lemma Sup_inter_less_eq: "\(A \ B) \ \A \ \B " by (auto intro: Sup_least Sup_upper) lemma Inf_union_distrib: "\(A \ B) = \A \ \B" by (rule order.antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2) lemma INF_union: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" by (auto intro!: order.antisym INF_mono intro: le_infI1 le_infI2 INF_greatest INF_lower) lemma Sup_union_distrib: "\(A \ B) = \A \ \B" by (rule order.antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2) lemma SUP_union: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" by (auto intro!: order.antisym SUP_mono intro: le_supI1 le_supI2 SUP_least SUP_upper) lemma INF_inf_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" by (rule order.antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono) lemma SUP_sup_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" (is "?L = ?R") proof (rule order.antisym) show "?L \ ?R" by (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono) show "?R \ ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper) qed lemma Inf_top_conv [simp]: "\A = \ \ (\x\A. x = \)" "\ = \A \ (\x\A. x = \)" proof - show "\A = \ \ (\x\A. x = \)" proof assume "\x\A. x = \" then have "A = {} \ A = {\}" by auto then show "\A = \" by auto next assume "\A = \" show "\x\A. x = \" proof (rule ccontr) assume "\ (\x\A. x = \)" then obtain x where "x \ A" and "x \ \" by blast then obtain B where "A = insert x B" by blast with \\A = \\ \x \ \\ show False by simp qed qed then show "\ = \A \ (\x\A. x = \)" by auto qed lemma INF_top_conv [simp]: "(\x\A. B x) = \ \ (\x\A. B x = \)" "\ = (\x\A. B x) \ (\x\A. B x = \)" using Inf_top_conv [of "B ` A"] by simp_all lemma Sup_bot_conv [simp]: "\A = \ \ (\x\A. x = \)" "\ = \A \ (\x\A. x = \)" using dual_complete_lattice by (rule complete_lattice.Inf_top_conv)+ lemma SUP_bot_conv [simp]: "(\x\A. B x) = \ \ (\x\A. B x = \)" "\ = (\x\A. B x) \ (\x\A. B x = \)" using Sup_bot_conv [of "B ` A"] by simp_all lemma INF_constant: "(\y\A. c) = (if A = {} then \ else c)" by (auto intro: order.antisym INF_lower INF_greatest) lemma SUP_constant: "(\y\A. c) = (if A = {} then \ else c)" by (auto intro: order.antisym SUP_upper SUP_least) lemma INF_const [simp]: "A \ {} \ (\i\A. f) = f" by (simp add: INF_constant) lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" by (simp add: SUP_constant) lemma INF_top [simp]: "(\x\A. \) = \" by (cases "A = {}") simp_all lemma SUP_bot [simp]: "(\x\A. \) = \" by (cases "A = {}") simp_all lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" by (iprover intro: INF_lower INF_greatest order_trans order.antisym) lemma SUP_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" by (iprover intro: SUP_upper SUP_least order_trans order.antisym) lemma INF_absorb: assumes "k \ I" shows "A k \ (\i\I. A i) = (\i\I. A i)" proof - from assms obtain J where "I = insert k J" by blast then show ?thesis by simp qed lemma SUP_absorb: assumes "k \ I" shows "A k \ (\i\I. A i) = (\i\I. A i)" proof - from assms obtain J where "I = insert k J" by blast then show ?thesis by simp qed lemma INF_inf_const1: "I \ {} \ (\i\I. inf x (f i)) = inf x (\i\I. f i)" by (intro order.antisym INF_greatest inf_mono order_refl INF_lower) (auto intro: INF_lower2 le_infI2 intro!: INF_mono) lemma INF_inf_const2: "I \ {} \ (\i\I. inf (f i) x) = inf (\i\I. f i) x" using INF_inf_const1[of I x f] by (simp add: inf_commute) lemma less_INF_D: assumes "y < (\i\A. f i)" "i \ A" shows "y < f i" proof - note \y < (\i\A. f i)\ also have "(\i\A. f i) \ f i" using \i \ A\ by (rule INF_lower) finally show "y < f i" . qed lemma SUP_lessD: assumes "(\i\A. f i) < y" "i \ A" shows "f i < y" proof - have "f i \ (\i\A. f i)" using \i \ A\ by (rule SUP_upper) also note \(\i\A. f i) < y\ finally show "f i < y" . qed lemma INF_UNIV_bool_expand: "(\b. A b) = A True \ A False" by (simp add: UNIV_bool inf_commute) lemma SUP_UNIV_bool_expand: "(\b. A b) = A True \ A False" by (simp add: UNIV_bool sup_commute) lemma Inf_le_Sup: "A \ {} \ Inf A \ Sup A" by (blast intro: Sup_upper2 Inf_lower ex_in_conv) lemma INF_le_SUP: "A \ {} \ \(f ` A) \ \(f ` A)" using Inf_le_Sup [of "f ` A"] by simp lemma INF_eq_const: "I \ {} \ (\i. i \ I \ f i = x) \ \(f ` I) = x" by (auto intro: INF_eqI) lemma SUP_eq_const: "I \ {} \ (\i. i \ I \ f i = x) \ \(f ` I) = x" by (auto intro: SUP_eqI) lemma INF_eq_iff: "I \ {} \ (\i. i \ I \ f i \ c) \ \(f ` I) = c \ (\i\I. f i = c)" by (auto intro: INF_eq_const INF_lower order.antisym) lemma SUP_eq_iff: "I \ {} \ (\i. i \ I \ c \ f i) \ \(f ` I) = c \ (\i\I. f i = c)" by (auto intro: SUP_eq_const SUP_upper order.antisym) end context complete_lattice begin lemma Sup_Inf_le: "Sup (Inf ` {f ` A | f . (\ Y \ A . f Y \ Y)}) \ Inf (Sup ` A)" by (rule SUP_least, clarify, rule INF_greatest, simp add: INF_lower2 Sup_upper) end class complete_distrib_lattice = complete_lattice + assumes Inf_Sup_le: "Inf (Sup ` A) \ Sup (Inf ` {f ` A | f . (\ Y \ A . f Y \ Y)})" begin lemma Inf_Sup: "Inf (Sup ` A) = Sup (Inf ` {f ` A | f . (\ Y \ A . f Y \ Y)})" by (rule order.antisym, rule Inf_Sup_le, rule Sup_Inf_le) subclass distrib_lattice proof fix a b c show "a \ b \ c = (a \ b) \ (a \ c)" proof (rule order.antisym, simp_all, safe) show "b \ c \ a \ b" by (rule le_infI1, simp) show "b \ c \ a \ c" by (rule le_infI2, simp) have [simp]: "a \ c \ a \ b \ c" by (rule le_infI1, simp) have [simp]: "b \ a \ a \ b \ c" by (rule le_infI2, simp) have "\(Sup ` {{a, b}, {a, c}}) = \(Inf ` {f ` {{a, b}, {a, c}} | f. \Y\{{a, b}, {a, c}}. f Y \ Y})" by (rule Inf_Sup) from this show "(a \ b) \ (a \ c) \ a \ b \ c" apply simp by (rule SUP_least, safe, simp_all) qed qed end context complete_lattice begin context fixes f :: "'a \ 'b::complete_lattice" assumes "mono f" begin lemma mono_Inf: "f (\A) \ (\x\A. f x)" using \mono f\ by (auto intro: complete_lattice_class.INF_greatest Inf_lower dest: monoD) lemma mono_Sup: "(\x\A. f x) \ f (\A)" using \mono f\ by (auto intro: complete_lattice_class.SUP_least Sup_upper dest: monoD) lemma mono_INF: "f (\i\I. A i) \ (\x\I. f (A x))" by (intro complete_lattice_class.INF_greatest monoD[OF \mono f\] INF_lower) lemma mono_SUP: "(\x\I. f (A x)) \ f (\i\I. A i)" by (intro complete_lattice_class.SUP_least monoD[OF \mono f\] SUP_upper) end end class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice begin lemma uminus_Inf: "- (\A) = \(uminus ` A)" proof (rule order.antisym) show "- \A \ \(uminus ` A)" by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp show "\(uminus ` A) \ - \A" by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto qed lemma uminus_INF: "- (\x\A. B x) = (\x\A. - B x)" by (simp add: uminus_Inf image_image) lemma uminus_Sup: "- (\A) = \(uminus ` A)" proof - have "\A = - \(uminus ` A)" by (simp add: image_image uminus_INF) then show ?thesis by simp qed lemma uminus_SUP: "- (\x\A. B x) = (\x\A. - B x)" by (simp add: uminus_Sup image_image) end class complete_linorder = linorder + complete_lattice begin lemma dual_complete_linorder: "class.complete_linorder Sup Inf sup (\) (>) inf \ \" by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder) lemma complete_linorder_inf_min: "inf = min" by (auto intro: order.antisym simp add: min_def fun_eq_iff) lemma complete_linorder_sup_max: "sup = max" by (auto intro: order.antisym simp add: max_def fun_eq_iff) lemma Inf_less_iff: "\S < a \ (\x\S. x < a)" by (simp add: not_le [symmetric] le_Inf_iff) lemma INF_less_iff: "(\i\A. f i) < a \ (\x\A. f x < a)" by (simp add: Inf_less_iff [of "f ` A"]) lemma less_Sup_iff: "a < \S \ (\x\S. a < x)" by (simp add: not_le [symmetric] Sup_le_iff) lemma less_SUP_iff: "a < (\i\A. f i) \ (\x\A. a < f x)" by (simp add: less_Sup_iff [of _ "f ` A"]) lemma Sup_eq_top_iff [simp]: "\A = \ \ (\x<\. \i\A. x < i)" proof assume *: "\A = \" show "(\x<\. \i\A. x < i)" unfolding * [symmetric] proof (intro allI impI) fix x assume "x < \A" then show "\i\A. x < i" by (simp add: less_Sup_iff) qed next assume *: "\x<\. \i\A. x < i" show "\A = \" proof (rule ccontr) assume "\A \ \" with top_greatest [of "\A"] have "\A < \" unfolding le_less by auto with * have "\A < \A" unfolding less_Sup_iff by auto then show False by auto qed qed lemma SUP_eq_top_iff [simp]: "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" using Sup_eq_top_iff [of "f ` A"] by simp lemma Inf_eq_bot_iff [simp]: "\A = \ \ (\x>\. \i\A. i < x)" using dual_complete_linorder by (rule complete_linorder.Sup_eq_top_iff) lemma INF_eq_bot_iff [simp]: "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" using Inf_eq_bot_iff [of "f ` A"] by simp lemma Inf_le_iff: "\A \ x \ (\y>x. \a\A. y > a)" proof safe fix y assume "x \ \A" "y > x" then have "y > \A" by auto then show "\a\A. y > a" unfolding Inf_less_iff . qed (auto elim!: allE[of _ "\A"] simp add: not_le[symmetric] Inf_lower) lemma INF_le_iff: "\(f ` A) \ x \ (\y>x. \i\A. y > f i)" using Inf_le_iff [of "f ` A"] by simp lemma le_Sup_iff: "x \ \A \ (\ya\A. y < a)" proof safe fix y assume "x \ \A" "y < x" then have "y < \A" by auto then show "\a\A. y < a" unfolding less_Sup_iff . qed (auto elim!: allE[of _ "\A"] simp add: not_le[symmetric] Sup_upper) lemma le_SUP_iff: "x \ \(f ` A) \ (\yi\A. y < f i)" using le_Sup_iff [of _ "f ` A"] by simp end subsection \Complete lattice on \<^typ>\bool\\ instantiation bool :: complete_lattice begin definition [simp, code]: "\A \ False \ A" definition [simp, code]: "\A \ True \ A" instance by standard (auto intro: bool_induct) end lemma not_False_in_image_Ball [simp]: "False \ P ` A \ Ball A P" by auto lemma True_in_image_Bex [simp]: "True \ P ` A \ Bex A P" by auto lemma INF_bool_eq [simp]: "(\A f. \(f ` A)) = Ball" by (simp add: fun_eq_iff) lemma SUP_bool_eq [simp]: "(\A f. \(f ` A)) = Bex" by (simp add: fun_eq_iff) instance bool :: complete_boolean_algebra by (standard, fastforce) subsection \Complete lattice on \<^typ>\_ \ _\\ instantiation "fun" :: (type, Inf) Inf begin definition "\A = (\x. \f\A. f x)" lemma Inf_apply [simp, code]: "(\A) x = (\f\A. f x)" by (simp add: Inf_fun_def) instance .. end instantiation "fun" :: (type, Sup) Sup begin definition "\A = (\x. \f\A. f x)" lemma Sup_apply [simp, code]: "(\A) x = (\f\A. f x)" by (simp add: Sup_fun_def) instance .. end instantiation "fun" :: (type, complete_lattice) complete_lattice begin instance by standard (auto simp add: le_fun_def intro: INF_lower INF_greatest SUP_upper SUP_least) end lemma INF_apply [simp]: "(\y\A. f y) x = (\y\A. f y x)" by (simp add: image_comp) lemma SUP_apply [simp]: "(\y\A. f y) x = (\y\A. f y x)" by (simp add: image_comp) subsection \Complete lattice on unary and binary predicates\ lemma Inf1_I: "(\P. P \ A \ P a) \ (\A) a" by auto lemma INF1_I: "(\x. x \ A \ B x b) \ (\x\A. B x) b" by simp lemma INF2_I: "(\x. x \ A \ B x b c) \ (\x\A. B x) b c" by simp lemma Inf2_I: "(\r. r \ A \ r a b) \ (\A) a b" by auto lemma Inf1_D: "(\A) a \ P \ A \ P a" by auto lemma INF1_D: "(\x\A. B x) b \ a \ A \ B a b" by simp lemma Inf2_D: "(\A) a b \ r \ A \ r a b" by auto lemma INF2_D: "(\x\A. B x) b c \ a \ A \ B a b c" by simp lemma Inf1_E: assumes "(\A) a" obtains "P a" | "P \ A" using assms by auto lemma INF1_E: assumes "(\x\A. B x) b" obtains "B a b" | "a \ A" using assms by auto lemma Inf2_E: assumes "(\A) a b" obtains "r a b" | "r \ A" using assms by auto lemma INF2_E: assumes "(\x\A. B x) b c" obtains "B a b c" | "a \ A" using assms by auto lemma Sup1_I: "P \ A \ P a \ (\A) a" by auto lemma SUP1_I: "a \ A \ B a b \ (\x\A. B x) b" by auto lemma Sup2_I: "r \ A \ r a b \ (\A) a b" by auto lemma SUP2_I: "a \ A \ B a b c \ (\x\A. B x) b c" by auto lemma Sup1_E: assumes "(\A) a" obtains P where "P \ A" and "P a" using assms by auto lemma SUP1_E: assumes "(\x\A. B x) b" obtains x where "x \ A" and "B x b" using assms by auto lemma Sup2_E: assumes "(\A) a b" obtains r where "r \ A" "r a b" using assms by auto lemma SUP2_E: assumes "(\x\A. B x) b c" obtains x where "x \ A" "B x b c" using assms by auto subsection \Complete lattice on \<^typ>\_ set\\ instantiation "set" :: (type) complete_lattice begin definition "\A = {x. \((\B. x \ B) ` A)}" definition "\A = {x. \((\B. x \ B) ` A)}" instance by standard (auto simp add: less_eq_set_def Inf_set_def Sup_set_def le_fun_def) end subsubsection \Inter\ abbreviation Inter :: "'a set set \ 'a set" ("\") where "\S \ \S" lemma Inter_eq: "\A = {x. \B \ A. x \ B}" proof (rule set_eqI) fix x have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" by auto then show "x \ \A \ x \ {x. \B \ A. x \ B}" by (simp add: Inf_set_def image_def) qed lemma Inter_iff [simp]: "A \ \C \ (\X\C. A \ X)" by (unfold Inter_eq) blast lemma InterI [intro!]: "(\X. X \ C \ A \ X) \ A \ \C" by (simp add: Inter_eq) text \ \<^medskip> A ``destruct'' rule -- every \<^term>\X\ in \<^term>\C\ contains \<^term>\A\ as an element, but \<^prop>\A \ X\ can hold when \<^prop>\X \ C\ does not! This rule is analogous to \spec\. \ lemma InterD [elim, Pure.elim]: "A \ \C \ X \ C \ A \ X" by auto lemma InterE [elim]: "A \ \C \ (X \ C \ R) \ (A \ X \ R) \ R" \ \``Classical'' elimination rule -- does not require proving \<^prop>\X \ C\.\ unfolding Inter_eq by blast lemma Inter_lower: "B \ A \ \A \ B" by (fact Inf_lower) lemma Inter_subset: "(\X. X \ A \ X \ B) \ A \ {} \ \A \ B" by (fact Inf_less_eq) lemma Inter_greatest: "(\X. X \ A \ C \ X) \ C \ \A" by (fact Inf_greatest) lemma Inter_empty: "\{} = UNIV" by (fact Inf_empty) (* already simp *) lemma Inter_UNIV: "\UNIV = {}" by (fact Inf_UNIV) (* already simp *) lemma Inter_insert: "\(insert a B) = a \ \B" by (fact Inf_insert) (* already simp *) lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" by (fact less_eq_Inf_inter) lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" by (fact Inf_union_distrib) lemma Inter_UNIV_conv [simp]: "\A = UNIV \ (\x\A. x = UNIV)" "UNIV = \A \ (\x\A. x = UNIV)" by (fact Inf_top_conv)+ lemma Inter_anti_mono: "B \ A \ \A \ \B" by (fact Inf_superset_mono) subsubsection \Intersections of families\ syntax (ASCII) "_INTER1" :: "pttrns \ 'b set \ 'b set" ("(3INT _./ _)" [0, 10] 10) "_INTER" :: "pttrn \ 'a set \ 'b set \ 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) syntax "_INTER1" :: "pttrns \ 'b set \ 'b set" ("(3\_./ _)" [0, 10] 10) "_INTER" :: "pttrn \ 'a set \ 'b set \ 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) syntax (latex output) "_INTER1" :: "pttrns \ 'b set \ 'b set" ("(3\(\unbreakable\\<^bsub>_\<^esub>)/ _)" [0, 10] 10) "_INTER" :: "pttrn \ 'a set \ 'b set \ 'b set" ("(3\(\unbreakable\\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) translations "\x y. f" \ "\x. \y. f" "\x. f" \ "\(CONST range (\x. f))" "\x\A. f" \ "CONST Inter ((\x. f) ` A)" lemma INTER_eq: "(\x\A. B x) = {y. \x\A. y \ B x}" by (auto intro!: INF_eqI) lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" using Inter_iff [of _ "B ` A"] by simp lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" by auto lemma INT_D [elim, Pure.elim]: "b \ (\x\A. B x) \ a \ A \ b \ B a" by auto lemma INT_E [elim]: "b \ (\x\A. B x) \ (b \ B a \ R) \ (a \ A \ R) \ R" \ \"Classical" elimination -- by the Excluded Middle on \<^prop>\a\A\.\ by auto lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" by blast lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" by blast lemma INT_lower: "a \ A \ (\x\A. B x) \ B a" by (fact INF_lower) lemma INT_greatest: "(\x. x \ A \ C \ B x) \ C \ (\x\A. B x)" by (fact INF_greatest) lemma INT_empty: "(\x\{}. B x) = UNIV" by (fact INF_empty) lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" by (fact INF_absorb) lemma INT_subset_iff: "B \ (\i\I. A i) \ (\i\I. B \ A i)" by (fact le_INF_iff) lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ \ (B ` A)" by (fact INF_insert) lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" by (fact INF_union) lemma INT_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" by blast lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" by (fact INF_constant) lemma INTER_UNIV_conv: "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" by (fact INF_top_conv)+ (* already simp *) lemma INT_bool_eq: "(\b. A b) = A True \ A False" by (fact INF_UNIV_bool_expand) lemma INT_anti_mono: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\A. g x)" \ \The last inclusion is POSITIVE!\ by (fact INF_superset_mono) lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" by blast lemma vimage_INT: "f -` (\x\A. B x) = (\x\A. f -` B x)" by blast subsubsection \Union\ abbreviation Union :: "'a set set \ 'a set" ("\") where "\S \ \S" lemma Union_eq: "\A = {x. \B \ A. x \ B}" proof (rule set_eqI) fix x have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" by auto then show "x \ \A \ x \ {x. \B\A. x \ B}" by (simp add: Sup_set_def image_def) qed lemma Union_iff [simp]: "A \ \C \ (\X\C. A\X)" by (unfold Union_eq) blast lemma UnionI [intro]: "X \ C \ A \ X \ A \ \C" \ \The order of the premises presupposes that \<^term>\C\ is rigid; \<^term>\A\ may be flexible.\ by auto lemma UnionE [elim!]: "A \ \C \ (\X. A \ X \ X \ C \ R) \ R" by auto lemma Union_upper: "B \ A \ B \ \A" by (fact Sup_upper) lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" by (fact Sup_least) lemma Union_empty: "\{} = {}" by (fact Sup_empty) (* already simp *) lemma Union_UNIV: "\UNIV = UNIV" by (fact Sup_UNIV) (* already simp *) lemma Union_insert: "\(insert a B) = a \ \B" by (fact Sup_insert) (* already simp *) lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" by (fact Sup_union_distrib) lemma Union_Int_subset: "\(A \ B) \ \A \ \B" by (fact Sup_inter_less_eq) lemma Union_empty_conv: "(\A = {}) \ (\x\A. x = {})" by (fact Sup_bot_conv) (* already simp *) lemma empty_Union_conv: "({} = \A) \ (\x\A. x = {})" by (fact Sup_bot_conv) (* already simp *) lemma subset_Pow_Union: "A \ Pow (\A)" by blast lemma Union_Pow_eq [simp]: "\(Pow A) = A" by blast lemma Union_mono: "A \ B \ \A \ \B" by (fact Sup_subset_mono) lemma Union_subsetI: "(\x. x \ A \ \y. y \ B \ x \ y) \ \A \ \B" by blast lemma disjnt_inj_on_iff: - "\inj_on f (\\); X \ \; Y \ \\ \ disjnt (f ` X) (f ` Y) \ disjnt X Y" - apply (auto simp: disjnt_def) - using inj_on_eq_iff by fastforce + "\inj_on f (\\); X \ \; Y \ \\ \ disjnt (f ` X) (f ` Y) \ disjnt X Y" + unfolding disjnt_def + by safe (use inj_on_eq_iff in \fastforce+\) lemma disjnt_Union1 [simp]: "disjnt (\\) B \ (\A \ \. disjnt A B)" by (auto simp: disjnt_def) lemma disjnt_Union2 [simp]: "disjnt B (\\) \ (\A \ \. disjnt B A)" by (auto simp: disjnt_def) subsubsection \Unions of families\ syntax (ASCII) "_UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 0, 10] 10) syntax "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\(\unbreakable\\<^bsub>_\<^esub>)/ _)" [0, 10] 10) "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(\unbreakable\\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) translations "\x y. f" \ "\x. \y. f" "\x. f" \ "\(CONST range (\x. f))" "\x\A. f" \ "CONST Union ((\x. f) ` A)" text \ Note the difference between ordinary syntax of indexed unions and intersections (e.g.\ \\a\<^sub>1\A\<^sub>1. B\) and their \LaTeX\ rendition: \<^term>\\a\<^sub>1\A\<^sub>1. B\. \ lemma disjoint_UN_iff: "disjnt A (\i\I. B i) \ (\i\I. disjnt A (B i))" by (auto simp: disjnt_def) lemma UNION_eq: "(\x\A. B x) = {y. \x\A. y \ B x}" by (auto intro!: SUP_eqI) lemma bind_UNION [code]: "Set.bind A f = \(f ` A)" by (simp add: bind_def UNION_eq) lemma member_bind [simp]: "x \ Set.bind A f \ x \ \(f ` A)" by (simp add: bind_UNION) lemma Union_SetCompr_eq: "\{f x| x. P x} = {a. \x. P x \ a \ f x}" by blast lemma UN_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" using Union_iff [of _ "B ` A"] by simp lemma UN_I [intro]: "a \ A \ b \ B a \ b \ (\x\A. B x)" \ \The order of the premises presupposes that \<^term>\A\ is rigid; \<^term>\b\ may be flexible.\ by auto lemma UN_E [elim!]: "b \ (\x\A. B x) \ (\x. x\A \ b \ B x \ R) \ R" by auto lemma UN_upper: "a \ A \ B a \ (\x\A. B x)" by (fact SUP_upper) lemma UN_least: "(\x. x \ A \ B x \ C) \ (\x\A. B x) \ C" by (fact SUP_least) lemma Collect_bex_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" by blast lemma UN_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" by blast lemma UN_empty: "(\x\{}. B x) = {}" by (fact SUP_empty) lemma UN_empty2: "(\x\A. {}) = {}" by (fact SUP_bot) (* already simp *) lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" by (fact SUP_absorb) lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ \(B ` A)" by (fact SUP_insert) lemma UN_Un [simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" by (fact SUP_union) lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" by blast lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" by (fact SUP_le_iff) lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" by (fact SUP_constant) lemma UNION_singleton_eq_range: "(\x\A. {f x}) = f ` A" by blast lemma image_Union: "f ` \S = (\x\S. f ` x)" by blast lemma UNION_empty_conv: "{} = (\x\A. B x) \ (\x\A. B x = {})" "(\x\A. B x) = {} \ (\x\A. B x = {})" by (fact SUP_bot_conv)+ (* already simp *) lemma Collect_ex_eq: "{x. \y. P x y} = (\y. {x. P x y})" by blast lemma ball_UN: "(\z \ \(B ` A). P z) \ (\x\A. \z \ B x. P z)" by blast lemma bex_UN: "(\z \ \(B ` A). P z) \ (\x\A. \z\B x. P z)" by blast lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" by safe (auto simp add: if_split_mem2) lemma UN_bool_eq: "(\b. A b) = (A True \ A False)" by (fact SUP_UNIV_bool_expand) lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" by blast lemma UN_mono: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" by (fact SUP_subset_mono) lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" by blast lemma vimage_UN: "f -` (\x\A. B x) = (\x\A. f -` B x)" by blast lemma vimage_eq_UN: "f -` B = (\y\B. f -` {y})" \ \NOT suitable for rewriting\ by blast lemma image_UN: "f ` \(B ` A) = (\x\A. f ` B x)" by blast lemma UN_singleton [simp]: "(\x\A. {x}) = A" by blast lemma inj_on_image: "inj_on f (\A) \ inj_on ((`) f) A" unfolding inj_on_def by blast subsubsection \Distributive laws\ lemma Int_Union: "A \ \B = (\C\B. A \ C)" by blast lemma Un_Inter: "A \ \B = (\C\B. A \ C)" by blast lemma Int_Union2: "\B \ A = (\C\B. C \ A)" by blast lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" by (rule sym) (rule INF_inf_distrib) lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" by (rule sym) (rule SUP_sup_distrib) lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" (* FIXME drop *) by (simp add: INT_Int_distrib) lemma Int_Inter_eq: "A \ \\ = (if \={} then A else (\B\\. A \ B))" "\\ \ A = (if \={} then A else (\B\\. B \ A))" by auto lemma Un_Union_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" (* FIXME drop *) \ \Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:\ \ \Union of a family of unions\ by (simp add: UN_Un_distrib) lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" by blast lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" \ \Halmos, Naive Set Theory, page 35.\ by blast lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" by blast lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" by blast lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" by blast lemma SUP_UNION: "(\x\(\y\A. g y). f x) = (\y\A. \x\g y. f x :: _ :: complete_lattice)" by (rule order_antisym) (blast intro: SUP_least SUP_upper2)+ subsection \Injections and bijections\ lemma inj_on_Inter: "S \ {} \ (\A. A \ S \ inj_on f A) \ inj_on f (\S)" unfolding inj_on_def by blast lemma inj_on_INTER: "I \ {} \ (\i. i \ I \ inj_on f (A i)) \ inj_on f (\i \ I. A i)" unfolding inj_on_def by safe simp lemma inj_on_UNION_chain: assumes chain: "\i j. i \ I \ j \ I \ A i \ A j \ A j \ A i" and inj: "\i. i \ I \ inj_on f (A i)" shows "inj_on f (\i \ I. A i)" proof - have "x = y" if *: "i \ I" "j \ I" and **: "x \ A i" "y \ A j" and ***: "f x = f y" for i j x y using chain [OF *] proof assume "A i \ A j" with ** have "x \ A j" by auto with inj * ** *** show ?thesis by (auto simp add: inj_on_def) next assume "A j \ A i" with ** have "y \ A i" by auto with inj * ** *** show ?thesis by (auto simp add: inj_on_def) qed then show ?thesis by (unfold inj_on_def UNION_eq) auto qed lemma bij_betw_UNION_chain: assumes chain: "\i j. i \ I \ j \ I \ A i \ A j \ A j \ A i" and bij: "\i. i \ I \ bij_betw f (A i) (A' i)" shows "bij_betw f (\i \ I. A i) (\i \ I. A' i)" unfolding bij_betw_def proof safe have "\i. i \ I \ inj_on f (A i)" using bij bij_betw_def[of f] by auto then show "inj_on f (\(A ` I))" using chain inj_on_UNION_chain[of I A f] by auto next fix i x assume *: "i \ I" "x \ A i" with bij have "f x \ A' i" by (auto simp: bij_betw_def) with * show "f x \ \(A' ` I)" by blast next fix i x' assume *: "i \ I" "x' \ A' i" with bij have "\x \ A i. x' = f x" unfolding bij_betw_def by blast with * have "\j \ I. \x \ A j. x' = f x" by blast then show "x' \ f ` \(A ` I)" by blast qed (*injectivity's required. Left-to-right inclusion holds even if A is empty*) lemma image_INT: "inj_on f C \ \x\A. B x \ C \ j \ A \ f ` (\(B ` A)) = (\x\A. f ` B x)" by (auto simp add: inj_on_def) blast lemma bij_image_INT: "bij f \ f ` (\(B ` A)) = (\x\A. f ` B x)" by (auto simp: bij_def inj_def surj_def) blast lemma UNION_fun_upd: "\(A(i := B) ` J) = \(A ` (J - {i})) \ (if i \ J then B else {})" by (auto simp add: set_eq_iff) lemma bij_betw_Pow: assumes "bij_betw f A B" shows "bij_betw (image f) (Pow A) (Pow B)" proof - from assms have "inj_on f A" by (rule bij_betw_imp_inj_on) then have "inj_on f (\(Pow A))" by simp then have "inj_on (image f) (Pow A)" by (rule inj_on_image) then have "bij_betw (image f) (Pow A) (image f ` Pow A)" by (rule inj_on_imp_bij_betw) moreover from assms have "f ` A = B" by (rule bij_betw_imp_surj_on) then have "image f ` Pow A = Pow B" by (rule image_Pow_surj) ultimately show ?thesis by simp qed subsubsection \Complement\ lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" by blast lemma Compl_UN [simp]: "- (\x\A. B x) = (\x\A. -B x)" by blast subsubsection \Miniscoping and maxiscoping\ text \\<^medskip> Miniscoping: pushing in quantifiers and big Unions and Intersections.\ lemma UN_simps [simp]: "\a B C. (\x\C. insert a (B x)) = (if C={} then {} else insert a (\x\C. B x))" "\A B C. (\x\C. A x \ B) = ((if C={} then {} else (\x\C. A x) \ B))" "\A B C. (\x\C. A \ B x) = ((if C={} then {} else A \ (\x\C. B x)))" "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" "\A B C. (\x\C. A \ B x) = (A \(\x\C. B x))" "\A B C. (\x\C. A x - B) = ((\x\C. A x) - B)" "\A B C. (\x\C. A - B x) = (A - (\x\C. B x))" "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" "\A B C. (\z\(\(B ` A)). C z) = (\x\A. \z\B x. C z)" "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" by auto lemma INT_simps [simp]: "\A B C. (\x\C. A x \ B) = (if C={} then UNIV else (\x\C. A x) \ B)" "\A B C. (\x\C. A \ B x) = (if C={} then UNIV else A \(\x\C. B x))" "\A B C. (\x\C. A x - B) = (if C={} then UNIV else (\x\C. A x) - B)" "\A B C. (\x\C. A - B x) = (if C={} then UNIV else A - (\x\C. B x))" "\a B C. (\x\C. insert a (B x)) = insert a (\x\C. B x)" "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" "\A B C. (\x\C. A \ B x) = (A \ (\x\C. B x))" "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" "\A B C. (\z\(\(B ` A)). C z) = (\x\A. \z\B x. C z)" "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" by auto lemma UN_ball_bex_simps [simp]: "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" "\A B P. (\x\(\(B ` A)). P x) = (\a\A. \x\ B a. P x)" "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" "\A B P. (\x\(\(B ` A)). P x) \ (\a\A. \x\B a. P x)" by auto text \\<^medskip> Maxiscoping: pulling out big Unions and Intersections.\ lemma UN_extend_simps: "\a B C. insert a (\x\C. B x) = (if C={} then {a} else (\x\C. insert a (B x)))" "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" "\A B C. (A \ (\x\C. B x)) = (\x\C. A \ B x)" "\A B C. ((\x\C. A x) - B) = (\x\C. A x - B)" "\A B C. (A - (\x\C. B x)) = (\x\C. A - B x)" "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" "\A B C. (\x\A. \z\B x. C z) = (\z\(\(B ` A)). C z)" "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" by auto lemma INT_extend_simps: "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" "\A B C. (\x\C. A x) - B = (if C={} then UNIV - B else (\x\C. A x - B))" "\A B C. A - (\x\C. B x) = (if C={} then A else (\x\C. A - B x))" "\a B C. insert a (\x\C. B x) = (\x\C. insert a (B x))" "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" "\A B C. A \ (\x\C. B x) = (\x\C. A \ B x)" "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" "\A B C. (\x\A. \z\B x. C z) = (\z\(\(B ` A)). C z)" "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" by auto text \Finally\ lemmas mem_simps = insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff \ \Each of these has ALREADY been added \[simp]\ above.\ end diff --git a/src/HOL/Conditionally_Complete_Lattices.thy b/src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy +++ b/src/HOL/Conditionally_Complete_Lattices.thy @@ -1,842 +1,840 @@ (* Title: HOL/Conditionally_Complete_Lattices.thy Author: Amine Chaieb and L C Paulson, University of Cambridge Author: Johannes Hölzl, TU München Author: Luke S. Serafin, Carnegie Mellon University *) section \Conditionally-complete Lattices\ theory Conditionally_Complete_Lattices imports Finite_Set Lattices_Big Set_Interval begin locale preordering_bdd = preordering begin definition bdd :: \'a set \ bool\ where unfold: \bdd A \ (\M. \x \ A. x \<^bold>\ M)\ lemma empty [simp, intro]: \bdd {}\ by (simp add: unfold) lemma I [intro]: \bdd A\ if \\x. x \ A \ x \<^bold>\ M\ using that by (auto simp add: unfold) lemma E: assumes \bdd A\ obtains M where \\x. x \ A \ x \<^bold>\ M\ using assms that by (auto simp add: unfold) lemma I2: \bdd (f ` A)\ if \\x. x \ A \ f x \<^bold>\ M\ using that by (auto simp add: unfold) lemma mono: \bdd A\ if \bdd B\ \A \ B\ using that by (auto simp add: unfold) lemma Int1 [simp]: \bdd (A \ B)\ if \bdd A\ using mono that by auto lemma Int2 [simp]: \bdd (A \ B)\ if \bdd B\ using mono that by auto end subsection \Preorders\ context preorder begin sublocale bdd_above: preordering_bdd \(\)\ \(<)\ defines bdd_above_primitive_def: bdd_above = bdd_above.bdd .. sublocale bdd_below: preordering_bdd \(\)\ \(>)\ defines bdd_below_primitive_def: bdd_below = bdd_below.bdd .. lemma bdd_above_def: \bdd_above A \ (\M. \x \ A. x \ M)\ by (fact bdd_above.unfold) lemma bdd_below_def: \bdd_below A \ (\M. \x \ A. M \ x)\ by (fact bdd_below.unfold) lemma bdd_aboveI: "(\x. x \ A \ x \ M) \ bdd_above A" by (fact bdd_above.I) lemma bdd_belowI: "(\x. x \ A \ m \ x) \ bdd_below A" by (fact bdd_below.I) lemma bdd_aboveI2: "(\x. x \ A \ f x \ M) \ bdd_above (f`A)" by (fact bdd_above.I2) lemma bdd_belowI2: "(\x. x \ A \ m \ f x) \ bdd_below (f`A)" by (fact bdd_below.I2) lemma bdd_above_empty: "bdd_above {}" by (fact bdd_above.empty) lemma bdd_below_empty: "bdd_below {}" by (fact bdd_below.empty) lemma bdd_above_mono: "bdd_above B \ A \ B \ bdd_above A" by (fact bdd_above.mono) lemma bdd_below_mono: "bdd_below B \ A \ B \ bdd_below A" by (fact bdd_below.mono) lemma bdd_above_Int1: "bdd_above A \ bdd_above (A \ B)" by (fact bdd_above.Int1) lemma bdd_above_Int2: "bdd_above B \ bdd_above (A \ B)" by (fact bdd_above.Int2) lemma bdd_below_Int1: "bdd_below A \ bdd_below (A \ B)" by (fact bdd_below.Int1) lemma bdd_below_Int2: "bdd_below B \ bdd_below (A \ B)" by (fact bdd_below.Int2) lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) end context order_top begin lemma bdd_above_top [simp, intro!]: "bdd_above A" by (rule bdd_aboveI [of _ top]) simp end context order_bot begin lemma bdd_below_bot [simp, intro!]: "bdd_below A" by (rule bdd_belowI [of _ bot]) simp end lemma bdd_above_image_mono: "mono f \ bdd_above A \ bdd_above (f`A)" by (auto simp: bdd_above_def mono_def) lemma bdd_below_image_mono: "mono f \ bdd_below A \ bdd_below (f`A)" by (auto simp: bdd_below_def mono_def) lemma bdd_above_image_antimono: "antimono f \ bdd_below A \ bdd_above (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma bdd_below_image_antimono: "antimono f \ bdd_above A \ bdd_below (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma fixes X :: "'a::ordered_ab_group_add set" shows bdd_above_uminus[simp]: "bdd_above (uminus ` X) \ bdd_below X" and bdd_below_uminus[simp]: "bdd_below (uminus ` X) \ bdd_above X" using bdd_above_image_antimono[of uminus X] bdd_below_image_antimono[of uminus "uminus`X"] using bdd_below_image_antimono[of uminus X] bdd_above_image_antimono[of uminus "uminus`X"] by (auto simp: antimono_def image_image) subsection \Lattices\ context lattice begin lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A" by (auto simp: bdd_above_def intro: le_supI2 sup_ge1) lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A" by (auto simp: bdd_below_def intro: le_infI2 inf_le1) lemma bdd_finite [simp]: assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A" using assms by (induct rule: finite_induct, auto) lemma bdd_above_Un [simp]: "bdd_above (A \ B) = (bdd_above A \ bdd_above B)" proof assume "bdd_above (A \ B)" thus "bdd_above A \ bdd_above B" unfolding bdd_above_def by auto next assume "bdd_above A \ bdd_above B" then obtain a b where "\x\A. x \ a" "\x\B. x \ b" unfolding bdd_above_def by auto hence "\x \ A \ B. x \ sup a b" by (auto intro: Un_iff le_supI1 le_supI2) thus "bdd_above (A \ B)" unfolding bdd_above_def .. qed lemma bdd_below_Un [simp]: "bdd_below (A \ B) = (bdd_below A \ bdd_below B)" proof assume "bdd_below (A \ B)" thus "bdd_below A \ bdd_below B" unfolding bdd_below_def by auto next assume "bdd_below A \ bdd_below B" then obtain a b where "\x\A. a \ x" "\x\B. b \ x" unfolding bdd_below_def by auto hence "\x \ A \ B. inf a b \ x" by (auto intro: Un_iff le_infI1 le_infI2) thus "bdd_below (A \ B)" unfolding bdd_below_def .. qed lemma bdd_above_image_sup[simp]: "bdd_above ((\x. sup (f x) (g x)) ` A) \ bdd_above (f`A) \ bdd_above (g`A)" by (auto simp: bdd_above_def intro: le_supI1 le_supI2) lemma bdd_below_image_inf[simp]: "bdd_below ((\x. inf (f x) (g x)) ` A) \ bdd_below (f`A) \ bdd_below (g`A)" by (auto simp: bdd_below_def intro: le_infI1 le_infI2) lemma bdd_below_UN[simp]: "finite I \ bdd_below (\i\I. A i) = (\i \ I. bdd_below (A i))" by (induction I rule: finite.induct) auto lemma bdd_above_UN[simp]: "finite I \ bdd_above (\i\I. A i) = (\i \ I. bdd_above (A i))" by (induction I rule: finite.induct) auto end text \ To avoid name classes with the \<^class>\complete_lattice\-class we prefix \<^const>\Sup\ and \<^const>\Inf\ in theorem names with c. \ subsection \Conditionally complete lattices\ class conditionally_complete_lattice = lattice + Sup + Inf + assumes cInf_lower: "x \ X \ bdd_below X \ Inf X \ x" and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" assumes cSup_upper: "x \ X \ bdd_above X \ x \ Sup X" and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" by (metis cSup_upper order_trans) lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" by (metis cInf_lower order_trans) lemma cSup_mono: "B \ {} \ bdd_above A \ (\b. b \ B \ \a\A. b \ a) \ Sup B \ Sup A" by (metis cSup_least cSup_upper2) lemma cInf_mono: "B \ {} \ bdd_below A \ (\b. b \ B \ \a\A. a \ b) \ Inf A \ Inf B" by (metis cInf_greatest cInf_lower2) lemma cSup_subset_mono: "A \ {} \ bdd_above B \ A \ B \ Sup A \ Sup B" by (metis cSup_least cSup_upper subsetD) lemma cInf_superset_mono: "A \ {} \ bdd_below B \ A \ B \ Inf B \ Inf A" by (metis cInf_greatest cInf_lower subsetD) lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" by (intro order.antisym cSup_upper[of z X] cSup_least[of X z]) auto lemma cInf_eq_minimum: "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" by (intro order.antisym cInf_lower[of z X] cInf_greatest[of X z]) auto lemma cSup_le_iff: "S \ {} \ bdd_above S \ Sup S \ a \ (\x\S. x \ a)" by (metis order_trans cSup_upper cSup_least) lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ x)" by (metis order_trans cInf_lower cInf_greatest) lemma cSup_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ x \ a" assumes 3: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" by (intro 3 1 order.antisym cSup_least) (auto intro: 2 1 cSup_upper) lemma cInf_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ a \ x" assumes 3: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" by (intro 3 1 order.antisym cInf_greatest) (auto intro: 2 1 cInf_lower) lemma cInf_cSup: "S \ {} \ bdd_below S \ Inf S = Sup {x. \s\S. x \ s}" by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) lemma cSup_cInf: "S \ {} \ bdd_above S \ Sup S = Inf {x. \s\S. s \ x}" by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) lemma cSup_insert: "X \ {} \ bdd_above X \ Sup (insert a X) = sup a (Sup X)" by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least) lemma cInf_insert: "X \ {} \ bdd_below X \ Inf (insert a X) = inf a (Inf X)" by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) lemma cSup_singleton [simp]: "Sup {x} = x" by (intro cSup_eq_maximum) auto lemma cInf_singleton [simp]: "Inf {x} = x" by (intro cInf_eq_minimum) auto lemma cSup_insert_If: "bdd_above X \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" using cSup_insert[of X] by simp lemma cInf_insert_If: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" using cInf_insert[of X] by simp lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2) qed simp lemma cInf_le_finite: "finite X \ x \ X \ Inf X \ x" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2) qed simp lemma cSup_eq_Sup_fin: "finite X \ X \ {} \ Sup X = Sup_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert) lemma cInf_eq_Inf_fin: "finite X \ X \ {} \ Inf X = Inf_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert) lemma cSup_atMost[simp]: "Sup {..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_greaterThanAtMost[simp]: "y < x \ Sup {y<..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = x" by (auto intro!: cSup_eq_maximum) lemma cInf_atLeast[simp]: "Inf {x..} = x" by (auto intro!: cInf_eq_minimum) lemma cInf_atLeastLessThan[simp]: "y < x \ Inf {y.. x \ Inf {y..x} = y" by (auto intro!: cInf_eq_minimum) lemma cINF_lower: "bdd_below (f ` A) \ x \ A \ \(f ` A) \ f x" using cInf_lower [of _ "f ` A"] by simp lemma cINF_greatest: "A \ {} \ (\x. x \ A \ m \ f x) \ m \ \(f ` A)" using cInf_greatest [of "f ` A"] by auto lemma cSUP_upper: "x \ A \ bdd_above (f ` A) \ f x \ \(f ` A)" using cSup_upper [of _ "f ` A"] by simp lemma cSUP_least: "A \ {} \ (\x. x \ A \ f x \ M) \ \(f ` A) \ M" using cSup_least [of "f ` A"] by auto lemma cINF_lower2: "bdd_below (f ` A) \ x \ A \ f x \ u \ \(f ` A) \ u" by (auto intro: cINF_lower order_trans) lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ \(f ` A)" by (auto intro: cSUP_upper order_trans) lemma cSUP_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro order.antisym cSUP_least) (auto intro: cSUP_upper) lemma cINF_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro order.antisym cINF_greatest) (auto intro: cINF_lower) lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ \(f ` A) \ (\x\A. u \ f x)" by (metis cINF_greatest cINF_lower order_trans) lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ \(f ` A) \ u \ (\x\A. f x \ u)" by (metis cSUP_least cSUP_upper order_trans) lemma less_cINF_D: "bdd_below (f`A) \ y < (\i\A. f i) \ i \ A \ y < f i" by (metis cINF_lower less_le_trans) lemma cSUP_lessD: "bdd_above (f`A) \ (\i\A. f i) < y \ i \ A \ f i < y" by (metis cSUP_upper le_less_trans) lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ \(f ` insert a A) = inf (f a) (\(f ` A))" by (simp add: cInf_insert) lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ \(f ` insert a A) = sup (f a) (\(f ` A))" by (simp add: cSup_insert) lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ \(f ` A) \ \(g ` B)" using cInf_mono [of "g ` B" "f ` A"] by auto lemma cSUP_mono: "A \ {} \ bdd_above (g ` B) \ (\n. n \ A \ \m\B. f n \ g m) \ \(f ` A) \ \(g ` B)" using cSup_mono [of "f ` A" "g ` B"] by auto lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ \(g ` B) \ \(f ` A)" by (rule cINF_mono) auto lemma cSUP_subset_mono: "\A \ {}; bdd_above (g ` B); A \ B; \x. x \ A \ f x \ g x\ \ \ (f ` A) \ \ (g ` B)" by (rule cSUP_mono) auto lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) lemma cSup_inter_less_eq: "bdd_above A \ bdd_above B \ A \ B \ {} \ Sup (A \ B) \ sup (Sup A) (Sup B) " by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) lemma cInf_union_distrib: "A \ {} \ bdd_below A \ B \ {} \ bdd_below B \ Inf (A \ B) = inf (Inf A) (Inf B)" by (intro order.antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) lemma cINF_union: "A \ {} \ bdd_below (f ` A) \ B \ {} \ bdd_below (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) lemma cSup_union_distrib: "A \ {} \ bdd_above A \ B \ {} \ bdd_above B \ Sup (A \ B) = sup (Sup A) (Sup B)" by (intro order.antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) lemma cSUP_union: "A \ {} \ bdd_above (f ` A) \ B \ {} \ bdd_above (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. inf (f a) (g a))" by (intro order.antisym le_infI cINF_greatest cINF_lower2) (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. sup (f a) (g a))" by (intro order.antisym le_supI cSUP_least cSUP_upper2) (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) lemma cInf_le_cSup: "A \ {} \ bdd_above A \ bdd_below A \ Inf A \ Sup A" by (auto intro!: cSup_upper2[of "SOME a. a \ A"] intro: someI cInf_lower) context fixes f :: "'a \ 'b::conditionally_complete_lattice" assumes "mono f" begin lemma mono_cInf: "\bdd_below A; A\{}\ \ f (Inf A) \ (INF x\A. f x)" by (simp add: \mono f\ conditionally_complete_lattice_class.cINF_greatest cInf_lower monoD) lemma mono_cSup: "\bdd_above A; A\{}\ \ (SUP x\A. f x) \ f (Sup A)" by (simp add: \mono f\ conditionally_complete_lattice_class.cSUP_least cSup_upper monoD) lemma mono_cINF: "\bdd_below (A`I); I\{}\ \ f (INF i\I. A i) \ (INF x\I. f (A x))" by (simp add: \mono f\ conditionally_complete_lattice_class.cINF_greatest cINF_lower monoD) lemma mono_cSUP: "\bdd_above (A`I); I\{}\ \ (SUP x\I. f (A x)) \ f (SUP i\I. A i)" by (simp add: \mono f\ conditionally_complete_lattice_class.cSUP_least cSUP_upper monoD) end end text \The special case of well-orderings\ lemma wellorder_InfI: fixes k :: "'a::{wellorder,conditionally_complete_lattice}" assumes "k \ A" shows "Inf A \ A" using wellorder_class.LeastI [of "\x. x \ A" k] by (simp add: Least_le assms cInf_eq_minimum) lemma wellorder_Inf_le1: fixes k :: "'a::{wellorder,conditionally_complete_lattice}" assumes "k \ A" shows "Inf A \ k" by (meson Least_le assms bdd_below.I cInf_lower) subsection \Complete lattices\ instance complete_lattice \ conditionally_complete_lattice by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) lemma cSup_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_bot}" assumes upper: "\x. x \ X \ x \ a" assumes least: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" proof cases assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cSup_eq_non_empty assms) lemma cInf_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_top}" assumes upper: "\x. x \ X \ a \ x" assumes least: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" proof cases assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cInf_eq_non_empty assms) class conditionally_complete_linorder = conditionally_complete_lattice + linorder begin lemma less_cSup_iff: "X \ {} \ bdd_above X \ y < Sup X \ (\x\X. y < x)" by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\X. x < y)" by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (\i\A. f i) < a \ (\x\A. f x < a)" using cInf_less_iff[of "f`A"] by auto lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (\i\A. f i) \ (\x\A. a < f x)" using less_cSup_iff[of "f`A"] by auto lemma less_cSupE: assumes "y < Sup X" "X \ {}" obtains x where "x \ X" "y < x" by (metis cSup_least assms not_le that) lemma less_cSupD: "X \ {} \ z < Sup X \ \x\X. z < x" by (metis less_cSup_iff not_le_imp_less bdd_above_def) lemma cInf_lessD: "X \ {} \ Inf X < z \ \x\X. x < z" by (metis cInf_less_iff not_le_imp_less bdd_below_def) lemma complete_interval: assumes "a < b" and "P a" and "\ P b" shows "\c. a \ c \ c \ b \ (\x. a \ x \ x < c \ P x) \ (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" -proof (rule exI [where x = "Sup {d. \x. a \ x \ x < d \ P x}"], auto) +proof (rule exI [where x = "Sup {d. \x. a \ x \ x < d \ P x}"], safe) show "a \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule cSup_upper, auto simp: bdd_above_def) (metis \a < b\ \\ P b\ linear less_le) next show "Sup {d. \c. a \ c \ c < d \ P c} \ b" - apply (rule cSup_least) - apply auto - apply (metis less_le_not_le) - apply (metis \a \\ P b\ linear less_le) - done + by (rule cSup_least) + (use \a \\ P b\ in \auto simp add: less_le_not_le\) next fix x assume x: "a \ x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" show "P x" - apply (rule less_cSupE [OF lt], auto) - apply (metis less_le_not_le) - apply (metis x) - done + by (rule less_cSupE [OF lt]) (use less_le_not_le x in \auto\) next fix d - assume 0: "\x. a \ x \ x < d \ P x" - thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" - by (rule_tac cSup_upper, auto simp: bdd_above_def) - (metis \a \\ P b\ linear less_le) + assume 0: "\x. a \ x \ x < d \ P x" + then have "d \ {d. \c. a \ c \ c < d \ P c}" + by auto + moreover have "bdd_above {d. \c. a \ c \ c < d \ P c}" + unfolding bdd_above_def using \a \\ P b\ linear + by (simp add: less_le) blast + ultimately show "d \ Sup {d. \c. a \ c \ c < d \ P c}" + by (auto simp: cSup_upper) qed end subsection \Instances\ instance complete_linorder < conditionally_complete_linorder .. lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Sup X = Max X" using cSup_eq_Sup_fin[of X] by (simp add: Sup_fin_Max) lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" using cInf_eq_Inf_fin[of X] by (simp add: Inf_fin_Min) lemma cSup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y" by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) lemma cInf_greaterThanLessThan[simp]: "y < x \ Inf {y<.. Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (simp add: cInf_eq_Min) lemma Sup_insert_finite: fixes S :: "'a::conditionally_complete_linorder set" shows "finite S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (simp add: cSup_insert sup_max) lemma finite_imp_less_Inf: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a < x\ \ a < Inf X" by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite) lemma finite_less_Inf_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a < Inf X \ (\x \ X. a < x)" by (auto simp: cInf_eq_Min) lemma finite_imp_Sup_less: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a > x\ \ a > Sup X" by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite) lemma finite_Sup_less_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a > Sup X \ (\x \ X. a > x)" by (auto simp: cSup_eq_Max) class linear_continuum = conditionally_complete_linorder + dense_linorder + assumes UNIV_not_singleton: "\a b::'a. a \ b" begin lemma ex_gt_or_lt: "\b. a < b \ b < a" by (metis UNIV_not_singleton neq_iff) end instantiation nat :: conditionally_complete_linorder begin definition "Sup (X::nat set) = (if X={} then 0 else Max X)" definition "Inf (X::nat set) = (LEAST n. n \ X)" lemma bdd_above_nat: "bdd_above X \ finite (X::nat set)" proof assume "bdd_above X" then obtain z where "X \ {.. z}" by (auto simp: bdd_above_def) then show "finite X" by (rule finite_subset) simp qed simp instance proof fix x :: nat fix X :: "nat set" show "Inf X \ x" if "x \ X" "bdd_below X" using that by (simp add: Inf_nat_def Least_le) show "x \ Inf X" if "X \ {}" "\y. y \ X \ x \ y" using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) show "x \ Sup X" if "x \ X" "bdd_above X" using that by (auto simp add: Sup_nat_def bdd_above_nat) show "Sup X \ x" if "X \ {}" "\y. y \ X \ y \ x" proof - from that have "bdd_above X" by (auto simp: bdd_above_def) with that show ?thesis by (simp add: Sup_nat_def bdd_above_nat) qed qed end lemma Inf_nat_def1: fixes K::"nat set" assumes "K \ {}" shows "Inf K \ K" by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI) lemma Sup_nat_empty [simp]: "Sup {} = (0::nat)" by (auto simp add: Sup_nat_def) instantiation int :: conditionally_complete_linorder begin definition "Sup (X::int set) = (THE x. x \ X \ (\y\X. y \ x))" definition "Inf (X::int set) = - (Sup (uminus ` X))" instance proof { fix x :: int and X :: "int set" assume "X \ {}" "bdd_above X" then obtain x y where "X \ {..y}" "x \ X" by (auto simp: bdd_above_def) then have *: "finite (X \ {x..y})" "X \ {x..y} \ {}" and "x \ y" by (auto simp: subset_eq) have "\!x\X. (\y\X. y \ x)" proof { fix z assume "z \ X" have "z \ Max (X \ {x..y})" proof cases assume "x \ z" with \z \ X\ \X \ {..y}\ *(1) show ?thesis by (auto intro!: Max_ge) next assume "\ x \ z" then have "z < x" by simp also have "x \ Max (X \ {x..y})" using \x \ X\ *(1) \x \ y\ by (intro Max_ge) auto finally show ?thesis by simp qed } note le = this with Max_in[OF *] show ex: "Max (X \ {x..y}) \ X \ (\z\X. z \ Max (X \ {x..y}))" by auto fix z assume *: "z \ X \ (\y\X. y \ z)" with le have "z \ Max (X \ {x..y})" by auto moreover have "Max (X \ {x..y}) \ z" using * ex by auto ultimately show "z = Max (X \ {x..y})" by auto qed then have "Sup X \ X \ (\y\X. y \ Sup X)" unfolding Sup_int_def by (rule theI') } note Sup_int = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_above X" then show "x \ Sup X" using Sup_int[of X] by auto } note le_Sup = this { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ y \ x" then show "Sup X \ x" using Sup_int[of X] by (auto simp: bdd_above_def) } note Sup_le = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_below X" then show "Inf X \ x" using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) } { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ x \ y" then show "x \ Inf X" using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) } qed end lemma interval_cases: fixes S :: "'a :: conditionally_complete_linorder set" assumes ivl: "\a b x. a \ S \ b \ S \ a \ x \ x \ b \ x \ S" shows "\a b. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" proof - define lower upper where "lower = {x. \s\S. s \ x}" and "upper = {x. \s\S. x \ s}" with ivl have "S = lower \ upper" by auto moreover have "\a. upper = UNIV \ upper = {} \ upper = {.. a} \ upper = {..< a}" proof cases assume *: "bdd_above S \ S \ {}" from * have "upper \ {.. Sup S}" by (auto simp: upper_def intro: cSup_upper2) moreover from * have "{..< Sup S} \ upper" by (force simp add: less_cSup_iff upper_def subset_eq Ball_def) ultimately have "upper = {.. Sup S} \ upper = {..< Sup S}" unfolding ivl_disj_un(2)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_above S \ S \ {})" then have "upper = UNIV \ upper = {}" by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le) then show ?thesis by auto qed moreover have "\b. lower = UNIV \ lower = {} \ lower = {b ..} \ lower = {b <..}" proof cases assume *: "bdd_below S \ S \ {}" from * have "lower \ {Inf S ..}" by (auto simp: lower_def intro: cInf_lower2) moreover from * have "{Inf S <..} \ lower" by (force simp add: cInf_less_iff lower_def subset_eq Ball_def) ultimately have "lower = {Inf S ..} \ lower = {Inf S <..}" unfolding ivl_disj_un(1)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_below S \ S \ {})" then have "lower = UNIV \ lower = {}" by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le) then show ?thesis by auto qed ultimately show ?thesis unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def by (metis inf_bot_left inf_bot_right inf_top.left_neutral inf_top.right_neutral) qed lemma cSUP_eq_cINF_D: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes eq: "(\x\A. f x) = (\x\A. f x)" and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)" and a: "a \ A" shows "f a = (\x\A. f x)" proof (rule antisym) show "f a \ \ (f ` A)" by (metis a bdd(1) eq cSUP_upper) show "\ (f ` A) \ f a" using a bdd by (auto simp: cINF_lower) qed lemma cSUP_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_above (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_above (f ` B x)" using bdd_UN by (meson UN_upper bdd_above_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_above_def) then have bdd2: "bdd_above ((\x. \z\B x. f z) ` A)" unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\ z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) ultimately show ?thesis by (rule order_antisym) qed lemma cINF_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_below (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_below (f ` B x)" using bdd_UN by (meson UN_upper bdd_below_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_below_def) then have bdd2: "bdd_below ((\x. \z\B x. f z) ` A)" unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) ultimately show ?thesis by (rule order_antisym) qed lemma cSup_abs_le: fixes S :: "('a::{linordered_idom,conditionally_complete_linorder}) set" shows "S \ {} \ (\x. x\S \ \x\ \ a) \ \Sup S\ \ a" apply (auto simp add: abs_le_iff intro: cSup_least) by (metis bdd_aboveI cSup_upper neg_le_iff_le order_trans) end diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,1244 +1,1263 @@ (* Title: HOL/Divides.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1999 University of Cambridge *) section \More on quotient and remainder\ theory Divides imports Parity begin subsection \More on division\ inductive eucl_rel_int :: "int \ int \ int \ int \ bool" where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)" | eucl_rel_int_dividesI: "l \ 0 \ k = q * l \ eucl_rel_int k l (q, 0)" | eucl_rel_int_remainderI: "sgn r = sgn l \ \r\ < \l\ \ k = q * l + r \ eucl_rel_int k l (q, r)" lemma eucl_rel_int_iff: "eucl_rel_int k l (q, r) \ k = l * q + r \ (if 0 < l then 0 \ r \ r < l else if l < 0 then l < r \ r \ 0 else q = 0)" by (cases "r = 0") (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI simp add: ac_simps sgn_1_pos sgn_1_neg) lemma unique_quotient_lemma: assumes "b * q' + r' \ b * q + r" "0 \ r'" "r' < b" "r < b" shows "q' \ (q::int)" proof - have "r' + b * (q'-q) \ r" using assms by (simp add: right_diff_distrib) moreover have "0 < b * (1 + q - q') " using assms by (simp add: right_diff_distrib distrib_left) moreover have "b * q' < b * (1 + q)" using assms by (simp add: right_diff_distrib distrib_left) ultimately show ?thesis using assms by (simp add: mult_less_cancel_left) qed lemma unique_quotient_lemma_neg: "b * q' + r' \ b*q + r \ r \ 0 \ b < r \ b < r' \ q \ (q'::int)" - by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma) auto + using unique_quotient_lemma[where b = "-b" and r = "-r'" and r'="-r"] by auto lemma unique_quotient: "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ q = q'" apply (rule order_antisym) apply (simp_all add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm) apply (blast intro: order_eq_refl [THEN unique_quotient_lemma] order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ done lemma unique_remainder: - "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ r = r'" -apply (subgoal_tac "q = q'") - apply (simp add: eucl_rel_int_iff) -apply (blast intro: unique_quotient) -done + assumes "eucl_rel_int a b (q, r)" + and "eucl_rel_int a b (q', r')" + shows "r = r'" +proof - + have "q = q'" + using assms by (blast intro: unique_quotient) + then show "r = r'" + using assms by (simp add: eucl_rel_int_iff) +qed lemma eucl_rel_int: "eucl_rel_int k l (k div l, k mod l)" proof (cases k rule: int_cases3) case zero then show ?thesis by (simp add: eucl_rel_int_iff divide_int_def modulo_int_def) next case (pos n) then show ?thesis using div_mult_mod_eq [of n] by (cases l rule: int_cases3) (auto simp del: of_nat_mult of_nat_add simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps eucl_rel_int_iff divide_int_def modulo_int_def) next case (neg n) then show ?thesis using div_mult_mod_eq [of n] by (cases l rule: int_cases3) (auto simp del: of_nat_mult of_nat_add simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps eucl_rel_int_iff divide_int_def modulo_int_def) qed lemma divmod_int_unique: assumes "eucl_rel_int k l (q, r)" shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" using assms eucl_rel_int [of k l] using unique_quotient [of k l] unique_remainder [of k l] by auto lemma div_abs_eq_div_nat: "\k\ div \l\ = int (nat \k\ div nat \l\)" by (simp add: divide_int_def) lemma mod_abs_eq_div_nat: "\k\ mod \l\ = int (nat \k\ mod nat \l\)" by (simp add: modulo_int_def) lemma zdiv_int: "int (a div b) = int a div int b" by (simp add: divide_int_def) lemma zmod_int: "int (a mod b) = int a mod int b" by (simp add: modulo_int_def) lemma div_sgn_abs_cancel: fixes k l v :: int assumes "v \ 0" shows "(sgn v * \k\) div (sgn v * \l\) = \k\ div \l\" proof - from assms have "sgn v = - 1 \ sgn v = 1" by (cases "v \ 0") auto then show ?thesis using assms unfolding divide_int_def [of "sgn v * \k\" "sgn v * \l\"] by (fastforce simp add: not_less div_abs_eq_div_nat) qed lemma div_eq_sgn_abs: fixes k l v :: int assumes "sgn k = sgn l" shows "k div l = \k\ div \l\" proof (cases "l = 0") case True then show ?thesis by simp next case False with assms have "(sgn k * \k\) div (sgn l * \l\) = \k\ div \l\" using div_sgn_abs_cancel [of l k l] by simp then show ?thesis by (simp add: sgn_mult_abs) qed lemma div_dvd_sgn_abs: fixes k l :: int assumes "l dvd k" shows "k div l = (sgn k * sgn l) * (\k\ div \l\)" proof (cases "k = 0 \ l = 0") case True then show ?thesis by auto next case False then have "k \ 0" and "l \ 0" by auto show ?thesis proof (cases "sgn l = sgn k") case True then show ?thesis by (auto simp add: div_eq_sgn_abs) next case False with \k \ 0\ \l \ 0\ have "sgn l * sgn k = - 1" by (simp add: sgn_if split: if_splits) with assms show ?thesis unfolding divide_int_def [of k l] by (auto simp add: zdiv_int ac_simps) qed qed lemma div_noneq_sgn_abs: fixes k l :: int assumes "l \ 0" assumes "sgn k \ sgn l" shows "k div l = - (\k\ div \l\) - of_bool (\ l dvd k)" using assms by (simp only: divide_int_def [of k l], auto simp add: not_less zdiv_int) subsubsection \Laws for div and mod with Unary Minus\ lemma zminus1_lemma: "eucl_rel_int a b (q, r) ==> b \ 0 ==> eucl_rel_int (-a) b (if r=0 then -q else -q - 1, if r=0 then 0 else b-r)" by (force simp add: eucl_rel_int_iff right_diff_distrib) lemma zdiv_zminus1_eq_if: "b \ (0::int) \ (-a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)" by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN div_int_unique]) lemma zmod_zminus1_eq_if: "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" proof (cases "b = 0") case False then show ?thesis by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN mod_int_unique]) qed auto lemma zmod_zminus1_not_zero: fixes k l :: int shows "- k mod l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zmod_zminus2_not_zero: fixes k l :: int shows "k mod - l \ 0 \ k mod l \ 0" by (simp add: mod_eq_0_iff_dvd) lemma zdiv_zminus2_eq_if: "b \ (0::int) ==> a div (-b) = (if a mod b = 0 then - (a div b) else - (a div b) - 1)" by (auto simp add: zdiv_zminus1_eq_if div_minus_right) lemma zmod_zminus2_eq_if: "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" by (auto simp add: zmod_zminus1_eq_if mod_minus_right) subsubsection \Monotonicity in the First Argument (Dividend)\ lemma zdiv_mono1: fixes b::int assumes "a \ a'" "0 < b" shows "a div b \ a' div b" proof (rule unique_quotient_lemma) show "b * (a div b) + a mod b \ b * (a' div b) + a' mod b" using assms(1) by auto qed (use assms in auto) lemma zdiv_mono1_neg: fixes b::int assumes "a \ a'" "b < 0" shows "a' div b \ a div b" proof (rule unique_quotient_lemma_neg) show "b * (a div b) + a mod b \ b * (a' div b) + a' mod b" using assms(1) by auto qed (use assms in auto) subsubsection \Monotonicity in the Second Argument (Divisor)\ lemma q_pos_lemma: fixes q'::int assumes "0 \ b'*q' + r'" "r' < b'" "0 < b'" shows "0 \ q'" proof - have "0 < b'* (q' + 1)" using assms by (simp add: distrib_left) with assms show ?thesis by (simp add: zero_less_mult_iff) qed lemma zdiv_mono2_lemma: fixes q'::int assumes eq: "b*q + r = b'*q' + r'" and le: "0 \ b'*q' + r'" and "r' < b'" "0 \ r" "0 < b'" "b' \ b" shows "q \ q'" proof - have "0 \ q'" using q_pos_lemma le \r' < b'\ \0 < b'\ by blast moreover have "b*q = r' - r + b'*q'" using eq by linarith ultimately have "b*q < b* (q' + 1)" using mult_right_mono assms unfolding distrib_left by fastforce with assms show ?thesis by (simp add: mult_less_cancel_left_pos) qed lemma zdiv_mono2: fixes a::int assumes "0 \ a" "0 < b'" "b' \ b" shows "a div b \ a div b'" proof (rule zdiv_mono2_lemma) have "b \ 0" using assms by linarith show "b * (a div b) + a mod b = b' * (a div b') + a mod b'" by simp qed (use assms in auto) lemma zdiv_mono2_neg_lemma: fixes q'::int assumes "b*q + r = b'*q' + r'" "b'*q' + r' < 0" "r < b" "0 \ r'" "0 < b'" "b' \ b" shows "q' \ q" proof - have "b'*q' < 0" using assms by linarith with assms have "q' \ 0" by (simp add: mult_less_0_iff) have "b*q' \ b'*q'" by (simp add: \q' \ 0\ assms(6) mult_right_mono_neg) then have "b*q' < b* (q + 1)" using assms by (simp add: distrib_left) then show ?thesis using assms by (simp add: mult_less_cancel_left) qed lemma zdiv_mono2_neg: fixes a::int assumes "a < 0" "0 < b'" "b' \ b" shows "a div b' \ a div b" proof (rule zdiv_mono2_neg_lemma) have "b \ 0" using assms by linarith show "b * (a div b) + a mod b = b' * (a div b') + a mod b'" by simp qed (use assms in auto) lemma div_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" shows "k div l = (k - l) div l + 1" proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with assms show ?thesis by (simp add: div_add_self2) qed lemma mod_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" shows "k mod l = (k - l) mod l" proof - have "k = (k - l) + l" by simp then obtain j where k: "k = j + l" .. with assms show ?thesis by simp qed subsubsection \Splitting Rules for div and mod\ text\The proofs of the two lemmas below are essentially identical\ lemma split_pos_lemma: "0 P(n div k :: int)(n mod k) = (\i j. 0\j \ j n = k*i + j \ P i j)" by auto lemma split_neg_lemma: "k<0 \ P(n div k :: int)(n mod k) = (\i j. k j\0 \ n = k*i + j \ P i j)" by auto lemma split_zdiv: "P(n div k :: int) = ((k = 0 \ P 0) \ (0 (\i j. 0\j \ j n = k*i + j \ P i)) \ (k<0 \ (\i j. k j\0 \ n = k*i + j \ P i)))" proof (cases "k = 0") case False then show ?thesis unfolding linorder_neq_iff by (auto simp add: split_pos_lemma [of concl: "\x y. P x"] split_neg_lemma [of concl: "\x y. P x"]) qed auto lemma split_zmod: "P(n mod k :: int) = ((k = 0 \ P n) \ (0 (\i j. 0\j \ j n = k*i + j \ P j)) \ (k<0 \ (\i j. k j\0 \ n = k*i + j \ P j)))" proof (cases "k = 0") case False then show ?thesis unfolding linorder_neq_iff by (auto simp add: split_pos_lemma [of concl: "\x y. P y"] split_neg_lemma [of concl: "\x y. P y"]) qed auto text \Enable (lin)arith to deal with \<^const>\divide\ and \<^const>\modulo\ when these are applied to some constant that is of the form \<^term>\numeral k\:\ declare split_zdiv [of _ _ "numeral k", arith_split] for k declare split_zmod [of _ _ "numeral k", arith_split] for k subsubsection \Computing \div\ and \mod\ with shifting\ lemma pos_eucl_rel_int_mult_2: assumes "0 \ b" assumes "eucl_rel_int a b (q, r)" shows "eucl_rel_int (1 + 2*a) (2*b) (q, 1 + 2*r)" using assms unfolding eucl_rel_int_iff by auto lemma neg_eucl_rel_int_mult_2: assumes "b \ 0" assumes "eucl_rel_int (a + 1) b (q, r)" shows "eucl_rel_int (1 + 2*a) (2*b) (q, 2*r - 1)" using assms unfolding eucl_rel_int_iff by auto text\computing div by shifting\ lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" using pos_eucl_rel_int_mult_2 [OF _ eucl_rel_int] by (rule div_int_unique) lemma neg_zdiv_mult_2: assumes A: "a \ (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" using neg_eucl_rel_int_mult_2 [OF A eucl_rel_int] by (rule div_int_unique) lemma zdiv_numeral_Bit0 [simp]: "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = numeral v div (numeral w :: int)" unfolding numeral.simps unfolding mult_2 [symmetric] by (rule div_mult_mult1, simp) lemma zdiv_numeral_Bit1 [simp]: "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = (numeral v div (numeral w :: int))" unfolding numeral.simps unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zdiv_mult_2, simp) lemma pos_zmod_mult_2: fixes a b :: int assumes "0 \ a" shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" using pos_eucl_rel_int_mult_2 [OF assms eucl_rel_int] by (rule mod_int_unique) lemma neg_zmod_mult_2: fixes a b :: int assumes "a \ 0" shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" using neg_eucl_rel_int_mult_2 [OF assms eucl_rel_int] by (rule mod_int_unique) lemma zmod_numeral_Bit0 [simp]: "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = (2::int) * (numeral v mod numeral w)" unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] by (rule mod_mult_mult1) lemma zmod_numeral_Bit1 [simp]: "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = 2 * (numeral v mod numeral w) + (1::int)" unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zmod_mult_2, simp) lemma zdiv_eq_0_iff: "i div k = 0 \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" (is "?L = ?R") for i k :: int proof assume ?L moreover have "?L \ ?R" by (rule split_zdiv [THEN iffD2]) simp ultimately show ?R by blast next assume ?R then show ?L by auto qed lemma zmod_trivial_iff: fixes i k :: int shows "i mod k = i \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" proof - have "i mod k = i \ i div k = 0" - by safe (insert div_mult_mod_eq [of i k], auto) + using div_mult_mod_eq [of i k] by safe auto with zdiv_eq_0_iff show ?thesis by simp qed subsubsection \Quotients of Signs\ lemma div_eq_minus1: "0 < b \ - 1 div b = - 1" for b :: int by (simp add: divide_int_def) lemma zmod_minus1: "0 < b \ - 1 mod b = b - 1" for b :: int by (auto simp add: modulo_int_def) lemma minus_mod_int_eq: \- k mod l = l - 1 - (k - 1) mod l\ if \l \ 0\ for k l :: int proof (cases \l = 0\) case True then show ?thesis by simp next case False with that have \l > 0\ by simp then show ?thesis proof (cases \l dvd k\) case True then obtain j where \k = l * j\ .. moreover have \(l * j mod l - 1) mod l = l - 1\ using \l > 0\ by (simp add: zmod_minus1) then have \(l * j - 1) mod l = l - 1\ by (simp only: mod_simps) ultimately show ?thesis by simp next case False - moreover have \0 < k mod l\ \k mod l < 1 + l\ - using \0 < l\ le_imp_0_less False apply auto - using le_less apply fastforce - using pos_mod_bound [of l k] apply linarith - done - with \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ + moreover have 1: \0 < k mod l\ + using \0 < l\ False le_less by fastforce + moreover have 2: \k mod l < 1 + l\ + using \0 < l\ pos_mod_bound[of l k] by linarith + from 1 2 \l > 0\ have \(k mod l - 1) mod l = k mod l - 1\ by (simp add: zmod_trivial_iff) ultimately show ?thesis - apply (simp only: zmod_zminus1_eq_if) - apply (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) - done + by (simp only: zmod_zminus1_eq_if) + (simp add: mod_eq_0_iff_dvd algebra_simps mod_simps) qed qed lemma div_neg_pos_less0: fixes a::int assumes "a < 0" "0 < b" shows "a div b < 0" proof - have "a div b \ - 1 div b" using zdiv_mono1 assms by auto also have "... \ -1" by (simp add: assms(2) div_eq_minus1) finally show ?thesis by force qed lemma div_nonneg_neg_le0: "[| (0::int) \ a; b < 0 |] ==> a div b \ 0" by (drule zdiv_mono1_neg, auto) lemma div_nonpos_pos_le0: "[| (a::int) \ 0; b > 0 |] ==> a div b \ 0" by (drule zdiv_mono1, auto) text\Now for some equivalences of the form \a div b >=< 0 \ \\ conditional upon the sign of \a\ or \b\. There are many more. They should all be simp rules unless that causes too much search.\ lemma pos_imp_zdiv_nonneg_iff: fixes a::int assumes "0 < b" shows "(0 \ a div b) = (0 \ a)" proof show "0 \ a div b \ 0 \ a" using assms by (simp add: linorder_not_less [symmetric]) (blast intro: div_neg_pos_less0) next assume "0 \ a" then have "0 div b \ a div b" using zdiv_mono1 assms by blast then show "0 \ a div b" by auto qed lemma pos_imp_zdiv_pos_iff: "0 0 < (i::int) div k \ k \ i" using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k] by arith lemma neg_imp_zdiv_nonneg_iff: fixes a::int assumes "b < 0" shows "(0 \ a div b) = (a \ 0)" using assms by (simp add: div_minus_minus [of a, symmetric] pos_imp_zdiv_nonneg_iff del: div_minus_minus) (*But not (a div b \ 0 iff a\0); consider a=1, b=2 when a div b = 0.*) lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) (*Again the law fails for \: consider a = -1, b = -2 when a div b = 0*) lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) lemma nonneg1_imp_zdiv_pos_iff: fixes a::int assumes "0 \ a" shows "a div b > 0 \ a \ b \ b>0" proof - have "0 < a div b \ b \ a" using div_pos_pos_trivial[of a b] assms by arith moreover have "0 < a div b \ b > 0" using assms div_nonneg_neg_le0[of a b] by(cases "b=0"; force) moreover have "b \ a \ 0 < b \ 0 < a div b" using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp ultimately show ?thesis by blast qed lemma zmod_le_nonneg_dividend: "(m::int) \ 0 \ m mod k \ m" by (rule split_zmod[THEN iffD2]) (fastforce dest: q_pos_lemma intro: split_mult_pos_le) subsubsection \Further properties\ lemma div_int_pos_iff: "k div l \ 0 \ k = 0 \ l = 0 \ k \ 0 \ l \ 0 \ k < 0 \ l < 0" for k l :: int proof (cases "k = 0 \ l = 0") case False + then have *: "k \ 0" "l \ 0" + by auto + then have "0 \ k div l \ \ k < 0 \ 0 \ l" + by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) then show ?thesis - apply (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) - by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) + using * by (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) qed auto lemma mod_int_pos_iff: "k mod l \ 0 \ l dvd k \ l = 0 \ k \ 0 \ l > 0" for k l :: int proof (cases "l > 0") case False then show ?thesis by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \auto simp add: le_less not_less\) qed auto text \Simplify expressions in which div and mod combine numerical constants\ lemma int_div_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a div b = q" by (rule div_int_unique [of a b q r]) (simp add: eucl_rel_int_iff) lemma int_div_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a div b = q" by (rule div_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma int_mod_pos_eq: "\(a::int) = b * q + r; 0 \ r; r < b\ \ a mod b = r" by (rule mod_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma int_mod_neg_eq: "\(a::int) = b * q + r; r \ 0; b < r\ \ a mod b = r" by (rule mod_int_unique [of a b q r], simp add: eucl_rel_int_iff) lemma abs_div: "(y::int) dvd x \ \x div y\ = \x\ div \y\" unfolding dvd_def by (cases "y=0") (auto simp add: abs_mult) text\Suggested by Matthias Daum\ lemma int_power_div_base: fixes k :: int assumes "0 < m" "0 < k" shows "k ^ m div k = (k::int) ^ (m - Suc 0)" proof - have eq: "k ^ m = k ^ ((m - Suc 0) + Suc 0)" by (simp add: assms) show ?thesis using assms by (simp only: power_add eq) auto qed text\Suggested by Matthias Daum\ lemma int_div_less_self: fixes x::int assumes "0 < x" "1 < k" shows "x div k < x" proof - have "nat x div nat k < nat x" by (simp add: assms) with assms show ?thesis by (simp add: nat_div_distrib [symmetric]) qed lemma mod_eq_dvd_iff_nat: "m mod q = n mod q \ q dvd m - n" if "m \ n" for m n q :: nat proof - have "int m mod int q = int n mod int q \ int q dvd int m - int n" by (simp add: mod_eq_dvd_iff) with that have "int (m mod q) = int (n mod q) \ int q dvd int (m - n)" by (simp only: of_nat_mod of_nat_diff) then show ?thesis by simp qed lemma mod_eq_nat1E: fixes m n q :: nat assumes "m mod q = n mod q" and "m \ n" obtains s where "m = n + q * s" proof - from assms have "q dvd m - n" by (simp add: mod_eq_dvd_iff_nat) then obtain s where "m - n = q * s" .. with \m \ n\ have "m = n + q * s" by simp with that show thesis . qed lemma mod_eq_nat2E: fixes m n q :: nat assumes "m mod q = n mod q" and "n \ m" obtains s where "n = m + q * s" using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) lemma nat_mod_eq_lemma: assumes "(x::nat) mod n = y mod n" and "y \ x" shows "\q. x = y + n * q" - using assms by (rule mod_eq_nat1E) rule + using assms by (rule mod_eq_nat1E) (rule exI) lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" (is "?lhs = ?rhs") proof assume H: "x mod n = y mod n" {assume xy: "x \ y" from H have th: "y mod n = x mod n" by simp from nat_mod_eq_lemma[OF th xy] have ?rhs - apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} + proof + fix q + assume "y = x + n * q" + then have "x + n * q = y + n * 0" + by simp + then show "\q1 q2. x + n * q1 = y + n * q2" + by blast + qed} moreover {assume xy: "y \ x" from nat_mod_eq_lemma[OF H xy] have ?rhs - apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} + proof + fix q + assume "x = y + n * q" + then have "x + n * 0 = y + n * q" + by simp + then show "\q1 q2. x + n * q1 = y + n * q2" + by blast + qed} ultimately show ?rhs using linear[of x y] by blast next assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp thus ?lhs by simp qed subsection \Numeral division with a pragmatic type class\ text \ The following type class contains everything necessary to formulate a division algorithm in ring structures with numerals, restricted to its positive segments. This is its primary motivation, and it could surely be formulated using a more fine-grained, more algebraic and less technical class hierarchy. \ class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom + assumes div_less: "0 \ a \ a < b \ a div b = 0" and mod_less: " 0 \ a \ a < b \ a mod b = a" and div_positive: "0 < b \ b \ a \ a div b > 0" and mod_less_eq_dividend: "0 \ a \ a mod b \ a" and pos_mod_bound: "0 < b \ a mod b < b" and pos_mod_sign: "0 < b \ 0 \ a mod b" and mod_mult2_eq: "0 \ c \ a mod (b * c) = b * (a div b mod c) + a mod b" and div_mult2_eq: "0 \ c \ a div (b * c) = a div b div c" assumes discrete: "a < b \ a + 1 \ b" fixes divmod :: "num \ num \ 'a \ 'a" and divmod_step :: "num \ 'a \ 'a \ 'a \ 'a" assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)" and divmod_step_def: "divmod_step l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" \ \These are conceptually definitions but force generated code to be monomorphic wrt. particular instances of this class which yields a significant speedup.\ begin lemma divmod_digit_1: assumes "0 \ a" "0 < b" and "b \ a mod (2 * b)" shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") and "a mod (2 * b) - b = a mod b" (is "?Q") proof - from assms mod_less_eq_dividend [of a "2 * b"] have "b \ a" by (auto intro: trans) with \0 < b\ have "0 < a div b" by (auto intro: div_positive) then have [simp]: "1 \ a div b" by (simp add: discrete) with \0 < b\ have mod_less: "a mod b < b" by (simp add: pos_mod_bound) define w where "w = a div b mod 2" then have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) from assms w_exhaust have "w = 1" - by (auto simp add: mod_w) (insert mod_less, auto) + using mod_less by (auto simp add: mod_w) with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp have "2 * (a div (2 * b)) = a div b - w" by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) with \w = 1\ have div: "2 * (a div (2 * b)) = a div b - 1" by simp then show ?P and ?Q by (simp_all add: div mod add_implies_diff [symmetric]) qed lemma divmod_digit_0: assumes "0 < b" and "a mod (2 * b) < b" shows "2 * (a div (2 * b)) = a div b" (is "?P") and "a mod (2 * b) = a mod b" (is "?Q") proof - define w where "w = a div b mod 2" then have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) moreover have "b \ a mod b + b" proof - from \0 < b\ pos_mod_sign have "0 \ a mod b" by blast then have "0 + b \ a mod b + b" by (rule add_right_mono) then show ?thesis by simp qed moreover note assms w_exhaust ultimately have "w = 0" by auto with mod_w have mod: "a mod (2 * b) = a mod b" by simp have "2 * (a div (2 * b)) = a div b - w" by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) with \w = 0\ have div: "2 * (a div (2 * b)) = a div b" by simp then show ?P and ?Q by (simp_all add: div mod) qed lemma mod_double_modulus: assumes "m > 0" "x \ 0" shows "x mod (2 * m) = x mod m \ x mod (2 * m) = x mod m + m" proof (cases "x mod (2 * m) < m") case True thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto next case False hence *: "x mod (2 * m) - m = x mod m" using assms by (intro divmod_digit_1) auto hence "x mod (2 * m) = x mod m + m" by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto) thus ?thesis by simp qed lemma fst_divmod: "fst (divmod m n) = numeral m div numeral n" by (simp add: divmod_def) lemma snd_divmod: "snd (divmod m n) = numeral m mod numeral n" by (simp add: divmod_def) text \ This is a formulation of one step (referring to one digit position) in school-method division: compare the dividend at the current digit position with the remainder from previous division steps and evaluate accordingly. \ lemma divmod_step_eq [simp]: "divmod_step l (q, r) = (if numeral l \ r then (2 * q + 1, r - numeral l) else (2 * q, r))" by (simp add: divmod_step_def) text \ This is a formulation of school-method division. If the divisor is smaller than the dividend, terminate. If not, shift the dividend to the right until termination occurs and then reiterate single division steps in the opposite direction. \ lemma divmod_divmod_step: "divmod m n = (if m < n then (0, numeral m) else divmod_step n (divmod m (Num.Bit0 n)))" proof (cases "m < n") case True then have "numeral m < numeral n" by simp then show ?thesis by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod) next case False have "divmod m n = divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n))" proof (cases "numeral n \ numeral m mod (2 * numeral n)") case True with divmod_step_eq have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)" by simp moreover from True divmod_digit_1 [of "numeral m" "numeral n"] have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n" and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n" by simp_all ultimately show ?thesis by (simp only: divmod_def) next case False then have *: "numeral m mod (2 * numeral n) < numeral n" by (simp add: not_le) with divmod_step_eq have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))" by auto moreover from * divmod_digit_0 [of "numeral n" "numeral m"] have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n" and "numeral m mod (2 * numeral n) = numeral m mod numeral n" by (simp_all only: zero_less_numeral) ultimately show ?thesis by (simp only: divmod_def) qed then have "divmod m n = divmod_step n (numeral m div numeral (Num.Bit0 n), numeral m mod numeral (Num.Bit0 n))" by (simp only: numeral.simps distrib mult_1) then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))" by (simp add: divmod_def) with False show ?thesis by simp qed text \The division rewrite proper -- first, trivial results involving \1\\ lemma divmod_trivial [simp]: "divmod m Num.One = (numeral m, 0)" "divmod num.One (num.Bit0 n) = (0, Numeral1)" "divmod num.One (num.Bit1 n) = (0, Numeral1)" using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) text \Division by an even number is a right-shift\ lemma divmod_cancel [simp]: "divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r))" (is ?P) "divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \ (q, 2 * r + 1))" (is ?Q) proof - have *: "\q. numeral (Num.Bit0 q) = 2 * numeral q" "\q. numeral (Num.Bit1 q) = 2 * numeral q + 1" by (simp_all only: numeral_mult numeral.simps distrib) simp_all have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less) then show ?P and ?Q by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1 div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] add.commute del: numeral_times_numeral) qed text \The really hard work\ lemma divmod_steps [simp]: "divmod (num.Bit0 m) (num.Bit1 n) = (if m \ n then (0, numeral (num.Bit0 m)) else divmod_step (num.Bit1 n) (divmod (num.Bit0 m) (num.Bit0 (num.Bit1 n))))" "divmod (num.Bit1 m) (num.Bit1 n) = (if m < n then (0, numeral (num.Bit1 m)) else divmod_step (num.Bit1 n) (divmod (num.Bit1 m) (num.Bit0 (num.Bit1 n))))" by (simp_all add: divmod_divmod_step) lemmas divmod_algorithm_code = divmod_step_eq divmod_trivial divmod_cancel divmod_steps text \Special case: divisibility\ definition divides_aux :: "'a \ 'a \ bool" where "divides_aux qr \ snd qr = 0" lemma divides_aux_eq [simp]: "divides_aux (q, r) \ r = 0" by (simp add: divides_aux_def) lemma dvd_numeral_simp [simp]: "numeral m dvd numeral n \ divides_aux (divmod n m)" by (simp add: divmod_def mod_eq_0_iff_dvd) text \Generic computation of quotient and remainder\ lemma numeral_div_numeral [simp]: "numeral k div numeral l = fst (divmod k l)" by (simp add: fst_divmod) lemma numeral_mod_numeral [simp]: "numeral k mod numeral l = snd (divmod k l)" by (simp add: snd_divmod) lemma one_div_numeral [simp]: "1 div numeral n = fst (divmod num.One n)" by (simp add: fst_divmod) lemma one_mod_numeral [simp]: "1 mod numeral n = snd (divmod num.One n)" by (simp add: snd_divmod) text \Computing congruences modulo \2 ^ q\\ lemma cong_exp_iff_simps: "numeral n mod numeral Num.One = 0 \ True" "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0 \ numeral n mod numeral q = 0" "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0 \ False" "numeral m mod numeral Num.One = (numeral n mod numeral Num.One) \ True" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ True" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ False" "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ (numeral n mod numeral q) = 0" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ numeral m mod numeral q = (numeral n mod numeral q)" "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) \ (numeral m mod numeral q) = 0" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) \ False" "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) \ numeral m mod numeral q = (numeral n mod numeral q)" by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even]) end hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq instantiation nat :: unique_euclidean_semiring_numeral begin definition divmod_nat :: "num \ num \ nat \ nat" where divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_nat :: "num \ nat \ nat \ nat \ nat" where "divmod_step_nat l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance by standard (auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq) end declare divmod_algorithm_code [where ?'a = nat, code] lemma Suc_0_div_numeral [simp]: fixes k l :: num shows "Suc 0 div numeral k = fst (divmod Num.One k)" by (simp_all add: fst_divmod) lemma Suc_0_mod_numeral [simp]: fixes k l :: num shows "Suc 0 mod numeral k = snd (divmod Num.One k)" by (simp_all add: snd_divmod) instantiation int :: unique_euclidean_semiring_numeral begin definition divmod_int :: "num \ num \ int \ int" where "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_int :: "num \ int \ int \ int \ int" where "divmod_step_int l qr = (let (q, r) = qr in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" instance by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) end declare divmod_algorithm_code [where ?'a = int, code] context begin qualified definition adjust_div :: "int \ int \ int" where "adjust_div qr = (let (q, r) = qr in q + of_bool (r \ 0))" qualified lemma adjust_div_eq [simp, code]: "adjust_div (q, r) = q + of_bool (r \ 0)" by (simp add: adjust_div_def) qualified definition adjust_mod :: "int \ int \ int" where [simp]: "adjust_mod l r = (if r = 0 then 0 else l - r)" lemma minus_numeral_div_numeral [simp]: "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma minus_numeral_mod_numeral [simp]: "- numeral m mod numeral n = adjust_mod (numeral n) (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma numeral_div_minus_numeral [simp]: "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" proof - have "int (fst (divmod m n)) = fst (divmod m n)" by (simp only: fst_divmod divide_int_def) auto then show ?thesis by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) qed lemma numeral_mod_minus_numeral [simp]: "numeral m mod - numeral n = - adjust_mod (numeral n) (snd (divmod m n) :: int)" proof (cases "snd (divmod m n) = (0::int)") case True then show ?thesis by (simp add: mod_eq_0_iff_dvd divides_aux_def) next case False then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \ (0::int)" by (simp only: snd_divmod modulo_int_def) auto then show ?thesis by (simp add: divides_aux_def adjust_div_def) (simp add: divides_aux_def modulo_int_def) qed lemma minus_one_div_numeral [simp]: "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" using minus_numeral_div_numeral [of Num.One n] by simp lemma minus_one_mod_numeral [simp]: "- 1 mod numeral n = adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" using minus_numeral_mod_numeral [of Num.One n] by simp lemma one_div_minus_numeral [simp]: "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" using numeral_div_minus_numeral [of Num.One n] by simp lemma one_mod_minus_numeral [simp]: "1 mod - numeral n = - adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" using numeral_mod_minus_numeral [of Num.One n] by simp end lemma divmod_BitM_2_eq [simp]: \divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\ by (cases m) simp_all lemma div_positive_int: "k div l > 0" if "k \ l" and "l > 0" for k l :: int using that div_positive [of l k] by blast subsubsection \Dedicated simproc for calculation\ text \ There is space for improvement here: the calculation itself could be carried out outside the logic, and a generic simproc (simplifier setup) for generic calculation would be helpful. \ simproc_setup numeral_divmod ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 div - 1 :: int" | "0 mod - 1 :: int" | "0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | "1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 div - 1 :: int" | "1 mod - 1 :: int" | "1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | "numeral a div 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \ let val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\If\); fun successful_rewrite ctxt ct = let val thm = Simplifier.rewrite ctxt ct in if Thm.is_reflexive thm then NONE else SOME thm end; in fn phi => let val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1 one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral one_div_minus_numeral one_mod_minus_numeral numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral numeral_div_minus_numeral numeral_mod_minus_numeral div_minus_minus mod_minus_minus Divides.adjust_div_eq of_bool_eq one_neq_zero numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial divmod_cancel divmod_steps divmod_step_eq fst_conv snd_conv numeral_One case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right minus_minus numeral_times_numeral mult_zero_right mult_1_right} @ [@{lemma "0 = 0 \ True" by simp}]); fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end end \ subsubsection \Code generation\ definition divmod_nat :: "nat \ nat \ nat \ nat" where "divmod_nat m n = (m div n, m mod n)" lemma fst_divmod_nat [simp]: "fst (divmod_nat m n) = m div n" by (simp add: divmod_nat_def) lemma snd_divmod_nat [simp]: "snd (divmod_nat m n) = m mod n" by (simp add: divmod_nat_def) lemma divmod_nat_if [code]: "Divides.divmod_nat m n = (if n = 0 \ m < n then (0, m) else let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))" by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) lemma [code]: "m div n = fst (divmod_nat m n)" "m mod n = snd (divmod_nat m n)" by simp_all lemma [code]: fixes k :: int shows "k div 0 = 0" "k mod 0 = k" "0 div k = 0" "0 mod k = 0" "k div Int.Pos Num.One = k" "k mod Int.Pos Num.One = 0" "k div Int.Neg Num.One = - k" "k mod Int.Neg Num.One = 0" "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" "Int.Neg m div Int.Pos n = - (Divides.adjust_div (divmod m n) :: int)" "Int.Neg m mod Int.Pos n = Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" "Int.Pos m div Int.Neg n = - (Divides.adjust_div (divmod m n) :: int)" "Int.Pos m mod Int.Neg n = - Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" by simp_all code_identifier code_module Divides \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Lemmas of doubtful value\ lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \ m < n" for m n :: nat by (rule le_div_geq) (use that in \simp_all add: not_less\) lemma mod_geq: "m mod n = (m - n) mod n" if "\ m < n" for m n :: nat by (rule le_mod_geq) (use that in \simp add: not_less\) lemma mod_eq_0D: "\q. m = d * q" if "m mod d = 0" for m d :: nat using that by (auto simp add: mod_eq_0_iff_dvd) lemma pos_mod_conj: "0 < b \ 0 \ a mod b \ a mod b < b" for a b :: int by simp lemma neg_mod_conj: "b < 0 \ a mod b \ 0 \ b < a mod b" for a b :: int by simp lemma zmod_eq_0_iff: "m mod d = 0 \ (\q. m = d * q)" for m d :: int by (auto simp add: mod_eq_0_iff_dvd) (* REVISIT: should this be generalized to all semiring_div types? *) lemma zmod_eq_0D [dest!]: "\q. m = d * q" if "m mod d = 0" for m d :: int using that by auto end diff --git a/src/HOL/Equiv_Relations.thy b/src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy +++ b/src/HOL/Equiv_Relations.thy @@ -1,588 +1,587 @@ (* Title: HOL/Equiv_Relations.thy Author: Lawrence C Paulson, 1996 Cambridge University Computer Laboratory *) section \Equivalence Relations in Higher-Order Set Theory\ theory Equiv_Relations imports BNF_Least_Fixpoint begin subsection \Equivalence relations -- set version\ definition equiv :: "'a set \ ('a \ 'a) set \ bool" where "equiv A r \ refl_on A r \ sym r \ trans r" lemma equivI: "refl_on A r \ sym r \ trans r \ equiv A r" by (simp add: equiv_def) lemma equivE: assumes "equiv A r" obtains "refl_on A r" and "sym r" and "trans r" using assms by (simp add: equiv_def) text \ Suppes, Theorem 70: \r\ is an equiv relation iff \r\ O r = r\. First half: \equiv A r \ r\ O r = r\. \ lemma sym_trans_comp_subset: "sym r \ trans r \ r\ O r \ r" unfolding trans_def sym_def converse_unfold by blast lemma refl_on_comp_subset: "refl_on A r \ r \ r\ O r" unfolding refl_on_def by blast lemma equiv_comp_eq: "equiv A r \ r\ O r = r" unfolding equiv_def by (iprover intro: sym_trans_comp_subset refl_on_comp_subset equalityI) text \Second half.\ lemma comp_equivI: assumes "r\ O r = r" "Domain r = A" shows "equiv A r" proof - have *: "\x y. (x, y) \ r \ (y, x) \ r" using assms by blast show ?thesis unfolding equiv_def refl_on_def sym_def trans_def using assms by (auto intro: *) qed subsection \Equivalence classes\ lemma equiv_class_subset: "equiv A r \ (a, b) \ r \ r``{a} \ r``{b}" \ \lemma for the next result\ unfolding equiv_def trans_def sym_def by blast theorem equiv_class_eq: "equiv A r \ (a, b) \ r \ r``{a} = r``{b}" by (intro equalityI equiv_class_subset; force simp add: equiv_def sym_def) lemma equiv_class_self: "equiv A r \ a \ A \ a \ r``{a}" unfolding equiv_def refl_on_def by blast lemma subset_equiv_class: "equiv A r \ r``{b} \ r``{a} \ b \ A \ (a, b) \ r" \ \lemma for the next result\ unfolding equiv_def refl_on_def by blast lemma eq_equiv_class: "r``{a} = r``{b} \ equiv A r \ b \ A \ (a, b) \ r" by (iprover intro: equalityD2 subset_equiv_class) lemma equiv_class_nondisjoint: "equiv A r \ x \ (r``{a} \ r``{b}) \ (a, b) \ r" unfolding equiv_def trans_def sym_def by blast lemma equiv_type: "equiv A r \ r \ A \ A" unfolding equiv_def refl_on_def by blast lemma equiv_class_eq_iff: "equiv A r \ (x, y) \ r \ r``{x} = r``{y} \ x \ A \ y \ A" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) lemma eq_equiv_class_iff: "equiv A r \ x \ A \ y \ A \ r``{x} = r``{y} \ (x, y) \ r" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) lemma disjnt_equiv_class: "equiv A r \ disjnt (r``{a}) (r``{b}) \ (a, b) \ r" by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def) subsection \Quotients\ definition quotient :: "'a set \ ('a \ 'a) set \ 'a set set" (infixl "'/'/" 90) where "A//r = (\x \ A. {r``{x}})" \ \set of equiv classes\ lemma quotientI: "x \ A \ r``{x} \ A//r" unfolding quotient_def by blast lemma quotientE: "X \ A//r \ (\x. X = r``{x} \ x \ A \ P) \ P" unfolding quotient_def by blast lemma Union_quotient: "equiv A r \ \(A//r) = A" unfolding equiv_def refl_on_def quotient_def by blast lemma quotient_disj: "equiv A r \ X \ A//r \ Y \ A//r \ X = Y \ X \ Y = {}" unfolding quotient_def equiv_def trans_def sym_def by blast lemma quotient_eqI: assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" "(x, y) \ r" shows "X = Y" proof - obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" using assms by (auto elim!: quotientE) then have "(a,b) \ r" using xy \equiv A r\ unfolding equiv_def sym_def trans_def by blast then show ?thesis unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) qed lemma quotient_eq_iff: assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" shows "X = Y \ (x, y) \ r" proof assume L: "X = Y" with assms show "(x, y) \ r" unfolding equiv_def sym_def trans_def by (blast elim!: quotientE) next assume \
: "(x, y) \ r" show "X = Y" by (rule quotient_eqI) (use \
assms in \blast+\) qed lemma eq_equiv_class_iff2: "equiv A r \ x \ A \ y \ A \ {x}//r = {y}//r \ (x, y) \ r" by (simp add: quotient_def eq_equiv_class_iff) lemma quotient_empty [simp]: "{}//r = {}" by (simp add: quotient_def) lemma quotient_is_empty [iff]: "A//r = {} \ A = {}" by (simp add: quotient_def) lemma quotient_is_empty2 [iff]: "{} = A//r \ A = {}" by (simp add: quotient_def) lemma singleton_quotient: "{x}//r = {r `` {x}}" by (simp add: quotient_def) lemma quotient_diff1: "inj_on (\a. {a}//r) A \ a \ A \ (A - {a})//r = A//r - {a}//r" unfolding quotient_def inj_on_def by blast subsection \Refinement of one equivalence relation WRT another\ lemma refines_equiv_class_eq: "R \ S \ equiv A R \ equiv A S \ R``(S``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_class_eq2: "R \ S \ equiv A R \ equiv A S \ S``(R``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_image_eq: "R \ S \ equiv A R \ equiv A S \ (\X. S``X) ` (A//R) = A//S" by (auto simp: quotient_def image_UN refines_equiv_class_eq2) lemma finite_refines_finite: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ finite (A//S)" by (erule finite_surj [where f = "\X. S``X"]) (simp add: refines_equiv_image_eq) lemma finite_refines_card_le: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ card (A//S) \ card (A//R)" by (subst refines_equiv_image_eq [of R S A, symmetric]) (auto simp: card_image_le [where f = "\X. S``X"]) subsection \Defining unary operations upon equivalence classes\ text \A congruence-preserving function.\ definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where "congruent r f \ (\(y, z) \ r. f y = f z)" lemma congruentI: "(\y z. (y, z) \ r \ f y = f z) \ congruent r f" by (auto simp add: congruent_def) lemma congruentD: "congruent r f \ (y, z) \ r \ f y = f z" by (auto simp add: congruent_def) abbreviation RESPECTS :: "('a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects" 80) where "f respects r \ congruent r f" lemma UN_constant_eq: "a \ A \ \y \ A. f y = c \ (\y \ A. f y) = c" \ \lemma required to prove \UN_equiv_class\\ by auto lemma UN_equiv_class: assumes "equiv A r" "f respects r" "a \ A" shows "(\x \ r``{a}. f x) = f a" \ \Conversion rule\ proof - have \
: "\x\r `` {a}. f x = f a" using assms unfolding equiv_def congruent_def sym_def by blast show ?thesis by (iprover intro: assms UN_constant_eq [OF equiv_class_self \
]) qed lemma UN_equiv_class_type: assumes r: "equiv A r" "f respects r" and X: "X \ A//r" and AB: "\x. x \ A \ f x \ B" shows "(\x \ X. f x) \ B" using assms unfolding quotient_def by (auto simp: UN_equiv_class [OF r]) text \ Sufficient conditions for injectiveness. Could weaken premises! major premise could be an inclusion; \bcong\ could be \\y. y \ A \ f y \ B\. \ lemma UN_equiv_class_inject: assumes "equiv A r" "f respects r" and eq: "(\x \ X. f x) = (\y \ Y. f y)" and X: "X \ A//r" and Y: "Y \ A//r" and fr: "\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r" shows "X = Y" proof - obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" using assms by (auto elim!: quotientE) then have "\ (f ` r `` {a}) = f a" "\ (f ` r `` {b}) = f b" by (iprover intro: UN_equiv_class [OF \equiv A r\] assms)+ then have "f a = f b" using eq unfolding a b by (iprover intro: trans sym) then have "(a,b) \ r" using fr \a \ A\ \b \ A\ by blast then show ?thesis unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) qed subsection \Defining binary operations upon equivalence classes\ text \A congruence-preserving function of two arguments.\ definition congruent2 :: "('a \ 'a) set \ ('b \ 'b) set \ ('a \ 'b \ 'c) \ bool" where "congruent2 r1 r2 f \ (\(y1, z1) \ r1. \(y2, z2) \ r2. f y1 y2 = f z1 z2)" lemma congruent2I': assumes "\y1 z1 y2 z2. (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" shows "congruent2 r1 r2 f" using assms by (auto simp add: congruent2_def) lemma congruent2D: "congruent2 r1 r2 f \ (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" by (auto simp add: congruent2_def) text \Abbreviation for the common case where the relations are identical.\ abbreviation RESPECTS2:: "('a \ 'a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects2" 80) where "f respects2 r \ congruent2 r r f" lemma congruent2_implies_congruent: "equiv A r1 \ congruent2 r1 r2 f \ a \ A \ congruent r2 (f a)" unfolding congruent_def congruent2_def equiv_def refl_on_def by blast lemma congruent2_implies_congruent_UN: assumes "equiv A1 r1" "equiv A2 r2" "congruent2 r1 r2 f" "a \ A2" shows "congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" unfolding congruent_def proof clarify fix c d assume cd: "(c,d) \ r1" then have "c \ A1" "d \ A1" using \equiv A1 r1\ by (auto elim!: equiv_type [THEN subsetD, THEN SigmaE2]) - with assms show "\ (f c ` r2 `` {a}) = \ (f d ` r2 `` {a})" - proof (simp add: UN_equiv_class congruent2_implies_congruent) - show "f c a = f d a" - using assms cd unfolding congruent2_def equiv_def refl_on_def by blast - qed + moreover have "f c a = f d a" + using assms cd unfolding congruent2_def equiv_def refl_on_def by blast + ultimately show "\ (f c ` r2 `` {a}) = \ (f d ` r2 `` {a})" + using assms by (simp add: UN_equiv_class congruent2_implies_congruent) qed lemma UN_equiv_class2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a1 \ A1 \ a2 \ A2 \ (\x1 \ r1``{a1}. \x2 \ r2``{a2}. f x1 x2) = f a1 a2" by (simp add: UN_equiv_class congruent2_implies_congruent congruent2_implies_congruent_UN) lemma UN_equiv_class_type2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ X1 \ A1//r1 \ X2 \ A2//r2 \ (\x1 x2. x1 \ A1 \ x2 \ A2 \ f x1 x2 \ B) \ (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" unfolding quotient_def by (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN congruent2_implies_congruent quotientI) lemma UN_UN_split_split_eq: "(\(x1, x2) \ X. \(y1, y2) \ Y. A x1 x2 y1 y2) = (\x \ X. \y \ Y. (\(x1, x2). (\(y1, y2). A x1 x2 y1 y2) y) x)" \ \Allows a natural expression of binary operators,\ \ \without explicit calls to \split\\ by auto lemma congruent2I: "equiv A1 r1 \ equiv A2 r2 \ (\y z w. w \ A2 \ (y,z) \ r1 \ f y w = f z w) \ (\y z w. w \ A1 \ (y,z) \ r2 \ f w y = f w z) \ congruent2 r1 r2 f" \ \Suggested by John Harrison -- the two subproofs may be\ \ \\<^emph>\much\ simpler than the direct proof.\ unfolding congruent2_def equiv_def refl_on_def by (blast intro: trans) lemma congruent2_commuteI: assumes equivA: "equiv A r" and commute: "\y z. y \ A \ z \ A \ f y z = f z y" and congt: "\y z w. w \ A \ (y,z) \ r \ f w y = f w z" shows "f respects2 r" proof (rule congruent2I [OF equivA equivA]) note eqv = equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2] show "\y z w. \w \ A; (y, z) \ r\ \ f y w = f z w" by (iprover intro: commute [THEN trans] sym congt elim: eqv) show "\y z w. \w \ A; (y, z) \ r\ \ f w y = f w z" by (iprover intro: congt elim: eqv) qed subsection \Quotients and finiteness\ text \Suggested by Florian Kammüller\ lemma finite_quotient: assumes "finite A" "r \ A \ A" shows "finite (A//r)" \ \recall @{thm equiv_type}\ proof - have "A//r \ Pow A" using assms unfolding quotient_def by blast moreover have "finite (Pow A)" using assms by simp ultimately show ?thesis by (iprover intro: finite_subset) qed lemma finite_equiv_class: "finite A \ r \ A \ A \ X \ A//r \ finite X" unfolding quotient_def by (erule rev_finite_subset) blast lemma equiv_imp_dvd_card: assumes "finite A" "equiv A r" "\X. X \ A//r \ k dvd card X" shows "k dvd card A" proof (rule Union_quotient [THEN subst]) show "k dvd card (\ (A // r))" apply (rule dvd_partition) using assms by (auto simp: Union_quotient dest: quotient_disj) qed (use assms in blast) subsection \Projection\ definition proj :: "('b \ 'a) set \ 'b \ 'a set" where "proj r x = r `` {x}" lemma proj_preserves: "x \ A \ proj r x \ A//r" unfolding proj_def by (rule quotientI) lemma proj_in_iff: assumes "equiv A r" shows "proj r x \ A//r \ x \ A" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: proj_preserves) next assume ?lhs then show ?rhs unfolding proj_def quotient_def - proof clarsimp + proof safe fix y assume y: "y \ A" and "r `` {x} = r `` {y}" moreover have "y \ r `` {y}" using assms y unfolding equiv_def refl_on_def by blast ultimately have "(x, y) \ r" by blast then show "x \ A" using assms unfolding equiv_def refl_on_def by blast qed qed lemma proj_iff: "equiv A r \ {x, y} \ A \ proj r x = proj r y \ (x, y) \ r" by (simp add: proj_def eq_equiv_class_iff) (* lemma in_proj: "\equiv A r; x \ A\ \ x \ proj r x" unfolding proj_def equiv_def refl_on_def by blast *) lemma proj_image: "proj r ` A = A//r" unfolding proj_def[abs_def] quotient_def by blast lemma in_quotient_imp_non_empty: "equiv A r \ X \ A//r \ X \ {}" unfolding quotient_def using equiv_class_self by fast lemma in_quotient_imp_in_rel: "equiv A r \ X \ A//r \ {x, y} \ X \ (x, y) \ r" using quotient_eq_iff[THEN iffD1] by fastforce lemma in_quotient_imp_closed: "equiv A r \ X \ A//r \ x \ X \ (x, y) \ r \ y \ X" unfolding quotient_def equiv_def trans_def by blast lemma in_quotient_imp_subset: "equiv A r \ X \ A//r \ X \ A" using in_quotient_imp_in_rel equiv_type by fastforce subsection \Equivalence relations -- predicate version\ text \Partial equivalences.\ definition part_equivp :: "('a \ 'a \ bool) \ bool" where "part_equivp R \ (\x. R x x) \ (\x y. R x y \ R x x \ R y y \ R x = R y)" \ \John-Harrison-style characterization\ lemma part_equivpI: "\x. R x x \ symp R \ transp R \ part_equivp R" by (auto simp add: part_equivp_def) (auto elim: sympE transpE) lemma part_equivpE: assumes "part_equivp R" obtains x where "R x x" and "symp R" and "transp R" proof - from assms have 1: "\x. R x x" and 2: "\x y. R x y \ R x x \ R y y \ R x = R y" unfolding part_equivp_def by blast+ from 1 obtain x where "R x x" .. moreover have "symp R" proof (rule sympI) fix x y assume "R x y" with 2 [of x y] show "R y x" by auto qed moreover have "transp R" proof (rule transpI) fix x y z assume "R x y" and "R y z" with 2 [of x y] 2 [of y z] show "R x z" by auto qed ultimately show thesis by (rule that) qed lemma part_equivp_refl_symp_transp: "part_equivp R \ (\x. R x x) \ symp R \ transp R" by (auto intro: part_equivpI elim: part_equivpE) lemma part_equivp_symp: "part_equivp R \ R x y \ R y x" by (erule part_equivpE, erule sympE) lemma part_equivp_transp: "part_equivp R \ R x y \ R y z \ R x z" by (erule part_equivpE, erule transpE) lemma part_equivp_typedef: "part_equivp R \ \d. d \ {c. \x. R x x \ c = Collect (R x)}" by (auto elim: part_equivpE) text \Total equivalences.\ definition equivp :: "('a \ 'a \ bool) \ bool" where "equivp R \ (\x y. R x y = (R x = R y))" \ \John-Harrison-style characterization\ lemma equivpI: "reflp R \ symp R \ transp R \ equivp R" by (auto elim: reflpE sympE transpE simp add: equivp_def) lemma equivpE: assumes "equivp R" obtains "reflp R" and "symp R" and "transp R" using assms by (auto intro!: that reflpI sympI transpI simp add: equivp_def) lemma equivp_implies_part_equivp: "equivp R \ part_equivp R" by (auto intro: part_equivpI elim: equivpE reflpE) lemma equivp_equiv: "equiv UNIV A \ equivp (\x y. (x, y) \ A)" by (auto intro!: equivI equivpI [to_set] elim!: equivE equivpE [to_set]) lemma equivp_reflp_symp_transp: "equivp R \ reflp R \ symp R \ transp R" by (auto intro: equivpI elim: equivpE) lemma identity_equivp: "equivp (=)" by (auto intro: equivpI reflpI sympI transpI) lemma equivp_reflp: "equivp R \ R x x" by (erule equivpE, erule reflpE) lemma equivp_symp: "equivp R \ R x y \ R y x" by (erule equivpE, erule sympE) lemma equivp_transp: "equivp R \ R x y \ R y z \ R x z" by (erule equivpE, erule transpE) lemma equivp_rtranclp: "symp r \ equivp r\<^sup>*\<^sup>*" by(intro equivpI reflpI sympI transpI)(auto dest: sympD[OF symp_rtranclp]) lemmas equivp_rtranclp_symclp [simp] = equivp_rtranclp[OF symp_symclp] lemma equivp_vimage2p: "equivp R \ equivp (vimage2p f f R)" by(auto simp add: equivp_def vimage2p_def dest: fun_cong) lemma equivp_imp_transp: "equivp R \ transp R" by(simp add: equivp_reflp_symp_transp) subsection \Equivalence closure\ definition equivclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" where "equivclp r = (symclp r)\<^sup>*\<^sup>*" lemma transp_equivclp [simp]: "transp (equivclp r)" by(simp add: equivclp_def) lemma reflp_equivclp [simp]: "reflp (equivclp r)" by(simp add: equivclp_def) lemma symp_equivclp [simp]: "symp (equivclp r)" by(simp add: equivclp_def) lemma equivp_evquivclp [simp]: "equivp (equivclp r)" by(simp add: equivpI) lemma tranclp_equivclp [simp]: "(equivclp r)\<^sup>+\<^sup>+ = equivclp r" by(simp add: equivclp_def) lemma rtranclp_equivclp [simp]: "(equivclp r)\<^sup>*\<^sup>* = equivclp r" by(simp add: equivclp_def) lemma symclp_equivclp [simp]: "symclp (equivclp r) = equivclp r" by(simp add: equivclp_def symp_symclp_eq) lemma equivclp_symclp [simp]: "equivclp (symclp r) = equivclp r" by(simp add: equivclp_def) lemma equivclp_conversep [simp]: "equivclp (conversep r) = equivclp r" by(simp add: equivclp_def) lemma equivclp_sym [sym]: "equivclp r x y \ equivclp r y x" by(rule sympD[OF symp_equivclp]) lemma equivclp_OO_equivclp_le_equivclp: "equivclp r OO equivclp r \ equivclp r" by(rule transp_relcompp_less_eq transp_equivclp)+ lemma rtranlcp_le_equivclp: "r\<^sup>*\<^sup>* \ equivclp r" unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) lemma rtranclp_conversep_le_equivclp: "r\\\<^sup>*\<^sup>* \ equivclp r" unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) lemma symclp_rtranclp_le_equivclp: "symclp r\<^sup>*\<^sup>* \ equivclp r" unfolding symclp_pointfree by(rule le_supI)(simp_all add: rtranclp_conversep[symmetric] rtranlcp_le_equivclp rtranclp_conversep_le_equivclp) lemma r_OO_conversep_into_equivclp: "r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>* \ equivclp r" by(blast intro: order_trans[OF _ equivclp_OO_equivclp_le_equivclp] relcompp_mono rtranlcp_le_equivclp rtranclp_conversep_le_equivclp del: predicate2I) lemma equivclp_induct [consumes 1, case_names base step, induct pred: equivclp]: assumes a: "equivclp r a b" and cases: "P a" "\y z. equivclp r a y \ r y z \ r z y \ P y \ P z" shows "P b" using a unfolding equivclp_def by(induction rule: rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) lemma converse_equivclp_induct [consumes 1, case_names base step]: assumes major: "equivclp r a b" and cases: "P b" "\y z. r y z \ r z y \ equivclp r z b \ P z \ P y" shows "P a" using major unfolding equivclp_def by(induction rule: converse_rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) lemma equivclp_refl [simp]: "equivclp r x x" by(rule reflpD[OF reflp_equivclp]) lemma r_into_equivclp [intro]: "r x y \ equivclp r x y" unfolding equivclp_def by(blast intro: symclpI) lemma converse_r_into_equivclp [intro]: "r y x \ equivclp r x y" unfolding equivclp_def by(blast intro: symclpI) lemma rtranclp_into_equivclp: "r\<^sup>*\<^sup>* x y \ equivclp r x y" using rtranlcp_le_equivclp[of r] by blast lemma converse_rtranclp_into_equivclp: "r\<^sup>*\<^sup>* y x \ equivclp r x y" by(blast intro: equivclp_sym rtranclp_into_equivclp) lemma equivclp_into_equivclp: "\ equivclp r a b; r b c \ r c b \ \ equivclp r a c" unfolding equivclp_def by(erule rtranclp.rtrancl_into_rtrancl)(auto intro: symclpI) lemma equivclp_trans [trans]: "\ equivclp r a b; equivclp r b c \ \ equivclp r a c" using equivclp_OO_equivclp_le_equivclp[of r] by blast hide_const (open) proj end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,2143 +1,2144 @@ (* Title: HOL/Euclidean_Division.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Division in euclidean (semi)rings\ theory Euclidean_Division imports Int Lattices_Big begin subsection \Euclidean (semi)rings with explicit division and remainder\ class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" assumes mod_size_less: "b \ 0 \ euclidean_size (a mod b) < euclidean_size b" assumes size_mult_mono: "b \ 0 \ euclidean_size a \ euclidean_size (a * b)" begin lemma euclidean_size_eq_0_iff [simp]: "euclidean_size b = 0 \ b = 0" proof assume "b = 0" then show "euclidean_size b = 0" by simp next assume "euclidean_size b = 0" show "b = 0" proof (rule ccontr) assume "b \ 0" with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . with \euclidean_size b = 0\ show False by simp qed qed lemma euclidean_size_greater_0_iff [simp]: "euclidean_size b > 0 \ b \ 0" using euclidean_size_eq_0_iff [symmetric, of b] by safe simp lemma size_mult_mono': "b \ 0 \ euclidean_size a \ euclidean_size (b * a)" by (subst mult.commute) (rule size_mult_mono) lemma dvd_euclidean_size_eq_imp_dvd: assumes "a \ 0" and "euclidean_size a = euclidean_size b" and "b dvd a" shows "a dvd b" proof (rule ccontr) assume "\ a dvd b" hence "b mod a \ 0" using mod_0_imp_dvd [of b a] by blast then have "b mod a \ 0" by (simp add: mod_eq_0_iff_dvd) from \b dvd a\ have "b dvd b mod a" by (simp add: dvd_mod_iff) then obtain c where "b mod a = b * c" unfolding dvd_def by blast with \b mod a \ 0\ have "c \ 0" by auto with \b mod a = b * c\ have "euclidean_size (b mod a) \ euclidean_size b" using size_mult_mono by force moreover from \\ a dvd b\ and \a \ 0\ have "euclidean_size (b mod a) < euclidean_size a" using mod_size_less by blast ultimately show False using \euclidean_size a = euclidean_size b\ by simp qed lemma euclidean_size_times_unit: assumes "is_unit a" shows "euclidean_size (a * b) = euclidean_size b" proof (rule antisym) from assms have [simp]: "a \ 0" by auto thus "euclidean_size (a * b) \ euclidean_size b" by (rule size_mult_mono') from assms have "is_unit (1 div a)" by simp hence "1 div a \ 0" by (intro notI) simp_all hence "euclidean_size (a * b) \ euclidean_size ((1 div a) * (a * b))" by (rule size_mult_mono') also from assms have "(1 div a) * (a * b) = b" by (simp add: algebra_simps unit_div_mult_swap) finally show "euclidean_size (a * b) \ euclidean_size b" . qed lemma euclidean_size_unit: "is_unit a \ euclidean_size a = euclidean_size 1" using euclidean_size_times_unit [of a 1] by simp lemma unit_iff_euclidean_size: "is_unit a \ euclidean_size a = euclidean_size 1 \ a \ 0" proof safe assume A: "a \ 0" and B: "euclidean_size a = euclidean_size 1" show "is_unit a" by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all qed (auto intro: euclidean_size_unit) lemma euclidean_size_times_nonunit: assumes "a \ 0" "b \ 0" "\ is_unit a" shows "euclidean_size b < euclidean_size (a * b)" proof (rule ccontr) assume "\euclidean_size b < euclidean_size (a * b)" with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" - by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (insert assms, simp_all) + by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) + (use assms in simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" - using assms by (auto elim!: dvdE simp: size_mult_mono) + using assms by (auto simp: size_mult_mono) lemma dvd_proper_imp_size_less: assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis by (auto intro!: euclidean_size_times_nonunit) qed lemma unit_imp_mod_eq_0: "a mod b = 0" if "is_unit b" using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" - by (rule; rule coprimeI) + by (rule iffI; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" shows "(c * b + a) div b = c + a div b" using assms by (simp add: add.commute) lemma div_mult_self4 [simp]: assumes "b \ 0" shows "(b * c + a) div b = c + a div b" using assms by (simp add: add.commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: div_mult_mod_eq) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: div_mult_mod_eq) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self4 [simp]: "(b * c + a) mod b = a mod b" by (simp add: add.commute) lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_add_self1: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_mult_eq) finally show ?thesis . qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" by simp_all lemma mod_mult_mult1: "(c * a) mod (c * b) = c * (a mod b)" proof (cases "c = 0") case True then show ?thesis by simp next case False from div_mult_mod_eq have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) = c * a + c * (a mod b)" by (simp add: algebra_simps) with div_mult_mod_eq show ?thesis by simp qed lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" by (fact mod_mult_mult1 [symmetric]) lemma dvd_mod: "k dvd m \ k dvd n \ k dvd (m mod n)" unfolding dvd_def by (auto simp add: mod_mult_mult1) lemma div_plus_div_distrib_dvd_left: "c dvd a \ (a + b) div c = a div c + b div c" by (cases "c = 0") (auto elim: dvdE) lemma div_plus_div_distrib_dvd_right: "c dvd b \ (a + b) div c = a div c + b div c" using div_plus_div_distrib_dvd_left [of c b a] by (simp add: ac_simps) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ lemma mod_add_left_eq [mod_simps]: "(a mod c + b) mod c = (a + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_add_right_eq [mod_simps]: "(a + b mod c) mod c = (a + b) mod c" using mod_add_left_eq [of b c a] by (simp add: ac_simps) lemma mod_add_eq: "(a mod c + b mod c) mod c = (a + b) mod c" by (simp add: mod_add_left_eq mod_add_right_eq) lemma mod_sum_eq [mod_simps]: "(\i\A. f i mod a) mod a = sum f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\A. f i) mod a) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. then show ?thesis by (simp add: mod_add_eq) qed text \Multiplication respects modular equivalence.\ lemma mod_mult_left_eq [mod_simps]: "((a mod c) * b) mod c = (a * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: div_mult_mod_eq) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_mult_right_eq [mod_simps]: "(a * (b mod c)) mod c = (a * b) mod c" using mod_mult_left_eq [of b c a] by (simp add: ac_simps) lemma mod_mult_eq: "((a mod c) * (b mod c)) mod c = (a * b) mod c" by (simp add: mod_mult_left_eq mod_mult_right_eq) lemma mod_prod_eq [mod_simps]: "(\i\A. f i mod a) mod a = prod f A mod a" proof (induct A rule: infinite_finite_induct) case (insert i A) then have "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\A. f i) mod a)) mod a" by (simp add: insert.hyps) finally show ?case by (simp add: insert.hyps mod_simps) qed simp_all lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. then show ?thesis by (simp add: mod_mult_eq) qed text \Exponentiation respects modular equivalence.\ lemma power_mod [mod_simps]: "((a mod b) ^ n) mod b = (a ^ n) mod b" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b" by (simp add: mod_mult_right_eq) with Suc show ?case by (simp add: mod_mult_left_eq mod_mult_right_eq) qed lemma power_diff_power_eq: \a ^ m div a ^ n = (if n \ m then a ^ (m - n) else 1 div a ^ (n - m))\ if \a \ 0\ proof (cases \n \ m\) case True with that power_diff [symmetric, of a n m] show ?thesis by simp next case False then obtain q where n: \n = m + Suc q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\ by (simp add: power_add ac_simps) moreover from that have \a ^ m \ 0\ by simp ultimately have \a ^ m div a ^ n = 1 div a ^ Suc q\ by (subst (asm) div_mult_mult1) simp with False n show ?thesis by simp qed end class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel begin subclass idom_divide .. lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" using div_mult_mult1 [of "- 1" a b] by simp lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)" using mod_mult_mult1 [of "- 1" a b] by simp lemma div_minus_right: "a div (- b) = (- a) div b" using div_minus_minus [of "- a" b] by simp lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)" using mod_minus_minus [of "- a" b] by simp lemma div_minus1_right [simp]: "a div (- 1) = - a" using div_minus_right [of a 1] by simp lemma mod_minus1_right [simp]: "a mod (- 1) = 0" using mod_minus_right [of a 1] by simp text \Negation respects modular equivalence.\ lemma mod_minus_eq [mod_simps]: "(- (a mod b)) mod b = (- a) mod b" proof - have "(- a) mod b = (- (a div b * b + a mod b)) mod b" by (simp only: div_mult_mod_eq) also have "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp add: ac_simps) also have "\ = (- (a mod b)) mod b" by (rule mod_mult_self1) finally show ?thesis by (rule sym) qed lemma mod_minus_cong: assumes "a mod b = a' mod b" shows "(- a) mod b = (- a') mod b" proof - have "(- (a mod b)) mod b = (- (a' mod b)) mod b" unfolding assms .. then show ?thesis by (simp add: mod_minus_eq) qed text \Subtraction respects modular equivalence.\ lemma mod_diff_left_eq [mod_simps]: "(a mod c - b) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp lemma mod_diff_right_eq [mod_simps]: "(a - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_eq: "(a mod c - b mod c) mod c = (a - b) mod c" using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp lemma mod_diff_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a - b) mod c = (a' - b') mod c" using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp lemma minus_mod_self2 [simp]: "(a - b) mod b = a mod b" using mod_diff_right_eq [of a b b] by (simp add: mod_diff_right_eq) lemma minus_mod_self1 [simp]: "(b - a) mod b = - a mod b" using mod_add_self2 [of "- a" b] by simp lemma mod_eq_dvd_iff: "a mod c = b mod c \ c dvd a - b" (is "?P \ ?Q") proof assume ?P then have "(a mod c - b mod c) mod c = 0" by simp then show ?Q by (simp add: dvd_eq_mod_eq_0 mod_simps) next assume ?Q then obtain d where d: "a - b = c * d" .. then have "a = c * d + b" by (simp add: algebra_simps) then show ?P by simp qed lemma mod_eqE: assumes "a mod c = b mod c" obtains d where "b = a + c * d" proof - from assms have "c dvd a - b" by (simp add: mod_eq_dvd_iff) then obtain d where "a - b = c * d" .. then have "b = a + c * - d" by (simp add: algebra_simps) with that show thesis . qed lemma invertible_coprime: "coprime a c" if "a * b mod c = 1" by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) end subsection \Uniquely determined division\ class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: "euclidean_size (a * b) = euclidean_size a * euclidean_size b" fixes division_segment :: "'a \ 'a" assumes is_unit_division_segment [simp]: "is_unit (division_segment a)" and division_segment_mult: "a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b" and division_segment_mod: "b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b" assumes div_bounded: "b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q" begin lemma division_segment_not_0 [simp]: "division_segment a \ 0" using is_unit_division_segment [of a] is_unitE [of "division_segment a"] by blast lemma divmod_cases [case_names divides remainder by0]: obtains (divides) q where "b \ 0" and "a div b = q" and "a mod b = 0" and "a = q * b" | (remainder) q r where "b \ 0" and "division_segment r = division_segment b" and "euclidean_size r < euclidean_size b" and "r \ 0" and "a div b = q" and "a mod b = r" and "a = q * b + r" | (by0) "b = 0" proof (cases "b = 0") case True then show thesis by (rule by0) next case False show thesis proof (cases "b dvd a") case True then obtain q where "a = b * q" .. with \b \ 0\ divides show thesis by (simp add: ac_simps) next case False then have "a mod b \ 0" by (simp add: mod_eq_0_iff_dvd) moreover from \b \ 0\ \\ b dvd a\ have "division_segment (a mod b) = division_segment b" by (rule division_segment_mod) moreover have "euclidean_size (a mod b) < euclidean_size b" using \b \ 0\ by (rule mod_size_less) moreover have "a = a div b * b + a mod b" by (simp add: div_mult_mod_eq) ultimately show thesis using \b \ 0\ by (blast intro!: remainder) qed qed lemma div_eqI: "a div b = q" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "(q * b + r) div b = q" by (auto intro: div_bounded) with that show ?thesis by simp qed lemma mod_eqI: "a mod b = r" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "a div b = q" by (rule div_eqI) moreover have "a div b * b + a mod b = a" by (fact div_mult_mod_eq) ultimately have "a div b * b + a mod b = a div b * b + r" using \q * b + r = a\ by simp then show ?thesis by simp qed subclass euclidean_semiring_cancel proof show "(a + c * b) div b = c + a div b" if "b \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 with \b \ 0\ show ?thesis by simp next case (divides q) then show ?thesis by (simp add: ac_simps) next case (remainder q r) then show ?thesis by (auto intro: div_eqI simp add: algebra_simps) qed next show"(c * a) div (c * b) = a div b" if "c \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 then show ?thesis by simp next case (divides q) with \c \ 0\ show ?thesis by (simp add: mult.left_commute [of c]) next case (remainder q r) from \b \ 0\ \c \ 0\ have "b * c \ 0" by simp from remainder \c \ 0\ have "division_segment (r * c) = division_segment (b * c)" and "euclidean_size (r * c) < euclidean_size (b * c)" by (simp_all add: division_segment_mult division_segment_mod euclidean_size_mult) with remainder show ?thesis by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps) (use \b * c \ 0\ in simp) qed qed lemma div_mult1_eq: "(a * b) div c = a * (b div c) + a * (b mod c) div c" proof (cases "a * (b mod c)" c rule: divmod_cases) case (divides q) have "a * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = (a * (b div c) + q) * c" using divides by (simp add: algebra_simps) finally have "(a * b) div c = \ div c" by simp with divides show ?thesis by simp next case (remainder q r) from remainder(1-3) show ?thesis proof (rule div_eqI) have "a * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = a * c * (b div c) + q * c + r" using remainder by (simp add: algebra_simps) finally show "(a * (b div c) + a * (b mod c) div c) * c + r = a * b" using remainder(5-7) by (simp add: algebra_simps) qed next case by0 then show ?thesis by simp qed lemma div_add1_eq: "(a + b) div c = a div c + b div c + (a mod c + b mod c) div c" proof (cases "a mod c + b mod c" c rule: divmod_cases) case (divides q) have "a + b = (a div c * c + a mod c) + (b div c * c + b mod c)" using mod_mult_div_eq [of a c] mod_mult_div_eq [of b c] by (simp add: ac_simps) also have "\ = (a div c + b div c) * c + (a mod c + b mod c)" by (simp add: algebra_simps) also have "\ = (a div c + b div c + q) * c" using divides by (simp add: algebra_simps) finally have "(a + b) div c = (a div c + b div c + q) * c div c" by simp with divides show ?thesis by simp next case (remainder q r) from remainder(1-3) show ?thesis proof (rule div_eqI) have "(a div c + b div c + q) * c + r + (a mod c + b mod c) = (a div c * c + a mod c) + (b div c * c + b mod c) + q * c + r" by (simp add: algebra_simps) also have "\ = a + b + (a mod c + b mod c)" by (simp add: div_mult_mod_eq remainder) (simp add: ac_simps) finally show "(a div c + b div c + (a mod c + b mod c) div c) * c + r = a + b" using remainder by simp qed next case by0 then show ?thesis by simp qed lemma div_eq_0_iff: "a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0" (is "_ \ ?P") if "division_segment a = division_segment b" proof assume ?P with that show "a div b = 0" by (cases "b = 0") (auto intro: div_eqI) next assume "a div b = 0" then have "a mod b = a" using div_mult_mod_eq [of a b] by simp with mod_size_less [of b a] show ?P by auto qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: "nat \ nat" where [simp]: "normalize = (id :: nat \ nat)" definition unit_factor_nat :: "nat \ nat" where "unit_factor n = (if n = 0 then 0 else 1 :: nat)" lemma unit_factor_simps [simp]: "unit_factor 0 = (0::nat)" "unit_factor (Suc n) = 1" by (simp_all add: unit_factor_nat_def) definition divide_nat :: "nat \ nat \ nat" where "m div n = (if n = 0 then 0 else Max {k::nat. k * n \ m})" instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: "nat \ nat" where [simp]: "euclidean_size_nat = id" definition division_segment_nat :: "nat \ nat" where [simp]: "division_segment_nat n = 1" definition modulo_nat :: "nat \ nat \ nat" where "m mod n = m - (m div n * (n::nat))" instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ l}" by (cases n) auto then show ?thesis by (rule finite_subset) simp qed have mult_div_unfold: "n * (m div n) = Max {l. l \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ n dvd l}" by (auto simp add: ac_simps elim!: dvdE) finally show ?thesis using False by (simp add: divide_nat_def ac_simps) qed have less_eq: "m div n * n \ m" by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI) then show "m div n * n + m mod n = m" by (simp add: modulo_nat_def) assume "n \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ k < Suc (m div n)" by auto ultimately have "Suc (m div n) < Suc (m div n)" by blast then show False by simp qed with \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ by (cases n) simp_all fix q r :: nat show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n" proof - from that have "r < n" by simp have "k \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end text \Tool support\ ML \ structure Cancel_Div_Mod_Nat = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val dest_plus = HOLogic.dest_bin \<^const_name>\Groups.plus\ HOLogic.natT; val mk_sum = Arith_Data.mk_sum; fun dest_sum tm = if HOLogic.is_zero tm then [] else (case try HOLogic.dest_Suc tm of SOME t => HOLogic.Suc_zero :: dest_sum t | NONE => (case try dest_plus tm of SOME (t, u) => dest_sum t @ dest_sum u | NONE => [tm])); val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_nat_eqI: "m div n = q" if "n * q \ m" and "m < n * Suc q" for m n q :: nat by (rule div_eqI [of _ "m - n * q"]) (use that in \simp_all add: algebra_simps\) lemma mod_nat_eqI: "m mod n = r" if "r < n" and "r \ m" and "n dvd m - r" for m n r :: nat by (rule mod_eqI [of _ _ "(m - r) div n"]) (use that in \simp_all add: algebra_simps\) lemma div_mult_self_is_m [simp]: "m * n div n = m" if "n > 0" for m n :: nat using that by simp lemma div_mult_self1_is_m [simp]: "n * m div n = m" if "n > 0" for m n :: nat using that by simp lemma mod_less_divisor [simp]: "m mod n < n" if "n > 0" for m n :: nat using mod_size_less [of n m] that by simp lemma mod_le_divisor [simp]: "m mod n \ n" if "n > 0" for m n :: nat using that by (auto simp add: le_less) lemma div_times_less_eq_dividend [simp]: "m div n * n \ m" for m n :: nat by (simp add: minus_mod_eq_div_mult [symmetric]) lemma times_div_less_eq_dividend [simp]: "n * (m div n) \ m" for m n :: nat using div_times_less_eq_dividend [of m n] by (simp add: ac_simps) lemma dividend_less_div_times: "m < n + (m div n) * n" if "0 < n" for m n :: nat proof - from that have "m mod n < n" by simp then show ?thesis by (simp add: minus_mod_eq_div_mult [symmetric]) qed lemma dividend_less_times_div: "m < n + n * (m div n)" if "0 < n" for m n :: nat using dividend_less_div_times [of n m] that by (simp add: ac_simps) lemma mod_Suc_le_divisor [simp]: "m mod Suc n \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ m" for m n :: nat proof (rule add_leD2) from div_mult_mod_eq have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma div_less [simp]: "m div n = 0" and mod_less [simp]: "m mod n = m" if "m < n" for m n :: nat using that by (auto intro: div_eqI mod_eqI) lemma le_div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) then show ?thesis by simp qed lemma div_if: "m div n = (if m < n \ n = 0 then 0 else Suc ((m - n) div n))" by (simp add: le_div_geq) lemma mod_if: "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat by (simp add: le_mod_geq) lemma div_eq_0_iff: "m div n = 0 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ n > 0" for m n :: nat using div_eq_0_iff [of m n] by auto lemma mod_greater_zero_iff_not_dvd: "m mod n > 0 \ \ n dvd m" for m n :: nat by (simp add: dvd_eq_mod_eq_0) lemma div_by_Suc_0 [simp]: "m div Suc 0 = m" using div_by_1 [of m] by simp lemma mod_by_Suc_0 [simp]: "m mod Suc 0 = 0" using mod_by_1 [of m] by simp lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" by (simp add: numeral_2_eq_2 le_div_geq) lemma Suc_n_div_2_gt_zero [simp]: "0 < Suc n div 2" if "n > 0" for n :: nat using that by (cases n) simp_all lemma div_2_gt_zero [simp]: "0 < n div 2" if "Suc 0 < n" for n :: nat using that Suc_n_div_2_gt_zero [of "n - 1"] by simp lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" by (simp add: numeral_2_eq_2 le_mod_geq) lemma add_self_div_2 [simp]: "(m + m) div 2 = m" for m :: nat by (simp add: mult_2 [symmetric]) lemma add_self_mod_2 [simp]: "(m + m) mod 2 = 0" for m :: nat by (simp add: mult_2 [symmetric]) lemma mod2_gr_0 [simp]: "0 < m mod 2 \ m mod 2 = 1" for m :: nat proof - have "m mod 2 < 2" by (rule mod_less_divisor) simp then have "m mod 2 = 0 \ m mod 2 = 1" by arith then show ?thesis by auto qed lemma mod_Suc_eq [mod_simps]: "Suc (m mod n) mod n = Suc m mod n" proof - have "(m mod n + 1) mod n = (m + 1) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma mod_Suc_Suc_eq [mod_simps]: "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n" proof - have "(m mod n + 2) mod n = (m + 2) mod n" by (simp only: mod_simps) then show ?thesis by simp qed lemma Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n" and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n" and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n" and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n" by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all context fixes m n q :: nat begin private lemma eucl_rel_mult2: "m mod n + n * (m div n mod q) < n * q" if "n > 0" and "q > 0" proof - from \n > 0\ have "m mod n < n" by (rule mod_less_divisor) from \q > 0\ have "m div n mod q < q" by (rule mod_less_divisor) then obtain s where "q = Suc (m div n mod q + s)" by (blast dest: less_imp_Suc_add) moreover have "m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)" using \m mod n < n\ by (simp add: add_mult_distrib2) ultimately show ?thesis by simp qed lemma div_mult2_eq: "m div (n * q) = (m div n) div q" proof (cases "n = 0 \ q = 0") case True then show ?thesis by auto next case False with eucl_rel_mult2 show ?thesis by (auto intro: div_eqI [of _ "n * (m div n mod q) + m mod n"] simp add: algebra_simps add_mult_distrib2 [symmetric]) qed lemma mod_mult2_eq: "m mod (n * q) = n * (m div n mod q) + m mod n" proof (cases "n = 0 \ q = 0") case True then show ?thesis by auto next case False with eucl_rel_mult2 show ?thesis by (auto intro: mod_eqI [of _ _ "(m div n) div q"] simp add: algebra_simps add_mult_distrib2 [symmetric]) qed end lemma div_le_mono: "m div k \ n div k" if "m \ n" for m n k :: nat proof - from that obtain q where "n = m + q" by (auto simp add: le_iff_add) then show ?thesis by (simp add: div_add1_eq [of m q k]) qed text \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ m" for m n :: nat using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all lemma div_less_dividend [simp]: "m div n < m" if "1 < n" and "0 < m" for m n :: nat using that proof (induct m rule: less_induct) case (less m) show ?case proof (cases "n < m") case False with less show ?thesis by (cases "n = m") simp_all next case True then show ?thesis using less.hyps [of "m - n"] less.prems by (simp add: le_div_geq) qed qed lemma div_eq_dividend_iff: "m div n = m \ n = 1" if "m > 0" for m n :: nat proof assume "n = 1" then show "m div n = m" by simp next assume P: "m div n = m" show "n = 1" proof (rule ccontr) have "n \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 1" ultimately have "n > 1" by simp with that have "m div n < m" by simp with P show False by simp qed qed lemma less_mult_imp_div_less: "m div n < i" if "m < i * n" for m n i :: nat proof - from that have "i * n > 0" by (cases "i * n = 0") simp_all then have "i > 0" and "n > 0" by simp_all have "m div n * n \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed lemma div_less_iff_less_mult: \m div q < n \ m < n * q\ (is \?P \ ?Q\) if \q > 0\ for m n q :: nat proof assume ?Q then show ?P by (rule less_mult_imp_div_less) next assume ?P then obtain h where \n = Suc (m div q + h)\ using less_natE by blast moreover have \m < m + (Suc h * q - m mod q)\ using that by (simp add: trans_less_add1) ultimately show ?Q by (simp add: algebra_simps flip: minus_mod_eq_mult_div) qed lemma less_eq_div_iff_mult_less_eq: \m \ n div q \ m * q \ n\ if \q > 0\ for m n q :: nat using div_less_iff_less_mult [of q n m] that by auto text \A fact for the mutilated chess board\ lemma mod_Suc: "Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))" (is "_ = ?rhs") proof (cases "n = 0") case True then show ?thesis by simp next case False have "Suc m mod n = Suc (m mod n) mod n" by (simp add: mod_simps) also have "\ = ?rhs" using False by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq) finally show ?thesis . qed lemma Suc_times_mod_eq: "Suc (m * n) mod m = 1" if "Suc 0 < m" using that by (simp add: mod_Suc) lemma Suc_times_numeral_mod_eq [simp]: "Suc (numeral k * n) mod numeral k = 1" if "numeral k \ (1::nat)" by (rule Suc_times_mod_eq) (use that in simp) lemma Suc_div_le_mono [simp]: "m div n \ Suc m div n" by (simp add: div_le_mono) text \These lemmas collapse some needless occurrences of Suc: at least three Sucs, since two and fewer are rewritten back to Suc again! We already have some rules to simplify operands smaller than 3.\ lemma div_Suc_eq_div_add3 [simp]: "m div Suc (Suc (Suc n)) = m div (3 + n)" by (simp add: Suc3_eq_add_3) lemma mod_Suc_eq_mod_add3 [simp]: "m mod Suc (Suc (Suc n)) = m mod (3 + n)" by (simp add: Suc3_eq_add_3) lemma Suc_div_eq_add3_div: "Suc (Suc (Suc m)) div n = (3 + m) div n" by (simp add: Suc3_eq_add_3) lemma Suc_mod_eq_add3_mod: "Suc (Suc (Suc m)) mod n = (3 + m) mod n" by (simp add: Suc3_eq_add_3) lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v lemma (in field_char_0) of_nat_div: "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)" proof - have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)" unfolding of_nat_add by (cases "n = 0") simp_all then show ?thesis by simp qed text \An ``induction'' law for modulus arithmetic.\ lemma mod_induct [consumes 3, case_names step]: "P m" if "P n" and "n < p" and "m < p" and step: "\n. n < p \ P n \ P (Suc n mod p)" using \m < p\ proof (induct m) case 0 show ?case proof (rule ccontr) assume "\ P 0" from \n < p\ have "0 < p" by simp from \n < p\ obtain m where "0 < m" and "p = n + m" by (blast dest: less_imp_add_positive) with \P n\ have "P (p - m)" by simp moreover have "\ P (p - m)" using \0 < m\ proof (induct m) case 0 then show ?case by simp next case (Suc m) show ?case proof assume P: "P (p - Suc m)" with \\ P 0\ have "Suc m < p" by (auto intro: ccontr) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ Suc.hyps by (cases "m = 0") simp_all qed qed ultimately show False by blast qed next case (Suc m) then have "m < p" and mod: "Suc m mod p = Suc m" by simp_all from \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma split_div: "P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P i))" (is "?P = ?Q") for m n :: nat proof (cases "n = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?P with False show ?Q by auto next assume ?Q with False have *: "\i j. j < n \ m = n * i + j \ P i" by simp with False show ?P by (auto intro: * [of "m mod n"]) qed qed lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ m div n = q" for q by (auto intro: div_nat_eqI dividend_less_times_div) then show ?thesis by auto qed lemma split_mod: "P (m mod n) \ (n = 0 \ P m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P j))" (is "?P \ ?Q") for m n :: nat proof (cases "n = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?P with False show ?Q by auto next assume ?Q with False have *: "\i j. j < n \ m = n * i + j \ P j" by simp with False show ?P by (auto intro: * [of _ "m div n"]) qed qed lemma funpow_mod_eq: \<^marker>\contributor \Lars Noschinski\\ \(f ^^ (m mod n)) x = (f ^^ m) x\ if \(f ^^ n) x = x\ proof - have \(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\ by simp also have \\ = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\ by (simp only: funpow_add funpow_mult ac_simps) simp also have \((f ^^ n) ^^ q) x = x\ for q by (induction q) (use \(f ^^ n) x = x\ in simp_all) finally show ?thesis by simp qed subsection \Euclidean division on \<^typ>\int\\ instantiation int :: normalization_semidom begin definition normalize_int :: "int \ int" where [simp]: "normalize = (abs :: int \ int)" definition unit_factor_int :: "int \ int" where [simp]: "unit_factor = (sgn :: int \ int)" definition divide_int :: "int \ int \ int" where "k div l = (if l = 0 then 0 else if sgn k = sgn l then int (nat \k\ div nat \l\) else - int (nat \k\ div nat \l\ + of_bool (\ l dvd k)))" lemma divide_int_unfold: "(sgn k * int m) div (sgn l * int n) = (if sgn l = 0 \ sgn k = 0 \ n = 0 then 0 else if sgn k = sgn l then int (m div n) else - int (m div n + of_bool (\ n dvd m)))" by (auto simp add: divide_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult nat_mult_distrib) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 0" obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then have "k * l = sgn (s * t) * int (n * m)" by (simp add: ac_simps sgn_mult) with k l \l \ 0\ show "k * l div l = k" by (simp only: divide_int_unfold) (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0) qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff') end lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation int :: idom_modulo begin definition modulo_int :: "int \ int \ int" where "k mod l = (if l = 0 then k else if sgn k = sgn l then sgn l * int (nat \k\ mod nat \l\) else sgn l * (\l\ * of_bool (\ l dvd k) - int (nat \k\ mod nat \l\)))" lemma modulo_int_unfold: "(sgn k * int m) mod (sgn l * int n) = (if sgn l = 0 \ sgn k = 0 \ n = 0 then sgn k * int m else if sgn k = sgn l then sgn l * int (m mod n) else sgn l * (int (n * of_bool (\ n dvd m)) - int (m mod n)))" by (auto simp add: modulo_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult nat_mult_distrib) instance proof fix k l :: int obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) then show "k div l * l + k mod l = k" by (auto simp add: divide_int_unfold modulo_int_unfold algebra_simps dest!: sgn_not_eq_imp) (simp_all add: of_nat_mult [symmetric] of_nat_add [symmetric] distrib_left [symmetric] minus_mult_right del: of_nat_mult minus_mult_right [symmetric]) qed end instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 0" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg abs_mult mod_greater_zero_iff_not_dvd) qed lemma sgn_mod: "sgn (k mod l) = sgn l" if "l \ 0" "\ l dvd k" for k l :: int proof - obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m" by (blast intro: int_sgnE elim: that) with that show ?thesis by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg sgn_mult) (simp add: dvd_eq_mod_eq_0) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) next fix l q r :: int obtain n m and s t where l: "l = sgn s * int n" and q: "q = sgn t * int m" by (blast intro: int_sgnE elim: that) assume \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ show ?thesis by (simp only: *, simp only: q l divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos) qed next case False with q r have "t \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ m * n" using mult_le_less_imp_less [of 1 m u n] by simp ultimately have *: "q * l + r = sgn (s * t) * int (if t < 0 then m * n - u else m * n + u)" using l q r by (simp add: sgn_mult algebra_simps of_nat_diff) have "(m * n - u) div n = m - 1" if "u > 0" using \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ by (simp only: *, simp only: l q divide_int_unfold) (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le) qed qed (use mult_le_mono2 [of 1] in \auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\) end lemma pos_mod_bound [simp]: "k mod l < l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) simp_all moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (simp add: mod_greater_zero_iff_not_dvd) qed lemma neg_mod_bound [simp]: "l < k mod l" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (simp add: mod_greater_zero_iff_not_dvd) qed lemma pos_mod_sign [simp]: "0 \ k mod l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) auto moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) simp qed lemma neg_mod_sign [simp]: "k mod l \ 0" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) simp qed lemma div_pos_pos_trivial [simp]: "k div l = 0" if "k \ 0" and "k < l" for k l :: int using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_pos_pos_trivial [simp]: "k mod l = k" if "k \ 0" and "k < l" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_neg_neg_trivial [simp]: "k div l = 0" if "k \ 0" and "l < k" for k l :: int using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_neg_neg_trivial [simp]: "k mod l = k" if "k \ 0" and "l < k" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_pos_neg_trivial: "k div l = - 1" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule div_eqI [of _ "k + l"]) using False that apply (simp_all add: division_segment_int_def) done qed lemma mod_pos_neg_trivial: "k mod l = k + l" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule mod_eqI [of _ _ \- 1\]) using False that apply (simp_all add: division_segment_int_def) done qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ text \Distributive laws for function \nat\.\ lemma nat_div_distrib: \nat (x div y) = nat x div nat y\ if \0 \ x\ using that by (simp add: divide_int_def sgn_if) lemma nat_div_distrib': \nat (x div y) = nat x div nat y\ if \0 \ y\ using that by (simp add: divide_int_def sgn_if) lemma nat_mod_distrib: \ \Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\ \nat (x mod y) = nat x mod nat y\ if \0 \ x\ \0 \ y\ using that by (simp add: modulo_int_def sgn_if) subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n" proof (cases a "of_nat m * of_nat n" rule: divmod_cases) case (divides q) then show ?thesis using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"] by (simp add: ac_simps) next case (remainder q r) then have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by simp with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with remainder(6) have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with remainder(1) have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp with remainder(5) remainder(7) show ?thesis using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r] by (simp add: ac_simps) next case by0 then show ?thesis by auto qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD min.absorb2) next case False then have \m < n\ by simp then obtain q where n: \n = Suc q + m\ by (auto dest: less_imp_Suc_add) then have \min m n = m\ by simp moreover have \(2::nat) ^ m \ 2 * 2 ^ q * 2 ^ m\ using mult_le_mono1 [of 1 \2 * 2 ^ q\ \2 ^ m\] by simp with n have \2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\ by (simp add: monoid_mult_class.power_add algebra_simps) ultimately show ?thesis by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp qed then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod of_nat_diff) qed lemma of_bool_half_eq_0 [simp]: \of_bool b div 2 = 0\ by simp end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) instance int :: unique_euclidean_ring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def) lemma zdiv_zmult2_eq: \a div (b * c) = (a div b) div c\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using div_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using div_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed lemma zmod_zmult2_eq: \a mod (b * c) = b * (a div b mod c) + a mod b\ if \c \ 0\ for a b c :: int proof (cases \b \ 0\) case True with that show ?thesis using mod_mult2_eq' [of a \nat b\ \nat c\] by simp next case False with that show ?thesis using mod_mult2_eq' [of \- a\ \nat (- b)\ \nat c\] by simp qed subsection \Code generation\ code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Fields.thy b/src/HOL/Fields.thy --- a/src/HOL/Fields.thy +++ b/src/HOL/Fields.thy @@ -1,1348 +1,1348 @@ (* Title: HOL/Fields.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Fields\ theory Fields imports Nat begin context idom begin lemma inj_mult_left [simp]: \inj ((*) a) \ a \ 0\ (is \?P \ ?Q\) proof assume ?P show ?Q proof assume \a = 0\ with \?P\ have "inj ((*) 0)" by simp moreover have "0 * 0 = 0 * 1" by simp ultimately have "0 = 1" by (rule injD) then show False by simp qed next assume ?Q then show ?P by (auto intro: injI) qed end subsection \Division rings\ text \ A division ring is like a field, but without the commutativity requirement. \ class inverse = divide + fixes inverse :: "'a \ 'a" begin abbreviation inverse_divide :: "'a \ 'a \ 'a" (infixl "'/" 70) where "inverse_divide \ divide" end text \Setup for linear arithmetic prover\ ML_file \~~/src/Provers/Arith/fast_lin_arith.ML\ ML_file \Tools/lin_arith.ML\ setup \Lin_Arith.global_setup\ declaration \K ( Lin_Arith.init_arith_data #> Lin_Arith.add_discrete_type \<^type_name>\nat\ #> Lin_Arith.add_lessD @{thm Suc_leI} #> Lin_Arith.add_simps @{thms simp_thms ring_distribs if_True if_False minus_diff_eq add_0_left add_0_right order_less_irrefl zero_neq_one zero_less_one zero_le_one zero_neq_one [THEN not_sym] not_one_le_zero not_one_less_zero add_Suc add_Suc_right nat.inject Suc_le_mono Suc_less_eq Zero_not_Suc Suc_not_Zero le_0_eq One_nat_def} #> Lin_Arith.add_simprocs [\<^simproc>\group_cancel_add\, \<^simproc>\group_cancel_diff\, \<^simproc>\group_cancel_eq\, \<^simproc>\group_cancel_le\, \<^simproc>\group_cancel_less\, \<^simproc>\nateq_cancel_sums\,\<^simproc>\natless_cancel_sums\, \<^simproc>\natle_cancel_sums\])\ simproc_setup fast_arith_nat ("(m::nat) < n" | "(m::nat) \ n" | "(m::nat) = n") = \K Lin_Arith.simproc\ \ \Because of this simproc, the arithmetic solver is really only useful to detect inconsistencies among the premises for subgoals which are \<^emph>\not\ themselves (in)equalities, because the latter activate \<^text>\fast_nat_arith_simproc\ anyway. However, it seems cheaper to activate the solver all the time rather than add the additional check.\ lemmas [arith_split] = nat_diff_split split_min split_max text\Lemmas \divide_simps\ move division to the outside and eliminates them on (in)equalities.\ named_theorems divide_simps "rewrite rules to eliminate divisions" class division_ring = ring_1 + inverse + assumes left_inverse [simp]: "a \ 0 \ inverse a * a = 1" assumes right_inverse [simp]: "a \ 0 \ a * inverse a = 1" assumes divide_inverse: "a / b = a * inverse b" assumes inverse_zero [simp]: "inverse 0 = 0" begin subclass ring_1_no_zero_divisors proof fix a b :: 'a assume a: "a \ 0" and b: "b \ 0" show "a * b \ 0" proof assume ab: "a * b = 0" hence "0 = inverse a * (a * b) * inverse b" by simp also have "\ = (inverse a * a) * (b * inverse b)" by (simp only: mult.assoc) also have "\ = 1" using a b by simp finally show False by simp qed qed lemma nonzero_imp_inverse_nonzero: "a \ 0 \ inverse a \ 0" proof assume ianz: "inverse a = 0" assume "a \ 0" hence "1 = a * inverse a" by simp also have "... = 0" by (simp add: ianz) finally have "1 = 0" . thus False by (simp add: eq_commute) qed lemma inverse_zero_imp_zero: assumes "inverse a = 0" shows "a = 0" proof (rule ccontr) assume "a \ 0" then have "inverse a \ 0" by (simp add: nonzero_imp_inverse_nonzero) with assms show False by auto qed lemma inverse_unique: assumes ab: "a * b = 1" shows "inverse a = b" proof - have "a \ 0" using ab by (cases "a = 0") simp_all moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) ultimately show ?thesis by (simp add: mult.assoc [symmetric]) qed lemma nonzero_inverse_minus_eq: "a \ 0 \ inverse (- a) = - inverse a" by (rule inverse_unique) simp lemma nonzero_inverse_inverse_eq: "a \ 0 \ inverse (inverse a) = a" by (rule inverse_unique) simp lemma nonzero_inverse_eq_imp_eq: assumes "inverse a = inverse b" and "a \ 0" and "b \ 0" shows "a = b" proof - from \inverse a = inverse b\ have "inverse (inverse a) = inverse (inverse b)" by (rule arg_cong) with \a \ 0\ and \b \ 0\ show "a = b" by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_1 [simp]: "inverse 1 = 1" by (rule inverse_unique) simp lemma nonzero_inverse_mult_distrib: assumes "a \ 0" and "b \ 0" shows "inverse (a * b) = inverse b * inverse a" proof - have "a * (b * inverse b) * inverse a = 1" using assms by simp hence "a * b * (inverse b * inverse a) = 1" by (simp only: mult.assoc) thus ?thesis by (rule inverse_unique) qed lemma division_ring_inverse_add: "a \ 0 \ b \ 0 \ inverse a + inverse b = inverse a * (a + b) * inverse b" by (simp add: algebra_simps) lemma division_ring_inverse_diff: "a \ 0 \ b \ 0 \ inverse a - inverse b = inverse a * (b - a) * inverse b" by (simp add: algebra_simps) lemma right_inverse_eq: "b \ 0 \ a / b = 1 \ a = b" proof assume neq: "b \ 0" { hence "a = (a / b) * b" by (simp add: divide_inverse mult.assoc) also assume "a / b = 1" finally show "a = b" by simp next assume "a = b" with neq show "a / b = 1" by (simp add: divide_inverse) } qed lemma nonzero_inverse_eq_divide: "a \ 0 \ inverse a = 1 / a" by (simp add: divide_inverse) lemma divide_self [simp]: "a \ 0 \ a / a = 1" by (simp add: divide_inverse) lemma inverse_eq_divide [field_simps, field_split_simps, divide_simps]: "inverse a = 1 / a" by (simp add: divide_inverse) lemma add_divide_distrib: "(a+b) / c = a/c + b/c" by (simp add: divide_inverse algebra_simps) lemma times_divide_eq_right [simp]: "a * (b / c) = (a * b) / c" by (simp add: divide_inverse mult.assoc) lemma minus_divide_left: "- (a / b) = (-a) / b" by (simp add: divide_inverse) lemma nonzero_minus_divide_right: "b \ 0 \ - (a / b) = a / (- b)" by (simp add: divide_inverse nonzero_inverse_minus_eq) lemma nonzero_minus_divide_divide: "b \ 0 \ (-a) / (-b) = a / b" by (simp add: divide_inverse nonzero_inverse_minus_eq) lemma divide_minus_left [simp]: "(-a) / b = - (a / b)" by (simp add: divide_inverse) lemma diff_divide_distrib: "(a - b) / c = a / c - b / c" using add_divide_distrib [of a "- b" c] by simp lemma nonzero_eq_divide_eq [field_simps]: "c \ 0 \ a = b / c \ a * c = b" proof - assume [simp]: "c \ 0" have "a = b / c \ a * c = (b / c) * c" by simp also have "... \ a * c = b" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed lemma nonzero_divide_eq_eq [field_simps]: "c \ 0 \ b / c = a \ b = a * c" proof - assume [simp]: "c \ 0" have "b / c = a \ (b / c) * c = a * c" by simp also have "... \ b = a * c" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed lemma nonzero_neg_divide_eq_eq [field_simps]: "b \ 0 \ - (a / b) = c \ - a = c * b" using nonzero_divide_eq_eq[of b "-a" c] by simp lemma nonzero_neg_divide_eq_eq2 [field_simps]: "b \ 0 \ c = - (a / b) \ c * b = - a" using nonzero_neg_divide_eq_eq[of b a c] by auto lemma divide_eq_imp: "c \ 0 \ b = a * c \ b / c = a" by (simp add: divide_inverse mult.assoc) lemma eq_divide_imp: "c \ 0 \ a * c = b \ a = b / c" by (drule sym) (simp add: divide_inverse mult.assoc) lemma add_divide_eq_iff [field_simps]: "z \ 0 \ x + y / z = (x * z + y) / z" by (simp add: add_divide_distrib nonzero_eq_divide_eq) lemma divide_add_eq_iff [field_simps]: "z \ 0 \ x / z + y = (x + y * z) / z" by (simp add: add_divide_distrib nonzero_eq_divide_eq) lemma diff_divide_eq_iff [field_simps]: "z \ 0 \ x - y / z = (x * z - y) / z" by (simp add: diff_divide_distrib nonzero_eq_divide_eq eq_diff_eq) lemma minus_divide_add_eq_iff [field_simps]: "z \ 0 \ - (x / z) + y = (- x + y * z) / z" by (simp add: add_divide_distrib diff_divide_eq_iff) lemma divide_diff_eq_iff [field_simps]: "z \ 0 \ x / z - y = (x - y * z) / z" by (simp add: field_simps) lemma minus_divide_diff_eq_iff [field_simps]: "z \ 0 \ - (x / z) - y = (- x - y * z) / z" by (simp add: divide_diff_eq_iff[symmetric]) lemma division_ring_divide_zero [simp]: "a / 0 = 0" by (simp add: divide_inverse) lemma divide_self_if [simp]: "a / a = (if a = 0 then 0 else 1)" by simp lemma inverse_nonzero_iff_nonzero [simp]: "inverse a = 0 \ a = 0" - by rule (fact inverse_zero_imp_zero, simp) + by (rule iffI) (fact inverse_zero_imp_zero, simp) lemma inverse_minus_eq [simp]: "inverse (- a) = - inverse a" proof cases assume "a=0" thus ?thesis by simp next assume "a\0" thus ?thesis by (simp add: nonzero_inverse_minus_eq) qed lemma inverse_inverse_eq [simp]: "inverse (inverse a) = a" proof cases assume "a=0" thus ?thesis by simp next assume "a\0" thus ?thesis by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_eq_imp_eq: "inverse a = inverse b \ a = b" by (drule arg_cong [where f="inverse"], simp) lemma inverse_eq_iff_eq [simp]: "inverse a = inverse b \ a = b" by (force dest!: inverse_eq_imp_eq) lemma mult_commute_imp_mult_inverse_commute: assumes "y * x = x * y" shows "inverse y * x = x * inverse y" proof (cases "y=0") case False hence "x * inverse y = inverse y * y * x * inverse y" by simp also have "\ = inverse y * (x * y * inverse y)" by (simp add: mult.assoc assms) finally show ?thesis by (simp add: mult.assoc False) qed simp lemmas mult_inverse_of_nat_commute = mult_commute_imp_mult_inverse_commute[OF mult_of_nat_commute] lemma divide_divide_eq_left': "(a / b) / c = a / (c * b)" by (cases "b = 0 \ c = 0") (auto simp: divide_inverse mult.assoc nonzero_inverse_mult_distrib) lemma add_divide_eq_if_simps [field_split_simps, divide_simps]: "a + b / z = (if z = 0 then a else (a * z + b) / z)" "a / z + b = (if z = 0 then b else (a + b * z) / z)" "- (a / z) + b = (if z = 0 then b else (-a + b * z) / z)" "a - b / z = (if z = 0 then a else (a * z - b) / z)" "a / z - b = (if z = 0 then -b else (a - b * z) / z)" "- (a / z) - b = (if z = 0 then -b else (- a - b * z) / z)" by (simp_all add: add_divide_eq_iff divide_add_eq_iff diff_divide_eq_iff divide_diff_eq_iff minus_divide_diff_eq_iff) lemma [field_split_simps, divide_simps]: shows divide_eq_eq: "b / c = a \ (if c \ 0 then b = a * c else a = 0)" and eq_divide_eq: "a = b / c \ (if c \ 0 then a * c = b else a = 0)" and minus_divide_eq_eq: "- (b / c) = a \ (if c \ 0 then - b = a * c else a = 0)" and eq_minus_divide_eq: "a = - (b / c) \ (if c \ 0 then a * c = - b else a = 0)" by (auto simp add: field_simps) end subsection \Fields\ class field = comm_ring_1 + inverse + assumes field_inverse: "a \ 0 \ inverse a * a = 1" assumes field_divide_inverse: "a / b = a * inverse b" assumes field_inverse_zero: "inverse 0 = 0" begin subclass division_ring proof fix a :: 'a assume "a \ 0" thus "inverse a * a = 1" by (rule field_inverse) thus "a * inverse a = 1" by (simp only: mult.commute) next fix a b :: 'a show "a / b = a * inverse b" by (rule field_divide_inverse) next show "inverse 0 = 0" by (fact field_inverse_zero) qed subclass idom_divide proof fix b a assume "b \ 0" then show "a * b / b = a" by (simp add: divide_inverse ac_simps) next fix a show "a / 0 = 0" by (simp add: divide_inverse) qed text\There is no slick version using division by zero.\ lemma inverse_add: "a \ 0 \ b \ 0 \ inverse a + inverse b = (a + b) * inverse a * inverse b" by (simp add: division_ring_inverse_add ac_simps) lemma nonzero_mult_divide_mult_cancel_left [simp]: assumes [simp]: "c \ 0" shows "(c * a) / (c * b) = a / b" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "(c*a)/(c*b) = c * a * (inverse b * inverse c)" by (simp add: divide_inverse nonzero_inverse_mult_distrib) also have "... = a * inverse b * (inverse c * c)" by (simp only: ac_simps) also have "... = a * inverse b" by simp finally show ?thesis by (simp add: divide_inverse) qed lemma nonzero_mult_divide_mult_cancel_right [simp]: "c \ 0 \ (a * c) / (b * c) = a / b" using nonzero_mult_divide_mult_cancel_left [of c a b] by (simp add: ac_simps) lemma times_divide_eq_left [simp]: "(b / c) * a = (b * a) / c" by (simp add: divide_inverse ac_simps) lemma divide_inverse_commute: "a / b = inverse b * a" by (simp add: divide_inverse mult.commute) lemma add_frac_eq: assumes "y \ 0" and "z \ 0" shows "x / y + w / z = (x * z + w * y) / (y * z)" proof - have "x / y + w / z = (x * z) / (y * z) + (y * w) / (y * z)" using assms by simp also have "\ = (x * z + y * w) / (y * z)" by (simp only: add_divide_distrib) finally show ?thesis by (simp only: mult.commute) qed text\Special Cancellation Simprules for Division\ lemma nonzero_divide_mult_cancel_right [simp]: "b \ 0 \ b / (a * b) = 1 / a" using nonzero_mult_divide_mult_cancel_right [of b 1 a] by simp lemma nonzero_divide_mult_cancel_left [simp]: "a \ 0 \ a / (a * b) = 1 / b" using nonzero_mult_divide_mult_cancel_left [of a 1 b] by simp lemma nonzero_mult_divide_mult_cancel_left2 [simp]: "c \ 0 \ (c * a) / (b * c) = a / b" using nonzero_mult_divide_mult_cancel_left [of c a b] by (simp add: ac_simps) lemma nonzero_mult_divide_mult_cancel_right2 [simp]: "c \ 0 \ (a * c) / (c * b) = a / b" using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: ac_simps) lemma diff_frac_eq: "y \ 0 \ z \ 0 \ x / y - w / z = (x * z - w * y) / (y * z)" by (simp add: field_simps) lemma frac_eq_eq: "y \ 0 \ z \ 0 \ (x / y = w / z) = (x * z = w * y)" by (simp add: field_simps) lemma divide_minus1 [simp]: "x / - 1 = - x" using nonzero_minus_divide_right [of "1" x] by simp text\This version builds in division by zero while also re-orienting the right-hand side.\ lemma inverse_mult_distrib [simp]: "inverse (a * b) = inverse a * inverse b" proof cases assume "a \ 0 \ b \ 0" thus ?thesis by (simp add: nonzero_inverse_mult_distrib ac_simps) next assume "\ (a \ 0 \ b \ 0)" thus ?thesis by force qed lemma inverse_divide [simp]: "inverse (a / b) = b / a" by (simp add: divide_inverse mult.commute) text \Calculations with fractions\ text\There is a whole bunch of simp-rules just for class \field\ but none for class \field\ and \nonzero_divides\ because the latter are covered by a simproc.\ lemmas mult_divide_mult_cancel_left = nonzero_mult_divide_mult_cancel_left lemmas mult_divide_mult_cancel_right = nonzero_mult_divide_mult_cancel_right lemma divide_divide_eq_right [simp]: "a / (b / c) = (a * c) / b" by (simp add: divide_inverse ac_simps) lemma divide_divide_eq_left [simp]: "(a / b) / c = a / (b * c)" by (simp add: divide_inverse mult.assoc) lemma divide_divide_times_eq: "(x / y) / (z / w) = (x * w) / (y * z)" by simp text \Special Cancellation Simprules for Division\ lemma mult_divide_mult_cancel_left_if [simp]: shows "(c * a) / (c * b) = (if c = 0 then 0 else a / b)" by simp text \Division and Unary Minus\ lemma minus_divide_right: "- (a / b) = a / - b" by (simp add: divide_inverse) lemma divide_minus_right [simp]: "a / - b = - (a / b)" by (simp add: divide_inverse) lemma minus_divide_divide: "(- a) / (- b) = a / b" by (cases "b=0") (simp_all add: nonzero_minus_divide_divide) lemma inverse_eq_1_iff [simp]: "inverse x = 1 \ x = 1" - by (insert inverse_eq_iff_eq [of x 1], simp) + using inverse_eq_iff_eq [of x 1] by simp lemma divide_eq_0_iff [simp]: "a / b = 0 \ a = 0 \ b = 0" by (simp add: divide_inverse) lemma divide_cancel_right [simp]: "a / c = b / c \ c = 0 \ a = b" by (cases "c=0") (simp_all add: divide_inverse) lemma divide_cancel_left [simp]: "c / a = c / b \ c = 0 \ a = b" by (cases "c=0") (simp_all add: divide_inverse) lemma divide_eq_1_iff [simp]: "a / b = 1 \ b \ 0 \ a = b" by (cases "b=0") (simp_all add: right_inverse_eq) lemma one_eq_divide_iff [simp]: "1 = a / b \ b \ 0 \ a = b" by (simp add: eq_commute [of 1]) lemma divide_eq_minus_1_iff: "(a / b = - 1) \ b \ 0 \ a = - b" using divide_eq_1_iff by fastforce lemma times_divide_times_eq: "(x / y) * (z / w) = (x * z) / (y * w)" by simp lemma add_frac_num: "y \ 0 \ x / y + z = (x + z * y) / y" by (simp add: add_divide_distrib) lemma add_num_frac: "y \ 0 \ z + x / y = (x + z * y) / y" by (simp add: add_divide_distrib add.commute) lemma dvd_field_iff: "a dvd b \ (a = 0 \ b = 0)" proof (cases "a = 0") case False then have "b = a * (b / a)" by (simp add: field_simps) then have "a dvd b" .. with False show ?thesis by simp qed simp lemma inj_divide_right [simp]: "inj (\b. b / a) \ a \ 0" proof - have "(\b. b / a) = (*) (inverse a)" by (simp add: field_simps fun_eq_iff) then have "inj (\y. y / a) \ inj ((*) (inverse a))" by simp also have "\ \ inverse a \ 0" by simp also have "\ \ a \ 0" by simp finally show ?thesis by simp qed end class field_char_0 = field + ring_char_0 subsection \Ordered fields\ class field_abs_sgn = field + idom_abs_sgn begin lemma sgn_inverse [simp]: "sgn (inverse a) = inverse (sgn a)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * inverse a = 1" by simp then have "sgn (a * inverse a) = sgn 1" by simp then have "sgn a * sgn (inverse a) = 1" by (simp add: sgn_mult) then have "inverse (sgn a) * (sgn a * sgn (inverse a)) = inverse (sgn a) * 1" by simp then have "(inverse (sgn a) * sgn a) * sgn (inverse a) = inverse (sgn a)" by (simp add: ac_simps) with False show ?thesis by (simp add: sgn_eq_0_iff) qed lemma abs_inverse [simp]: "\inverse a\ = inverse \a\" proof - from sgn_mult_abs [of "inverse a"] sgn_mult_abs [of a] have "inverse (sgn a) * \inverse a\ = inverse (sgn a * \a\)" by simp then show ?thesis by (auto simp add: sgn_eq_0_iff) qed lemma sgn_divide [simp]: "sgn (a / b) = sgn a / sgn b" unfolding divide_inverse sgn_mult by simp lemma abs_divide [simp]: "\a / b\ = \a\ / \b\" unfolding divide_inverse abs_mult by simp end class linordered_field = field + linordered_idom begin lemma positive_imp_inverse_positive: assumes a_gt_0: "0 < a" shows "0 < inverse a" proof - have "0 < a * inverse a" by (simp add: a_gt_0 [THEN less_imp_not_eq2]) thus "0 < inverse a" by (simp add: a_gt_0 [THEN less_not_sym] zero_less_mult_iff) qed lemma negative_imp_inverse_negative: "a < 0 \ inverse a < 0" - by (insert positive_imp_inverse_positive [of "-a"], - simp add: nonzero_inverse_minus_eq less_imp_not_eq) + using positive_imp_inverse_positive [of "-a"] + by (simp add: nonzero_inverse_minus_eq less_imp_not_eq) lemma inverse_le_imp_le: assumes invle: "inverse a \ inverse b" and apos: "0 < a" shows "b \ a" proof (rule classical) assume "\ b \ a" hence "a < b" by (simp add: linorder_not_le) hence bpos: "0 < b" by (blast intro: apos less_trans) hence "a * inverse a \ a * inverse b" by (simp add: apos invle less_imp_le mult_left_mono) hence "(a * inverse a) * b \ (a * inverse b) * b" by (simp add: bpos less_imp_le mult_right_mono) thus "b \ a" by (simp add: mult.assoc apos bpos less_imp_not_eq2) qed lemma inverse_positive_imp_positive: assumes inv_gt_0: "0 < inverse a" and nz: "a \ 0" shows "0 < a" proof - have "0 < inverse (inverse a)" using inv_gt_0 by (rule positive_imp_inverse_positive) thus "0 < a" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_negative_imp_negative: assumes inv_less_0: "inverse a < 0" and nz: "a \ 0" shows "a < 0" proof - have "inverse (inverse a) < 0" using inv_less_0 by (rule negative_imp_inverse_negative) thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma linordered_field_no_lb: "\x. \y. y < x" proof fix x::'a have m1: "- (1::'a) < 0" by simp from add_strict_right_mono[OF m1, where c=x] have "(- 1) + x < x" by simp thus "\y. y < x" by blast qed lemma linordered_field_no_ub: "\ x. \y. y > x" proof fix x::'a have m1: " (1::'a) > 0" by simp from add_strict_right_mono[OF m1, where c=x] have "1 + x > x" by simp thus "\y. y > x" by blast qed lemma less_imp_inverse_less: assumes less: "a < b" and apos: "0 < a" shows "inverse b < inverse a" proof (rule ccontr) assume "\ inverse b < inverse a" hence "inverse a \ inverse b" by simp hence "\ (a < b)" by (simp add: not_less inverse_le_imp_le [OF _ apos]) thus False by (rule notE [OF _ less]) qed lemma inverse_less_imp_less: assumes "inverse a < inverse b" "0 < a" shows "b < a" proof - have "a \ b" using assms by (simp add: less_le) moreover have "b \ a" using assms by (force simp: less_le dest: inverse_le_imp_le) ultimately show ?thesis by (simp add: less_le) qed text\Both premises are essential. Consider -1 and 1.\ lemma inverse_less_iff_less [simp]: "0 < a \ 0 < b \ inverse a < inverse b \ b < a" by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) lemma le_imp_inverse_le: "a \ b \ 0 < a \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less) lemma inverse_le_iff_le [simp]: "0 < a \ 0 < b \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) text\These results refer to both operands being negative. The opposite-sign case is trivial, since inverse preserves signs.\ lemma inverse_le_imp_le_neg: assumes "inverse a \ inverse b" "b < 0" shows "b \ a" proof (rule classical) assume "\ b \ a" with \b < 0\ have "a < 0" by force with assms show "b \ a" using inverse_le_imp_le [of "-b" "-a"] by (simp add: nonzero_inverse_minus_eq) qed lemma less_imp_inverse_less_neg: assumes "a < b" "b < 0" shows "inverse b < inverse a" proof - have "a < 0" using assms by (blast intro: less_trans) with less_imp_inverse_less [of "-b" "-a"] show ?thesis by (simp add: nonzero_inverse_minus_eq assms) qed lemma inverse_less_imp_less_neg: assumes "inverse a < inverse b" "b < 0" shows "b < a" proof (rule classical) assume "\ b < a" with \b < 0\ have "a < 0" by force with inverse_less_imp_less [of "-b" "-a"] show ?thesis by (simp add: nonzero_inverse_minus_eq assms) qed lemma inverse_less_iff_less_neg [simp]: "a < 0 \ b < 0 \ inverse a < inverse b \ b < a" using inverse_less_iff_less [of "-b" "-a"] by (simp del: inverse_less_iff_less add: nonzero_inverse_minus_eq) lemma le_imp_inverse_le_neg: "a \ b \ b < 0 \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less_neg) lemma inverse_le_iff_le_neg [simp]: "a < 0 \ b < 0 \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) lemma one_less_inverse: "0 < a \ a < 1 \ 1 < inverse a" using less_imp_inverse_less [of a 1, unfolded inverse_1] . lemma one_le_inverse: "0 < a \ a \ 1 \ 1 \ inverse a" using le_imp_inverse_le [of a 1, unfolded inverse_1] . lemma pos_le_divide_eq [field_simps]: assumes "0 < c" shows "a \ b / c \ a * c \ b" proof - from assms have "a \ b / c \ a * c \ (b / c) * c" using mult_le_cancel_right [of a c "b * inverse c"] by (auto simp add: field_simps) also have "... \ a * c \ b" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma pos_less_divide_eq [field_simps]: assumes "0 < c" shows "a < b / c \ a * c < b" proof - from assms have "a < b / c \ a * c < (b / c) * c" using mult_less_cancel_right [of a c "b / c"] by auto also have "... = (a*c < b)" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_less_divide_eq [field_simps]: assumes "c < 0" shows "a < b / c \ b < a * c" proof - from assms have "a < b / c \ (b / c) * c < a * c" using mult_less_cancel_right [of "b / c" c a] by auto also have "... \ b < a * c" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_le_divide_eq [field_simps]: assumes "c < 0" shows "a \ b / c \ b \ a * c" proof - from assms have "a \ b / c \ (b / c) * c \ a * c" using mult_le_cancel_right [of "b * inverse c" c a] by (auto simp add: field_simps) also have "... \ b \ a * c" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma pos_divide_le_eq [field_simps]: assumes "0 < c" shows "b / c \ a \ b \ a * c" proof - from assms have "b / c \ a \ (b / c) * c \ a * c" using mult_le_cancel_right [of "b / c" c a] by auto also have "... \ b \ a * c" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma pos_divide_less_eq [field_simps]: assumes "0 < c" shows "b / c < a \ b < a * c" proof - from assms have "b / c < a \ (b / c) * c < a * c" using mult_less_cancel_right [of "b / c" c a] by auto also have "... \ b < a * c" by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_divide_le_eq [field_simps]: assumes "c < 0" shows "b / c \ a \ a * c \ b" proof - from assms have "b / c \ a \ a * c \ (b / c) * c" using mult_le_cancel_right [of a c "b / c"] by auto also have "... \ a * c \ b" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed lemma neg_divide_less_eq [field_simps]: assumes "c < 0" shows "b / c < a \ a * c < b" proof - from assms have "b / c < a \ a * c < b / c * c" using mult_less_cancel_right [of a c "b / c"] by auto also have "... \ a * c < b" by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed text\The following \field_simps\ rules are necessary, as minus is always moved atop of division but we want to get rid of division.\ lemma pos_le_minus_divide_eq [field_simps]: "0 < c \ a \ - (b / c) \ a * c \ - b" unfolding minus_divide_left by (rule pos_le_divide_eq) lemma neg_le_minus_divide_eq [field_simps]: "c < 0 \ a \ - (b / c) \ - b \ a * c" unfolding minus_divide_left by (rule neg_le_divide_eq) lemma pos_less_minus_divide_eq [field_simps]: "0 < c \ a < - (b / c) \ a * c < - b" unfolding minus_divide_left by (rule pos_less_divide_eq) lemma neg_less_minus_divide_eq [field_simps]: "c < 0 \ a < - (b / c) \ - b < a * c" unfolding minus_divide_left by (rule neg_less_divide_eq) lemma pos_minus_divide_less_eq [field_simps]: "0 < c \ - (b / c) < a \ - b < a * c" unfolding minus_divide_left by (rule pos_divide_less_eq) lemma neg_minus_divide_less_eq [field_simps]: "c < 0 \ - (b / c) < a \ a * c < - b" unfolding minus_divide_left by (rule neg_divide_less_eq) lemma pos_minus_divide_le_eq [field_simps]: "0 < c \ - (b / c) \ a \ - b \ a * c" unfolding minus_divide_left by (rule pos_divide_le_eq) lemma neg_minus_divide_le_eq [field_simps]: "c < 0 \ - (b / c) \ a \ a * c \ - b" unfolding minus_divide_left by (rule neg_divide_le_eq) lemma frac_less_eq: "y \ 0 \ z \ 0 \ x / y < w / z \ (x * z - w * y) / (y * z) < 0" by (subst less_iff_diff_less_0) (simp add: diff_frac_eq ) lemma frac_le_eq: "y \ 0 \ z \ 0 \ x / y \ w / z \ (x * z - w * y) / (y * z) \ 0" by (subst le_iff_diff_le_0) (simp add: diff_frac_eq ) lemma divide_pos_pos[simp]: "0 < x \ 0 < y \ 0 < x / y" by(simp add:field_simps) lemma divide_nonneg_pos: "0 \ x \ 0 < y \ 0 \ x / y" by(simp add:field_simps) lemma divide_neg_pos: "x < 0 \ 0 < y \ x / y < 0" by(simp add:field_simps) lemma divide_nonpos_pos: "x \ 0 \ 0 < y \ x / y \ 0" by(simp add:field_simps) lemma divide_pos_neg: "0 < x \ y < 0 \ x / y < 0" by(simp add:field_simps) lemma divide_nonneg_neg: "0 \ x \ y < 0 \ x / y \ 0" by(simp add:field_simps) lemma divide_neg_neg: "x < 0 \ y < 0 \ 0 < x / y" by(simp add:field_simps) lemma divide_nonpos_neg: "x \ 0 \ y < 0 \ 0 \ x / y" by(simp add:field_simps) lemma divide_strict_right_mono: "\a < b; 0 < c\ \ a / c < b / c" by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono positive_imp_inverse_positive) lemma divide_strict_right_mono_neg: assumes "b < a" "c < 0" shows "a / c < b / c" proof - have "b / - c < a / - c" by (rule divide_strict_right_mono) (use assms in auto) then show ?thesis by (simp add: less_imp_not_eq) qed text\The last premise ensures that \<^term>\a\ and \<^term>\b\ have the same sign\ lemma divide_strict_left_mono: "\b < a; 0 < c; 0 < a*b\ \ c / a < c / b" by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono) lemma divide_left_mono: "\b \ a; 0 \ c; 0 < a*b\ \ c / a \ c / b" by (auto simp: field_simps zero_less_mult_iff mult_right_mono) lemma divide_strict_left_mono_neg: "\a < b; c < 0; 0 < a*b\ \ c / a < c / b" by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono_neg) lemma mult_imp_div_pos_le: "0 < y \ x \ z * y \ x / y \ z" by (subst pos_divide_le_eq, assumption+) lemma mult_imp_le_div_pos: "0 < y \ z * y \ x \ z \ x / y" by(simp add:field_simps) lemma mult_imp_div_pos_less: "0 < y \ x < z * y \ x / y < z" by(simp add:field_simps) lemma mult_imp_less_div_pos: "0 < y \ z * y < x \ z < x / y" by(simp add:field_simps) lemma frac_le: assumes "0 \ y" "x \ y" "0 < w" "w \ z" shows "x / z \ y / w" proof (rule mult_imp_div_pos_le) show "z > 0" using assms by simp have "x \ y * z / w" proof (rule mult_imp_le_div_pos [OF \0 < w\]) show "x * w \ y * z" using assms by (auto intro: mult_mono) qed also have "... = y / w * z" by simp finally show "x \ y / w * z" . qed lemma frac_less: assumes "0 \ x" "x < y" "0 < w" "w \ z" shows "x / z < y / w" proof (rule mult_imp_div_pos_less) show "z > 0" using assms by simp have "x < y * z / w" proof (rule mult_imp_less_div_pos [OF \0 < w\]) show "x * w < y * z" using assms by (auto intro: mult_less_le_imp_less) qed also have "... = y / w * z" by simp finally show "x < y / w * z" . qed lemma frac_less2: assumes "0 < x" "x \ y" "0 < w" "w < z" shows "x / z < y / w" proof (rule mult_imp_div_pos_less) show "z > 0" using assms by simp show "x < y / w * z" using assms by (force intro: mult_imp_less_div_pos mult_le_less_imp_less) qed lemma less_half_sum: "a < b \ a < (a+b) / (1+1)" by (simp add: field_simps zero_less_two) lemma gt_half_sum: "a < b \ (a+b)/(1+1) < b" by (simp add: field_simps zero_less_two) subclass unbounded_dense_linorder proof fix x y :: 'a from less_add_one show "\y. x < y" .. from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono) then have "x - 1 < x + 1 - 1" by simp then have "x - 1 < x" by (simp add: algebra_simps) then show "\y. y < x" .. show "x < y \ \z>x. z < y" by (blast intro!: less_half_sum gt_half_sum) qed subclass field_abs_sgn .. lemma inverse_sgn [simp]: "inverse (sgn a) = sgn a" by (cases a 0 rule: linorder_cases) simp_all lemma divide_sgn [simp]: "a / sgn b = a * sgn b" by (cases b 0 rule: linorder_cases) simp_all lemma nonzero_abs_inverse: "a \ 0 \ \inverse a\ = inverse \a\" by (rule abs_inverse) lemma nonzero_abs_divide: "b \ 0 \ \a / b\ = \a\ / \b\" by (rule abs_divide) lemma field_le_epsilon: assumes e: "\e. 0 < e \ x \ y + e" shows "x \ y" proof (rule dense_le) fix t assume "t < x" hence "0 < x - t" by (simp add: less_diff_eq) from e [OF this] have "x + 0 \ x + (y - t)" by (simp add: algebra_simps) then have "0 \ y - t" by (simp only: add_le_cancel_left) then show "t \ y" by (simp add: algebra_simps) qed lemma inverse_positive_iff_positive [simp]: "(0 < inverse a) = (0 < a)" proof (cases "a = 0") case False then show ?thesis by (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive) qed auto lemma inverse_negative_iff_negative [simp]: "(inverse a < 0) = (a < 0)" proof (cases "a = 0") case False then show ?thesis by (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative) qed auto lemma inverse_nonnegative_iff_nonnegative [simp]: "0 \ inverse a \ 0 \ a" by (simp add: not_less [symmetric]) lemma inverse_nonpositive_iff_nonpositive [simp]: "inverse a \ 0 \ a \ 0" by (simp add: not_less [symmetric]) lemma one_less_inverse_iff: "1 < inverse x \ 0 < x \ x < 1" using less_trans[of 1 x 0 for x] by (cases x 0 rule: linorder_cases) (auto simp add: field_simps) lemma one_le_inverse_iff: "1 \ inverse x \ 0 < x \ x \ 1" proof (cases "x = 1") case True then show ?thesis by simp next case False then have "inverse x \ 1" by simp then have "1 \ inverse x" by blast then have "1 \ inverse x \ 1 < inverse x" by (simp add: le_less) with False show ?thesis by (auto simp add: one_less_inverse_iff) qed lemma inverse_less_1_iff: "inverse x < 1 \ x \ 0 \ 1 < x" by (simp add: not_le [symmetric] one_le_inverse_iff) lemma inverse_le_1_iff: "inverse x \ 1 \ x \ 0 \ 1 \ x" by (simp add: not_less [symmetric] one_less_inverse_iff) lemma [field_split_simps, divide_simps]: shows le_divide_eq: "a \ b / c \ (if 0 < c then a * c \ b else if c < 0 then b \ a * c else a \ 0)" and divide_le_eq: "b / c \ a \ (if 0 < c then b \ a * c else if c < 0 then a * c \ b else 0 \ a)" and less_divide_eq: "a < b / c \ (if 0 < c then a * c < b else if c < 0 then b < a * c else a < 0)" and divide_less_eq: "b / c < a \ (if 0 < c then b < a * c else if c < 0 then a * c < b else 0 < a)" and le_minus_divide_eq: "a \ - (b / c) \ (if 0 < c then a * c \ - b else if c < 0 then - b \ a * c else a \ 0)" and minus_divide_le_eq: "- (b / c) \ a \ (if 0 < c then - b \ a * c else if c < 0 then a * c \ - b else 0 \ a)" and less_minus_divide_eq: "a < - (b / c) \ (if 0 < c then a * c < - b else if c < 0 then - b < a * c else a < 0)" and minus_divide_less_eq: "- (b / c) < a \ (if 0 < c then - b < a * c else if c < 0 then a * c < - b else 0 < a)" by (auto simp: field_simps not_less dest: order.antisym) text \Division and Signs\ lemma shows zero_less_divide_iff: "0 < a / b \ 0 < a \ 0 < b \ a < 0 \ b < 0" and divide_less_0_iff: "a / b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" and zero_le_divide_iff: "0 \ a / b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" and divide_le_0_iff: "a / b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" by (auto simp add: field_split_simps) text \Division and the Number One\ text\Simplify expressions equated with 1\ lemma zero_eq_1_divide_iff [simp]: "0 = 1 / a \ a = 0" by (cases "a = 0") (auto simp: field_simps) lemma one_divide_eq_0_iff [simp]: "1 / a = 0 \ a = 0" using zero_eq_1_divide_iff[of a] by simp text\Simplify expressions such as \0 < 1/x\ to \0 < x\\ lemma zero_le_divide_1_iff [simp]: "0 \ 1 / a \ 0 \ a" by (simp add: zero_le_divide_iff) lemma zero_less_divide_1_iff [simp]: "0 < 1 / a \ 0 < a" by (simp add: zero_less_divide_iff) lemma divide_le_0_1_iff [simp]: "1 / a \ 0 \ a \ 0" by (simp add: divide_le_0_iff) lemma divide_less_0_1_iff [simp]: "1 / a < 0 \ a < 0" by (simp add: divide_less_0_iff) lemma divide_right_mono: "\a \ b; 0 \ c\ \ a/c \ b/c" by (force simp add: divide_strict_right_mono le_less) lemma divide_right_mono_neg: "a \ b \ c \ 0 \ b / c \ a / c" by (auto dest: divide_right_mono [of _ _ "- c"]) lemma divide_left_mono_neg: "a \ b \ c \ 0 \ 0 < a * b \ c / a \ c / b" by (auto simp add: mult.commute dest: divide_left_mono [of _ _ "- c"]) lemma inverse_le_iff: "inverse a \ inverse b \ (0 < a * b \ b \ a) \ (a * b \ 0 \ a \ b)" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: field_simps zero_less_mult_iff mult_le_0_iff) lemma inverse_less_iff: "inverse a < inverse b \ (0 < a * b \ b < a) \ (a * b \ 0 \ a < b)" by (subst less_le) (auto simp: inverse_le_iff) lemma divide_le_cancel: "a / c \ b / c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: divide_inverse mult_le_cancel_right) lemma divide_less_cancel: "a / c < b / c \ (0 < c \ a < b) \ (c < 0 \ b < a) \ c \ 0" by (auto simp add: divide_inverse mult_less_cancel_right) text\Simplify quotients that are compared with the value 1.\ lemma le_divide_eq_1: "(1 \ b / a) = ((0 < a \ a \ b) \ (a < 0 \ b \ a))" by (auto simp add: le_divide_eq) lemma divide_le_eq_1: "(b / a \ 1) = ((0 < a \ b \ a) \ (a < 0 \ a \ b) \ a=0)" by (auto simp add: divide_le_eq) lemma less_divide_eq_1: "(1 < b / a) = ((0 < a \ a < b) \ (a < 0 \ b < a))" by (auto simp add: less_divide_eq) lemma divide_less_eq_1: "(b / a < 1) = ((0 < a \ b < a) \ (a < 0 \ a < b) \ a=0)" by (auto simp add: divide_less_eq) lemma divide_nonneg_nonneg [simp]: "0 \ x \ 0 \ y \ 0 \ x / y" by (auto simp add: field_split_simps) lemma divide_nonpos_nonpos: "x \ 0 \ y \ 0 \ 0 \ x / y" by (auto simp add: field_split_simps) lemma divide_nonneg_nonpos: "0 \ x \ y \ 0 \ x / y \ 0" by (auto simp add: field_split_simps) lemma divide_nonpos_nonneg: "x \ 0 \ 0 \ y \ x / y \ 0" by (auto simp add: field_split_simps) text \Conditional Simplification Rules: No Case Splits\ lemma le_divide_eq_1_pos [simp]: "0 < a \ (1 \ b/a) = (a \ b)" by (auto simp add: le_divide_eq) lemma le_divide_eq_1_neg [simp]: "a < 0 \ (1 \ b/a) = (b \ a)" by (auto simp add: le_divide_eq) lemma divide_le_eq_1_pos [simp]: "0 < a \ (b/a \ 1) = (b \ a)" by (auto simp add: divide_le_eq) lemma divide_le_eq_1_neg [simp]: "a < 0 \ (b/a \ 1) = (a \ b)" by (auto simp add: divide_le_eq) lemma less_divide_eq_1_pos [simp]: "0 < a \ (1 < b/a) = (a < b)" by (auto simp add: less_divide_eq) lemma less_divide_eq_1_neg [simp]: "a < 0 \ (1 < b/a) = (b < a)" by (auto simp add: less_divide_eq) lemma divide_less_eq_1_pos [simp]: "0 < a \ (b/a < 1) = (b < a)" by (auto simp add: divide_less_eq) lemma divide_less_eq_1_neg [simp]: "a < 0 \ b/a < 1 \ a < b" by (auto simp add: divide_less_eq) lemma eq_divide_eq_1 [simp]: "(1 = b/a) = ((a \ 0 \ a = b))" by (auto simp add: eq_divide_eq) lemma divide_eq_eq_1 [simp]: "(b/a = 1) = ((a \ 0 \ a = b))" by (auto simp add: divide_eq_eq) lemma abs_div_pos: "0 < y \ \x\ / y = \x / y\" by (simp add: order_less_imp_le) lemma zero_le_divide_abs_iff [simp]: "(0 \ a / \b\) = (0 \ a \ b = 0)" by (auto simp: zero_le_divide_iff) lemma divide_le_0_abs_iff [simp]: "(a / \b\ \ 0) = (a \ 0 \ b = 0)" by (auto simp: divide_le_0_iff) lemma field_le_mult_one_interval: assumes *: "\z. \ 0 < z ; z < 1 \ \ z * x \ y" shows "x \ y" proof (cases "0 < x") assume "0 < x" thus ?thesis using dense_le_bounded[of 0 1 "y/x"] * unfolding le_divide_eq if_P[OF \0 < x\] by simp next assume "\0 < x" hence "x \ 0" by simp obtain s::'a where s: "0 < s" "s < 1" using dense[of 0 "1::'a"] by auto hence "x \ s * x" using mult_le_cancel_right[of 1 x s] \x \ 0\ by auto also note *[OF s] finally show ?thesis . qed text\For creating values between \<^term>\u\ and \<^term>\v\.\ lemma scaling_mono: assumes "u \ v" "0 \ r" "r \ s" shows "u + r * (v - u) / s \ v" proof - have "r/s \ 1" using assms using divide_le_eq_1 by fastforce moreover have "0 \ v - u" using assms by simp ultimately have "(r/s) * (v - u) \ 1 * (v - u)" by (rule mult_right_mono) then show ?thesis by (simp add: field_simps) qed end text \Min/max Simplification Rules\ lemma min_mult_distrib_left: fixes x::"'a::linordered_idom" shows "p * min x y = (if 0 \ p then min (p*x) (p*y) else max (p*x) (p*y))" by (auto simp add: min_def max_def mult_le_cancel_left) lemma min_mult_distrib_right: fixes x::"'a::linordered_idom" shows "min x y * p = (if 0 \ p then min (x*p) (y*p) else max (x*p) (y*p))" by (auto simp add: min_def max_def mult_le_cancel_right) lemma min_divide_distrib_right: fixes x::"'a::linordered_field" shows "min x y / p = (if 0 \ p then min (x/p) (y/p) else max (x/p) (y/p))" by (simp add: min_mult_distrib_right divide_inverse) lemma max_mult_distrib_left: fixes x::"'a::linordered_idom" shows "p * max x y = (if 0 \ p then max (p*x) (p*y) else min (p*x) (p*y))" by (auto simp add: min_def max_def mult_le_cancel_left) lemma max_mult_distrib_right: fixes x::"'a::linordered_idom" shows "max x y * p = (if 0 \ p then max (x*p) (y*p) else min (x*p) (y*p))" by (auto simp add: min_def max_def mult_le_cancel_right) lemma max_divide_distrib_right: fixes x::"'a::linordered_field" shows "max x y / p = (if 0 \ p then max (x/p) (y/p) else min (x/p) (y/p))" by (simp add: max_mult_distrib_right divide_inverse) hide_fact (open) field_inverse field_divide_inverse field_inverse_zero code_identifier code_module Fields \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Finite_Set.thy b/src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy +++ b/src/HOL/Finite_Set.thy @@ -1,2606 +1,2607 @@ (* Title: HOL/Finite_Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad Author: Andrei Popescu *) section \Finite sets\ theory Finite_Set imports Product_Type Sum_Type Fields begin subsection \Predicate for finite sets\ context notes [[inductive_internals]] begin inductive finite :: "'a set \ bool" where emptyI [simp, intro!]: "finite {}" | insertI [simp, intro!]: "finite A \ finite (insert a A)" end simproc_setup finite_Collect ("finite (Collect P)") = \K Set_Comprehension_Pointfree.simproc\ declare [[simproc del: finite_Collect]] lemma finite_induct [case_names empty insert, induct set: finite]: \ \Discharging \x \ F\ entails extra work.\ assumes "finite F" assumes "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P F" using \finite F\ proof induct show "P {}" by fact next fix x F assume F: "finite F" and P: "P F" show "P (insert x F)" proof cases assume "x \ F" then have "insert x F = F" by (rule insert_absorb) with P show ?thesis by (simp only:) next assume "x \ F" from F this P show ?thesis by (rule insert) qed qed lemma infinite_finite_induct [case_names infinite empty insert]: assumes infinite: "\A. \ finite A \ P A" and empty: "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P A" proof (cases "finite A") case False with infinite show ?thesis . next case True then show ?thesis by (induct A) (fact empty insert)+ qed subsubsection \Choice principles\ lemma ex_new_if_finite: \ \does not depend on def of finite at all\ assumes "\ finite (UNIV :: 'a set)" and "finite A" shows "\a::'a. a \ A" proof - from assms have "A \ UNIV" by blast then show ?thesis by blast qed text \A finite choice principle. Does not need the SOME choice operator.\ lemma finite_set_choice: "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert a A) then obtain f b where f: "\x\A. P x (f x)" and ab: "P a b" by auto show ?case (is "\f. ?P f") proof show "?P (\x. if x = a then b else f x)" using f ab by auto qed qed subsubsection \Finite sets are the images of initial segments of natural numbers\ lemma finite_imp_nat_seg_image_inj_on: assumes "finite A" shows "\(n::nat) f. A = f ` {i. i < n} \ inj_on f {i. i < n}" using assms proof induct case empty show ?case proof show "\f. {} = f ` {i::nat. i < 0} \ inj_on f {i. i < 0}" by simp qed next case (insert a A) have notinA: "a \ A" by fact from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast then have "insert a A = f(n:=a) ` {i. i < Suc n}" and "inj_on (f(n:=a)) {i. i < Suc n}" using notinA by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq) then show ?case by blast qed lemma nat_seg_image_imp_finite: "A = f ` {i::nat. i < n} \ finite A" proof (induct n arbitrary: A) case 0 then show ?case by simp next case (Suc n) let ?B = "f ` {i. i < n}" have finB: "finite ?B" by (rule Suc.hyps[OF refl]) show ?case proof (cases "\k (\n f. A = f ` {i::nat. i < n})" by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on) lemma finite_imp_inj_to_nat_seg: assumes "finite A" shows "\f n. f ` A = {i::nat. i < n} \ inj_on f A" proof - from finite_imp_nat_seg_image_inj_on [OF \finite A\] obtain f and n :: nat where bij: "bij_betw f {i. i ?f ` A = {i. i k}" by (simp add: le_eq_less_or_eq Collect_disj_eq) subsection \Finiteness and common set operations\ lemma rev_finite_subset: "finite B \ A \ B \ finite A" proof (induct arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert x F A) have A: "A \ insert x F" and r: "A - {x} \ F \ finite (A - {x})" by fact+ show "finite A" proof cases assume x: "x \ A" with A have "A - {x} \ F" by (simp add: subset_insert_iff) with r have "finite (A - {x})" . then have "finite (insert x (A - {x}))" .. also have "insert x (A - {x}) = A" using x by (rule insert_Diff) finally show ?thesis . next show ?thesis when "A \ F" using that by fact assume "x \ A" with A show "A \ F" by (simp add: subset_insert_iff) qed qed lemma finite_subset: "A \ B \ finite B \ finite A" by (rule rev_finite_subset) lemma finite_UnI: assumes "finite F" and "finite G" shows "finite (F \ G)" using assms by induct simp_all lemma finite_Un [iff]: "finite (F \ G) \ finite F \ finite G" by (blast intro: finite_UnI finite_subset [of _ "F \ G"]) lemma finite_insert [simp]: "finite (insert a A) \ finite A" proof - have "finite {a} \ finite A \ finite A" by simp then have "finite ({a} \ A) \ finite A" by (simp only: finite_Un) then show ?thesis by simp qed lemma finite_Int [simp, intro]: "finite F \ finite G \ finite (F \ G)" by (blast intro: finite_subset) lemma finite_Collect_conjI [simp, intro]: "finite {x. P x} \ finite {x. Q x} \ finite {x. P x \ Q x}" by (simp add: Collect_conj_eq) lemma finite_Collect_disjI [simp]: "finite {x. P x \ Q x} \ finite {x. P x} \ finite {x. Q x}" by (simp add: Collect_disj_eq) lemma finite_Diff [simp, intro]: "finite A \ finite (A - B)" by (rule finite_subset, rule Diff_subset) lemma finite_Diff2 [simp]: assumes "finite B" shows "finite (A - B) \ finite A" proof - have "finite A \ finite ((A - B) \ (A \ B))" by (simp add: Un_Diff_Int) also have "\ \ finite (A - B)" using \finite B\ by simp finally show ?thesis .. qed lemma finite_Diff_insert [iff]: "finite (A - insert a B) \ finite (A - B)" proof - have "finite (A - B) \ finite (A - B - {a})" by simp moreover have "A - insert a B = A - B - {a}" by auto ultimately show ?thesis by simp qed lemma finite_compl [simp]: "finite (A :: 'a set) \ finite (- A) \ finite (UNIV :: 'a set)" by (simp add: Compl_eq_Diff_UNIV) lemma finite_Collect_not [simp]: "finite {x :: 'a. P x} \ finite {x. \ P x} \ finite (UNIV :: 'a set)" by (simp add: Collect_neg_eq) lemma finite_Union [simp, intro]: "finite A \ (\M. M \ A \ finite M) \ finite (\A)" by (induct rule: finite_induct) simp_all lemma finite_UN_I [intro]: "finite A \ (\a. a \ A \ finite (B a)) \ finite (\a\A. B a)" by (induct rule: finite_induct) simp_all lemma finite_UN [simp]: "finite A \ finite (\(B ` A)) \ (\x\A. finite (B x))" by (blast intro: finite_subset) lemma finite_Inter [intro]: "\A\M. finite A \ finite (\M)" by (blast intro: Inter_lower finite_subset) lemma finite_INT [intro]: "\x\I. finite (A x) \ finite (\x\I. A x)" by (blast intro: INT_lower finite_subset) lemma finite_imageI [simp, intro]: "finite F \ finite (h ` F)" by (induct rule: finite_induct) simp_all lemma finite_image_set [simp]: "finite {x. P x} \ finite {f x |x. P x}" by (simp add: image_Collect [symmetric]) lemma finite_image_set2: "finite {x. P x} \ finite {y. Q y} \ finite {f x y |x y. P x \ Q y}" by (rule finite_subset [where B = "\x \ {x. P x}. \y \ {y. Q y}. {f x y}"]) auto lemma finite_imageD: assumes "finite (f ` A)" and "inj_on f A" shows "finite A" using assms proof (induct "f ` A" arbitrary: A) case empty then show ?case by simp next case (insert x B) then have B_A: "insert x B = f ` A" by simp then obtain y where "x = f y" and "y \ A" by blast from B_A \x \ B\ have "B = f ` A - {x}" by blast with B_A \x \ B\ \x = f y\ \inj_on f A\ \y \ A\ have "B = f ` (A - {y})" by (simp add: inj_on_image_set_diff) moreover from \inj_on f A\ have "inj_on f (A - {y})" by (rule inj_on_diff) ultimately have "finite (A - {y})" by (rule insert.hyps) then show "finite A" by simp qed lemma finite_image_iff: "inj_on f A \ finite (f ` A) \ finite A" using finite_imageD by blast lemma finite_surj: "finite A \ B \ f ` A \ finite B" by (erule finite_subset) (rule finite_imageI) lemma finite_range_imageI: "finite (range g) \ finite (range (\x. f (g x)))" by (drule finite_imageI) (simp add: range_composition) lemma finite_subset_image: assumes "finite B" shows "B \ f ` A \ \C\A. finite C \ B = f ` C" using assms proof induct case empty then show ?case by simp next case insert then show ?case by (clarsimp simp del: image_insert simp add: image_insert [symmetric]) blast qed lemma all_subset_image: "(\B. B \ f ` A \ P B) \ (\B. B \ A \ P(f ` B))" by (safe elim!: subset_imageE) (use image_mono in \blast+\) (* slow *) lemma all_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and P: "\B. finite B \ B \ A \ P (f ` B)" show "P B" using finite_subset_image [OF B] P by blast qed blast lemma ex_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and "P B" show "\B. finite B \ B \ A \ P (f ` B)" using finite_subset_image [OF B] \P B\ by blast qed blast lemma finite_vimage_IntI: "finite F \ inj_on h A \ finite (h -` F \ A)" proof (induct rule: finite_induct) case (insert x F) then show ?case by (simp add: vimage_insert [of h x F] finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2) qed simp lemma finite_finite_vimage_IntI: assumes "finite F" and "\y. y \ F \ finite ((h -` {y}) \ A)" shows "finite (h -` F \ A)" proof - have *: "h -` F \ A = (\ y\F. (h -` {y}) \ A)" by blast show ?thesis by (simp only: * assms finite_UN_I) qed lemma finite_vimageI: "finite F \ inj h \ finite (h -` F)" using finite_vimage_IntI[of F h UNIV] by auto lemma finite_vimageD': "finite (f -` A) \ A \ range f \ finite A" by (auto simp add: subset_image_iff intro: finite_subset[rotated]) lemma finite_vimageD: "finite (h -` F) \ surj h \ finite F" by (auto dest: finite_vimageD') lemma finite_vimage_iff: "bij h \ finite (h -` F) \ finite F" unfolding bij_def by (auto elim: finite_vimageD finite_vimageI) lemma finite_inverse_image_gen: assumes "finite A" "inj_on f D" shows "finite {j\D. f j \ A}" using finite_vimage_IntI [OF assms] by (simp add: Collect_conj_eq inf_commute vimage_def) lemma finite_inverse_image: assumes "finite A" "inj f" shows "finite {j. f j \ A}" using finite_inverse_image_gen [OF assms] by simp lemma finite_Collect_bex [simp]: assumes "finite A" shows "finite {x. \y\A. Q x y} \ (\y\A. finite {x. Q x y})" proof - have "{x. \y\A. Q x y} = (\y\A. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Collect_bounded_ex [simp]: assumes "finite {y. P y}" shows "finite {x. \y. P y \ Q x y} \ (\y. P y \ finite {x. Q x y})" proof - have "{x. \y. P y \ Q x y} = (\y\{y. P y}. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Plus: "finite A \ finite B \ finite (A <+> B)" by (simp add: Plus_def) lemma finite_PlusD: fixes A :: "'a set" and B :: "'b set" assumes fin: "finite (A <+> B)" shows "finite A" "finite B" proof - have "Inl ` A \ A <+> B" by auto then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite A" by (rule finite_imageD) (auto intro: inj_onI) next have "Inr ` B \ A <+> B" by auto then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite B" by (rule finite_imageD) (auto intro: inj_onI) qed lemma finite_Plus_iff [simp]: "finite (A <+> B) \ finite A \ finite B" by (auto intro: finite_PlusD finite_Plus) lemma finite_Plus_UNIV_iff [simp]: "finite (UNIV :: ('a + 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff) lemma finite_SigmaI [simp, intro]: "finite A \ (\a. a\A \ finite (B a)) \ finite (SIGMA a:A. B a)" unfolding Sigma_def by blast lemma finite_SigmaI2: assumes "finite {x\A. B x \ {}}" and "\a. a \ A \ finite (B a)" shows "finite (Sigma A B)" proof - from assms have "finite (Sigma {x\A. B x \ {}} B)" by auto also have "Sigma {x:A. B x \ {}} B = Sigma A B" by auto finally show ?thesis . qed lemma finite_cartesian_product: "finite A \ finite B \ finite (A \ B)" by (rule finite_SigmaI) lemma finite_Prod_UNIV: "finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ finite (UNIV :: ('a \ 'b) set)" by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product) lemma finite_cartesian_productD1: assumes "finite (A \ B)" and "B \ {}" shows "finite A" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "fst ` (A \ B) = fst ` f ` {i::nat. i < n}" by simp with \B \ {}\ have "A = (fst \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. A = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_productD2: assumes "finite (A \ B)" and "A \ {}" shows "finite B" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "snd ` (A \ B) = snd ` f ` {i::nat. i < n}" by simp with \A \ {}\ have "B = (snd \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. B = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_product_iff: "finite (A \ B) \ (A = {} \ B = {} \ (finite A \ finite B))" by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product) lemma finite_prod: "finite (UNIV :: ('a \ 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" using finite_cartesian_product_iff[of UNIV UNIV] by simp lemma finite_Pow_iff [iff]: "finite (Pow A) \ finite A" proof assume "finite (Pow A)" then have "finite ((\x. {x}) ` A)" by (blast intro: finite_subset) (* somewhat slow *) then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp next assume "finite A" then show "finite (Pow A)" by induct (simp_all add: Pow_insert) qed corollary finite_Collect_subsets [simp, intro]: "finite A \ finite {B. B \ A}" by (simp add: Pow_def [symmetric]) lemma finite_set: "finite (UNIV :: 'a set set) \ finite (UNIV :: 'a set)" by (simp only: finite_Pow_iff Pow_UNIV[symmetric]) lemma finite_UnionD: "finite (\A) \ finite A" by (blast intro: finite_subset [OF subset_Pow_Union]) lemma finite_bind: assumes "finite S" assumes "\x \ S. finite (f x)" shows "finite (Set.bind S f)" using assms by (simp add: bind_UNION) lemma finite_filter [simp]: "finite S \ finite (Set.filter P S)" unfolding Set.filter_def by simp lemma finite_set_of_finite_funs: assumes "finite A" "finite B" shows "finite {f. \x. (x \ A \ f x \ B) \ (x \ A \ f x = d)}" (is "finite ?S") proof - let ?F = "\f. {(a,b). a \ A \ b = f a}" have "?F ` ?S \ Pow(A \ B)" by auto from finite_subset[OF this] assms have 1: "finite (?F ` ?S)" by simp have 2: "inj_on ?F ?S" by (fastforce simp add: inj_on_def set_eq_iff fun_eq_iff) (* somewhat slow *) show ?thesis by (rule finite_imageD [OF 1 2]) qed lemma not_finite_existsD: assumes "\ finite {a. P a}" shows "\a. P a" proof (rule classical) assume "\ ?thesis" with assms show ?thesis by auto qed subsection \Further induction rules on finite sets\ lemma finite_ne_induct [case_names singleton insert, consumes 2]: assumes "finite F" and "F \ {}" assumes "\x. P {x}" and "\x F. finite F \ F \ {} \ x \ F \ P F \ P (insert x F)" shows "P F" using assms proof induct case empty then show ?case by simp next case (insert x F) then show ?case by cases auto qed lemma finite_subset_induct [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. finite F \ a \ A \ a \ F \ P F \ P (insert a F)" shows "P F" using \finite F\ \F \ A\ proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact qed qed lemma finite_empty_induct: assumes "finite A" and "P A" and remove: "\a A. finite A \ a \ A \ P A \ P (A - {a})" shows "P {}" proof - have "P (A - B)" if "B \ A" for B :: "'a set" proof - from \finite A\ that have "finite B" by (rule rev_finite_subset) from this \B \ A\ show "P (A - B)" proof induct case empty from \P A\ show ?case by simp next case (insert b B) have "P (A - B - {b})" proof (rule remove) from \finite A\ show "finite (A - B)" by induct auto from insert show "b \ A - B" by simp from insert show "P (A - B)" by simp qed also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric]) finally show ?case . qed qed then have "P (A - A)" by blast then show ?thesis by simp qed lemma finite_update_induct [consumes 1, case_names const update]: assumes finite: "finite {a. f a \ c}" and const: "P (\a. c)" and update: "\a b f. finite {a. f a \ c} \ f a = c \ b \ c \ P f \ P (f(a := b))" shows "P f" using finite proof (induct "{a. f a \ c}" arbitrary: f) case empty with const show ?case by simp next case (insert a A) then have "A = {a'. (f(a := c)) a' \ c}" and "f a \ c" by auto with \finite A\ have "finite {a'. (f(a := c)) a' \ c}" by simp have "(f(a := c)) a = c" by simp from insert \A = {a'. (f(a := c)) a' \ c}\ have "P (f(a := c))" by simp with \finite {a'. (f(a := c)) a' \ c}\ \(f(a := c)) a = c\ \f a \ c\ have "P ((f(a := c))(a := f a))" by (rule update) then show ?case by simp qed lemma finite_subset_induct' [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. \finite F; a \ A; F \ A; a \ F; P F \ \ P (insert a F)" shows "P F" using assms(1,2) proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact show "F \ A" by fact qed qed subsection \Class \finite\\ class finite = assumes finite_UNIV: "finite (UNIV :: 'a set)" begin lemma finite [simp]: "finite (A :: 'a set)" by (rule subset_UNIV finite_UNIV finite_subset)+ lemma finite_code [code]: "finite (A :: 'a set) \ True" by simp end instance prod :: (finite, finite) finite by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite) lemma inj_graph: "inj (\f. {(x, y). y = f x})" by (rule inj_onI) (auto simp add: set_eq_iff fun_eq_iff) instance "fun" :: (finite, finite) finite proof show "finite (UNIV :: ('a \ 'b) set)" proof (rule finite_imageD) let ?graph = "\f::'a \ 'b. {(x, y). y = f x}" have "range ?graph \ Pow UNIV" by simp moreover have "finite (Pow (UNIV :: ('a * 'b) set))" by (simp only: finite_Pow_iff finite) ultimately show "finite (range ?graph)" by (rule finite_subset) show "inj ?graph" by (rule inj_graph) qed qed instance bool :: finite by standard (simp add: UNIV_bool) instance set :: (finite) finite by standard (simp only: Pow_UNIV [symmetric] finite_Pow_iff finite) instance unit :: finite by standard (simp add: UNIV_unit) instance sum :: (finite, finite) finite by standard (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite) subsection \A basic fold functional for finite sets\ text \ The intended behaviour is \fold f z {x\<^sub>1, \, x\<^sub>n} = f x\<^sub>1 (\ (f x\<^sub>n z)\)\ if \f\ is ``left-commutative''. The commutativity requirement is relativised to the carrier set \S\: \ locale comp_fun_commute_on = fixes S :: "'a set" fixes f :: "'a \ 'b \ 'b" assumes comp_fun_commute_on: "x \ S \ y \ S \ f y \ f x = f x \ f y" begin lemma fun_left_comm: "x \ S \ y \ S \ f y (f x z) = f x (f y z)" using comp_fun_commute_on by (simp add: fun_eq_iff) lemma commute_left_comp: "x \ S \ y \ S \ f y \ (f x \ g) = f x \ (f y \ g)" by (simp add: o_assoc comp_fun_commute_on) end inductive fold_graph :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" for f :: "'a \ 'b \ 'b" and z :: 'b where emptyI [intro]: "fold_graph f z {} z" | insertI [intro]: "x \ A \ fold_graph f z A y \ fold_graph f z (insert x A) (f x y)" inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x" lemma fold_graph_closed_lemma: "fold_graph f z A x \ x \ B" if "fold_graph g z A x" "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using that(1-3) proof (induction rule: fold_graph.induct) case (insertI x A y) have "fold_graph f z A y" "y \ B" unfolding atomize_conj by (rule insertI.IH) (auto intro: insertI.prems) then have "g x y \ B" and f_eq: "f x y = g x y" by (auto simp: insertI.prems) moreover have "fold_graph f z (insert x A) (f x y)" by (rule fold_graph.insertI; fact) ultimately show ?case by (simp add: f_eq) qed (auto intro!: that) lemma fold_graph_closed_eq: "fold_graph f z A = fold_graph g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using fold_graph_closed_lemma[of f z A _ B g] fold_graph_closed_lemma[of g z A _ B f] that by auto definition fold :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" where "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)" lemma fold_closed_eq: "fold f z A = fold g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" unfolding Finite_Set.fold_def by (subst fold_graph_closed_eq[where B=B and g=g]) (auto simp: that) text \ A tempting alternative for the definition is \<^term>\if finite A then THE y. fold_graph f z A y else e\. It allows the removal of finiteness assumptions from the theorems \fold_comm\, \fold_reindex\ and \fold_distrib\. The proofs become ugly. It is not worth the effort. (???) \ lemma finite_imp_fold_graph: "finite A \ \x. fold_graph f z A x" by (induct rule: finite_induct) auto subsubsection \From \<^const>\fold_graph\ to \<^term>\fold\\ context comp_fun_commute_on begin lemma fold_graph_finite: assumes "fold_graph f z A y" shows "finite A" using assms by induct simp_all lemma fold_graph_insertE_aux: assumes "A \ S" assumes "fold_graph f z A y" "a \ A" shows "\y'. y = f a y' \ fold_graph f z (A - {a}) y'" using assms(2-,1) proof (induct set: fold_graph) case emptyI then show ?case by simp next case (insertI x A y) show ?case proof (cases "x = a") case True with insertI show ?thesis by auto next case False then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'" using insertI by auto from insertI have "x \ S" "a \ S" by auto then have "f x y = f a (f x y')" unfolding y by (intro fun_left_comm; simp) moreover have "fold_graph f z (insert x A - {a}) (f x y')" using y' and \x \ a\ and \x \ A\ by (simp add: insert_Diff_if fold_graph.insertI) ultimately show ?thesis by fast qed qed lemma fold_graph_insertE: assumes "insert x A \ S" assumes "fold_graph f z (insert x A) v" and "x \ A" obtains y where "v = f x y" and "fold_graph f z A y" using assms by (auto dest: fold_graph_insertE_aux[OF \insert x A \ S\ _ insertI1]) lemma fold_graph_determ: assumes "A \ S" assumes "fold_graph f z A x" "fold_graph f z A y" shows "y = x" using assms(2-,1) proof (induct arbitrary: y set: fold_graph) case emptyI then show ?case by fast next case (insertI x A y v) from \insert x A \ S\ and \fold_graph f z (insert x A) v\ and \x \ A\ obtain y' where "v = f x y'" and "fold_graph f z A y'" by (rule fold_graph_insertE) from \fold_graph f z A y'\ insertI have "y' = y" by simp with \v = f x y'\ show "v = f x y" by simp qed lemma fold_equality: "A \ S \ fold_graph f z A y \ fold f z A = y" by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite) lemma fold_graph_fold: assumes "A \ S" assumes "finite A" shows "fold_graph f z A (fold f z A)" proof - from \finite A\ have "\x. fold_graph f z A x" by (rule finite_imp_fold_graph) moreover note fold_graph_determ[OF \A \ S\] ultimately have "\!x. fold_graph f z A x" by (rule ex_ex1I) then have "fold_graph f z A (The (fold_graph f z A))" by (rule theI') with assms show ?thesis by (simp add: fold_def) qed text \The base case for \fold\:\ lemma (in -) fold_infinite [simp]: "\ finite A \ fold f z A = z" by (auto simp: fold_def) lemma (in -) fold_empty [simp]: "fold f z {} = z" by (auto simp: fold_def) text \The various recursion equations for \<^const>\fold\:\ lemma fold_insert [simp]: assumes "insert x A \ S" assumes "finite A" and "x \ A" shows "fold f z (insert x A) = f x (fold f z A)" proof (rule fold_equality[OF \insert x A \ S\]) fix z from \insert x A \ S\ \finite A\ have "fold_graph f z A (fold f z A)" by (blast intro: fold_graph_fold) with \x \ A\ have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI) then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp qed declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del] \ \No more proofs involve these.\ lemma fold_fun_left_comm: assumes "insert x A \ S" "finite A" shows "f x (fold f z A) = fold f (f x z) A" using assms(2,1) proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert y F) then have "fold f (f x z) (insert y F) = f y (fold f (f x z) F)" by simp also have "\ = f x (f y (fold f z F))" using insert by (simp add: fun_left_comm[where ?y=x]) also have "\ = f x (fold f z (insert y F))" proof - from insert have "insert y F \ S" by simp from fold_insert[OF this] insert show ?thesis by simp qed finally show ?case .. qed lemma fold_insert2: "insert x A \ S \ finite A \ x \ A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) lemma fold_rec: assumes "A \ S" assumes "finite A" and "x \ A" shows "fold f z A = f x (fold f z (A - {x}))" proof - have A: "A = insert x (A - {x})" using \x \ A\ by blast then have "fold f z A = fold f z (insert x (A - {x}))" by simp also have "\ = f x (fold f z (A - {x}))" by (rule fold_insert) (use assms in \auto\) finally show ?thesis . qed lemma fold_insert_remove: assumes "insert x A \ S" assumes "finite A" shows "fold f z (insert x A) = f x (fold f z (A - {x}))" proof - from \finite A\ have "finite (insert x A)" by auto moreover have "x \ insert x A" by auto ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))" using \insert x A \ S\ by (blast intro: fold_rec) then show ?thesis by simp qed lemma fold_set_union_disj: assumes "A \ S" "B \ S" assumes "finite A" "finite B" "A \ B = {}" shows "Finite_Set.fold f z (A \ B) = Finite_Set.fold f (Finite_Set.fold f z A) B" using \finite B\ assms(1,2,3,5) proof induct case (insert x F) have "fold f z (A \ insert x F) = f x (fold f (fold f z A) F)" using insert by auto also have "\ = fold f (fold f z A) (insert x F)" using insert by (blast intro: fold_insert[symmetric]) finally show ?case . qed simp end text \Other properties of \<^const>\fold\:\ lemma fold_graph_image: assumes "inj_on g A" shows "fold_graph f z (g ` A) = fold_graph (f \ g) z A" proof fix w show "fold_graph f z (g ` A) w = fold_graph (f o g) z A w" proof assume "fold_graph f z (g ` A) w" then show "fold_graph (f \ g) z A w" using assms proof (induct "g ` A" w arbitrary: A) case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r B) from \inj_on g B\ \x \ A\ \insert x A = image g B\ obtain x' A' where "x' \ A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'" by (rule inj_img_insertE) from insertI.prems have "fold_graph (f \ g) z A' r" by (auto intro: insertI.hyps) with \x' \ A'\ have "fold_graph (f \ g) z (insert x' A') ((f \ g) x' r)" by (rule fold_graph.insertI) then show ?case by simp qed next assume "fold_graph (f \ g) z A w" then show "fold_graph f z (g ` A) w" using assms proof induct case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r) from \x \ A\ insertI.prems have "g x \ g ` A" by auto moreover from insertI have "fold_graph f z (g ` A) r" by simp ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)" by (rule fold_graph.insertI) then show ?case by simp qed qed qed lemma fold_image: assumes "inj_on g A" shows "fold f z (g ` A) = fold (f \ g) z A" proof (cases "finite A") case False with assms show ?thesis by (auto dest: finite_imageD simp add: fold_def) next case True then show ?thesis by (auto simp add: fold_def fold_graph_image[OF assms]) qed lemma fold_cong: assumes "comp_fun_commute_on S f" "comp_fun_commute_on S g" and "A \ S" "finite A" and cong: "\x. x \ A \ f x = g x" and "s = t" and "A = B" shows "fold f s A = fold g t B" proof - have "fold f s A = fold g s A" using \finite A\ \A \ S\ cong proof (induct A) case empty then show ?case by simp next case insert interpret f: comp_fun_commute_on S f by (fact \comp_fun_commute_on S f\) interpret g: comp_fun_commute_on S g by (fact \comp_fun_commute_on S g\) from insert show ?case by simp qed with assms show ?thesis by simp qed text \A simplified version for idempotent functions:\ locale comp_fun_idem_on = comp_fun_commute_on + assumes comp_fun_idem_on: "x \ S \ f x \ f x = f x" begin lemma fun_left_idem: "x \ S \ f x (f x z) = f x z" using comp_fun_idem_on by (simp add: fun_eq_iff) lemma fold_insert_idem: assumes "insert x A \ S" assumes fin: "finite A" shows "fold f z (insert x A) = f x (fold f z A)" proof cases assume "x \ A" then obtain B where "A = insert x B" and "x \ B" by (rule set_insert) then show ?thesis using assms by (simp add: comp_fun_idem_on fun_left_idem) next assume "x \ A" then show ?thesis using assms by auto qed declare fold_insert [simp del] fold_insert_idem [simp] lemma fold_insert_idem2: "insert x A \ S \ finite A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) end subsubsection \Liftings to \comp_fun_commute_on\ etc.\ lemma (in comp_fun_commute_on) comp_comp_fun_commute_on: "range g \ S \ comp_fun_commute_on R (f \ g)" by standard (force intro: comp_fun_commute_on) lemma (in comp_fun_idem_on) comp_comp_fun_idem_on: assumes "range g \ S" shows "comp_fun_idem_on R (f \ g)" proof interpret f_g: comp_fun_commute_on R "f o g" by (fact comp_comp_fun_commute_on[OF \range g \ S\]) show "x \ R \ y \ R \ (f \ g) y \ (f \ g) x = (f \ g) x \ (f \ g) y" for x y by (fact f_g.comp_fun_commute_on) qed (use \range g \ S\ in \force intro: comp_fun_idem_on\) lemma (in comp_fun_commute_on) comp_fun_commute_on_funpow: "comp_fun_commute_on S (\x. f x ^^ g x)" proof fix x y assume "x \ S" "y \ S" show "f y ^^ g y \ f x ^^ g x = f x ^^ g x \ f y ^^ g y" proof (cases "x = y") case True then show ?thesis by simp next case False show ?thesis proof (induct "g x" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) have hyp1: "f y ^^ g y \ f x = f x \ f y ^^ g y" proof (induct "g y" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) define h where "h z = g z - 1" for z with Suc have "n = h y" by simp with Suc have hyp: "f y ^^ h y \ f x = f x \ f y ^^ h y" by auto from Suc h_def have "g y = Suc (h y)" by simp with \x \ S\ \y \ S\ show ?case by (simp add: comp_assoc hyp) (simp add: o_assoc comp_fun_commute_on) qed define h where "h z = (if z = x then g x - 1 else g z)" for z with Suc have "n = h x" by simp with Suc have "f y ^^ h y \ f x ^^ h x = f x ^^ h x \ f y ^^ h y" by auto with False h_def have hyp2: "f y ^^ g y \ f x ^^ h x = f x ^^ h x \ f y ^^ g y" by simp from Suc h_def have "g x = Suc (h x)" by simp then show ?case by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2) (simp add: comp_assoc hyp1) qed qed qed subsubsection \\<^term>\UNIV\ as carrier set\ locale comp_fun_commute = fixes f :: "'a \ 'b \ 'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin lemma (in -) comp_fun_commute_def': "comp_fun_commute f = comp_fun_commute_on UNIV f" unfolding comp_fun_commute_def comp_fun_commute_on_def by blast text \ We abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale comp_fun_commute_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "comp_fun_commute_on UNIV f" by standard (simp add: comp_fun_commute) qed simp_all end lemma (in comp_fun_commute) comp_comp_fun_commute: "comp_fun_commute (f o g)" unfolding comp_fun_commute_def' by (fact comp_comp_fun_commute_on) lemma (in comp_fun_commute) comp_fun_commute_funpow: "comp_fun_commute (\x. f x ^^ g x)" unfolding comp_fun_commute_def' by (fact comp_fun_commute_on_funpow) locale comp_fun_idem = comp_fun_commute + assumes comp_fun_idem: "f x o f x = f x" begin lemma (in -) comp_fun_idem_def': "comp_fun_idem f = comp_fun_idem_on UNIV f" unfolding comp_fun_idem_on_def comp_fun_idem_def comp_fun_commute_def' unfolding comp_fun_idem_axioms_def comp_fun_idem_on_axioms_def by blast text \ Again, we abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale comp_fun_idem_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "comp_fun_idem_on UNIV f" by standard (simp_all add: comp_fun_idem comp_fun_commute) qed simp_all end lemma (in comp_fun_idem) comp_comp_fun_idem: "comp_fun_idem (f o g)" unfolding comp_fun_idem_def' by (fact comp_comp_fun_idem_on) subsubsection \Expressing set operations via \<^const>\fold\\ lemma comp_fun_commute_const: "comp_fun_commute (\_. f)" - by standard rule + by standard (rule refl) lemma comp_fun_idem_insert: "comp_fun_idem insert" by standard auto lemma comp_fun_idem_remove: "comp_fun_idem Set.remove" by standard auto lemma (in semilattice_inf) comp_fun_idem_inf: "comp_fun_idem inf" by standard (auto simp add: inf_left_commute) lemma (in semilattice_sup) comp_fun_idem_sup: "comp_fun_idem sup" by standard (auto simp add: sup_left_commute) lemma union_fold_insert: assumes "finite A" shows "A \ B = fold insert B A" proof - interpret comp_fun_idem insert by (fact comp_fun_idem_insert) from \finite A\ show ?thesis by (induct A arbitrary: B) simp_all qed lemma minus_fold_remove: assumes "finite A" shows "B - A = fold Set.remove B A" proof - interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove) from \finite A\ have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto (* slow *) then show ?thesis .. qed lemma comp_fun_commute_filter_fold: "comp_fun_commute (\x A'. if P x then Set.insert x A' else A')" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff) qed lemma Set_filter_fold: assumes "finite A" shows "Set.filter P A = fold (\x A'. if P x then Set.insert x A' else A') {} A" using assms proof - interpret commute_insert: comp_fun_commute "(\x A'. if P x then Set.insert x A' else A')" by (fact comp_fun_commute_filter_fold) from \finite A\ show ?thesis by induct (auto simp add: Set.filter_def) qed lemma inter_Set_filter: assumes "finite B" shows "A \ B = Set.filter (\x. x \ A) B" using assms by induct (auto simp: Set.filter_def) lemma image_fold_insert: assumes "finite A" shows "image f A = fold (\k A. Set.insert (f k) A) {} A" proof - interpret comp_fun_commute "\k A. Set.insert (f k) A" by standard auto show ?thesis using assms by (induct A) auto qed lemma Ball_fold: assumes "finite A" shows "Ball A P = fold (\k s. s \ P k) True A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma Bex_fold: assumes "finite A" shows "Bex A P = fold (\k s. s \ P k) False A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma comp_fun_commute_Pow_fold: "comp_fun_commute (\x A. A \ Set.insert x ` A)" by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast lemma Pow_fold: assumes "finite A" shows "Pow A = fold (\x A. A \ Set.insert x ` A) {{}} A" proof - interpret comp_fun_commute "\x A. A \ Set.insert x ` A" by (rule comp_fun_commute_Pow_fold) show ?thesis using assms by (induct A) (auto simp: Pow_insert) qed lemma fold_union_pair: assumes "finite B" shows "(\y\B. {(x, y)}) \ A = fold (\y. Set.insert (x, y)) A B" proof - interpret comp_fun_commute "\y. Set.insert (x, y)" by standard auto show ?thesis using assms by (induct arbitrary: A) simp_all qed lemma comp_fun_commute_product_fold: "finite B \ comp_fun_commute (\x z. fold (\y. Set.insert (x, y)) z B)" by standard (auto simp: fold_union_pair [symmetric]) lemma product_fold: assumes "finite A" "finite B" shows "A \ B = fold (\x z. fold (\y. Set.insert (x, y)) z B) {} A" proof - interpret commute_product: comp_fun_commute "(\x z. fold (\y. Set.insert (x, y)) z B)" by (fact comp_fun_commute_product_fold[OF \finite B\]) from assms show ?thesis unfolding Sigma_def by (induct A) (simp_all add: fold_union_pair) qed context complete_lattice begin lemma inf_Inf_fold_inf: assumes "finite A" shows "inf (Inf A) B = fold inf B A" proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: inf_commute fun_eq_iff) qed lemma sup_Sup_fold_sup: assumes "finite A" shows "sup (Sup A) B = fold sup B A" proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: sup_commute fun_eq_iff) qed lemma Inf_fold_inf: "finite A \ Inf A = fold inf top A" using inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2) lemma Sup_fold_sup: "finite A \ Sup A = fold sup bot A" using sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2) lemma inf_INF_fold_inf: assumes "finite A" shows "inf B (\(f ` A)) = fold (inf \ f) B A" (is "?inf = ?fold") proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) interpret comp_fun_idem "inf \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?inf" by (induct A arbitrary: B) (simp_all add: inf_left_commute) then show ?thesis .. qed lemma sup_SUP_fold_sup: assumes "finite A" shows "sup B (\(f ` A)) = fold (sup \ f) B A" (is "?sup = ?fold") proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) interpret comp_fun_idem "sup \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?sup" by (induct A arbitrary: B) (simp_all add: sup_left_commute) then show ?thesis .. qed lemma INF_fold_inf: "finite A \ \(f ` A) = fold (inf \ f) top A" using inf_INF_fold_inf [of A top] by simp lemma SUP_fold_sup: "finite A \ \(f ` A) = fold (sup \ f) bot A" using sup_SUP_fold_sup [of A bot] by simp lemma finite_Inf_in: assumes "finite A" "A\{}" and inf: "\x y. \x \ A; y \ A\ \ inf x y \ A" shows "Inf A \ A" proof - have "Inf B \ A" if "B \ A" "B\{}" for B using finite_subset [OF \B \ A\ \finite A\] that by (induction B) (use inf in \force+\) then show ?thesis by (simp add: assms) qed lemma finite_Sup_in: assumes "finite A" "A\{}" and sup: "\x y. \x \ A; y \ A\ \ sup x y \ A" shows "Sup A \ A" proof - have "Sup B \ A" if "B \ A" "B\{}" for B using finite_subset [OF \B \ A\ \finite A\] that by (induction B) (use sup in \force+\) then show ?thesis by (simp add: assms) qed end subsection \Locales as mini-packages for fold operations\ subsubsection \The natural case\ locale folding_on = fixes S :: "'a set" fixes f :: "'a \ 'b \ 'b" and z :: "'b" assumes comp_fun_commute_on: "x \ S \ y \ S \ f y o f x = f x o f y" begin interpretation fold?: comp_fun_commute_on S f by standard (simp add: comp_fun_commute_on) definition F :: "'a set \ 'b" where eq_fold: "F A = Finite_Set.fold f z A" lemma empty [simp]: "F {} = z" by (simp add: eq_fold) lemma infinite [simp]: "\ finite A \ F A = z" by (simp add: eq_fold) lemma insert [simp]: assumes "insert x A \ S" and "finite A" and "x \ A" shows "F (insert x A) = f x (F A)" proof - from fold_insert assms have "Finite_Set.fold f z (insert x A) = f x (Finite_Set.fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed lemma remove: assumes "A \ S" and "finite A" and "x \ A" shows "F A = f x (F (A - {x}))" proof - from \x \ A\ obtain B where A: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ A have "finite B" by simp ultimately show ?thesis using \A \ S\ by auto qed lemma insert_remove: assumes "insert x A \ S" and "finite A" shows "F (insert x A) = f x (F (A - {x}))" using assms by (cases "x \ A") (simp_all add: remove insert_absorb) end subsubsection \With idempotency\ locale folding_idem_on = folding_on + assumes comp_fun_idem_on: "x \ S \ y \ S \ f x \ f x = f x" begin declare insert [simp del] interpretation fold?: comp_fun_idem_on S f by standard (simp_all add: comp_fun_commute_on comp_fun_idem_on) lemma insert_idem [simp]: assumes "insert x A \ S" and "finite A" shows "F (insert x A) = f x (F A)" proof - from fold_insert_idem assms have "fold f z (insert x A) = f x (fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed end subsubsection \\<^term>\UNIV\ as the carrier set\ locale folding = fixes f :: "'a \ 'b \ 'b" and z :: "'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin lemma (in -) folding_def': "folding f = folding_on UNIV f" unfolding folding_def folding_on_def by blast text \ Again, we abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale folding_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "folding_on UNIV f" by standard (simp add: comp_fun_commute) qed simp_all end locale folding_idem = folding + assumes comp_fun_idem: "f x \ f x = f x" begin lemma (in -) folding_idem_def': "folding_idem f = folding_idem_on UNIV f" unfolding folding_idem_def folding_def' folding_idem_on_def unfolding folding_idem_axioms_def folding_idem_on_axioms_def by blast text \ Again, we abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale folding_idem_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "folding_idem_on UNIV f" by standard (simp add: comp_fun_idem) qed simp_all end subsection \Finite cardinality\ text \ The traditional definition \<^prop>\card A \ LEAST n. \f. A = {f i |i. i < n}\ is ugly to work with. But now that we have \<^const>\fold\ things are easy: \ global_interpretation card: folding "\_. Suc" 0 defines card = "folding_on.F (\_. Suc) 0" - by standard rule + by standard (rule refl) lemma card_insert_disjoint: "finite A \ x \ A \ card (insert x A) = Suc (card A)" by (fact card.insert) lemma card_insert_if: "finite A \ card (insert x A) = (if x \ A then card A else Suc (card A))" by auto (simp add: card.insert_remove card.remove) lemma card_ge_0_finite: "card A > 0 \ finite A" by (rule ccontr) simp lemma card_0_eq [simp]: "finite A \ card A = 0 \ A = {}" by (auto dest: mk_disjoint_insert) lemma finite_UNIV_card_ge_0: "finite (UNIV :: 'a set) \ card (UNIV :: 'a set) > 0" by (rule ccontr) simp lemma card_eq_0_iff: "card A = 0 \ A = {} \ \ finite A" by auto lemma card_range_greater_zero: "finite (range f) \ card (range f) > 0" by (rule ccontr) (simp add: card_eq_0_iff) lemma card_gt_0_iff: "0 < card A \ A \ {} \ finite A" by (simp add: neq0_conv [symmetric] card_eq_0_iff) lemma card_Suc_Diff1: assumes "finite A" "x \ A" shows "Suc (card (A - {x})) = card A" proof - have "Suc (card (A - {x})) = card (insert x (A - {x}))" using assms by (simp add: card.insert_remove) also have "... = card A" using assms by (simp add: card_insert_if) finally show ?thesis . qed lemma card_insert_le_m1: assumes "n > 0" "card y \ n - 1" shows "card (insert x y) \ n" using assms by (cases "finite y") (auto simp: card_insert_if) lemma card_Diff_singleton: assumes "x \ A" shows "card (A - {x}) = card A - 1" proof (cases "finite A") case True with assms show ?thesis by (simp add: card_Suc_Diff1 [symmetric]) qed auto lemma card_Diff_singleton_if: "card (A - {x}) = (if x \ A then card A - 1 else card A)" by (simp add: card_Diff_singleton) lemma card_Diff_insert[simp]: assumes "a \ A" and "a \ B" shows "card (A - insert a B) = card (A - B) - 1" proof - have "A - insert a B = (A - B) - {a}" using assms by blast then show ?thesis using assms by (simp add: card_Diff_singleton) qed lemma card_insert_le: "card A \ card (insert x A)" proof (cases "finite A") case True then show ?thesis by (simp add: card_insert_if) qed auto lemma card_Collect_less_nat[simp]: "card {i::nat. i < n} = n" by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq) lemma card_Collect_le_nat[simp]: "card {i::nat. i \ n} = Suc n" using card_Collect_less_nat[of "Suc n"] by (simp add: less_Suc_eq_le) lemma card_mono: assumes "finite B" and "A \ B" shows "card A \ card B" proof - from assms have "finite A" by (auto intro: finite_subset) then show ?thesis using assms proof (induct A arbitrary: B) case empty then show ?case by simp next case (insert x A) then have "x \ B" by simp from insert have "A \ B - {x}" and "finite (B - {x})" by auto with insert.hyps have "card A \ card (B - {x})" by auto with \finite A\ \x \ A\ \finite B\ \x \ B\ show ?case by simp (simp only: card.remove) qed qed lemma card_seteq: assumes "finite B" and A: "A \ B" "card B \ card A" shows "A = B" using assms proof (induction arbitrary: A rule: finite_induct) case (insert b B) then have A: "finite A" "A - {b} \ B" by force+ then have "card B \ card (A - {b})" using insert by (auto simp add: card_Diff_singleton_if) then have "A - {b} = B" using A insert.IH by auto then show ?case using insert.hyps insert.prems by auto qed auto lemma psubset_card_mono: "finite B \ A < B \ card A < card B" using card_seteq [of B A] by (auto simp add: psubset_eq) lemma card_Un_Int: assumes "finite A" "finite B" shows "card A + card B = card (A \ B) + card (A \ B)" using assms proof (induct A) case empty then show ?case by simp next case insert then show ?case by (auto simp add: insert_absorb Int_insert_left) qed lemma card_Un_disjoint: "finite A \ finite B \ A \ B = {} \ card (A \ B) = card A + card B" using card_Un_Int [of A B] by simp lemma card_Un_disjnt: "\finite A; finite B; disjnt A B\ \ card (A \ B) = card A + card B" by (simp add: card_Un_disjoint disjnt_def) lemma card_Un_le: "card (A \ B) \ card A + card B" proof (cases "finite A \ finite B") case True then show ?thesis using le_iff_add card_Un_Int [of A B] by auto qed auto lemma card_Diff_subset: assumes "finite B" and "B \ A" shows "card (A - B) = card A - card B" using assms proof (cases "finite A") case False with assms show ?thesis by simp next case True with assms show ?thesis by (induct B arbitrary: A) simp_all qed lemma card_Diff_subset_Int: assumes "finite (A \ B)" shows "card (A - B) = card A - card (A \ B)" proof - have "A - B = A - A \ B" by auto with assms show ?thesis by (simp add: card_Diff_subset) qed lemma diff_card_le_card_Diff: assumes "finite B" shows "card A - card B \ card (A - B)" proof - have "card A - card B \ card A - card (A \ B)" using card_mono[OF assms Int_lower2, of A] by arith also have "\ = card (A - B)" using assms by (simp add: card_Diff_subset_Int) finally show ?thesis . qed lemma card_le_sym_Diff: assumes "finite A" "finite B" "card A \ card B" shows "card(A - B) \ card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ \ card B - card (A \ B)" using assms(3) by linarith also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_less_sym_Diff: assumes "finite A" "finite B" "card A < card B" shows "card(A - B) < card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ < card B - card (A \ B)" using assms(1,3) by (simp add: card_mono diff_less_mono) also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_Diff1_less_iff: "card (A - {x}) < card A \ finite A \ x \ A" proof (cases "finite A \ x \ A") case True then show ?thesis by (auto simp: card_gt_0_iff intro: diff_less) qed auto lemma card_Diff1_less: "finite A \ x \ A \ card (A - {x}) < card A" unfolding card_Diff1_less_iff by auto lemma card_Diff2_less: assumes "finite A" "x \ A" "y \ A" shows "card (A - {x} - {y}) < card A" proof (cases "x = y") case True with assms show ?thesis by (simp add: card_Diff1_less del: card_Diff_insert) next case False then have "card (A - {x} - {y}) < card (A - {x})" "card (A - {x}) < card A" using assms by (intro card_Diff1_less; simp)+ then show ?thesis by (blast intro: less_trans) qed lemma card_Diff1_le: "card (A - {x}) \ card A" proof (cases "finite A") case True then show ?thesis by (cases "x \ A") (simp_all add: card_Diff1_less less_imp_le) qed auto lemma card_psubset: "finite B \ A \ B \ card A < card B \ A < B" by (erule psubsetI) blast lemma card_le_inj: assumes fA: "finite A" and fB: "finite B" and c: "card A \ card B" shows "\f. f ` A \ B \ inj_on f A" using fA fB c proof (induct arbitrary: B rule: finite_induct) case empty then show ?case by simp next case (insert x s t) then show ?case proof (induct rule: finite_induct [OF insert.prems(1)]) case 1 then show ?case by simp next case (2 y t) from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \ card t" by simp from "2.prems"(3) [OF "2.hyps"(1) cst] - obtain f where "f ` s \ t" "inj_on f s" + obtain f where *: "f ` s \ t" "inj_on f s" by blast - with "2.prems"(2) "2.hyps"(2) show ?case - unfolding inj_on_def - by (rule_tac x = "\z. if z = x then y else f z" in exI) auto + let ?g = "(\a. if a = x then y else f a)" + have "?g ` insert x s \ insert y t \ inj_on ?g (insert x s)" + using * "2.prems"(2) "2.hyps"(2) unfolding inj_on_def by auto + then show ?case by (rule exI[where ?x="?g"]) qed qed lemma card_subset_eq: assumes fB: "finite B" and AB: "A \ B" and c: "card A = card B" shows "A = B" proof - from fB AB have fA: "finite A" by (auto intro: finite_subset) from fA fB have fBA: "finite (B - A)" by auto have e: "A \ (B - A) = {}" by blast have eq: "A \ (B - A) = B" using AB by blast from card_Un_disjoint[OF fA fBA e, unfolded eq c] have "card (B - A) = 0" by arith then have "B - A = {}" unfolding card_eq_0_iff using fA fB by simp with AB show "A = B" by blast qed lemma insert_partition: "x \ F \ \c1 \ insert x F. \c2 \ insert x F. c1 \ c2 \ c1 \ c2 = {} \ x \ \F = {}" by auto lemma finite_psubset_induct [consumes 1, case_names psubset]: assumes finite: "finite A" and major: "\A. finite A \ (\B. B \ A \ P B) \ P A" shows "P A" using finite proof (induct A taking: card rule: measure_induct_rule) case (less A) have fin: "finite A" by fact have ih: "card B < card A \ finite B \ P B" for B by fact have "P B" if "B \ A" for B proof - from that have "card B < card A" using psubset_card_mono fin by blast moreover from that have "B \ A" by auto then have "finite B" using fin finite_subset by blast ultimately show ?thesis using ih by simp qed with fin show "P A" using major by blast qed lemma finite_induct_select [consumes 1, case_names empty select]: assumes "finite S" and "P {}" and select: "\T. T \ S \ P T \ \s\S - T. P (insert s T)" shows "P S" proof - have "0 \ card S" by simp then have "\T \ S. card T = card S \ P T" proof (induct rule: dec_induct) case base with \P {}\ show ?case by (intro exI[of _ "{}"]) auto next case (step n) then obtain T where T: "T \ S" "card T = n" "P T" by auto with \n < card S\ have "T \ S" "P T" by auto with select[of T] obtain s where "s \ S" "s \ T" "P (insert s T)" by auto with step(2) T \finite S\ show ?case by (intro exI[of _ "insert s T"]) (auto dest: finite_subset) qed with \finite S\ show "P S" by (auto dest: card_subset_eq) qed lemma remove_induct [case_names empty infinite remove]: assumes empty: "P ({} :: 'a set)" and infinite: "\ finite B \ P B" and remove: "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" shows "P B" proof (cases "finite B") case False then show ?thesis by (rule infinite) next case True define A where "A = B" with True have "finite A" "A \ B" by simp_all then show "P A" proof (induct "card A" arbitrary: A) case 0 then have "A = {}" by auto with empty show ?case by simp next case (Suc n A) from \A \ B\ and \finite B\ have "finite A" by (rule finite_subset) moreover from Suc.hyps have "A \ {}" by auto moreover note \A \ B\ moreover have "P (A - {x})" if x: "x \ A" for x using x Suc.prems \Suc n = card A\ by (intro Suc) auto ultimately show ?case by (rule remove) qed qed lemma finite_remove_induct [consumes 1, case_names empty remove]: fixes P :: "'a set \ bool" assumes "finite B" and "P {}" and "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" defines "B' \ B" shows "P B'" by (induct B' rule: remove_induct) (simp_all add: assms) text \Main cardinality theorem.\ lemma card_partition [rule_format]: "finite C \ finite (\C) \ (\c\C. card c = k) \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}) \ k * card C = card (\C)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) then show ?case by (simp add: card_Un_disjoint insert_partition finite_subset [of _ "\(insert _ _)"]) qed lemma card_eq_UNIV_imp_eq_UNIV: assumes fin: "finite (UNIV :: 'a set)" and card: "card A = card (UNIV :: 'a set)" shows "A = (UNIV :: 'a set)" proof show "A \ UNIV" by simp show "UNIV \ A" proof show "x \ A" for x proof (rule ccontr) assume "x \ A" then have "A \ UNIV" by auto with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono) with card show False by simp qed qed qed text \The form of a finite set of given cardinality\ lemma card_eq_SucD: assumes "card A = Suc k" shows "\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {})" proof - have fin: "finite A" using assms by (auto intro: ccontr) moreover have "card A \ 0" using assms by auto ultimately obtain b where b: "b \ A" by auto show ?thesis proof (intro exI conjI) show "A = insert b (A - {b})" using b by blast show "b \ A - {b}" by blast show "card (A - {b}) = k" and "k = 0 \ A - {b} = {}" using assms b fin by (fastforce dest: mk_disjoint_insert)+ qed qed lemma card_Suc_eq: "card A = Suc k \ (\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {}))" by (auto simp: card_insert_if card_gt_0_iff elim!: card_eq_SucD) lemma card_Suc_eq_finite: "card A = Suc k \ (\b B. A = insert b B \ b \ B \ card B = k \ finite B)" unfolding card_Suc_eq using card_gt_0_iff by fastforce lemma card_1_singletonE: assumes "card A = 1" obtains x where "A = {x}" using assms by (auto simp: card_Suc_eq) lemma is_singleton_altdef: "is_singleton A \ card A = 1" unfolding is_singleton_def by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def) lemma card_1_singleton_iff: "card A = Suc 0 \ (\x. A = {x})" by (simp add: card_Suc_eq) lemma card_le_Suc0_iff_eq: assumes "finite A" shows "card A \ Suc 0 \ (\a1 \ A. \a2 \ A. a1 = a2)" (is "?C = ?A") proof assume ?C thus ?A using assms by (auto simp: le_Suc_eq dest: card_eq_SucD) next assume ?A show ?C proof cases assume "A = {}" thus ?C using \?A\ by simp next assume "A \ {}" then obtain a where "A = {a}" using \?A\ by blast thus ?C by simp qed qed lemma card_le_Suc_iff: "Suc n \ card A = (\a B. A = insert a B \ a \ B \ n \ card B \ finite B)" proof (cases "finite A") case True then show ?thesis by (fastforce simp: card_Suc_eq less_eq_nat.simps split: nat.splits) qed auto lemma finite_fun_UNIVD2: assumes fin: "finite (UNIV :: ('a \ 'b) set)" shows "finite (UNIV :: 'b set)" proof - from fin have "finite (range (\f :: 'a \ 'b. f arbitrary))" for arbitrary by (rule finite_imageI) moreover have "UNIV = range (\f :: 'a \ 'b. f arbitrary)" for arbitrary by (rule UNIV_eq_I) auto ultimately show "finite (UNIV :: 'b set)" by simp qed lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1" unfolding UNIV_unit by simp lemma infinite_arbitrarily_large: assumes "\ finite A" shows "\B. finite B \ card B = n \ B \ A" proof (induction n) case 0 show ?case by (intro exI[of _ "{}"]) auto next case (Suc n) then obtain B where B: "finite B \ card B = n \ B \ A" .. with \\ finite A\ have "A \ B" by auto with B have "B \ A" by auto then have "\x. x \ A - B" by (elim psubset_imp_ex_mem) then obtain x where x: "x \ A - B" .. with B have "finite (insert x B) \ card (insert x B) = Suc n \ insert x B \ A" by auto then show "\B. finite B \ card B = Suc n \ B \ A" .. qed text \Sometimes, to prove that a set is finite, it is convenient to work with finite subsets and to show that their cardinalities are uniformly bounded. This possibility is formalized in the next criterion.\ lemma finite_if_finite_subsets_card_bdd: assumes "\G. G \ F \ finite G \ card G \ C" shows "finite F \ card F \ C" proof (cases "finite F") case False obtain n::nat where n: "n > max C 0" by auto obtain G where G: "G \ F" "card G = n" using infinite_arbitrarily_large[OF False] by auto hence "finite G" using \n > max C 0\ using card.infinite gr_implies_not0 by blast hence False using assms G n not_less by auto thus ?thesis .. next case True thus ?thesis using assms[of F] by auto qed lemma obtain_subset_with_card_n: assumes "n \ card S" obtains T where "T \ S" "card T = n" "finite T" proof - obtain n' where "card S = n + n'" using le_Suc_ex[OF assms] by blast with that show thesis proof (induct n' arbitrary: S) case 0 thus ?case by (cases "finite S") auto next case Suc thus ?case by (auto simp add: card_Suc_eq) qed qed lemma exists_subset_between: assumes "card A \ n" "n \ card C" "A \ C" "finite C" shows "\B. A \ B \ B \ C \ card B = n" using assms proof (induct n arbitrary: A C) case 0 thus ?case using finite_subset[of A C] by (intro exI[of _ "{}"], auto) next case (Suc n A C) show ?case proof (cases "A = {}") case True from obtain_subset_with_card_n[OF Suc(3)] obtain B where "B \ C" "card B = Suc n" by blast thus ?thesis unfolding True by blast next case False then obtain a where a: "a \ A" by auto let ?A = "A - {a}" let ?C = "C - {a}" have 1: "card ?A \ n" using Suc(2-) a using finite_subset by fastforce have 2: "card ?C \ n" using Suc(2-) a by auto from Suc(1)[OF 1 2 _ finite_subset[OF _ Suc(5)]] Suc(2-) obtain B where "?A \ B" "B \ ?C" "card B = n" by blast thus ?thesis using a Suc(2-) by (intro exI[of _ "insert a B"], auto intro!: card_insert_disjoint finite_subset[of B C]) qed qed subsubsection \Cardinality of image\ lemma card_image_le: "finite A \ card (f ` A) \ card A" by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if) lemma card_image: "inj_on f A \ card (f ` A) = card A" proof (induct A rule: infinite_finite_induct) case (infinite A) then have "\ finite (f ` A)" by (auto dest: finite_imageD) with infinite show ?case by simp qed simp_all lemma bij_betw_same_card: "bij_betw f A B \ card A = card B" by (auto simp: card_image bij_betw_def) lemma endo_inj_surj: "finite A \ f ` A \ A \ inj_on f A \ f ` A = A" by (simp add: card_seteq card_image) lemma eq_card_imp_inj_on: assumes "finite A" "card(f ` A) = card A" shows "inj_on f A" using assms proof (induct rule:finite_induct) case empty show ?case by simp next case (insert x A) then show ?case using card_image_le [of A f] by (simp add: card_insert_if split: if_splits) qed lemma inj_on_iff_eq_card: "finite A \ inj_on f A \ card (f ` A) = card A" by (blast intro: card_image eq_card_imp_inj_on) lemma card_inj_on_le: assumes "inj_on f A" "f ` A \ B" "finite B" shows "card A \ card B" proof - have "finite A" using assms by (blast intro: finite_imageD dest: finite_subset) then show ?thesis using assms by (force intro: card_mono simp: card_image [symmetric]) qed lemma inj_on_iff_card_le: "\ finite A; finite B \ \ (\f. inj_on f A \ f ` A \ B) = (card A \ card B)" using card_inj_on_le[of _ A B] card_le_inj[of A B] by blast lemma surj_card_le: "finite A \ B \ f ` A \ card B \ card A" by (blast intro: card_image_le card_mono le_trans) lemma card_bij_eq: "inj_on f A \ f ` A \ B \ inj_on g B \ g ` B \ A \ finite A \ finite B \ card A = card B" by (auto intro: le_antisym card_inj_on_le) lemma bij_betw_finite: "bij_betw f A B \ finite A \ finite B" unfolding bij_betw_def using finite_imageD [of f A] by auto lemma inj_on_finite: "inj_on f A \ f ` A \ B \ finite B \ finite A" using finite_imageD finite_subset by blast lemma card_vimage_inj_on_le: assumes "inj_on f D" "finite A" shows "card (f-`A \ D) \ card A" proof (rule card_inj_on_le) show "inj_on f (f -` A \ D)" by (blast intro: assms inj_on_subset) qed (use assms in auto) lemma card_vimage_inj: "inj f \ A \ range f \ card (f -` A) = card A" by (auto 4 3 simp: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on]) subsubsection \Pigeonhole Principles\ lemma pigeonhole: "card A > card (f ` A) \ \ inj_on f A " by (auto dest: card_image less_irrefl_nat) lemma pigeonhole_infinite: assumes "\ finite A" and "finite (f`A)" shows "\a0\A. \ finite {a\A. f a = f a0}" using assms(2,1) proof (induct "f`A" arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert b F) show ?case proof (cases "finite {a\A. f a = b}") case True with \\ finite A\ have "\ finite (A - {a\A. f a = b})" by simp also have "A - {a\A. f a = b} = {a\A. f a \ b}" by blast finally have "\ finite {a\A. f a \ b}" . from insert(3)[OF _ this] insert(2,4) show ?thesis by simp (blast intro: rev_finite_subset) next case False then have "{a \ A. f a = b} \ {}" by force with False show ?thesis by blast qed qed lemma pigeonhole_infinite_rel: assumes "\ finite A" and "finite B" and "\a\A. \b\B. R a b" shows "\b\B. \ finite {a:A. R a b}" proof - let ?F = "\a. {b\B. R a b}" from finite_Pow_iff[THEN iffD2, OF \finite B\] have "finite (?F ` A)" by (blast intro: rev_finite_subset) from pigeonhole_infinite [where f = ?F, OF assms(1) this] obtain a0 where "a0 \ A" and infinite: "\ finite {a\A. ?F a = ?F a0}" .. obtain b0 where "b0 \ B" and "R a0 b0" using \a0 \ A\ assms(3) by blast have "finite {a\A. ?F a = ?F a0}" if "finite {a\A. R a b0}" using \b0 \ B\ \R a0 b0\ that by (blast intro: rev_finite_subset) with infinite \b0 \ B\ show ?thesis by blast qed subsubsection \Cardinality of sums\ lemma card_Plus: assumes "finite A" "finite B" shows "card (A <+> B) = card A + card B" proof - have "Inl`A \ Inr`B = {}" by fast with assms show ?thesis by (simp add: Plus_def card_Un_disjoint card_image) qed lemma card_Plus_conv_if: "card (A <+> B) = (if finite A \ finite B then card A + card B else 0)" by (auto simp add: card_Plus) text \Relates to equivalence classes. Based on a theorem of F. Kammüller.\ lemma dvd_partition: assumes f: "finite (\C)" and "\c\C. k dvd card c" "\c1\C. \c2\C. c1 \ c2 \ c1 \ c2 = {}" shows "k dvd card (\C)" proof - have "finite C" by (rule finite_UnionD [OF f]) then show ?thesis using assms proof (induct rule: finite_induct) case empty show ?case by simp next case (insert c C) then have "c \ \C = {}" by auto with insert show ?case by (simp add: card_Un_disjoint) qed qed subsubsection \Finite orders\ context order begin lemma finite_has_maximal: "\ finite A; A \ {} \ \ \ m \ A. \ b \ A. m \ b \ m = b" proof (induction rule: finite_psubset_induct) case (psubset A) from \A \ {}\ obtain a where "a \ A" by auto let ?B = "{b \ A. a < b}" show ?case proof cases assume "?B = {}" hence "\ b \ A. a \ b \ a = b" using le_neq_trans by blast thus ?thesis using \a \ A\ by blast next assume "?B \ {}" have "a \ ?B" by auto hence "?B \ A" using \a \ A\ by blast from psubset.IH[OF this \?B \ {}\] show ?thesis using order.strict_trans2 by blast qed qed lemma finite_has_maximal2: "\ finite A; a \ A \ \ \ m \ A. a \ m \ (\ b \ A. m \ b \ m = b)" using finite_has_maximal[of "{b \ A. a \ b}"] by fastforce lemma finite_has_minimal: "\ finite A; A \ {} \ \ \ m \ A. \ b \ A. b \ m \ m = b" proof (induction rule: finite_psubset_induct) case (psubset A) from \A \ {}\ obtain a where "a \ A" by auto let ?B = "{b \ A. b < a}" show ?case proof cases assume "?B = {}" hence "\ b \ A. b \ a \ a = b" using le_neq_trans by blast thus ?thesis using \a \ A\ by blast next assume "?B \ {}" have "a \ ?B" by auto hence "?B \ A" using \a \ A\ by blast from psubset.IH[OF this \?B \ {}\] show ?thesis using order.strict_trans1 by blast qed qed lemma finite_has_minimal2: "\ finite A; a \ A \ \ \ m \ A. m \ a \ (\ b \ A. b \ m \ m = b)" using finite_has_minimal[of "{b \ A. b \ a}"] by fastforce end subsubsection \Relating injectivity and surjectivity\ lemma finite_surj_inj: assumes "finite A" "A \ f ` A" shows "inj_on f A" proof - have "f ` A = A" by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le) then show ?thesis using assms by (simp add: eq_card_imp_inj_on) qed lemma finite_UNIV_surj_inj: "finite(UNIV:: 'a set) \ surj f \ inj f" for f :: "'a \ 'a" by (blast intro: finite_surj_inj subset_UNIV) lemma finite_UNIV_inj_surj: "finite(UNIV:: 'a set) \ inj f \ surj f" for f :: "'a \ 'a" by (fastforce simp:surj_def dest!: endo_inj_surj) lemma surjective_iff_injective_gen: assumes fS: "finite S" and fT: "finite T" and c: "card S = card T" and ST: "f ` S \ T" shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" (is "?lhs \ ?rhs") proof assume h: "?lhs" { fix x y assume x: "x \ S" assume y: "y \ S" assume f: "f x = f y" from x fS have S0: "card S \ 0" by auto have "x = y" proof (rule ccontr) assume xy: "\ ?thesis" have th: "card S \ card (f ` (S - {y}))" unfolding c proof (rule card_mono) show "finite (f ` (S - {y}))" by (simp add: fS) have "\x \ y; x \ S; z \ S; f x = f y\ \ \x \ S. x \ y \ f z = f x" for z - by (case_tac "z = y \ z = x") auto + by (cases "z = y \ z = x") auto then show "T \ f ` (S - {y})" using h xy x y f by fastforce qed also have " \ \ card (S - {y})" by (simp add: card_image_le fS) also have "\ \ card S - 1" using y fS by simp finally show False using S0 by arith qed } then show ?rhs unfolding inj_on_def by blast next assume h: ?rhs have "f ` S = T" by (simp add: ST c card_image card_subset_eq fT h) then show ?lhs by blast qed hide_const (open) Finite_Set.fold subsection \Infinite Sets\ text \ Some elementary facts about infinite sets, mostly by Stephan Merz. Beware! Because "infinite" merely abbreviates a negation, these lemmas may not work well with \blast\. \ abbreviation infinite :: "'a set \ bool" where "infinite S \ \ finite S" text \ Infinite sets are non-empty, and if we remove some elements from an infinite set, the result is still infinite. \ lemma infinite_UNIV_nat [iff]: "infinite (UNIV :: nat set)" proof assume "finite (UNIV :: nat set)" with finite_UNIV_inj_surj [of Suc] show False by simp (blast dest: Suc_neq_Zero surjD) qed lemma infinite_UNIV_char_0: "infinite (UNIV :: 'a::semiring_char_0 set)" proof assume "finite (UNIV :: 'a set)" with subset_UNIV have "finite (range of_nat :: 'a set)" by (rule finite_subset) moreover have "inj (of_nat :: nat \ 'a)" by (simp add: inj_on_def) ultimately have "finite (UNIV :: nat set)" by (rule finite_imageD) then show False by simp qed lemma infinite_imp_nonempty: "infinite S \ S \ {}" by auto lemma infinite_remove: "infinite S \ infinite (S - {a})" by simp lemma Diff_infinite_finite: assumes "finite T" "infinite S" shows "infinite (S - T)" using \finite T\ proof induct from \infinite S\ show "infinite (S - {})" by auto next fix T x assume ih: "infinite (S - T)" have "S - (insert x T) = (S - T) - {x}" by (rule Diff_insert) with ih show "infinite (S - (insert x T))" by (simp add: infinite_remove) qed lemma Un_infinite: "infinite S \ infinite (S \ T)" by simp lemma infinite_Un: "infinite (S \ T) \ infinite S \ infinite T" by simp lemma infinite_super: assumes "S \ T" and "infinite S" shows "infinite T" proof assume "finite T" with \S \ T\ have "finite S" by (simp add: finite_subset) with \infinite S\ show False by simp qed proposition infinite_coinduct [consumes 1, case_names infinite]: assumes "X A" and step: "\A. X A \ \x\A. X (A - {x}) \ infinite (A - {x})" shows "infinite A" proof assume "finite A" then show False using \X A\ proof (induction rule: finite_psubset_induct) case (psubset A) then obtain x where "x \ A" "X (A - {x}) \ infinite (A - {x})" using local.step psubset.prems by blast then have "X (A - {x})" using psubset.hyps by blast show False proof (rule psubset.IH [where B = "A - {x}"]) show "A - {x} \ A" using \x \ A\ by blast qed fact qed qed text \ For any function with infinite domain and finite range there is some element that is the image of infinitely many domain elements. In particular, any infinite sequence of elements from a finite set contains some element that occurs infinitely often. \ lemma inf_img_fin_dom': assumes img: "finite (f ` A)" and dom: "infinite A" shows "\y \ f ` A. infinite (f -` {y} \ A)" proof (rule ccontr) have "A \ (\y\f ` A. f -` {y} \ A)" by auto moreover assume "\ ?thesis" with img have "finite (\y\f ` A. f -` {y} \ A)" by blast ultimately have "finite A" by (rule finite_subset) with dom show False by contradiction qed lemma inf_img_fin_domE': assumes "finite (f ` A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y} \ A)" using assms by (blast dest: inf_img_fin_dom') lemma inf_img_fin_dom: assumes img: "finite (f`A)" and dom: "infinite A" shows "\y \ f`A. infinite (f -` {y})" using inf_img_fin_dom'[OF assms] by auto lemma inf_img_fin_domE: assumes "finite (f`A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) proposition finite_image_absD: "finite (abs ` S) \ finite S" for S :: "'a::linordered_ring set" by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom) subsection \The finite powerset operator\ definition Fpow :: "'a set \ 'a set set" where "Fpow A \ {X. X \ A \ finite X}" lemma Fpow_mono: "A \ B \ Fpow A \ Fpow B" unfolding Fpow_def by auto lemma empty_in_Fpow: "{} \ Fpow A" unfolding Fpow_def by auto lemma Fpow_not_empty: "Fpow A \ {}" using empty_in_Fpow by blast lemma Fpow_subset_Pow: "Fpow A \ Pow A" unfolding Fpow_def by auto lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}" unfolding Fpow_def Pow_def by blast lemma inj_on_image_Fpow: assumes "inj_on f A" shows "inj_on (image f) (Fpow A)" using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"] inj_on_image_Pow by blast lemma image_Fpow_mono: assumes "f ` A \ B" shows "(image f) ` (Fpow A) \ Fpow B" using assms by(unfold Fpow_def, auto) end diff --git a/src/HOL/Fun.thy b/src/HOL/Fun.thy --- a/src/HOL/Fun.thy +++ b/src/HOL/Fun.thy @@ -1,1128 +1,1136 @@ (* Title: HOL/Fun.thy Author: Tobias Nipkow, Cambridge University Computer Laboratory Author: Andrei Popescu, TU Muenchen Copyright 1994, 2012 *) section \Notions about functions\ theory Fun imports Set keywords "functor" :: thy_goal_defn begin lemma apply_inverse: "f x = u \ (\x. P x \ g (f x) = x) \ P x \ x = g u" by auto text \Uniqueness, so NOT the axiom of choice.\ lemma uniq_choice: "\x. \!y. Q x y \ \f. \x. Q x (f x)" by (force intro: theI') lemma b_uniq_choice: "\x\S. \!y. Q x y \ \f. \x\S. Q x (f x)" by (force intro: theI') subsection \The Identity Function \id\\ definition id :: "'a \ 'a" where "id = (\x. x)" lemma id_apply [simp]: "id x = x" by (simp add: id_def) lemma image_id [simp]: "image id = id" by (simp add: id_def fun_eq_iff) lemma vimage_id [simp]: "vimage id = id" by (simp add: id_def fun_eq_iff) lemma eq_id_iff: "(\x. f x = x) \ f = id" by auto code_printing constant id \ (Haskell) "id" subsection \The Composition Operator \f \ g\\ definition comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "\" 55) where "f \ g = (\x. f (g x))" notation (ASCII) comp (infixl "o" 55) lemma comp_apply [simp]: "(f \ g) x = f (g x)" by (simp add: comp_def) lemma comp_assoc: "(f \ g) \ h = f \ (g \ h)" by (simp add: fun_eq_iff) lemma id_comp [simp]: "id \ g = g" by (simp add: fun_eq_iff) lemma comp_id [simp]: "f \ id = f" by (simp add: fun_eq_iff) lemma comp_eq_dest: "a \ b = c \ d \ a (b v) = c (d v)" by (simp add: fun_eq_iff) lemma comp_eq_elim: "a \ b = c \ d \ ((\v. a (b v) = c (d v)) \ R) \ R" by (simp add: fun_eq_iff) lemma comp_eq_dest_lhs: "a \ b = c \ a (b v) = c v" by clarsimp lemma comp_eq_id_dest: "a \ b = id \ c \ a (b v) = c v" by clarsimp lemma image_comp: "f ` (g ` r) = (f \ g) ` r" by auto lemma vimage_comp: "f -` (g -` x) = (g \ f) -` x" by auto lemma image_eq_imp_comp: "f ` A = g ` B \ (h \ f) ` A = (h \ g) ` B" by (auto simp: comp_def elim!: equalityE) lemma image_bind: "f ` (Set.bind A g) = Set.bind A ((`) f \ g)" by (auto simp add: Set.bind_def) lemma bind_image: "Set.bind (f ` A) g = Set.bind A (g \ f)" by (auto simp add: Set.bind_def) lemma (in group_add) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) code_printing constant comp \ (SML) infixl 5 "o" and (Haskell) infixr 9 "." subsection \The Forward Composition Operator \fcomp\\ definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) where "f \> g = (\x. g (f x))" lemma fcomp_apply [simp]: "(f \> g) x = g (f x)" by (simp add: fcomp_def) lemma fcomp_assoc: "(f \> g) \> h = f \> (g \> h)" by (simp add: fcomp_def) lemma id_fcomp [simp]: "id \> g = g" by (simp add: fcomp_def) lemma fcomp_id [simp]: "f \> id = f" by (simp add: fcomp_def) lemma fcomp_comp: "fcomp f g = comp g f" by (simp add: ext) code_printing constant fcomp \ (Eval) infixl 1 "#>" no_notation fcomp (infixl "\>" 60) subsection \Mapping functions\ definition map_fun :: "('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" where "map_fun f g h = g \ h \ f" lemma map_fun_apply [simp]: "map_fun f g h x = g (h (f x))" by (simp add: map_fun_def) subsection \Injectivity and Bijectivity\ definition inj_on :: "('a \ 'b) \ 'a set \ bool" \ \injective\ where "inj_on f A \ (\x\A. \y\A. f x = f y \ x = y)" definition bij_betw :: "('a \ 'b) \ 'a set \ 'b set \ bool" \ \bijective\ where "bij_betw f A B \ inj_on f A \ f ` A = B" text \ A common special case: functions injective, surjective or bijective over the entire domain type. \ abbreviation inj :: "('a \ 'b) \ bool" where "inj f \ inj_on f UNIV" abbreviation surj :: "('a \ 'b) \ bool" where "surj f \ range f = UNIV" translations \ \The negated case:\ "\ CONST surj f" \ "CONST range f \ CONST UNIV" abbreviation bij :: "('a \ 'b) \ bool" where "bij f \ bij_betw f UNIV UNIV" lemma inj_def: "inj f \ (\x y. f x = f y \ x = y)" unfolding inj_on_def by blast lemma injI: "(\x y. f x = f y \ x = y) \ inj f" unfolding inj_def by blast theorem range_ex1_eq: "inj f \ b \ range f \ (\!x. b = f x)" unfolding inj_def by blast lemma injD: "inj f \ f x = f y \ x = y" by (simp add: inj_def) lemma inj_on_eq_iff: "inj_on f A \ x \ A \ y \ A \ f x = f y \ x = y" by (auto simp: inj_on_def) lemma inj_on_cong: "(\a. a \ A \ f a = g a) \ inj_on f A \ inj_on g A" by (auto simp: inj_on_def) lemma inj_on_strict_subset: "inj_on f B \ A \ B \ f ` A \ f ` B" unfolding inj_on_def by blast lemma inj_compose: "inj f \ inj g \ inj (f \ g)" by (simp add: inj_def) lemma inj_fun: "inj f \ inj (\x y. f x)" by (simp add: inj_def fun_eq_iff) lemma inj_eq: "inj f \ f x = f y \ x = y" by (simp add: inj_on_eq_iff) lemma inj_on_iff_Uniq: "inj_on f A \ (\x\A. \\<^sub>\\<^sub>1y. y\A \ f x = f y)" by (auto simp: Uniq_def inj_on_def) lemma inj_on_id[simp]: "inj_on id A" by (simp add: inj_on_def) lemma inj_on_id2[simp]: "inj_on (\x. x) A" by (simp add: inj_on_def) lemma inj_on_Int: "inj_on f A \ inj_on f B \ inj_on f (A \ B)" unfolding inj_on_def by blast lemma surj_id: "surj id" by simp lemma bij_id[simp]: "bij id" by (simp add: bij_betw_def) lemma bij_uminus: "bij (uminus :: 'a \ 'a::ab_group_add)" unfolding bij_betw_def inj_on_def by (force intro: minus_minus [symmetric]) lemma bij_betwE: "bij_betw f A B \ \a\A. f a \ B" unfolding bij_betw_def by auto lemma inj_onI [intro?]: "(\x y. x \ A \ y \ A \ f x = f y \ x = y) \ inj_on f A" by (simp add: inj_on_def) lemma inj_on_inverseI: "(\x. x \ A \ g (f x) = x) \ inj_on f A" by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) lemma inj_onD: "inj_on f A \ f x = f y \ x \ A \ y \ A \ x = y" unfolding inj_on_def by blast lemma inj_on_subset: assumes "inj_on f A" and "B \ A" shows "inj_on f B" proof (rule inj_onI) fix a b assume "a \ B" and "b \ B" with assms have "a \ A" and "b \ A" by auto moreover assume "f a = f b" ultimately show "a = b" using assms by (auto dest: inj_onD) qed lemma comp_inj_on: "inj_on f A \ inj_on g (f ` A) \ inj_on (g \ f) A" by (simp add: comp_def inj_on_def) lemma inj_on_imageI: "inj_on (g \ f) A \ inj_on g (f ` A)" by (auto simp add: inj_on_def) lemma inj_on_image_iff: "\x\A. \y\A. g (f x) = g (f y) \ g x = g y \ inj_on f A \ inj_on g (f ` A) \ inj_on g A" unfolding inj_on_def by blast lemma inj_on_contraD: "inj_on f A \ x \ y \ x \ A \ y \ A \ f x \ f y" unfolding inj_on_def by blast lemma inj_singleton [simp]: "inj_on (\x. {x}) A" by (simp add: inj_on_def) lemma inj_on_empty[iff]: "inj_on f {}" by (simp add: inj_on_def) lemma subset_inj_on: "inj_on f B \ A \ B \ inj_on f A" unfolding inj_on_def by blast lemma inj_on_Un: "inj_on f (A \ B) \ inj_on f A \ inj_on f B \ f ` (A - B) \ f ` (B - A) = {}" unfolding inj_on_def by (blast intro: sym) lemma inj_on_insert [iff]: "inj_on f (insert a A) \ inj_on f A \ f a \ f ` (A - {a})" unfolding inj_on_def by (blast intro: sym) lemma inj_on_diff: "inj_on f A \ inj_on f (A - B)" unfolding inj_on_def by blast lemma comp_inj_on_iff: "inj_on f A \ inj_on f' (f ` A) \ inj_on (f' \ f) A" by (auto simp: comp_inj_on inj_on_def) lemma inj_on_imageI2: "inj_on (f' \ f) A \ inj_on f A" by (auto simp: comp_inj_on inj_on_def) lemma inj_img_insertE: assumes "inj_on f A" assumes "x \ B" and "insert x B = f ` A" obtains x' A' where "x' \ A'" and "A = insert x' A'" and "x = f x'" and "B = f ` A'" proof - from assms have "x \ f ` A" by auto then obtain x' where *: "x' \ A" "x = f x'" by auto then have A: "A = insert x' (A - {x'})" by auto with assms * have B: "B = f ` (A - {x'})" by (auto dest: inj_on_contraD) have "x' \ A - {x'}" by simp from this A \x = f x'\ B show ?thesis .. qed lemma linorder_inj_onI: fixes A :: "'a::order set" assumes ne: "\x y. \x < y; x\A; y\A\ \ f x \ f y" and lin: "\x y. \x\A; y\A\ \ x\y \ y\x" shows "inj_on f A" proof (rule inj_onI) fix x y assume eq: "f x = f y" and "x\A" "y\A" then show "x = y" using lin [of x y] ne by (force simp: dual_order.order_iff_strict) qed lemma linorder_injI: assumes "\x y::'a::linorder. x < y \ f x \ f y" shows "inj f" \ \Courtesy of Stephan Merz\ using assms by (auto intro: linorder_inj_onI linear) lemma inj_on_image_Pow: "inj_on f A \inj_on (image f) (Pow A)" unfolding Pow_def inj_on_def by blast lemma bij_betw_image_Pow: "bij_betw f A B \ bij_betw (image f) (Pow A) (Pow B)" by (auto simp add: bij_betw_def inj_on_image_Pow image_Pow_surj) lemma surj_def: "surj f \ (\y. \x. y = f x)" by auto lemma surjI: assumes "\x. g (f x) = x" shows "surj g" using assms [symmetric] by auto lemma surjD: "surj f \ \x. y = f x" by (simp add: surj_def) lemma surjE: "surj f \ (\x. y = f x \ C) \ C" by (simp add: surj_def) blast lemma comp_surj: "surj f \ surj g \ surj (g \ f)" using image_comp [of g f UNIV] by simp lemma bij_betw_imageI: "inj_on f A \ f ` A = B \ bij_betw f A B" unfolding bij_betw_def by clarify lemma bij_betw_imp_surj_on: "bij_betw f A B \ f ` A = B" unfolding bij_betw_def by clarify lemma bij_betw_imp_surj: "bij_betw f A UNIV \ surj f" unfolding bij_betw_def by auto lemma bij_betw_empty1: "bij_betw f {} A \ A = {}" unfolding bij_betw_def by blast lemma bij_betw_empty2: "bij_betw f A {} \ A = {}" unfolding bij_betw_def by blast lemma inj_on_imp_bij_betw: "inj_on f A \ bij_betw f A (f ` A)" unfolding bij_betw_def by simp lemma bij_betw_apply: "\bij_betw f A B; a \ A\ \ f a \ B" unfolding bij_betw_def by auto lemma bij_def: "bij f \ inj f \ surj f" by (rule bij_betw_def) lemma bijI: "inj f \ surj f \ bij f" by (rule bij_betw_imageI) lemma bij_is_inj: "bij f \ inj f" by (simp add: bij_def) lemma bij_is_surj: "bij f \ surj f" by (simp add: bij_def) lemma bij_betw_imp_inj_on: "bij_betw f A B \ inj_on f A" by (simp add: bij_betw_def) lemma bij_betw_trans: "bij_betw f A B \ bij_betw g B C \ bij_betw (g \ f) A C" by (auto simp add:bij_betw_def comp_inj_on) lemma bij_comp: "bij f \ bij g \ bij (g \ f)" by (rule bij_betw_trans) lemma bij_betw_comp_iff: "bij_betw f A A' \ bij_betw f' A' A'' \ bij_betw (f' \ f) A A''" by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_comp_iff2: assumes bij: "bij_betw f' A' A''" and img: "f ` A \ A'" - shows "bij_betw f A A' \ bij_betw (f' \ f) A A''" - using assms -proof (auto simp add: bij_betw_comp_iff) - assume *: "bij_betw (f' \ f) A A''" - then show "bij_betw f A A'" - using img - proof (auto simp add: bij_betw_def) - assume "inj_on (f' \ f) A" - then show "inj_on f A" - using inj_on_imageI2 by blast + shows "bij_betw f A A' \ bij_betw (f' \ f) A A''" (is "?L \ ?R") +proof + assume "?L" + then show "?R" + using assms by (auto simp add: bij_betw_comp_iff) next - fix a' - assume **: "a' \ A'" - with bij have "f' a' \ A''" - unfolding bij_betw_def by auto - with * obtain a where 1: "a \ A \ f' (f a) = f' a'" - unfolding bij_betw_def by force - with img have "f a \ A'" by auto - with bij ** 1 have "f a = a'" - unfolding bij_betw_def inj_on_def by auto - with 1 show "a' \ f ` A" by auto - qed + assume *: "?R" + have "inj_on (f' \ f) A \ inj_on f A" + using inj_on_imageI2 by blast + moreover have "A' \ f ` A" + proof + fix a' + assume **: "a' \ A'" + with bij have "f' a' \ A''" + unfolding bij_betw_def by auto + with * obtain a where 1: "a \ A \ f' (f a) = f' a'" + unfolding bij_betw_def by force + with img have "f a \ A'" by auto + with bij ** 1 have "f a = a'" + unfolding bij_betw_def inj_on_def by auto + with 1 show "a' \ f ` A" by auto + qed + ultimately show "?L" + using img * by (auto simp add: bij_betw_def) qed lemma bij_betw_inv: assumes "bij_betw f A B" shows "\g. bij_betw g B A" proof - have i: "inj_on f A" and s: "f ` A = B" using assms by (auto simp: bij_betw_def) let ?P = "\b a. a \ A \ f a = b" let ?g = "\b. The (?P b)" have g: "?g b = a" if P: "?P b a" for a b proof - from that s have ex1: "\a. ?P b a" by blast then have uex1: "\!a. ?P b a" by (blast dest:inj_onD[OF i]) then show ?thesis using the1_equality[OF uex1, OF P] P by simp qed have "inj_on ?g B" proof (rule inj_onI) fix x y assume "x \ B" "y \ B" "?g x = ?g y" from s \x \ B\ obtain a1 where a1: "?P x a1" by blast from s \y \ B\ obtain a2 where a2: "?P y a2" by blast from g [OF a1] a1 g [OF a2] a2 \?g x = ?g y\ show "x = y" by simp qed moreover have "?g ` B = A" - proof (auto simp: image_def) + proof safe fix b assume "b \ B" with s obtain a where P: "?P b a" by blast with g[OF P] show "?g b \ A" by auto next fix a assume "a \ A" with s obtain b where P: "?P b a" by blast with s have "b \ B" by blast - with g[OF P] show "\b\B. a = ?g b" by blast + with g[OF P] have "\b\B. a = ?g b" by blast + then show "a \ ?g ` B" + by auto qed ultimately show ?thesis by (auto simp: bij_betw_def) qed lemma bij_betw_cong: "(\a. a \ A \ f a = g a) \ bij_betw f A A' = bij_betw g A A'" unfolding bij_betw_def inj_on_def by safe force+ (* somewhat slow *) lemma bij_betw_id[intro, simp]: "bij_betw id A A" unfolding bij_betw_def id_def by auto lemma bij_betw_id_iff: "bij_betw id A B \ A = B" by (auto simp add: bij_betw_def) lemma bij_betw_combine: "bij_betw f A B \ bij_betw f C D \ B \ D = {} \ bij_betw f (A \ C) (B \ D)" unfolding bij_betw_def inj_on_Un image_Un by auto lemma bij_betw_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw f B B'" by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_ball: "bij_betw f A B \ (\b \ B. phi b) = (\a \ A. phi (f a))" unfolding bij_betw_def inj_on_def by blast lemma bij_pointE: assumes "bij f" obtains x where "y = f x" and "\x'. y = f x' \ x' = x" proof - from assms have "inj f" by (rule bij_is_inj) moreover from assms have "surj f" by (rule bij_is_surj) then have "y \ range f" by simp ultimately have "\!x. y = f x" by (simp add: range_ex1_eq) with that show thesis by blast qed lemma bij_iff: \<^marker>\contributor \Amine Chaieb\\ \bij f \ (\x. \!y. f y = x)\ (is \?P \ ?Q\) proof assume ?P then have \inj f\ \surj f\ by (simp_all add: bij_def) show ?Q proof fix y from \surj f\ obtain x where \y = f x\ by (auto simp add: surj_def) with \inj f\ show \\!x. f x = y\ by (auto simp add: inj_def) qed next assume ?Q then have \inj f\ by (auto simp add: inj_def) moreover have \\x. y = f x\ for y proof - from \?Q\ obtain x where \f x = y\ by blast then have \y = f x\ by simp then show ?thesis .. qed then have \surj f\ by (auto simp add: surj_def) ultimately show ?P by (rule bijI) qed lemma bij_betw_partition: \bij_betw f A B\ if \bij_betw f (A \ C) (B \ D)\ \bij_betw f C D\ \A \ C = {}\ \B \ D = {}\ proof - from that have \inj_on f (A \ C)\ \inj_on f C\ \f ` (A \ C) = B \ D\ \f ` C = D\ by (simp_all add: bij_betw_def) then have \inj_on f A\ and \f ` (A - C) \ f ` (C - A) = {}\ by (simp_all add: inj_on_Un) with \A \ C = {}\ have \f ` A \ f ` C = {}\ by auto with \f ` (A \ C) = B \ D\ \f ` C = D\ \B \ D = {}\ have \f ` A = B\ by blast with \inj_on f A\ show ?thesis by (simp add: bij_betw_def) qed lemma surj_image_vimage_eq: "surj f \ f ` (f -` A) = A" by simp lemma surj_vimage_empty: assumes "surj f" shows "f -` A = {} \ A = {}" using surj_image_vimage_eq [OF \surj f\, of A] by (intro iffI) fastforce+ lemma inj_vimage_image_eq: "inj f \ f -` (f ` A) = A" unfolding inj_def by blast lemma vimage_subsetD: "surj f \ f -` B \ A \ B \ f ` A" by (blast intro: sym) lemma vimage_subsetI: "inj f \ B \ f ` A \ f -` B \ A" unfolding inj_def by blast lemma vimage_subset_eq: "bij f \ f -` B \ A \ B \ f ` A" unfolding bij_def by (blast del: subsetI intro: vimage_subsetI vimage_subsetD) lemma inj_on_image_eq_iff: "inj_on f C \ A \ C \ B \ C \ f ` A = f ` B \ A = B" by (fastforce simp: inj_on_def) lemma inj_on_Un_image_eq_iff: "inj_on f (A \ B) \ f ` A = f ` B \ A = B" by (erule inj_on_image_eq_iff) simp_all lemma inj_on_image_Int: "inj_on f C \ A \ C \ B \ C \ f ` (A \ B) = f ` A \ f ` B" unfolding inj_on_def by blast lemma inj_on_image_set_diff: "inj_on f C \ A - B \ C \ B \ C \ f ` (A - B) = f ` A - f ` B" unfolding inj_on_def by blast lemma image_Int: "inj f \ f ` (A \ B) = f ` A \ f ` B" unfolding inj_def by blast lemma image_set_diff: "inj f \ f ` (A - B) = f ` A - f ` B" unfolding inj_def by blast lemma inj_on_image_mem_iff: "inj_on f B \ a \ B \ A \ B \ f a \ f ` A \ a \ A" by (auto simp: inj_on_def) lemma inj_image_mem_iff: "inj f \ f a \ f ` A \ a \ A" by (blast dest: injD) lemma inj_image_subset_iff: "inj f \ f ` A \ f ` B \ A \ B" by (blast dest: injD) lemma inj_image_eq_iff: "inj f \ f ` A = f ` B \ A = B" by (blast dest: injD) lemma surj_Compl_image_subset: "surj f \ - (f ` A) \ f ` (- A)" by auto lemma inj_image_Compl_subset: "inj f \ f ` (- A) \ - (f ` A)" by (auto simp: inj_def) lemma bij_image_Compl_eq: "bij f \ f ` (- A) = - (f ` A)" by (simp add: bij_def inj_image_Compl_subset surj_Compl_image_subset equalityI) lemma inj_vimage_singleton: "inj f \ f -` {a} \ {THE x. f x = a}" \ \The inverse image of a singleton under an injective function is included in a singleton.\ by (simp add: inj_def) (blast intro: the_equality [symmetric]) lemma inj_on_vimage_singleton: "inj_on f A \ f -` {a} \ A \ {THE x. x \ A \ f x = a}" by (auto simp add: inj_on_def intro: the_equality [symmetric]) lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A" by (auto intro!: inj_onI) lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \ inj_on f A" by (auto intro!: inj_onI dest: strict_mono_eq) lemma bij_betw_byWitness: assumes left: "\a \ A. f' (f a) = a" and right: "\a' \ A'. f (f' a') = a'" and "f ` A \ A'" and img2: "f' ` A' \ A" shows "bij_betw f A A'" using assms unfolding bij_betw_def inj_on_def proof safe fix a b assume "a \ A" "b \ A" with left have "a = f' (f a) \ b = f' (f b)" by simp moreover assume "f a = f b" ultimately show "a = b" by simp next fix a' assume *: "a' \ A'" with img2 have "f' a' \ A" by blast moreover from * right have "a' = f (f' a')" by simp ultimately show "a' \ f ` A" by blast qed corollary notIn_Un_bij_betw: assumes "b \ A" and "f b \ A'" and "bij_betw f A A'" shows "bij_betw f (A \ {b}) (A' \ {f b})" proof - have "bij_betw f {b} {f b}" unfolding bij_betw_def inj_on_def by simp with assms show ?thesis using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast qed lemma notIn_Un_bij_betw3: assumes "b \ A" and "f b \ A'" shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" proof assume "bij_betw f A A'" then show "bij_betw f (A \ {b}) (A' \ {f b})" using assms notIn_Un_bij_betw [of b A f A'] by blast next assume *: "bij_betw f (A \ {b}) (A' \ {f b})" have "f ` A = A'" - proof auto + proof safe fix a assume **: "a \ A" then have "f a \ A' \ {f b}" using * unfolding bij_betw_def by blast moreover have False if "f a = f b" proof - have "a = b" using * ** that unfolding bij_betw_def inj_on_def by blast with \b \ A\ ** show ?thesis by blast qed ultimately show "f a \ A'" by blast next fix a' assume **: "a' \ A'" then have "a' \ f ` (A \ {b})" using * by (auto simp add: bij_betw_def) then obtain a where 1: "a \ A \ {b} \ f a = a'" by blast moreover have False if "a = b" using 1 ** \f b \ A'\ that by blast ultimately have "a \ A" by blast with 1 show "a' \ f ` A" by blast qed then show "bij_betw f A A'" using * bij_betw_subset[of f "A \ {b}" _ A] by blast qed lemma inj_on_disjoint_Un: assumes "inj_on f A" and "inj_on g B" and "f ` A \ g ` B = {}" shows "inj_on (\x. if x \ A then f x else g x) (A \ B)" using assms by (simp add: inj_on_def disjoint_iff) (blast) lemma bij_betw_disjoint_Un: assumes "bij_betw f A C" and "bij_betw g B D" and "A \ B = {}" and "C \ D = {}" shows "bij_betw (\x. if x \ A then f x else g x) (A \ B) (C \ D)" using assms by (auto simp: inj_on_disjoint_Un bij_betw_def) lemma involuntory_imp_bij: \bij f\ if \\x. f (f x) = x\ proof (rule bijI) from that show \surj f\ by (rule surjI) show \inj f\ proof (rule injI) fix x y assume \f x = f y\ then have \f (f x) = f (f y)\ by simp then show \x = y\ by (simp add: that) qed qed subsubsection \Important examples\ context cancel_semigroup_add begin lemma inj_on_add [simp]: "inj_on ((+) a) A" by (rule inj_onI) simp lemma inj_add_left: \inj ((+) a)\ by simp lemma inj_on_add' [simp]: "inj_on (\b. b + a) A" by (rule inj_onI) simp lemma bij_betw_add [simp]: "bij_betw ((+) a) A B \ (+) a ` A = B" by (simp add: bij_betw_def) end context ab_group_add begin lemma surj_plus [simp]: "surj ((+) a)" by (auto intro!: range_eqI [of b "(+) a" "b - a" for b]) (simp add: algebra_simps) lemma inj_diff_right [simp]: \inj (\b. b - a)\ proof - have \inj ((+) (- a))\ by (fact inj_add_left) also have \(+) (- a) = (\b. b - a)\ by (simp add: fun_eq_iff) finally show ?thesis . qed lemma surj_diff_right [simp]: "surj (\x. x - a)" using surj_plus [of "- a"] by (simp cong: image_cong_simp) lemma translation_Compl: "(+) a ` (- t) = - ((+) a ` t)" proof (rule set_eqI) fix b show "b \ (+) a ` (- t) \ b \ - (+) a ` t" by (auto simp: image_iff algebra_simps intro!: bexI [of _ "b - a"]) qed lemma translation_subtract_Compl: "(\x. x - a) ` (- t) = - ((\x. x - a) ` t)" using translation_Compl [of "- a" t] by (simp cong: image_cong_simp) lemma translation_diff: "(+) a ` (s - t) = ((+) a ` s) - ((+) a ` t)" by auto lemma translation_subtract_diff: "(\x. x - a) ` (s - t) = ((\x. x - a) ` s) - ((\x. x - a) ` t)" using translation_diff [of "- a"] by (simp cong: image_cong_simp) lemma translation_Int: "(+) a ` (s \ t) = ((+) a ` s) \ ((+) a ` t)" by auto lemma translation_subtract_Int: "(\x. x - a) ` (s \ t) = ((\x. x - a) ` s) \ ((\x. x - a) ` t)" using translation_Int [of " -a"] by (simp cong: image_cong_simp) end subsection \Function Updating\ definition fun_upd :: "('a \ 'b) \ 'a \ 'b \ ('a \ 'b)" where "fun_upd f a b = (\x. if x = a then b else f x)" nonterminal updbinds and updbind syntax "_updbind" :: "'a \ 'a \ updbind" ("(2_ :=/ _)") "" :: "updbind \ updbinds" ("_") "_updbinds":: "updbind \ updbinds \ updbinds" ("_,/ _") "_Update" :: "'a \ updbinds \ 'a" ("_/'((_)')" [1000, 0] 900) translations "_Update f (_updbinds b bs)" \ "_Update (_Update f b) bs" "f(x:=y)" \ "CONST fun_upd f x y" (* Hint: to define the sum of two functions (or maps), use case_sum. A nice infix syntax could be defined by notation case_sum (infixr "'(+')"80) *) lemma fun_upd_idem_iff: "f(x:=y) = f \ f x = y" unfolding fun_upd_def apply safe apply (erule subst) - apply (rule_tac [2] ext) apply auto done lemma fun_upd_idem: "f x = y \ f(x := y) = f" by (simp only: fun_upd_idem_iff) lemma fun_upd_triv [iff]: "f(x := f x) = f" by (simp only: fun_upd_idem) lemma fun_upd_apply [simp]: "(f(x := y)) z = (if z = x then y else f z)" by (simp add: fun_upd_def) (* fun_upd_apply supersedes these two, but they are useful if fun_upd_apply is intentionally removed from the simpset *) lemma fun_upd_same: "(f(x := y)) x = y" by simp lemma fun_upd_other: "z \ x \ (f(x := y)) z = f z" by simp lemma fun_upd_upd [simp]: "f(x := y, x := z) = f(x := z)" by (simp add: fun_eq_iff) lemma fun_upd_twist: "a \ c \ (m(a := b))(c := d) = (m(c := d))(a := b)" by auto lemma inj_on_fun_updI: "inj_on f A \ y \ f ` A \ inj_on (f(x := y)) A" by (auto simp: inj_on_def) lemma fun_upd_image: "f(x := y) ` A = (if x \ A then insert y (f ` (A - {x})) else f ` A)" by auto lemma fun_upd_comp: "f \ (g(x := y)) = (f \ g)(x := f y)" by auto lemma fun_upd_eqD: "f(x := y) = g(x := z) \ y = z" by (simp add: fun_eq_iff split: if_split_asm) subsection \\override_on\\ definition override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" where "override_on f g A = (\a. if a \ A then g a else f a)" lemma override_on_emptyset[simp]: "override_on f g {} = f" by (simp add: override_on_def) lemma override_on_apply_notin[simp]: "a \ A \ (override_on f g A) a = f a" by (simp add: override_on_def) lemma override_on_apply_in[simp]: "a \ A \ (override_on f g A) a = g a" by (simp add: override_on_def) lemma override_on_insert: "override_on f g (insert x X) = (override_on f g X)(x:=g x)" by (simp add: override_on_def fun_eq_iff) lemma override_on_insert': "override_on f g (insert x X) = (override_on (f(x:=g x)) g X)" by (simp add: override_on_def fun_eq_iff) subsection \Inversion of injective functions\ definition the_inv_into :: "'a set \ ('a \ 'b) \ ('b \ 'a)" where "the_inv_into A f = (\x. THE y. y \ A \ f y = x)" lemma the_inv_into_f_f: "inj_on f A \ x \ A \ the_inv_into A f (f x) = x" unfolding the_inv_into_def inj_on_def by blast lemma f_the_inv_into_f: "inj_on f A \ y \ f ` A \ f (the_inv_into A f y) = y" unfolding the_inv_into_def by (rule the1I2; blast dest: inj_onD) lemma f_the_inv_into_f_bij_betw: "bij_betw f A B \ (bij_betw f A B \ x \ B) \ f (the_inv_into A f x) = x" unfolding bij_betw_def by (blast intro: f_the_inv_into_f) lemma the_inv_into_into: "inj_on f A \ x \ f ` A \ A \ B \ the_inv_into A f x \ B" unfolding the_inv_into_def by (rule the1I2; blast dest: inj_onD) lemma the_inv_into_onto [simp]: "inj_on f A \ the_inv_into A f ` (f ` A) = A" by (fast intro: the_inv_into_into the_inv_into_f_f [symmetric]) lemma the_inv_into_f_eq: "inj_on f A \ f x = y \ x \ A \ the_inv_into A f y = x" by (force simp add: the_inv_into_f_f) lemma the_inv_into_comp: "inj_on f (g ` A) \ inj_on g A \ x \ f ` g ` A \ the_inv_into A (f \ g) x = (the_inv_into A g \ the_inv_into (g ` A) f) x" apply (rule the_inv_into_f_eq) apply (fast intro: comp_inj_on) apply (simp add: f_the_inv_into_f the_inv_into_into) apply (simp add: the_inv_into_into) done lemma inj_on_the_inv_into: "inj_on f A \ inj_on (the_inv_into A f) (f ` A)" by (auto intro: inj_onI simp: the_inv_into_f_f) lemma bij_betw_the_inv_into: "bij_betw f A B \ bij_betw (the_inv_into A f) B A" by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) lemma bij_betw_iff_bijections: "bij_betw f A B \ (\g. (\x \ A. f x \ B \ g(f x) = x) \ (\y \ B. g y \ A \ f(g y) = y))" (is "?lhs = ?rhs") proof - assume L: ?lhs - then show ?rhs - apply (rule_tac x="the_inv_into A f" in exI) - apply (auto simp: bij_betw_def f_the_inv_into_f the_inv_into_f_f the_inv_into_into) - done -qed (force intro: bij_betw_byWitness) + show "?lhs \ ?rhs" + by (auto simp: bij_betw_def f_the_inv_into_f the_inv_into_f_f the_inv_into_into + exI[where ?x="the_inv_into A f"]) +next + show "?rhs \ ?lhs" + by (force intro: bij_betw_byWitness) +qed abbreviation the_inv :: "('a \ 'b) \ ('b \ 'a)" where "the_inv f \ the_inv_into UNIV f" lemma the_inv_f_f: "the_inv f (f x) = x" if "inj f" using that UNIV_I by (rule the_inv_into_f_f) subsection \Cantor's Paradox\ theorem Cantors_paradox: "\f. f ` A = Pow A" proof assume "\f. f ` A = Pow A" then obtain f where f: "f ` A = Pow A" .. let ?X = "{a \ A. a \ f a}" have "?X \ Pow A" by blast then have "?X \ f ` A" by (simp only: f) then obtain x where "x \ A" and "f x = ?X" by blast then show False by blast qed subsection \Monotonic functions over a set\ definition monotone_on :: "'a set \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" where "monotone_on A orda ordb f \ (\x\A. \y\A. orda x y \ ordb (f x) (f y))" abbreviation monotone :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" where "monotone \ monotone_on UNIV" lemma monotone_def[no_atp]: "monotone orda ordb f \ (\x y. orda x y \ ordb (f x) (f y))" by (simp add: monotone_on_def) text \Lemma @{thm [source] monotone_def} is provided for backward compatibility.\ lemma monotone_onI: "(\x y. x \ A \ y \ A \ orda x y \ ordb (f x) (f y)) \ monotone_on A orda ordb f" by (simp add: monotone_on_def) lemma monotoneI[intro?]: "(\x y. orda x y \ ordb (f x) (f y)) \ monotone orda ordb f" by (rule monotone_onI) lemma monotone_onD: "monotone_on A orda ordb f \ x \ A \ y \ A \ orda x y \ ordb (f x) (f y)" by (simp add: monotone_on_def) lemma monotoneD[dest?]: "monotone orda ordb f \ orda x y \ ordb (f x) (f y)" by (rule monotone_onD[of UNIV, simplified]) lemma monotone_on_subset: "monotone_on A orda ordb f \ B \ A \ monotone_on B orda ordb f" by (auto intro: monotone_onI dest: monotone_onD) lemma monotone_on_empty[simp]: "monotone_on {} orda ordb f" by (auto intro: monotone_onI dest: monotone_onD) lemma monotone_on_o: assumes mono_f: "monotone_on A orda ordb f" and mono_g: "monotone_on B ordc orda g" and "g ` B \ A" shows "monotone_on B ordc ordb (f \ g)" proof (rule monotone_onI) fix x y assume "x \ B" and "y \ B" and "ordc x y" hence "orda (g x) (g y)" by (rule mono_g[THEN monotone_onD]) moreover from \g ` B \ A\ \x \ B\ \y \ B\ have "g x \ A" and "g y \ A" unfolding image_subset_iff by simp_all ultimately show "ordb ((f \ g) x) ((f \ g) y)" using mono_f[THEN monotone_onD] by simp qed abbreviation mono_on :: "('a :: ord) set \ ('a \ 'b :: ord) \ bool" where "mono_on A \ monotone_on A (\) (\)" lemma mono_on_def: "mono_on A f \ (\r s. r \ A \ s \ A \ r \ s \ f r \ f s)" by (auto simp add: monotone_on_def) lemma mono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on A f" by (rule monotone_onI) lemma mono_onD: "\mono_on A f; r \ A; s \ A; r \ s\ \ f r \ f s" by (rule monotone_onD) lemma mono_imp_mono_on: "mono f \ mono_on A f" unfolding mono_def mono_on_def by auto lemma mono_on_subset: "mono_on A f \ B \ A \ mono_on B f" by (rule monotone_on_subset) abbreviation strict_mono_on :: "('a :: ord) set \ ('a \ 'b :: ord) \ bool" where "strict_mono_on A \ monotone_on A (<) (<)" lemma strict_mono_on_def: "strict_mono_on A f \ (\r s. r \ A \ s \ A \ r < s \ f r < f s)" by (auto simp add: monotone_on_def) lemma strict_mono_onI: "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on A f" by (rule monotone_onI) lemma strict_mono_onD: "\strict_mono_on A f; r \ A; s \ A; r < s\ \ f r < f s" by (rule monotone_onD) lemma mono_on_greaterD: assumes "mono_on A g" "x \ A" "y \ A" "g x > (g (y::_::linorder) :: _ :: linorder)" shows "x > y" proof (rule ccontr) assume "\x > y" hence "x \ y" by (simp add: not_less) from assms(1-3) and this have "g x \ g y" by (rule mono_onD) with assms(4) show False by simp qed lemma strict_mono_inv: fixes f :: "('a::linorder) \ ('b::linorder)" assumes "strict_mono f" and "surj f" and inv: "\x. g (f x) = x" shows "strict_mono g" proof fix x y :: 'b assume "x < y" from \surj f\ obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast with \x < y\ and \strict_mono f\ have "x' < y'" by (simp add: strict_mono_less) with inv show "g x < g y" by simp qed lemma strict_mono_on_imp_inj_on: assumes "strict_mono_on A (f :: (_ :: linorder) \ (_ :: preorder))" shows "inj_on f A" proof (rule inj_onI) fix x y assume "x \ A" "y \ A" "f x = f y" thus "x = y" by (cases x y rule: linorder_cases) (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) qed lemma strict_mono_on_leD: assumes "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder)" "x \ A" "y \ A" "x \ y" shows "f x \ f y" -proof (insert le_less_linear[of y x], elim disjE) - assume "x < y" - with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all - thus ?thesis by (rule less_imp_le) -qed (insert assms, simp) +proof (cases "x = y") + case True + then show ?thesis by simp +next + case False + with assms have "f x < f y" + using strict_mono_onD[OF assms(1)] by simp + then show ?thesis by (rule less_imp_le) +qed lemma strict_mono_on_eqD: fixes f :: "(_ :: linorder) \ (_ :: preorder)" assumes "strict_mono_on A f" "f x = f y" "x \ A" "y \ A" shows "y = x" - using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD) + using assms by (cases rule: linorder_cases) (auto dest: strict_mono_onD) lemma strict_mono_on_imp_mono_on: "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder) \ mono_on A f" by (rule mono_onI, rule strict_mono_on_leD) subsection \Setup\ subsubsection \Proof tools\ text \Simplify terms of the form \f(\,x:=y,\,x:=z,\)\ to \f(\,x:=z,\)\\ simproc_setup fun_upd2 ("f(v := w, x := y)") = \fn _ => let fun gen_fun_upd NONE T _ _ = NONE | gen_fun_upd (SOME f) T x y = SOME (Const (\<^const_name>\fun_upd\, T) $ f $ x $ y) fun dest_fun_T1 (Type (_, T :: Ts)) = T fun find_double (t as Const (\<^const_name>\fun_upd\,T) $ f $ x $ y) = let fun find (Const (\<^const_name>\fun_upd\,T) $ g $ v $ w) = if v aconv x then SOME g else gen_fun_upd (find g) T v w | find t = NONE in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end val ss = simpset_of \<^context> fun proc ctxt ct = let val t = Thm.term_of ct in (case find_double t of (T, NONE) => NONE | (T, SOME rhs) => SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) (fn _ => resolve_tac ctxt [eq_reflection] 1 THEN resolve_tac ctxt @{thms ext} 1 THEN simp_tac (put_simpset ss ctxt) 1))) end in proc end \ subsubsection \Functorial structure of types\ ML_file \Tools/functor.ML\ functor map_fun: map_fun by (simp_all add: fun_eq_iff) functor vimage by (simp_all add: fun_eq_iff vimage_comp) text \Legacy theorem names\ lemmas o_def = comp_def lemmas o_apply = comp_apply lemmas o_assoc = comp_assoc [symmetric] lemmas id_o = id_comp lemmas o_id = comp_id lemmas o_eq_dest = comp_eq_dest lemmas o_eq_elim = comp_eq_elim lemmas o_eq_dest_lhs = comp_eq_dest_lhs lemmas o_eq_id_dest = comp_eq_id_dest end diff --git a/src/HOL/Fun_Def.thy b/src/HOL/Fun_Def.thy --- a/src/HOL/Fun_Def.thy +++ b/src/HOL/Fun_Def.thy @@ -1,310 +1,330 @@ (* Title: HOL/Fun_Def.thy Author: Alexander Krauss, TU Muenchen *) section \Function Definitions and Termination Proofs\ theory Fun_Def imports Basic_BNF_LFPs Partial_Function SAT keywords "function" "termination" :: thy_goal_defn and "fun" "fun_cases" :: thy_defn begin subsection \Definitions with default value\ definition THE_default :: "'a \ ('a \ bool) \ 'a" where "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" lemma THE_defaultI': "\!x. P x \ P (THE_default d P)" by (simp add: theI' THE_default_def) lemma THE_default1_equality: "\!x. P x \ P a \ THE_default d P = a" by (simp add: the1_equality THE_default_def) lemma THE_default_none: "\ (\!x. P x) \ THE_default d P = d" by (simp add: THE_default_def) lemma fundef_ex1_existence: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "G x (f x)" apply (simp only: f_def) apply (rule THE_defaultI') apply (rule ex1) done lemma fundef_ex1_uniqueness: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" assumes elm: "G x (h x)" shows "h x = f x" - apply (simp only: f_def) - apply (rule THE_default1_equality [symmetric]) - apply (rule ex1) - apply (rule elm) - done + by (auto simp add: f_def ex1 elm THE_default1_equality[symmetric]) lemma fundef_ex1_iff: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "(G x y) = (f x = y)" - apply (auto simp:ex1 f_def THE_default1_equality) - apply (rule THE_defaultI') - apply (rule ex1) - done + by (auto simp add: ex1 f_def THE_default1_equality THE_defaultI') + lemma fundef_default_value: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes graph: "\x y. G x y \ D x" assumes "\ D x" shows "f x = d x" proof - have "\(\y. G x y)" proof assume "\y. G x y" then have "D x" using graph .. with \\ D x\ show False .. qed then have "\(\!y. G x y)" by blast then show ?thesis unfolding f_def by (rule THE_default_none) qed definition in_rel_def[simp]: "in_rel R x y \ (x, y) \ R" lemma wf_in_rel: "wf R \ wfP (in_rel R)" by (simp add: wfP_def) ML_file \Tools/Function/function_core.ML\ ML_file \Tools/Function/mutual.ML\ ML_file \Tools/Function/pattern_split.ML\ ML_file \Tools/Function/relation.ML\ ML_file \Tools/Function/function_elims.ML\ method_setup relation = \ Args.term >> (fn t => fn ctxt => SIMPLE_METHOD' (Function_Relation.relation_infer_tac ctxt t)) \ "prove termination using a user-specified wellfounded relation" ML_file \Tools/Function/function.ML\ ML_file \Tools/Function/pat_completeness.ML\ method_setup pat_completeness = \ Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac) \ "prove completeness of (co)datatype patterns" ML_file \Tools/Function/fun.ML\ ML_file \Tools/Function/induction_schema.ML\ method_setup induction_schema = \ Scan.succeed (CONTEXT_TACTIC oo Induction_Schema.induction_schema_tac) \ "prove an induction principle" subsection \Measure functions\ inductive is_measure :: "('a \ nat) \ bool" where is_measure_trivial: "is_measure f" named_theorems measure_function "rules that guide the heuristic generation of measure functions" ML_file \Tools/Function/measure_functions.ML\ lemma measure_size[measure_function]: "is_measure size" by (rule is_measure_trivial) lemma measure_fst[measure_function]: "is_measure f \ is_measure (\p. f (fst p))" by (rule is_measure_trivial) lemma measure_snd[measure_function]: "is_measure f \ is_measure (\p. f (snd p))" by (rule is_measure_trivial) ML_file \Tools/Function/lexicographic_order.ML\ method_setup lexicographic_order = \ Method.sections clasimp_modifiers >> (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false)) \ "termination prover for lexicographic orderings" subsection \Congruence rules\ lemma let_cong [fundef_cong]: "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" unfolding Let_def by blast lemmas [fundef_cong] = if_cong image_cong bex_cong ball_cong imp_cong map_option_cong Option.bind_cong lemma split_cong [fundef_cong]: "(\x y. (x, y) = q \ f x y = g x y) \ p = q \ case_prod f p = case_prod g q" by (auto simp: split_def) lemma comp_cong [fundef_cong]: "f (g x) = f' (g' x') \ (f \ g) x = (f' \ g') x'" by (simp only: o_apply) subsection \Simp rules for termination proofs\ declare trans_less_add1[termination_simp] trans_less_add2[termination_simp] trans_le_add1[termination_simp] trans_le_add2[termination_simp] less_imp_le_nat[termination_simp] le_imp_less_Suc[termination_simp] lemma size_prod_simp[termination_simp]: "size_prod f g p = f (fst p) + g (snd p) + Suc 0" by (induct p) auto subsection \Decomposition\ lemma less_by_empty: "A = {} \ A \ B" and union_comp_emptyL: "A O C = {} \ B O C = {} \ (A \ B) O C = {}" and union_comp_emptyR: "A O B = {} \ A O C = {} \ A O (B \ C) = {}" and wf_no_loop: "R O R = {} \ wf R" by (auto simp add: wf_comp_self [of R]) subsection \Reduction pairs\ definition "reduction_pair P \ wf (fst P) \ fst P O snd P \ fst P" lemma reduction_pairI[intro]: "wf R \ R O S \ R \ reduction_pair (R, S)" by (auto simp: reduction_pair_def) lemma reduction_pair_lemma: assumes rp: "reduction_pair P" assumes "R \ fst P" assumes "S \ snd P" assumes "wf S" shows "wf (R \ S)" proof - from rp \S \ snd P\ have "wf (fst P)" "fst P O S \ fst P" unfolding reduction_pair_def by auto with \wf S\ have "wf (fst P \ S)" by (auto intro: wf_union_compatible) moreover from \R \ fst P\ have "R \ S \ fst P \ S" by auto ultimately show ?thesis by (rule wf_subset) qed definition "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" lemma rp_inv_image_rp: "reduction_pair P \ reduction_pair (rp_inv_image P f)" unfolding reduction_pair_def rp_inv_image_def split_def by force subsection \Concrete orders for SCNP termination proofs\ definition "pair_less = less_than <*lex*> less_than" definition "pair_leq = pair_less\<^sup>=" definition "max_strict = max_ext pair_less" definition "max_weak = max_ext pair_leq \ {({}, {})}" definition "min_strict = min_ext pair_less" definition "min_weak = min_ext pair_leq \ {({}, {})}" lemma wf_pair_less[simp]: "wf pair_less" by (auto simp: pair_less_def) lemma total_pair_less [iff]: "total_on A pair_less" and trans_pair_less [iff]: "trans pair_less" by (auto simp: total_on_def pair_less_def) text \Introduction rules for \pair_less\/\pair_leq\\ lemma pair_leqI1: "a < b \ ((a, s), (b, t)) \ pair_leq" and pair_leqI2: "a \ b \ s \ t \ ((a, s), (b, t)) \ pair_leq" and pair_lessI1: "a < b \ ((a, s), (b, t)) \ pair_less" and pair_lessI2: "a \ b \ s < t \ ((a, s), (b, t)) \ pair_less" by (auto simp: pair_leq_def pair_less_def) lemma pair_less_iff1 [simp]: "((x,y), (x,z)) \ pair_less \ yIntroduction rules for max\ lemma smax_emptyI: "finite Y \ Y \ {} \ ({}, Y) \ max_strict" and smax_insertI: "y \ Y \ (x, y) \ pair_less \ (X, Y) \ max_strict \ (insert x X, Y) \ max_strict" and wmax_emptyI: "finite X \ ({}, X) \ max_weak" and wmax_insertI: "y \ YS \ (x, y) \ pair_leq \ (XS, YS) \ max_weak \ (insert x XS, YS) \ max_weak" by (auto simp: max_strict_def max_weak_def elim!: max_ext.cases) text \Introduction rules for min\ lemma smin_emptyI: "X \ {} \ (X, {}) \ min_strict" and smin_insertI: "x \ XS \ (x, y) \ pair_less \ (XS, YS) \ min_strict \ (XS, insert y YS) \ min_strict" and wmin_emptyI: "(X, {}) \ min_weak" and wmin_insertI: "x \ XS \ (x, y) \ pair_leq \ (XS, YS) \ min_weak \ (XS, insert y YS) \ min_weak" by (auto simp: min_strict_def min_weak_def min_ext_def) text \Reduction Pairs.\ lemma max_ext_compat: assumes "R O S \ R" shows "max_ext R O (max_ext S \ {({}, {})}) \ max_ext R" - using assms - apply auto - apply (elim max_ext.cases) - apply rule - apply auto[3] - apply (drule_tac x=xa in meta_spec) - apply simp - apply (erule bexE) - apply (drule_tac x=xb in meta_spec) - apply auto - done +proof - + have "\X Y Z. (X, Y) \ max_ext R \ (Y, Z) \ max_ext S \ (X, Z) \ max_ext R" + proof - + fix X Y Z + assume "(X,Y)\max_ext R" + "(Y, Z)\max_ext S" + then have *: "finite X" "finite Y" "finite Z" "Y\{}" "Z\{}" + "(\x. x\X \ \y\Y. (x, y)\R)" + "(\y. y\Y \ \z\Z. (y, z)\S)" + by (auto elim: max_ext.cases) + moreover have "\x. x\X \ \z\Z. (x, z)\R" + proof - + fix x + assume "x\X" + then obtain y where 1: "y\Y" "(x, y)\R" + using * by auto + then obtain z where "z\Z" "(y, z)\S" + using * by auto + then show "\z\Z. (x, z)\R" + using assms 1 by (auto elim: max_ext.cases) + qed + ultimately show "(X,Z)\max_ext R" + by auto + qed + then show "max_ext R O (max_ext S \ {({}, {})}) \ max_ext R" + by auto +qed lemma max_rpair_set: "reduction_pair (max_strict, max_weak)" unfolding max_strict_def max_weak_def apply (intro reduction_pairI max_ext_wf) apply simp apply (rule max_ext_compat) apply (auto simp: pair_less_def pair_leq_def) done lemma min_ext_compat: assumes "R O S \ R" - shows "min_ext R O (min_ext S \ {({},{})}) \ min_ext R" - using assms - apply (auto simp: min_ext_def) - apply (drule_tac x=ya in bspec, assumption) - apply (erule bexE) - apply (drule_tac x=xc in bspec) - apply assumption - apply auto - done + shows "min_ext R O (min_ext S \ {({},{})}) \ min_ext R" +proof - + have "\X Y Z z. \y\Y. \x\X. (x, y) \ R \ \z\Z. \y\Y. (y, z) \ S + \ z \ Z \ \x\X. (x, z) \ R" + proof - + fix X Y Z z + assume *: "\y\Y. \x\X. (x, y) \ R" + "\z\Z. \y\Y. (y, z) \ S" + "z\Z" + then obtain y' where 1: "y'\Y" "(y', z) \ S" + by auto + then obtain x' where 2: "x'\X" "(x', y') \ R" + using * by auto + show "\x\X. (x, z) \ R" + using 1 2 assms by auto + qed + then show ?thesis + using assms by (auto simp: min_ext_def) +qed lemma min_rpair_set: "reduction_pair (min_strict, min_weak)" unfolding min_strict_def min_weak_def apply (intro reduction_pairI min_ext_wf) apply simp apply (rule min_ext_compat) apply (auto simp: pair_less_def pair_leq_def) done subsection \Yet another induction principle on the natural numbers\ lemma nat_descend_induct [case_names base descend]: fixes P :: "nat \ bool" assumes H1: "\k. k > n \ P k" assumes H2: "\k. k \ n \ (\i. i > k \ P i) \ P k" shows "P m" using assms by induction_schema (force intro!: wf_measure [of "\k. Suc n - k"])+ subsection \Tool setup\ ML_file \Tools/Function/termination.ML\ ML_file \Tools/Function/scnp_solve.ML\ ML_file \Tools/Function/scnp_reconstruct.ML\ ML_file \Tools/Function/fun_cases.ML\ ML_val \ \setup inactive\ \ Context.theory_map (Function_Common.set_termination_prover (K (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS]))) \ end diff --git a/src/HOL/Groups_Big.thy b/src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy +++ b/src/HOL/Groups_Big.thy @@ -1,1737 +1,1740 @@ (* Title: HOL/Groups_Big.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Big sum and product over finite (non-empty) sets\ theory Groups_Big imports Power Equiv_Relations begin subsection \Generic monoid operation over a set\ locale comm_monoid_set = comm_monoid begin subsubsection \Standard sum or product indexed by a finite set\ interpretation comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) interpretation comp?: comp_fun_commute "f \ g" by (fact comp_comp_fun_commute) definition F :: "('b \ 'a) \ 'b set \ 'a" where eq_fold: "F g A = Finite_Set.fold (f \ g) \<^bold>1 A" lemma infinite [simp]: "\ finite A \ F g A = \<^bold>1" by (simp add: eq_fold) lemma empty [simp]: "F g {} = \<^bold>1" by (simp add: eq_fold) lemma insert [simp]: "finite A \ x \ A \ F g (insert x A) = g x \<^bold>* F g A" by (simp add: eq_fold) lemma remove: assumes "finite A" and "x \ A" shows "F g A = g x \<^bold>* F g (A - {x})" proof - from \x \ A\ obtain B where B: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ B have "finite B" by simp ultimately show ?thesis by simp qed lemma insert_remove: "finite A \ F g (insert x A) = g x \<^bold>* F g (A - {x})" by (cases "x \ A") (simp_all add: remove insert_absorb) lemma insert_if: "finite A \ F g (insert x A) = (if x \ A then F g A else g x \<^bold>* F g A)" by (cases "x \ A") (simp_all add: insert_absorb) lemma neutral: "\x\A. g x = \<^bold>1 \ F g A = \<^bold>1" by (induct A rule: infinite_finite_induct) simp_all lemma neutral_const [simp]: "F (\_. \<^bold>1) A = \<^bold>1" by (simp add: neutral) lemma union_inter: assumes "finite A" and "finite B" shows "F g (A \ B) \<^bold>* F g (A \ B) = F g A \<^bold>* F g B" \ \The reversed orientation looks more natural, but LOOPS as a simprule!\ using assms proof (induct A) case empty then show ?case by simp next case (insert x A) then show ?case by (auto simp: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) qed corollary union_inter_neutral: assumes "finite A" and "finite B" and "\x \ A \ B. g x = \<^bold>1" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter [symmetric] neutral) corollary union_disjoint: assumes "finite A" and "finite B" assumes "A \ B = {}" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter_neutral) lemma union_diff2: assumes "finite A" and "finite B" shows "F g (A \ B) = F g (A - B) \<^bold>* F g (B - A) \<^bold>* F g (A \ B)" proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto with assms show ?thesis by simp (subst union_disjoint, auto)+ qed lemma subset_diff: assumes "B \ A" and "finite A" shows "F g A = F g (A - B) \<^bold>* F g B" proof - from assms have "finite (A - B)" by auto moreover from assms have "finite B" by (rule finite_subset) moreover from assms have "(A - B) \ B = {}" by auto ultimately have "F g (A - B \ B) = F g (A - B) \<^bold>* F g B" by (rule union_disjoint) moreover from assms have "A \ B = A" by auto ultimately show ?thesis by simp qed lemma Int_Diff: assumes "finite A" shows "F g A = F g (A \ B) \<^bold>* F g (A - B)" by (subst subset_diff [where B = "A - B"]) (auto simp: Diff_Diff_Int assms) lemma setdiff_irrelevant: assumes "finite A" shows "F g (A - {x. g x = z}) = F g A" using assms by (induct A) (simp_all add: insert_Diff_if) lemma not_neutral_contains_not_neutral: assumes "F g A \ \<^bold>1" obtains a where "a \ A" and "g a \ \<^bold>1" proof - from assms have "\a\A. g a \ \<^bold>1" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert a A) then show ?case by fastforce qed with that show thesis by blast qed lemma reindex: assumes "inj_on h A" shows "F g (h ` A) = F (g \ h) A" proof (cases "finite A") case True with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc) next case False with assms have "\ finite (h ` A)" by (blast dest: finite_imageD) with False show ?thesis by simp qed lemma cong [fundef_cong]: assumes "A = B" assumes g_h: "\x. x \ B \ g x = h x" shows "F g A = F h B" using g_h unfolding \A = B\ by (induct B rule: infinite_finite_induct) auto lemma cong_simp [cong]: "\ A = B; \x. x \ B =simp=> g x = h x \ \ F (\x. g x) A = F (\x. h x) B" by (rule cong) (simp_all add: simp_implies_def) lemma reindex_cong: assumes "inj_on l B" assumes "A = l ` B" assumes "\x. x \ B \ g (l x) = h x" shows "F g A = F h B" using assms by (simp add: reindex) lemma image_eq: assumes "inj_on g A" shows "F (\x. x) (g ` A) = F g A" using assms reindex_cong by fastforce lemma UNION_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "F g (\(A ` I)) = F (\x. F g (A x)) I" using assms proof (induction rule: finite_induct) case (insert i I) then have "\j\I. j \ i" by blast with insert.prems have "A i \ \(A ` I) = {}" by blast with insert show ?case by (simp add: union_disjoint) qed auto lemma Union_disjoint: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A \ B = {}" shows "F g (\C) = (F \ F) g C" proof (cases "finite C") case True from UNION_disjoint [OF this assms] show ?thesis by simp next case False then show ?thesis by (auto dest: finite_UnionD intro: infinite) qed lemma distrib: "F (\x. g x \<^bold>* h x) A = F g A \<^bold>* F h A" by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute) lemma Sigma: assumes "finite A" "\x\A. finite (B x)" shows "F (\x. F (g x) (B x)) A = F (case_prod g) (SIGMA x:A. B x)" unfolding Sigma_def proof (subst UNION_disjoint) show "F (\x. F (g x) (B x)) A = F (\x. F (\(x, y). g x y) (\y\B x. {(x, y)})) A" proof (rule cong [OF refl]) show "F (g x) (B x) = F (\(x, y). g x y) (\y\B x. {(x, y)})" if "x \ A" for x using that assms by (simp add: UNION_disjoint) qed qed (use assms in auto) lemma related: assumes Re: "R \<^bold>1 \<^bold>1" and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 \<^bold>* y1) (x2 \<^bold>* y2)" and fin: "finite S" and R_h_g: "\x\S. R (h x) (g x)" shows "R (F h S) (F g S)" using fin by (rule finite_subset_induct) (use assms in auto) lemma mono_neutral_cong_left: assumes "finite T" and "S \ T" and "\i \ T - S. h i = \<^bold>1" and "\x. x \ S \ g x = h x" shows "F g S = F h T" proof- have eq: "T = S \ (T - S)" using \S \ T\ by blast have d: "S \ (T - S) = {}" using \S \ T\ by blast from \finite T\ \S \ T\ have f: "finite S" "finite (T - S)" by (auto intro: finite_subset) show ?thesis using assms(4) by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)]) qed lemma mono_neutral_cong_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ F g T = F h S" by (auto intro!: mono_neutral_cong_left [symmetric]) lemma mono_neutral_left: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g S = F g T" by (blast intro: mono_neutral_cong_left) lemma mono_neutral_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g T = F g S" by (blast intro!: mono_neutral_left [symmetric]) lemma mono_neutral_cong: assumes [simp]: "finite T" "finite S" and *: "\i. i \ T - S \ h i = \<^bold>1" "\i. i \ S - T \ g i = \<^bold>1" and gh: "\x. x \ S \ T \ g x = h x" shows "F g S = F h T" proof- have "F g S = F g (S \ T)" by(rule mono_neutral_right)(auto intro: *) also have "\ = F h (S \ T)" using refl gh by(rule cong) also have "\ = F h T" by(rule mono_neutral_left)(auto intro: *) finally show ?thesis . qed lemma reindex_bij_betw: "bij_betw h S T \ F (\x. g (h x)) S = F g T" by (auto simp: bij_betw_def reindex) lemma reindex_bij_witness: assumes witness: "\a. a \ S \ i (j a) = a" "\a. a \ S \ j a \ T" "\b. b \ T \ j (i b) = b" "\b. b \ T \ i b \ S" assumes eq: "\a. a \ S \ h (j a) = g a" shows "F g S = F h T" proof - have "bij_betw j S T" using bij_betw_byWitness[where A=S and f=j and f'=i and A'=T] witness by auto moreover have "F g S = F (\x. h (j x)) S" by (intro cong) (auto simp: eq) ultimately show ?thesis by (simp add: reindex_bij_betw) qed lemma reindex_bij_betw_not_neutral: assumes fin: "finite S'" "finite T'" assumes bij: "bij_betw h (S - S') (T - T')" assumes nn: "\a. a \ S' \ g (h a) = z" "\b. b \ T' \ g b = z" shows "F (\x. g (h x)) S = F g T" proof - have [simp]: "finite S \ finite T" using bij_betw_finite[OF bij] fin by auto show ?thesis proof (cases "finite S") case True with nn have "F (\x. g (h x)) S = F (\x. g (h x)) (S - S')" by (intro mono_neutral_cong_right) auto also have "\ = F g (T - T')" using bij by (rule reindex_bij_betw) also have "\ = F g T" using nn \finite S\ by (intro mono_neutral_cong_left) auto finally show ?thesis . next case False then show ?thesis by simp qed qed lemma reindex_nontrivial: assumes "finite A" and nz: "\x y. x \ A \ y \ A \ x \ y \ h x = h y \ g (h x) = \<^bold>1" shows "F g (h ` A) = F (g \ h) A" proof (subst reindex_bij_betw_not_neutral [symmetric]) show "bij_betw h (A - {x \ A. (g \ h) x = \<^bold>1}) (h ` A - h ` {x \ A. (g \ h) x = \<^bold>1})" using nz by (auto intro!: inj_onI simp: bij_betw_def) qed (use \finite A\ in auto) lemma reindex_bij_witness_not_neutral: assumes fin: "finite S'" "finite T'" assumes witness: "\a. a \ S - S' \ i (j a) = a" "\a. a \ S - S' \ j a \ T - T'" "\b. b \ T - T' \ j (i b) = b" "\b. b \ T - T' \ i b \ S - S'" assumes nn: "\a. a \ S' \ g a = z" "\b. b \ T' \ h b = z" assumes eq: "\a. a \ S \ h (j a) = g a" shows "F g S = F h T" proof - have bij: "bij_betw j (S - (S' \ S)) (T - (T' \ T))" using witness by (intro bij_betw_byWitness[where f'=i]) auto have F_eq: "F g S = F (\x. h (j x)) S" by (intro cong) (auto simp: eq) show ?thesis unfolding F_eq using fin nn eq by (intro reindex_bij_betw_not_neutral[OF _ _ bij]) auto qed lemma delta_remove: assumes fS: "finite S" shows "F (\k. if k = a then b k else c k) S = (if a \ S then b a \<^bold>* F c (S-{a}) else F c (S-{a}))" proof - let ?f = "(\k. if k = a then b k else c k)" show ?thesis proof (cases "a \ S") case False then have "\k\S. ?f k = c k" by simp with False show ?thesis by simp next case True let ?A = "S - {a}" let ?B = "{a}" from True have eq: "S = ?A \ ?B" by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have "F ?f S = F ?f ?A \<^bold>* F ?f ?B" using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] by simp with True show ?thesis using comm_monoid_set.remove comm_monoid_set_axioms fS by fastforce qed qed lemma delta [simp]: assumes fS: "finite S" shows "F (\k. if k = a then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" by (simp add: delta_remove [OF assms]) lemma delta' [simp]: assumes fin: "finite S" shows "F (\k. if a = k then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" using delta [OF fin, of a b, symmetric] by (auto intro: cong) lemma If_cases: fixes P :: "'b \ bool" and g h :: "'b \ 'a" assumes fin: "finite A" shows "F (\x. if P x then h x else g x) A = F h (A \ {x. P x}) \<^bold>* F g (A \ - {x. P x})" proof - have a: "A = A \ {x. P x} \ A \ -{x. P x}" "(A \ {x. P x}) \ (A \ -{x. P x}) = {}" by blast+ from fin have f: "finite (A \ {x. P x})" "finite (A \ -{x. P x})" by auto let ?g = "\x. if P x then h x else g x" from union_disjoint [OF f a(2), of ?g] a(1) show ?thesis by (subst (1 2) cong) simp_all qed lemma cartesian_product: "F (\x. F (g x) B) A = F (case_prod g) (A \ B)" proof (cases "A = {} \ B = {}") case True then show ?thesis by auto next case False then have "A \ {}" "B \ {}" by auto show ?thesis proof (cases "finite A \ finite B") case True then show ?thesis by (simp add: Sigma) next case False then consider "infinite A" | "infinite B" by auto then have "infinite (A \ B)" by cases (use \A \ {}\ \B \ {}\ in \auto dest: finite_cartesian_productD1 finite_cartesian_productD2\) then show ?thesis using False by auto qed qed lemma inter_restrict: assumes "finite A" shows "F g (A \ B) = F (\x. if x \ B then g x else \<^bold>1) A" proof - let ?g = "\x. if x \ A \ B then g x else \<^bold>1" have "\i\A - A \ B. (if i \ A \ B then g i else \<^bold>1) = \<^bold>1" by simp moreover have "A \ B \ A" by blast ultimately have "F ?g (A \ B) = F ?g A" using \finite A\ by (intro mono_neutral_left) auto then show ?thesis by simp qed lemma inter_filter: "finite A \ F g {x \ A. P x} = F (\x. if P x then g x else \<^bold>1) A" by (simp add: inter_restrict [symmetric, of A "{x. P x}" g, simplified mem_Collect_eq] Int_def) lemma Union_comp: assumes "\A \ B. finite A" and "\A1 A2 x. A1 \ B \ A2 \ B \ A1 \ A2 \ x \ A1 \ x \ A2 \ g x = \<^bold>1" shows "F g (\B) = (F \ F) g B" using assms proof (induct B rule: infinite_finite_induct) case (infinite A) then have "\ finite (\A)" by (blast dest: finite_UnionD) with infinite show ?case by simp next case empty then show ?case by simp next case (insert A B) then have "finite A" "finite B" "finite (\B)" "A \ B" and "\x\A \ \B. g x = \<^bold>1" and H: "F g (\B) = (F \ F) g B" by auto then have "F g (A \ \B) = F g A \<^bold>* F g (\B)" by (simp add: union_inter_neutral) with \finite B\ \A \ B\ show ?case by (simp add: H) qed lemma swap: "F (\i. F (g i) B) A = F (\j. F (\i. g i j) A) B" unfolding cartesian_product by (rule reindex_bij_witness [where i = "\(i, j). (j, i)" and j = "\(i, j). (j, i)"]) auto lemma swap_restrict: "finite A \ finite B \ F (\x. F (g x) {y. y \ B \ R x y}) A = F (\y. F (\x. g x y) {x. x \ A \ R x y}) B" by (simp add: inter_filter) (rule swap) lemma image_gen: assumes fin: "finite S" shows "F h S = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" proof - have "{y. y\ g`S \ g x = y} = {g x}" if "x \ S" for x using that by auto then have "F h S = F (\x. F (\y. h x) {y. y\ g`S \ g x = y}) S" by simp also have "\ = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" by (rule swap_restrict [OF fin finite_imageI [OF fin]]) finally show ?thesis . qed lemma group: assumes fS: "finite S" and fT: "finite T" and fST: "g ` S \ T" shows "F (\y. F h {x. x \ S \ g x = y}) T = F h S" unfolding image_gen[OF fS, of h g] by (auto intro: neutral mono_neutral_right[OF fT fST]) lemma Plus: fixes A :: "'b set" and B :: "'c set" assumes fin: "finite A" "finite B" shows "F g (A <+> B) = F (g \ Inl) A \<^bold>* F (g \ Inr) B" proof - have "A <+> B = Inl ` A \ Inr ` B" by auto moreover from fin have "finite (Inl ` A)" "finite (Inr ` B)" by auto moreover have "Inl ` A \ Inr ` B = {}" by auto moreover have "inj_on Inl A" "inj_on Inr B" by (auto intro: inj_onI) ultimately show ?thesis using fin by (simp add: union_disjoint reindex) qed lemma same_carrier: assumes "finite C" assumes subset: "A \ C" "B \ C" assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" shows "F g A = F h B \ F g C = F h C" proof - have "finite A" and "finite B" and "finite (C - A)" and "finite (C - B)" using \finite C\ subset by (auto elim: finite_subset) from subset have [simp]: "A - (C - A) = A" by auto from subset have [simp]: "B - (C - B) = B" by auto from subset have "C = A \ (C - A)" by auto then have "F g C = F g (A \ (C - A))" by simp also have "\ = F g (A - (C - A)) \<^bold>* F g (C - A - A) \<^bold>* F g (A \ (C - A))" using \finite A\ \finite (C - A)\ by (simp only: union_diff2) finally have *: "F g C = F g A" using trivial by simp from subset have "C = B \ (C - B)" by auto then have "F h C = F h (B \ (C - B))" by simp also have "\ = F h (B - (C - B)) \<^bold>* F h (C - B - B) \<^bold>* F h (B \ (C - B))" using \finite B\ \finite (C - B)\ by (simp only: union_diff2) finally have "F h C = F h B" using trivial by simp with * show ?thesis by simp qed lemma same_carrierI: assumes "finite C" assumes subset: "A \ C" "B \ C" assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" assumes "F g C = F h C" shows "F g A = F h B" using assms same_carrier [of C A B] by simp lemma eq_general: assumes B: "\y. y \ B \ \!x. x \ A \ h x = y" and A: "\x. x \ A \ h x \ B \ \(h x) = \ x" shows "F \ A = F \ B" proof - have eq: "B = h ` A" by (auto dest: assms) have h: "inj_on h A" using assms by (blast intro: inj_onI) have "F \ A = F (\ \ h) A" using A by auto also have "\ = F \ B" by (simp add: eq reindex h) finally show ?thesis . qed lemma eq_general_inverses: assumes B: "\y. y \ B \ k y \ A \ h(k y) = y" and A: "\x. x \ A \ h x \ B \ k(h x) = x \ \(h x) = \ x" shows "F \ A = F \ B" by (rule eq_general [where h=h]) (force intro: dest: A B)+ subsubsection \HOL Light variant: sum/product indexed by the non-neutral subset\ text \NB only a subset of the properties above are proved\ definition G :: "['b \ 'a,'b set] \ 'a" where "G p I \ if finite {x \ I. p x \ \<^bold>1} then F p {x \ I. p x \ \<^bold>1} else \<^bold>1" lemma finite_Collect_op: shows "\finite {i \ I. x i \ \<^bold>1}; finite {i \ I. y i \ \<^bold>1}\ \ finite {i \ I. x i \<^bold>* y i \ \<^bold>1}" apply (rule finite_subset [where B = "{i \ I. x i \ \<^bold>1} \ {i \ I. y i \ \<^bold>1}"]) using left_neutral by force+ lemma empty' [simp]: "G p {} = \<^bold>1" by (auto simp: G_def) lemma eq_sum [simp]: "finite I \ G p I = F p I" by (auto simp: G_def intro: mono_neutral_cong_left) lemma insert' [simp]: assumes "finite {x \ I. p x \ \<^bold>1}" shows "G p (insert i I) = (if i \ I then G p I else p i \<^bold>* G p I)" proof - have "{x. x = i \ p x \ \<^bold>1 \ x \ I \ p x \ \<^bold>1} = (if p i = \<^bold>1 then {x \ I. p x \ \<^bold>1} else insert i {x \ I. p x \ \<^bold>1})" by auto then show ?thesis using assms by (simp add: G_def conj_disj_distribR insert_absorb) qed lemma distrib_triv': assumes "finite I" shows "G (\i. g i \<^bold>* h i) I = G g I \<^bold>* G h I" by (simp add: assms local.distrib) lemma non_neutral': "G g {x \ I. g x \ \<^bold>1} = G g I" by (simp add: G_def) lemma distrib': assumes "finite {x \ I. g x \ \<^bold>1}" "finite {x \ I. h x \ \<^bold>1}" shows "G (\i. g i \<^bold>* h i) I = G g I \<^bold>* G h I" proof - have "a \<^bold>* a \ a \ a \ \<^bold>1" for a by auto then have "G (\i. g i \<^bold>* h i) I = G (\i. g i \<^bold>* h i) ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1})" using assms by (force simp: G_def finite_Collect_op intro!: mono_neutral_cong) also have "\ = G g I \<^bold>* G h I" proof - have "F g ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1}) = G g I" "F h ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1}) = G h I" by (auto simp: G_def assms intro: mono_neutral_right) then show ?thesis using assms by (simp add: distrib) qed finally show ?thesis . qed lemma cong': assumes "A = B" assumes g_h: "\x. x \ B \ g x = h x" shows "G g A = G h B" using assms by (auto simp: G_def cong: conj_cong intro: cong) lemma mono_neutral_cong_left': assumes "S \ T" and "\i. i \ T - S \ h i = \<^bold>1" and "\x. x \ S \ g x = h x" shows "G g S = G h T" proof - have *: "{x \ S. g x \ \<^bold>1} = {x \ T. h x \ \<^bold>1}" using assms by (metis DiffI subset_eq) then have "finite {x \ S. g x \ \<^bold>1} = finite {x \ T. h x \ \<^bold>1}" by simp then show ?thesis using assms by (auto simp add: G_def * intro: cong) qed lemma mono_neutral_cong_right': "S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ G g T = G h S" by (auto intro!: mono_neutral_cong_left' [symmetric]) lemma mono_neutral_left': "S \ T \ \i \ T - S. g i = \<^bold>1 \ G g S = G g T" by (blast intro: mono_neutral_cong_left') lemma mono_neutral_right': "S \ T \ \i \ T - S. g i = \<^bold>1 \ G g T = G g S" by (blast intro!: mono_neutral_left' [symmetric]) end subsection \Generalized summation over a set\ context comm_monoid_add begin sublocale sum: comm_monoid_set plus 0 defines sum = sum.F and sum' = sum.G .. abbreviation Sum ("\") where "\ \ sum (\x. x)" end text \Now: lots of fancy syntax. First, \<^term>\sum (\x. e) A\ is written \\x\A. e\.\ syntax (ASCII) "_sum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(3SUM (_/:_)./ _)" [0, 51, 10] 10) syntax "_sum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(2\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\i\A. b" \ "CONST sum (\i. b) A" text \Instead of \<^term>\\x\{x. P}. e\ we introduce the shorter \\x|P. e\.\ syntax (ASCII) "_qsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0, 0, 10] 10) syntax "_qsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations "\x|P. t" => "CONST sum (\x. t) {x. P}" print_translation \ let fun sum_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\Collect\, _) $ Abs (y, Ty, P)] = if x <> y then raise Match else let val x' = Syntax_Trans.mark_bound_body (x, Tx); val t' = subst_bound (x', t); val P' = subst_bound (x', P); in Syntax.const \<^syntax_const>\_qsum\ $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t' end | sum_tr' _ = raise Match; in [(\<^const_syntax>\sum\, K sum_tr')] end \ subsubsection \Properties in more restricted classes of structures\ lemma sum_Un: "finite A \ finite B \ sum f (A \ B) = sum f A + sum f B - sum f (A \ B)" for f :: "'b \ 'a::ab_group_add" by (subst sum.union_inter [symmetric]) (auto simp add: algebra_simps) lemma sum_Un2: assumes "finite (A \ B)" shows "sum f (A \ B) = sum f (A - B) + sum f (B - A) + sum f (A \ B)" proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto with assms show ?thesis by simp (subst sum.union_disjoint, auto)+ qed (*Like sum.subset_diff but expressed perhaps more conveniently using subtraction*) lemma sum_diff: fixes f :: "'b \ 'a::ab_group_add" assumes "finite A" "B \ A" shows "sum f (A - B) = sum f A - sum f B" using sum.subset_diff [of B A f] assms by simp lemma sum_diff1: fixes f :: "'b \ 'a::ab_group_add" assumes "finite A" shows "sum f (A - {a}) = (if a \ A then sum f A - f a else sum f A)" using assms by (simp add: sum_diff) lemma sum_diff1'_aux: fixes f :: "'a \ 'b::ab_group_add" assumes "finite F" "{i \ I. f i \ 0} \ F" shows "sum' f (I - {i}) = (if i \ I then sum' f I - f i else sum' f I)" using assms proof induct case (insert x F) have 1: "finite {x \ I. f x \ 0} \ finite {x \ I. x \ i \ f x \ 0}" by (erule rev_finite_subset) auto have 2: "finite {x \ I. x \ i \ f x \ 0} \ finite {x \ I. f x \ 0}" apply (drule finite_insert [THEN iffD2]) by (erule rev_finite_subset) auto have 3: "finite {i \ I. f i \ 0}" using finite_subset insert by blast show ?case using insert sum_diff1 [of "{i \ I. f i \ 0}" f i] by (auto simp: sum.G_def 1 2 3 set_diff_eq conj_ac) qed (simp add: sum.G_def) lemma sum_diff1': fixes f :: "'a \ 'b::ab_group_add" assumes "finite {i \ I. f i \ 0}" shows "sum' f (I - {i}) = (if i \ I then sum' f I - f i else sum' f I)" by (rule sum_diff1'_aux [OF assms order_refl]) lemma (in ordered_comm_monoid_add) sum_mono: "(\i. i\K \ f i \ g i) \ (\i\K. f i) \ (\i\K. g i)" by (induct K rule: infinite_finite_induct) (use add_mono in auto) lemma (in strict_ordered_comm_monoid_add) sum_strict_mono: assumes "finite A" "A \ {}" and "\x. x \ A \ f x < g x" shows "sum f A < sum g A" using assms proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case insert then show ?case by (auto simp: add_strict_mono) qed lemma sum_strict_mono_ex1: fixes f g :: "'i \ 'a::ordered_cancel_comm_monoid_add" assumes "finite A" and "\x\A. f x \ g x" and "\a\A. f a < g a" shows "sum f A < sum g A" proof- from assms(3) obtain a where a: "a \ A" "f a < g a" by blast have "sum f A = sum f ((A - {a}) \ {a})" by(simp add: insert_absorb[OF \a \ A\]) also have "\ = sum f (A - {a}) + sum f {a}" using \finite A\ by(subst sum.union_disjoint) auto also have "sum f (A - {a}) \ sum g (A - {a})" by (rule sum_mono) (simp add: assms(2)) also from a have "sum f {a} < sum g {a}" by simp also have "sum g (A - {a}) + sum g {a} = sum g((A - {a}) \ {a})" using \finite A\ by (subst sum.union_disjoint[symmetric]) auto also have "\ = sum g A" by (simp add: insert_absorb[OF \a \ A\]) finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono) qed lemma sum_mono_inv: fixes f g :: "'i \ 'a :: ordered_cancel_comm_monoid_add" assumes eq: "sum f I = sum g I" assumes le: "\i. i \ I \ f i \ g i" assumes i: "i \ I" assumes I: "finite I" shows "f i = g i" proof (rule ccontr) assume "\ ?thesis" with le[OF i] have "f i < g i" by simp with i have "\i\I. f i < g i" .. from sum_strict_mono_ex1[OF I _ this] le have "sum f I < sum g I" by blast with eq show False by simp qed lemma member_le_sum: fixes f :: "_ \ 'b::{semiring_1, ordered_comm_monoid_add}" assumes "i \ A" and le: "\x. x \ A - {i} \ 0 \ f x" and "finite A" shows "f i \ sum f A" proof - have "f i \ sum f (A \ {i})" by (simp add: assms) also have "... = (\x\A. if x \ {i} then f x else 0)" using assms sum.inter_restrict by blast also have "... \ sum f A" apply (rule sum_mono) apply (auto simp: le) done finally show ?thesis . qed lemma sum_negf: "(\x\A. - f x) = - (\x\A. f x)" for f :: "'b \ 'a::ab_group_add" by (induct A rule: infinite_finite_induct) auto lemma sum_subtractf: "(\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" for f g :: "'b \'a::ab_group_add" using sum.distrib [of f "- g" A] by (simp add: sum_negf) lemma sum_subtractf_nat: "(\x. x \ A \ g x \ f x) \ (\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" for f g :: "'a \ nat" by (induct A rule: infinite_finite_induct) (auto simp: sum_mono) context ordered_comm_monoid_add begin lemma sum_nonneg: "(\x. x \ A \ 0 \ f x) \ 0 \ sum f A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then have "0 + 0 \ f x + sum f F" by (blast intro: add_mono) with insert show ?case by simp qed lemma sum_nonpos: "(\x. x \ A \ f x \ 0) \ sum f A \ 0" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then have "f x + sum f F \ 0 + 0" by (blast intro: add_mono) with insert show ?case by simp qed lemma sum_nonneg_eq_0_iff: "finite A \ (\x. x \ A \ 0 \ f x) \ sum f A = 0 \ (\x\A. f x = 0)" by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff sum_nonneg) lemma sum_nonneg_0: "finite s \ (\i. i \ s \ f i \ 0) \ (\ i \ s. f i) = 0 \ i \ s \ f i = 0" by (simp add: sum_nonneg_eq_0_iff) lemma sum_nonneg_leq_bound: assumes "finite s" "\i. i \ s \ f i \ 0" "(\i \ s. f i) = B" "i \ s" shows "f i \ B" proof - from assms have "f i \ f i + (\i \ s - {i}. f i)" by (intro add_increasing2 sum_nonneg) auto also have "\ = B" using sum.remove[of s i f] assms by simp finally show ?thesis by auto qed lemma sum_mono2: assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 0 \ f b" shows "sum f A \ sum f B" proof - have "sum f A \ sum f A + sum f (B-A)" by (auto intro: add_increasing2 [OF sum_nonneg] nn) also from fin finite_subset[OF sub fin] have "\ = sum f (A \ (B-A))" by (simp add: sum.union_disjoint del: Un_Diff_cancel) also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed lemma sum_le_included: assumes "finite s" "finite t" and "\y\t. 0 \ g y" "(\x\s. \y\t. i y = x \ f x \ g y)" shows "sum f s \ sum g t" proof - have "sum f s \ sum (\y. sum g {x. x\t \ i x = y}) s" proof (rule sum_mono) fix y assume "y \ s" with assms obtain z where z: "z \ t" "y = i z" "f y \ g z" by auto with assms show "f y \ sum g {x \ t. i x = y}" (is "?A y \ ?B y") using order_trans[of "?A (i z)" "sum g {z}" "?B (i z)", intro] by (auto intro!: sum_mono2) qed also have "\ \ sum (\y. sum g {x. x\t \ i x = y}) (i ` t)" using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg) also have "\ \ sum g t" using assms by (auto simp: sum.image_gen[symmetric]) finally show ?thesis . qed end lemma (in canonically_ordered_monoid_add) sum_eq_0_iff [simp]: "finite F \ (sum f F = 0) = (\a\F. f a = 0)" by (intro ballI sum_nonneg_eq_0_iff zero_le) context semiring_0 begin lemma sum_distrib_left: "r * sum f A = (\n\A. r * f n)" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) lemma sum_distrib_right: "sum f A * r = (\n\A. f n * r)" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) end lemma sum_divide_distrib: "sum f A / r = (\n\A. f n / r)" for r :: "'a::field" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case by (simp add: add_divide_distrib) qed lemma sum_abs[iff]: "\sum f A\ \ sum (\i. \f i\) A" for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case by (auto intro: abs_triangle_ineq order_trans) qed lemma sum_abs_ge_zero[iff]: "0 \ sum (\i. \f i\) A" for f :: "'a \ 'b::ordered_ab_group_add_abs" by (simp add: sum_nonneg) lemma abs_sum_abs[simp]: "\\a\A. \f a\\ = (\a\A. \f a\)" for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert a A) then have "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp also from insert have "\ = \\f a\ + \\a\A. \f a\\\" by simp also have "\ = \f a\ + \\a\A. \f a\\" by (simp del: abs_of_nonneg) also from insert have "\ = (\a\insert a A. \f a\)" by simp finally show ?case . qed lemma sum_product: fixes f :: "'a \ 'b::semiring_0" shows "sum f A * sum g B = (\i\A. \j\B. f i * g j)" by (simp add: sum_distrib_left sum_distrib_right) (rule sum.swap) lemma sum_mult_sum_if_inj: fixes f :: "'a \ 'b::semiring_0" shows "inj_on (\(a, b). f a * g b) (A \ B) \ sum f A * sum g B = sum id {f a * g b |a b. a \ A \ b \ B}" by(auto simp: sum_product sum.cartesian_product intro!: sum.reindex_cong[symmetric]) lemma sum_SucD: "sum f A = Suc n \ \a\A. 0 < f a" by (induct A rule: infinite_finite_induct) auto lemma sum_eq_Suc0_iff: "finite A \ sum f A = Suc 0 \ (\a\A. f a = Suc 0 \ (\b\A. a \ b \ f b = 0))" by (induct A rule: finite_induct) (auto simp add: add_is_1) lemmas sum_eq_1_iff = sum_eq_Suc0_iff[simplified One_nat_def[symmetric]] lemma sum_Un_nat: "finite A \ finite B \ sum f (A \ B) = sum f A + sum f B - sum f (A \ B)" for f :: "'a \ nat" \ \For the natural numbers, we have subtraction.\ by (subst sum.union_inter [symmetric]) (auto simp: algebra_simps) lemma sum_diff1_nat: "sum f (A - {a}) = (if a \ A then sum f A - f a else sum f A)" for f :: "'a \ nat" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next - case insert + case (insert x F) then show ?case - apply (auto simp: insert_Diff_if) - apply (drule mk_disjoint_insert) - apply auto - done + proof (cases "a \ F") + case True + then have "\B. F = insert a B \ a \ B" + by (auto simp: mk_disjoint_insert) + then show ?thesis using insert + by (auto simp: insert_Diff_if) + qed (auto) qed lemma sum_diff_nat: fixes f :: "'a \ nat" assumes "finite B" and "B \ A" shows "sum f (A - B) = sum f A - sum f B" using assms proof induct case empty then show ?case by simp next case (insert x F) note IH = \F \ A \ sum f (A - F) = sum f A - sum f F\ from \x \ F\ \insert x F \ A\ have "x \ A - F" by simp then have A: "sum f ((A - F) - {x}) = sum f (A - F) - f x" by (simp add: sum_diff1_nat) from \insert x F \ A\ have "F \ A" by simp with IH have "sum f (A - F) = sum f A - sum f F" by simp with A have B: "sum f ((A - F) - {x}) = sum f A - sum f F - f x" by simp from \x \ F\ have "A - insert x F = (A - F) - {x}" by auto with B have C: "sum f (A - insert x F) = sum f A - sum f F - f x" by simp from \finite F\ \x \ F\ have "sum f (insert x F) = sum f F + f x" by simp with C have "sum f (A - insert x F) = sum f A - sum f (insert x F)" by simp then show ?case by simp qed lemma sum_comp_morphism: "h 0 = 0 \ (\x y. h (x + y) = h x + h y) \ sum (h \ g) A = h (sum g A)" by (induct A rule: infinite_finite_induct) simp_all lemma (in comm_semiring_1) dvd_sum: "(\a. a \ A \ d dvd f a) \ d dvd sum f A" by (induct A rule: infinite_finite_induct) simp_all lemma (in ordered_comm_monoid_add) sum_pos: "finite I \ I \ {} \ (\i. i \ I \ 0 < f i) \ 0 < sum f I" by (induct I rule: finite_ne_induct) (auto intro: add_pos_pos) lemma (in ordered_comm_monoid_add) sum_pos2: assumes I: "finite I" "i \ I" "0 < f i" "\i. i \ I \ 0 \ f i" shows "0 < sum f I" proof - have "0 < f i + sum f (I - {i})" using assms by (intro add_pos_nonneg sum_nonneg) auto also have "\ = sum f I" using assms by (simp add: sum.remove) finally show ?thesis . qed lemma sum_strict_mono2: fixes f :: "'a \ 'b::ordered_cancel_comm_monoid_add" assumes "finite B" "A \ B" "b \ B-A" "f b > 0" and "\x. x \ B \ f x \ 0" shows "sum f A < sum f B" proof - have "B - A \ {}" using assms(3) by blast have "sum f (B-A) > 0" by (rule sum_pos2) (use assms in auto) moreover have "sum f B = sum f (B-A) + sum f A" by (rule sum.subset_diff) (use assms in auto) ultimately show ?thesis using add_strict_increasing by auto qed lemma sum_cong_Suc: assumes "0 \ A" "\x. Suc x \ A \ f (Suc x) = g (Suc x)" shows "sum f A = sum g A" proof (rule sum.cong) fix x assume "x \ A" with assms(1) show "f x = g x" by (cases x) (auto intro!: assms(2)) qed simp_all subsubsection \Cardinality as special case of \<^const>\sum\\ lemma card_eq_sum: "card A = sum (\x. 1) A" proof - have "plus \ (\_. Suc 0) = (\_. Suc)" by (simp add: fun_eq_iff) then have "Finite_Set.fold (plus \ (\_. Suc 0)) = Finite_Set.fold (\_. Suc)" by (rule arg_cong) then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" by (blast intro: fun_cong) then show ?thesis by (simp add: card.eq_fold sum.eq_fold) qed context semiring_1 begin lemma sum_constant [simp]: "(\x \ A. y) = of_nat (card A) * y" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) context fixes A assumes \finite A\ begin lemma sum_of_bool_eq [simp]: \(\x \ A. of_bool (P x)) = of_nat (card (A \ {x. P x}))\ if \finite A\ using \finite A\ by induction simp_all lemma sum_mult_of_bool_eq [simp]: \(\x \ A. f x * of_bool (P x)) = (\x \ (A \ {x. P x}). f x)\ by (rule sum.mono_neutral_cong) (use \finite A\ in auto) lemma sum_of_bool_mult_eq [simp]: \(\x \ A. of_bool (P x) * f x) = (\x \ (A \ {x. P x}). f x)\ by (rule sum.mono_neutral_cong) (use \finite A\ in auto) end end lemma sum_Suc: "sum (\x. Suc(f x)) A = sum f A + card A" using sum.distrib[of f "\_. 1" A] by simp lemma sum_bounded_above: fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" assumes le: "\i. i\A \ f i \ K" shows "sum f A \ of_nat (card A) * K" proof (cases "finite A") case True then show ?thesis using le sum_mono[where K=A and g = "\x. K"] by simp next case False then show ?thesis by simp qed lemma sum_bounded_above_divide: fixes K :: "'a::linordered_field" assumes le: "\i. i\A \ f i \ K / of_nat (card A)" and fin: "finite A" "A \ {}" shows "sum f A \ K" using sum_bounded_above [of A f "K / of_nat (card A)", OF le] fin by simp lemma sum_bounded_above_strict: fixes K :: "'a::{ordered_cancel_comm_monoid_add,semiring_1}" assumes "\i. i\A \ f i < K" "card A > 0" shows "sum f A < of_nat (card A) * K" using assms sum_strict_mono[where A=A and g = "\x. K"] by (simp add: card_gt_0_iff) lemma sum_bounded_below: fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" assumes le: "\i. i\A \ K \ f i" shows "of_nat (card A) * K \ sum f A" proof (cases "finite A") case True then show ?thesis using le sum_mono[where K=A and f = "\x. K"] by simp next case False then show ?thesis by simp qed lemma convex_sum_bound_le: fixes x :: "'a \ 'b::linordered_idom" assumes 0: "\i. i \ I \ 0 \ x i" and 1: "sum x I = 1" and \: "\i. i \ I \ \a i - b\ \ \" shows "\(\i\I. a i * x i) - b\ \ \" proof - have [simp]: "(\i\I. c * x i) = c" for c by (simp flip: sum_distrib_left 1) then have "\(\i\I. a i * x i) - b\ = \\i\I. (a i - b) * x i\" by (simp add: sum_subtractf left_diff_distrib) also have "\ \ (\i\I. \(a i - b) * x i\)" using abs_abs abs_of_nonneg by blast also have "\ \ (\i\I. \(a i - b)\ * x i)" by (simp add: abs_mult 0) also have "\ \ (\i\I. \ * x i)" by (rule sum_mono) (use \ "0" mult_right_mono in blast) also have "\ = \" by simp finally show ?thesis . qed lemma card_UN_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "card (\(A ` I)) = (\i\I. card(A i))" proof - have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" by simp with assms show ?thesis by (simp add: card_eq_sum sum.UNION_disjoint del: sum_constant) qed lemma card_Union_disjoint: assumes "pairwise disjnt C" and fin: "\A. A \ C \ finite A" shows "card (\C) = sum card C" proof (cases "finite C") case True then show ?thesis using card_UN_disjoint [OF True, of "\x. x"] assms by (simp add: disjnt_def fin pairwise_def) next case False then show ?thesis using assms card_eq_0_iff finite_UnionD by fastforce qed lemma card_Union_le_sum_card_weak: fixes U :: "'a set set" assumes "\u \ U. finite u" shows "card (\U) \ sum card U" proof (cases "finite U") case False then show "card (\U) \ sum card U" using card_eq_0_iff finite_UnionD by auto next case True then show "card (\U) \ sum card U" proof (induct U rule: finite_induct) case empty then show ?case by auto next case (insert x F) then have "card(\(insert x F)) \ card(x) + card (\F)" using card_Un_le by auto also have "... \ card(x) + sum card F" using insert.hyps by auto also have "... = sum card (insert x F)" using sum.insert_if and insert.hyps by auto finally show ?case . qed qed lemma card_Union_le_sum_card: fixes U :: "'a set set" shows "card (\U) \ sum card U" by (metis Union_upper card.infinite card_Union_le_sum_card_weak finite_subset zero_le) lemma card_UN_le: assumes "finite I" shows "card(\i\I. A i) \ (\i\I. card(A i))" using assms proof induction case (insert i I) then show ?case using card_Un_le nat_add_left_cancel_le by (force intro: order_trans) qed auto lemma card_quotient_disjoint: assumes "finite A" "inj_on (\x. {x} // r) A" shows "card (A//r) = card A" proof - have "\i\A. \j\A. i \ j \ r `` {j} \ r `` {i}" using assms by (fastforce simp add: quotient_def inj_on_def) with assms show ?thesis by (simp add: quotient_def card_UN_disjoint) qed lemma sum_multicount_gen: assumes "finite s" "finite t" "\j\t. (card {i\s. R i j} = k j)" shows "sum (\i. (card {j\t. R i j})) s = sum k t" (is "?l = ?r") proof- have "?l = sum (\i. sum (\x.1) {j\t. R i j}) s" by auto also have "\ = ?r" unfolding sum.swap_restrict [OF assms(1-2)] using assms(3) by auto finally show ?thesis . qed lemma sum_multicount: assumes "finite S" "finite T" "\j\T. (card {i\S. R i j} = k)" shows "sum (\i. card {j\T. R i j}) S = k * card T" (is "?l = ?r") proof- have "?l = sum (\i. k) T" by (rule sum_multicount_gen) (auto simp: assms) also have "\ = ?r" by (simp add: mult.commute) finally show ?thesis by auto qed lemma sum_card_image: assumes "finite A" assumes "pairwise (\s t. disjnt (f s) (f t)) A" shows "sum card (f ` A) = sum (\a. card (f a)) A" using assms proof (induct A) case (insert a A) show ?case proof cases assume "f a = {}" with insert show ?case by (subst sum.mono_neutral_right[where S="f ` A"]) (auto simp: pairwise_insert) next assume "f a \ {}" then have "sum card (insert (f a) (f ` A)) = card (f a) + sum card (f ` A)" using insert by (subst sum.insert) (auto simp: pairwise_insert) with insert show ?case by (simp add: pairwise_insert) qed qed simp text \By Jakub Kądziołka:\ lemma sum_fun_comp: assumes "finite S" "finite R" "g ` S \ R" shows "(\x \ S. f (g x)) = (\y \ R. of_nat (card {x \ S. g x = y}) * f y)" proof - let ?r = "relation_of (\p q. g p = g q) S" have eqv: "equiv S ?r" unfolding relation_of_def by (auto intro: comp_equivI) have finite: "C \ S//?r \ finite C" for C by (fact finite_equiv_class[OF `finite S` equiv_type[OF `equiv S ?r`]]) have disjoint: "A \ S//?r \ B \ S//?r \ A \ B \ A \ B = {}" for A B using eqv quotient_disj by blast let ?cls = "\y. {x \ S. y = g x}" have quot_as_img: "S//?r = ?cls ` g ` S" by (auto simp add: relation_of_def quotient_def) have cls_inj: "inj_on ?cls (g ` S)" by (auto intro: inj_onI) have rest_0: "(\y \ R - g ` S. of_nat (card (?cls y)) * f y) = 0" proof - have "of_nat (card (?cls y)) * f y = 0" if asm: "y \ R - g ` S" for y proof - from asm have *: "?cls y = {}" by auto show ?thesis unfolding * by simp qed thus ?thesis by simp qed have "(\x \ S. f (g x)) = (\C \ S//?r. \x \ C. f (g x))" using eqv finite disjoint by (simp flip: sum.Union_disjoint[simplified] add: Union_quotient) also have "... = (\y \ g ` S. \x \ ?cls y. f (g x))" unfolding quot_as_img by (simp add: sum.reindex[OF cls_inj]) also have "... = (\y \ g ` S. \x \ ?cls y. f y)" by auto also have "... = (\y \ g ` S. of_nat (card (?cls y)) * f y)" by (simp flip: sum_constant) also have "... = (\y \ R. of_nat (card (?cls y)) * f y)" using rest_0 by (simp add: sum.subset_diff[OF \g ` S \ R\ \finite R\]) finally show ?thesis by (simp add: eq_commute) qed subsubsection \Cardinality of products\ lemma card_SigmaI [simp]: "finite A \ \a\A. finite (B a) \ card (SIGMA x: A. B x) = (\a\A. card (B a))" by (simp add: card_eq_sum sum.Sigma del: sum_constant) (* lemma SigmaI_insert: "y \ A ==> (SIGMA x:(insert y A). B x) = (({y} \ (B y)) \ (SIGMA x: A. B x))" by auto *) lemma card_cartesian_product: "card (A \ B) = card A * card B" by (cases "finite A \ finite B") (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) lemma card_cartesian_product_singleton: "card ({x} \ A) = card A" by (simp add: card_cartesian_product) subsection \Generalized product over a set\ context comm_monoid_mult begin sublocale prod: comm_monoid_set times 1 defines prod = prod.F and prod' = prod.G .. abbreviation Prod ("\_" [1000] 999) where "\A \ prod (\x. x) A" end syntax (ASCII) "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(4PROD (_/:_)./ _)" [0, 51, 10] 10) syntax "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(2\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\i\A. b" == "CONST prod (\i. b) A" text \Instead of \<^term>\\x\{x. P}. e\ we introduce the shorter \\x|P. e\.\ syntax (ASCII) "_qprod" :: "pttrn \ bool \ 'a \ 'a" ("(4PROD _ |/ _./ _)" [0, 0, 10] 10) syntax "_qprod" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations "\x|P. t" => "CONST prod (\x. t) {x. P}" context comm_monoid_mult begin lemma prod_dvd_prod: "(\a. a \ A \ f a dvd g a) \ prod f A dvd prod g A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by (auto intro: dvdI) next case empty then show ?case by (auto intro: dvdI) next case (insert a A) then have "f a dvd g a" and "prod f A dvd prod g A" by simp_all then obtain r s where "g a = f a * r" and "prod g A = prod f A * s" by (auto elim!: dvdE) then have "g a * prod g A = f a * prod f A * (r * s)" by (simp add: ac_simps) with insert.hyps show ?case by (auto intro: dvdI) qed lemma prod_dvd_prod_subset: "finite B \ A \ B \ prod f A dvd prod f B" by (auto simp add: prod.subset_diff ac_simps intro: dvdI) end subsubsection \Properties in more restricted classes of structures\ context linordered_nonzero_semiring begin lemma prod_ge_1: "(\x. x \ A \ 1 \ f x) \ 1 \ prod f A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) have "1 * 1 \ f x * prod f F" by (rule mult_mono') (use insert in auto) with insert show ?case by simp qed lemma prod_le_1: fixes f :: "'b \ 'a" assumes "\x. x \ A \ 0 \ f x \ f x \ 1" shows "prod f A \ 1" using assms proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then show ?case by (force simp: mult.commute intro: dest: mult_le_one) qed end context comm_semiring_1 begin lemma dvd_prod_eqI [intro]: assumes "finite A" and "a \ A" and "b = f a" shows "b dvd prod f A" proof - from \finite A\ have "prod f (insert a (A - {a})) = f a * prod f (A - {a})" by (intro prod.insert) auto also from \a \ A\ have "insert a (A - {a}) = A" by blast finally have "prod f A = f a * prod f (A - {a})" . with \b = f a\ show ?thesis by simp qed lemma dvd_prodI [intro]: "finite A \ a \ A \ f a dvd prod f A" by auto lemma prod_zero: assumes "finite A" and "\a\A. f a = 0" shows "prod f A = 0" using assms proof (induct A) case empty then show ?case by simp next case (insert a A) then have "f a = 0 \ (\a\A. f a = 0)" by simp - then have "f a * prod f A = 0" by rule (simp_all add: insert) + then have "f a * prod f A = 0" by (rule disjE) (simp_all add: insert) with insert show ?case by simp qed lemma prod_dvd_prod_subset2: assumes "finite B" and "A \ B" and "\a. a \ A \ f a dvd g a" shows "prod f A dvd prod g B" proof - from assms have "prod f A dvd prod g A" by (auto intro: prod_dvd_prod) moreover from assms have "prod g A dvd prod g B" by (auto intro: prod_dvd_prod_subset) ultimately show ?thesis by (rule dvd_trans) qed end lemma (in semidom) prod_zero_iff [simp]: fixes f :: "'b \ 'a" assumes "finite A" shows "prod f A = 0 \ (\a\A. f a = 0)" using assms by (induct A) (auto simp: no_zero_divisors) lemma (in semidom_divide) prod_diff1: assumes "finite A" and "f a \ 0" shows "prod f (A - {a}) = (if a \ A then prod f A div f a else prod f A)" proof (cases "a \ A") case True then show ?thesis by simp next case False with assms show ?thesis proof induct case empty then show ?case by simp next case (insert b B) then show ?case proof (cases "a = b") case True with insert show ?thesis by simp next case False with insert have "a \ B" by simp define C where "C = B - {a}" with \finite B\ \a \ B\ have "B = insert a C" "finite C" "a \ C" by auto with insert show ?thesis by (auto simp add: insert_commute ac_simps) qed qed qed lemma sum_zero_power [simp]: "(\i\A. c i * 0^i) = (if finite A \ 0 \ A then c 0 else 0)" for c :: "nat \ 'a::division_ring" by (induct A rule: infinite_finite_induct) auto lemma sum_zero_power' [simp]: "(\i\A. c i * 0^i / d i) = (if finite A \ 0 \ A then c 0 / d 0 else 0)" for c :: "nat \ 'a::field" using sum_zero_power [of "\i. c i / d i" A] by auto lemma (in field) prod_inversef: "prod (inverse \ f) A = inverse (prod f A)" proof (cases "finite A") case True then show ?thesis by (induct A rule: finite_induct) simp_all next case False then show ?thesis by auto qed lemma (in field) prod_dividef: "(\x\A. f x / g x) = prod f A / prod g A" using prod_inversef [of g A] by (simp add: divide_inverse prod.distrib) lemma prod_Un: fixes f :: "'b \ 'a :: field" assumes "finite A" and "finite B" and "\x\A \ B. f x \ 0" shows "prod f (A \ B) = prod f A * prod f B / prod f (A \ B)" proof - from assms have "prod f A * prod f B = prod f (A \ B) * prod f (A \ B)" by (simp add: prod.union_inter [symmetric, of A B]) with assms show ?thesis by simp qed context linordered_semidom begin lemma prod_nonneg: "(\a\A. 0 \ f a) \ 0 \ prod f A" by (induct A rule: infinite_finite_induct) simp_all lemma prod_pos: "(\a\A. 0 < f a) \ 0 < prod f A" by (induct A rule: infinite_finite_induct) simp_all lemma prod_mono: "(\i. i \ A \ 0 \ f i \ f i \ g i) \ prod f A \ prod g A" by (induct A rule: infinite_finite_induct) (force intro!: prod_nonneg mult_mono)+ lemma prod_mono_strict: assumes "finite A" "\i. i \ A \ 0 \ f i \ f i < g i" "A \ {}" shows "prod f A < prod g A" using assms proof (induct A rule: finite_induct) case empty then show ?case by simp next case insert then show ?case by (force intro: mult_strict_mono' prod_nonneg) qed lemma prod_le_power: assumes A: "\i. i \ A \ 0 \ f i \ f i \ n" "card A \ k" and "n \ 1" shows "prod f A \ n ^ k" using A proof (induction A arbitrary: k rule: infinite_finite_induct) case (insert i A) then obtain k' where k': "card A \ k'" "k = Suc k'" using Suc_le_D by force have "f i * prod f A \ n * n ^ k'" using insert \n \ 1\ k' by (intro prod_nonneg mult_mono; force) then show ?case by (auto simp: \k = Suc k'\ insert.hyps) qed (use \n \ 1\ in auto) end lemma prod_mono2: fixes f :: "'a \ 'b :: linordered_idom" assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 1 \ f b" and A: "\a. a \ A \ 0 \ f a" shows "prod f A \ prod f B" proof - have "prod f A \ prod f A * prod f (B-A)" by (metis prod_ge_1 A mult_le_cancel_left1 nn not_less prod_nonneg) also from fin finite_subset[OF sub fin] have "\ = prod f (A \ (B-A))" by (simp add: prod.union_disjoint del: Un_Diff_cancel) also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed lemma less_1_prod: fixes f :: "'a \ 'b::linordered_idom" shows "finite I \ I \ {} \ (\i. i \ I \ 1 < f i) \ 1 < prod f I" by (induct I rule: finite_ne_induct) (auto intro: less_1_mult) lemma less_1_prod2: fixes f :: "'a \ 'b::linordered_idom" assumes I: "finite I" "i \ I" "1 < f i" "\i. i \ I \ 1 \ f i" shows "1 < prod f I" proof - have "1 < f i * prod f (I - {i})" using assms by (meson DiffD1 leI less_1_mult less_le_trans mult_le_cancel_left1 prod_ge_1) also have "\ = prod f I" using assms by (simp add: prod.remove) finally show ?thesis . qed lemma (in linordered_field) abs_prod: "\prod f A\ = (\x\A. \f x\)" by (induct A rule: infinite_finite_induct) (simp_all add: abs_mult) lemma prod_eq_1_iff [simp]: "finite A \ prod f A = 1 \ (\a\A. f a = 1)" for f :: "'a \ nat" by (induct A rule: finite_induct) simp_all lemma prod_pos_nat_iff [simp]: "finite A \ prod f A > 0 \ (\a\A. f a > 0)" for f :: "'a \ nat" using prod_zero_iff by (simp del: neq0_conv add: zero_less_iff_neq_zero) lemma prod_constant [simp]: "(\x\ A. y) = y ^ card A" for y :: "'a::comm_monoid_mult" by (induct A rule: infinite_finite_induct) simp_all lemma prod_power_distrib: "prod f A ^ n = prod (\x. (f x) ^ n) A" for f :: "'a \ 'b::comm_semiring_1" by (induct A rule: infinite_finite_induct) (auto simp add: power_mult_distrib) lemma power_sum: "c ^ (\a\A. f a) = (\a\A. c ^ f a)" by (induct A rule: infinite_finite_induct) (simp_all add: power_add) lemma prod_gen_delta: fixes b :: "'b \ 'a::comm_monoid_mult" assumes fin: "finite S" shows "prod (\k. if k = a then b k else c) S = (if a \ S then b a * c ^ (card S - 1) else c ^ card S)" proof - let ?f = "(\k. if k=a then b k else c)" show ?thesis proof (cases "a \ S") case False then have "\ k\ S. ?f k = c" by simp with False show ?thesis by (simp add: prod_constant) next case True let ?A = "S - {a}" let ?B = "{a}" from True have eq: "S = ?A \ ?B" by blast have disjoint: "?A \ ?B = {}" by simp from fin have fin': "finite ?A" "finite ?B" by auto have f_A0: "prod ?f ?A = prod (\i. c) ?A" by (rule prod.cong) auto from fin True have card_A: "card ?A = card S - 1" by auto have f_A1: "prod ?f ?A = c ^ card ?A" unfolding f_A0 by (rule prod_constant) have "prod ?f ?A * prod ?f ?B = prod ?f S" using prod.union_disjoint[OF fin' disjoint, of ?f, unfolded eq[symmetric]] by simp with True card_A show ?thesis by (simp add: f_A1 field_simps cong add: prod.cong cong del: if_weak_cong) qed qed lemma sum_image_le: fixes g :: "'a \ 'b::ordered_comm_monoid_add" assumes "finite I" "\i. i \ I \ 0 \ g(f i)" shows "sum g (f ` I) \ sum (g \ f) I" using assms proof induction case empty then show ?case by auto next case (insert x F) from insertI1 have "0 \ g (f x)" by (rule insert) hence 1: "sum g (f ` F) \ g (f x) + sum g (f ` F)" using add_increasing by blast have 2: "sum g (f ` F) \ sum (g \ f) F" using insert by blast have "sum g (f ` insert x F) = sum g (insert (f x) (f ` F))" by simp also have "\ \ g (f x) + sum g (f ` F)" by (simp add: 1 insert sum.insert_if) also from 2 have "\ \ g (f x) + sum (g \ f) F" by (rule add_left_mono) also from insert(1, 2) have "\ = sum (g \ f) (insert x F)" by (simp add: sum.insert_if) finally show ?case . qed end diff --git a/src/HOL/HOL.thy b/src/HOL/HOL.thy --- a/src/HOL/HOL.thy +++ b/src/HOL/HOL.thy @@ -1,2177 +1,2177 @@ (* Title: HOL/HOL.thy Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson *) section \The basis of Higher-Order Logic\ theory HOL imports Pure Tools.Code_Generator keywords "try" "solve_direct" "quickcheck" "print_coercions" "print_claset" "print_induct_rules" :: diag and "quickcheck_params" :: thy_decl abbrevs "?<" = "\\<^sub>\\<^sub>1" begin ML_file \~~/src/Tools/misc_legacy.ML\ ML_file \~~/src/Tools/try.ML\ ML_file \~~/src/Tools/quickcheck.ML\ ML_file \~~/src/Tools/solve_direct.ML\ ML_file \~~/src/Tools/IsaPlanner/zipper.ML\ ML_file \~~/src/Tools/IsaPlanner/isand.ML\ ML_file \~~/src/Tools/IsaPlanner/rw_inst.ML\ ML_file \~~/src/Provers/hypsubst.ML\ ML_file \~~/src/Provers/splitter.ML\ ML_file \~~/src/Provers/classical.ML\ ML_file \~~/src/Provers/blast.ML\ ML_file \~~/src/Provers/clasimp.ML\ ML_file \~~/src/Tools/eqsubst.ML\ ML_file \~~/src/Provers/quantifier1.ML\ ML_file \~~/src/Tools/atomize_elim.ML\ ML_file \~~/src/Tools/cong_tac.ML\ ML_file \~~/src/Tools/intuitionistic.ML\ setup \Intuitionistic.method_setup \<^binding>\iprover\\ ML_file \~~/src/Tools/project_rule.ML\ ML_file \~~/src/Tools/subtyping.ML\ ML_file \~~/src/Tools/case_product.ML\ ML \Plugin_Name.declare_setup \<^binding>\extraction\\ ML \ Plugin_Name.declare_setup \<^binding>\quickcheck_random\; Plugin_Name.declare_setup \<^binding>\quickcheck_exhaustive\; Plugin_Name.declare_setup \<^binding>\quickcheck_bounded_forall\; Plugin_Name.declare_setup \<^binding>\quickcheck_full_exhaustive\; Plugin_Name.declare_setup \<^binding>\quickcheck_narrowing\; \ ML \ Plugin_Name.define_setup \<^binding>\quickcheck\ [\<^plugin>\quickcheck_exhaustive\, \<^plugin>\quickcheck_random\, \<^plugin>\quickcheck_bounded_forall\, \<^plugin>\quickcheck_full_exhaustive\, \<^plugin>\quickcheck_narrowing\] \ subsection \Primitive logic\ text \ The definition of the logic is based on Mike Gordon's technical report @{cite "Gordon-TR68"} that describes the first implementation of HOL. However, there are a number of differences. In particular, we start with the definite description operator and introduce Hilbert's \\\ operator only much later. Moreover, axiom \(P \ Q) \ (Q \ P) \ (P = Q)\ is derived from the other axioms. The fact that this axiom is derivable was first noticed by Bruno Barras (for Mike Gordon's line of HOL systems) and later independently by Alexander Maletzky (for Isabelle/HOL). \ subsubsection \Core syntax\ setup \Axclass.class_axiomatization (\<^binding>\type\, [])\ default_sort type setup \Object_Logic.add_base_sort \<^sort>\type\\ setup \Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\ axiomatization where fun_arity: "OFCLASS('a \ 'b, type_class)" instance "fun" :: (type, type) type by (rule fun_arity) axiomatization where itself_arity: "OFCLASS('a itself, type_class)" instance itself :: (type) type by (rule itself_arity) typedecl bool judgment Trueprop :: "bool \ prop" ("(_)" 5) axiomatization implies :: "[bool, bool] \ bool" (infixr "\" 25) and eq :: "['a, 'a] \ bool" and The :: "('a \ bool) \ 'a" notation (input) eq (infixl "=" 50) notation (output) eq (infix "=" 50) text \The input syntax for \eq\ is more permissive than the output syntax because of the large amount of material that relies on infixl.\ subsubsection \Defined connectives and quantifiers\ definition True :: bool where "True \ ((\x::bool. x) = (\x. x))" definition All :: "('a \ bool) \ bool" (binder "\" 10) where "All P \ (P = (\x. True))" definition Ex :: "('a \ bool) \ bool" (binder "\" 10) where "Ex P \ \Q. (\x. P x \ Q) \ Q" definition False :: bool where "False \ (\P. P)" definition Not :: "bool \ bool" ("\ _" [40] 40) where not_def: "\ P \ P \ False" definition conj :: "[bool, bool] \ bool" (infixr "\" 35) where and_def: "P \ Q \ \R. (P \ Q \ R) \ R" definition disj :: "[bool, bool] \ bool" (infixr "\" 30) where or_def: "P \ Q \ \R. (P \ R) \ (Q \ R) \ R" definition Uniq :: "('a \ bool) \ bool" where "Uniq P \ (\x y. P x \ P y \ y = x)" definition Ex1 :: "('a \ bool) \ bool" where "Ex1 P \ \x. P x \ (\y. P y \ y = x)" subsubsection \Additional concrete syntax\ syntax (ASCII) "_Uniq" :: "pttrn \ bool \ bool" ("(4?< _./ _)" [0, 10] 10) syntax "_Uniq" :: "pttrn \ bool \ bool" ("(2\\<^sub>\\<^sub>1 _./ _)" [0, 10] 10) translations "\\<^sub>\\<^sub>1x. P" \ "CONST Uniq (\x. P)" print_translation \ [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\Uniq\ \<^syntax_const>\_Uniq\] \ \ \to avoid eta-contraction of body\ syntax (ASCII) "_Ex1" :: "pttrn \ bool \ bool" ("(3EX! _./ _)" [0, 10] 10) syntax (input) "_Ex1" :: "pttrn \ bool \ bool" ("(3?! _./ _)" [0, 10] 10) syntax "_Ex1" :: "pttrn \ bool \ bool" ("(3\!_./ _)" [0, 10] 10) translations "\!x. P" \ "CONST Ex1 (\x. P)" print_translation \ [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\Ex1\ \<^syntax_const>\_Ex1\] \ \ \to avoid eta-contraction of body\ syntax "_Not_Ex" :: "idts \ bool \ bool" ("(3\_./ _)" [0, 10] 10) "_Not_Ex1" :: "pttrn \ bool \ bool" ("(3\!_./ _)" [0, 10] 10) translations "\x. P" \ "\ (\x. P)" "\!x. P" \ "\ (\!x. P)" abbreviation not_equal :: "['a, 'a] \ bool" (infix "\" 50) where "x \ y \ \ (x = y)" notation (ASCII) Not ("~ _" [40] 40) and conj (infixr "&" 35) and disj (infixr "|" 30) and implies (infixr "-->" 25) and not_equal (infix "~=" 50) abbreviation (iff) iff :: "[bool, bool] \ bool" (infixr "\" 25) where "A \ B \ A = B" syntax "_The" :: "[pttrn, bool] \ 'a" ("(3THE _./ _)" [0, 10] 10) translations "THE x. P" \ "CONST The (\x. P)" print_translation \ [(\<^const_syntax>\The\, fn _ => fn [Abs abs] => let val (x, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_The\ $ x $ t end)] \ \ \To avoid eta-contraction of body\ nonterminal letbinds and letbind syntax "_bind" :: "[pttrn, 'a] \ letbind" ("(2_ =/ _)" 10) "" :: "letbind \ letbinds" ("_") "_binds" :: "[letbind, letbinds] \ letbinds" ("_;/ _") "_Let" :: "[letbinds, 'a] \ 'a" ("(let (_)/ in (_))" [0, 10] 10) nonterminal case_syn and cases_syn syntax "_case_syntax" :: "['a, cases_syn] \ 'b" ("(case _ of/ _)" 10) "_case1" :: "['a, 'b] \ case_syn" ("(2_ \/ _)" 10) "" :: "case_syn \ cases_syn" ("_") "_case2" :: "[case_syn, cases_syn] \ cases_syn" ("_/ | _") syntax (ASCII) "_case1" :: "['a, 'b] \ case_syn" ("(2_ =>/ _)" 10) notation (ASCII) All (binder "ALL " 10) and Ex (binder "EX " 10) notation (input) All (binder "! " 10) and Ex (binder "? " 10) subsubsection \Axioms and basic definitions\ axiomatization where refl: "t = (t::'a)" and subst: "s = t \ P s \ P t" and ext: "(\x::'a. (f x ::'b) = g x) \ (\x. f x) = (\x. g x)" \ \Extensionality is built into the meta-logic, and this rule expresses a related property. It is an eta-expanded version of the traditional rule, and similar to the ABS rule of HOL\ and the_eq_trivial: "(THE x. x = a) = (a::'a)" axiomatization where impI: "(P \ Q) \ P \ Q" and mp: "\P \ Q; P\ \ Q" and True_or_False: "(P = True) \ (P = False)" definition If :: "bool \ 'a \ 'a \ 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where "If P x y \ (THE z::'a. (P = True \ z = x) \ (P = False \ z = y))" definition Let :: "'a \ ('a \ 'b) \ 'b" where "Let s f \ f s" translations "_Let (_binds b bs) e" \ "_Let b (_Let bs e)" "let x = a in e" \ "CONST Let a (\x. e)" axiomatization undefined :: 'a class default = fixes default :: 'a subsection \Fundamental rules\ subsubsection \Equality\ lemma sym: "s = t \ t = s" by (erule subst) (rule refl) lemma ssubst: "t = s \ P s \ P t" by (drule sym) (erule subst) lemma trans: "\r = s; s = t\ \ r = t" by (erule subst) lemma trans_sym [Pure.elim?]: "r = s \ t = s \ r = t" by (rule trans [OF _ sym]) lemma meta_eq_to_obj_eq: assumes "A \ B" shows "A = B" unfolding assms by (rule refl) text \Useful with \erule\ for proving equalities from known equalities.\ (* a = b | | c = d *) lemma box_equals: "\a = b; a = c; b = d\ \ c = d" by (iprover intro: sym trans) text \For calculational reasoning:\ lemma forw_subst: "a = b \ P b \ P a" by (rule ssubst) lemma back_subst: "P a \ a = b \ P b" by (rule subst) subsubsection \Congruence rules for application\ text \Similar to \AP_THM\ in Gordon's HOL.\ lemma fun_cong: "(f :: 'a \ 'b) = g \ f x = g x" by (iprover intro: refl elim: subst) text \Similar to \AP_TERM\ in Gordon's HOL and FOL's \subst_context\.\ lemma arg_cong: "x = y \ f x = f y" by (iprover intro: refl elim: subst) lemma arg_cong2: "\a = b; c = d\ \ f a c = f b d" by (iprover intro: refl elim: subst) lemma cong: "\f = g; (x::'a) = y\ \ f x = g y" by (iprover intro: refl elim: subst) ML \fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\ subsubsection \Equality of booleans -- iff\ lemma iffD2: "\P = Q; Q\ \ P" by (erule ssubst) lemma rev_iffD2: "\Q; P = Q\ \ P" by (erule iffD2) lemma iffD1: "Q = P \ Q \ P" by (drule sym) (rule iffD2) lemma rev_iffD1: "Q \ Q = P \ P" by (drule sym) (rule rev_iffD2) lemma iffE: assumes major: "P = Q" and minor: "\P \ Q; Q \ P\ \ R" shows R by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1]) subsubsection \True (1)\ lemma TrueI: True unfolding True_def by (rule refl) lemma eqTrueE: "P = True \ P" by (erule iffD2) (rule TrueI) subsubsection \Universal quantifier (1)\ lemma spec: "\x::'a. P x \ P x" unfolding All_def by (iprover intro: eqTrueE fun_cong) lemma allE: assumes major: "\x. P x" and minor: "P x \ R" shows R by (iprover intro: minor major [THEN spec]) lemma all_dupE: assumes major: "\x. P x" and minor: "\P x; \x. P x\ \ R" shows R by (iprover intro: minor major major [THEN spec]) subsubsection \False\ text \ Depends upon \spec\; it is impossible to do propositional logic before quantifiers! \ lemma FalseE: "False \ P" unfolding False_def by (erule spec) lemma False_neq_True: "False = True \ P" by (erule eqTrueE [THEN FalseE]) subsubsection \Negation\ lemma notI: assumes "P \ False" shows "\ P" unfolding not_def by (iprover intro: impI assms) lemma False_not_True: "False \ True" by (iprover intro: notI elim: False_neq_True) lemma True_not_False: "True \ False" by (iprover intro: notI dest: sym elim: False_neq_True) lemma notE: "\\ P; P\ \ R" unfolding not_def by (iprover intro: mp [THEN FalseE]) subsubsection \Implication\ lemma impE: assumes "P \ Q" P "Q \ R" shows R by (iprover intro: assms mp) text \Reduces \Q\ to \P \ Q\, allowing substitution in \P\.\ lemma rev_mp: "\P; P \ Q\ \ Q" by (rule mp) lemma contrapos_nn: assumes major: "\ Q" and minor: "P \ Q" shows "\ P" by (iprover intro: notI minor major [THEN notE]) text \Not used at all, but we already have the other 3 combinations.\ lemma contrapos_pn: assumes major: "Q" and minor: "P \ \ Q" shows "\ P" by (iprover intro: notI minor major notE) lemma not_sym: "t \ s \ s \ t" by (erule contrapos_nn) (erule sym) lemma eq_neq_eq_imp_neq: "\x = a; a \ b; b = y\ \ x \ y" by (erule subst, erule ssubst, assumption) subsubsection \Disjunction (1)\ lemma disjE: assumes major: "P \ Q" and minorP: "P \ R" and minorQ: "Q \ R" shows R by (iprover intro: minorP minorQ impI major [unfolded or_def, THEN spec, THEN mp, THEN mp]) subsubsection \Derivation of \iffI\\ text \In an intuitionistic version of HOL \iffI\ needs to be an axiom.\ lemma iffI: assumes "P \ Q" and "Q \ P" shows "P = Q" proof (rule disjE[OF True_or_False[of P]]) assume 1: "P = True" note Q = assms(1)[OF eqTrueE[OF this]] from 1 show ?thesis proof (rule ssubst) from True_or_False[of Q] show "True = Q" proof (rule disjE) assume "Q = True" thus ?thesis by(rule sym) next assume "Q = False" with Q have False by (rule rev_iffD1) thus ?thesis by (rule FalseE) qed qed next assume 2: "P = False" thus ?thesis proof (rule ssubst) from True_or_False[of Q] show "False = Q" proof (rule disjE) assume "Q = True" from 2 assms(2)[OF eqTrueE[OF this]] have False by (rule iffD1) thus ?thesis by (rule FalseE) next assume "Q = False" thus ?thesis by(rule sym) qed qed qed subsubsection \True (2)\ lemma eqTrueI: "P \ P = True" by (iprover intro: iffI TrueI) subsubsection \Universal quantifier (2)\ lemma allI: assumes "\x::'a. P x" shows "\x. P x" unfolding All_def by (iprover intro: ext eqTrueI assms) subsubsection \Existential quantifier\ lemma exI: "P x \ \x::'a. P x" unfolding Ex_def by (iprover intro: allI allE impI mp) lemma exE: assumes major: "\x::'a. P x" and minor: "\x. P x \ Q" shows "Q" by (rule major [unfolded Ex_def, THEN spec, THEN mp]) (iprover intro: impI [THEN allI] minor) subsubsection \Conjunction\ lemma conjI: "\P; Q\ \ P \ Q" unfolding and_def by (iprover intro: impI [THEN allI] mp) lemma conjunct1: "\P \ Q\ \ P" unfolding and_def by (iprover intro: impI dest: spec mp) lemma conjunct2: "\P \ Q\ \ Q" unfolding and_def by (iprover intro: impI dest: spec mp) lemma conjE: assumes major: "P \ Q" and minor: "\P; Q\ \ R" shows R proof (rule minor) show P by (rule major [THEN conjunct1]) show Q by (rule major [THEN conjunct2]) qed lemma context_conjI: assumes P "P \ Q" shows "P \ Q" by (iprover intro: conjI assms) subsubsection \Disjunction (2)\ lemma disjI1: "P \ P \ Q" unfolding or_def by (iprover intro: allI impI mp) lemma disjI2: "Q \ P \ Q" unfolding or_def by (iprover intro: allI impI mp) subsubsection \Classical logic\ lemma classical: assumes "\ P \ P" shows P proof (rule True_or_False [THEN disjE]) show P if "P = True" using that by (iprover intro: eqTrueE) show P if "P = False" proof (intro notI assms) assume P with that show False by (iprover elim: subst) qed qed lemmas ccontr = FalseE [THEN classical] text \\notE\ with premises exchanged; it discharges \\ R\ so that it can be used to make elimination rules.\ lemma rev_notE: assumes premp: P and premnot: "\ R \ \ P" shows R by (iprover intro: ccontr notE [OF premnot premp]) text \Double negation law.\ lemma notnotD: "\\ P \ P" by (iprover intro: ccontr notE ) lemma contrapos_pp: assumes p1: Q and p2: "\ P \ \ Q" shows P by (iprover intro: classical p1 p2 notE) subsubsection \Unique existence\ lemma Uniq_I [intro?]: assumes "\x y. \P x; P y\ \ y = x" shows "Uniq P" unfolding Uniq_def by (iprover intro: assms allI impI) lemma Uniq_D [dest?]: "\Uniq P; P a; P b\ \ a=b" unfolding Uniq_def by (iprover dest: spec mp) lemma ex1I: assumes "P a" "\x. P x \ x = a" shows "\!x. P x" unfolding Ex1_def by (iprover intro: assms exI conjI allI impI) text \Sometimes easier to use: the premises have no shared variables. Safe!\ lemma ex_ex1I: assumes ex_prem: "\x. P x" and eq: "\x y. \P x; P y\ \ x = y" shows "\!x. P x" by (iprover intro: ex_prem [THEN exE] ex1I eq) lemma ex1E: assumes major: "\!x. P x" and minor: "\x. \P x; \y. P y \ y = x\ \ R" shows R proof (rule major [unfolded Ex1_def, THEN exE]) show "\x. P x \ (\y. P y \ y = x) \ R" by (iprover intro: minor elim: conjE) qed lemma ex1_implies_ex: "\!x. P x \ \x. P x" by (iprover intro: exI elim: ex1E) subsubsection \Classical intro rules for disjunction and existential quantifiers\ lemma disjCI: assumes "\ Q \ P" shows "P \ Q" by (rule classical) (iprover intro: assms disjI1 disjI2 notI elim: notE) lemma excluded_middle: "\ P \ P" by (iprover intro: disjCI) text \ case distinction as a natural deduction rule. Note that \\ P\ is the second case, not the first. \ lemma case_split [case_names True False]: assumes "P \ Q" "\ P \ Q" shows Q using excluded_middle [of P] by (iprover intro: assms elim: disjE) text \Classical implies (\\\) elimination.\ lemma impCE: assumes major: "P \ Q" and minor: "\ P \ R" "Q \ R" shows R using excluded_middle [of P] by (iprover intro: minor major [THEN mp] elim: disjE)+ text \ This version of \\\ elimination works on \Q\ before \P\. It works best for those cases in which \P\ holds "almost everywhere". Can't install as default: would break old proofs. \ lemma impCE': assumes major: "P \ Q" and minor: "Q \ R" "\ P \ R" shows R using assms by (elim impCE) text \Classical \\\ elimination.\ lemma iffCE: assumes major: "P = Q" and minor: "\P; Q\ \ R" "\\ P; \ Q\ \ R" shows R by (rule major [THEN iffE]) (iprover intro: minor elim: impCE notE) lemma exCI: assumes "\x. \ P x \ P a" shows "\x. P x" by (rule ccontr) (iprover intro: assms exI allI notI notE [of "\x. P x"]) subsubsection \Intuitionistic Reasoning\ lemma impE': assumes 1: "P \ Q" and 2: "Q \ R" and 3: "P \ Q \ P" shows R proof - from 3 and 1 have P . with 1 have Q by (rule impE) with 2 show R . qed lemma allE': assumes 1: "\x. P x" and 2: "P x \ \x. P x \ Q" shows Q proof - from 1 have "P x" by (rule spec) from this and 1 show Q by (rule 2) qed lemma notE': assumes 1: "\ P" and 2: "\ P \ P" shows R proof - from 2 and 1 have P . with 1 show R by (rule notE) qed lemma TrueE: "True \ P \ P" . lemma notFalseE: "\ False \ P \ P" . lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE and [Pure.intro!] = iffI conjI impI TrueI notI allI refl and [Pure.elim 2] = allE notE' impE' and [Pure.intro] = exI disjI2 disjI1 lemmas [trans] = trans and [sym] = sym not_sym and [Pure.elim?] = iffD1 iffD2 impE subsubsection \Atomizing meta-level connectives\ axiomatization where eq_reflection: "x = y \ x \ y" \ \admissible axiom\ lemma atomize_all [atomize]: "(\x. P x) \ Trueprop (\x. P x)" proof assume "\x. P x" then show "\x. P x" .. next assume "\x. P x" then show "\x. P x" by (rule allE) qed lemma atomize_imp [atomize]: "(A \ B) \ Trueprop (A \ B)" proof assume r: "A \ B" show "A \ B" by (rule impI) (rule r) next assume "A \ B" and A then show B by (rule mp) qed lemma atomize_not: "(A \ False) \ Trueprop (\ A)" proof assume r: "A \ False" show "\ A" by (rule notI) (rule r) next assume "\ A" and A then show False by (rule notE) qed lemma atomize_eq [atomize, code]: "(x \ y) \ Trueprop (x = y)" proof assume "x \ y" show "x = y" by (unfold \x \ y\) (rule refl) next assume "x = y" then show "x \ y" by (rule eq_reflection) qed lemma atomize_conj [atomize]: "(A &&& B) \ Trueprop (A \ B)" proof assume conj: "A &&& B" show "A \ B" proof (rule conjI) from conj show A by (rule conjunctionD1) from conj show B by (rule conjunctionD2) qed next assume conj: "A \ B" show "A &&& B" proof - from conj show A .. from conj show B .. qed qed lemmas [symmetric, rulify] = atomize_all atomize_imp and [symmetric, defn] = atomize_all atomize_imp atomize_eq subsubsection \Atomizing elimination rules\ lemma atomize_exL[atomize_elim]: "(\x. P x \ Q) \ ((\x. P x) \ Q)" - by rule iprover+ + by (rule equal_intr_rule) iprover+ lemma atomize_conjL[atomize_elim]: "(A \ B \ C) \ (A \ B \ C)" - by rule iprover+ + by (rule equal_intr_rule) iprover+ lemma atomize_disjL[atomize_elim]: "((A \ C) \ (B \ C) \ C) \ ((A \ B \ C) \ C)" - by rule iprover+ + by (rule equal_intr_rule) iprover+ lemma atomize_elimL[atomize_elim]: "(\B. (A \ B) \ B) \ Trueprop A" .. subsection \Package setup\ ML_file \Tools/hologic.ML\ ML_file \Tools/rewrite_hol_proof.ML\ setup \Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc Rewrite_HOL_Proof.rews)\ subsubsection \Sledgehammer setup\ text \ Theorems blacklisted to Sledgehammer. These theorems typically produce clauses that are prolific (match too many equality or membership literals) and relate to seldom-used facts. Some duplicate other rules. \ named_theorems no_atp "theorems that should be filtered out by Sledgehammer" subsubsection \Classical Reasoner setup\ lemma imp_elim: "P \ Q \ (\ R \ P) \ (Q \ R) \ R" by (rule classical) iprover lemma swap: "\ P \ (\ R \ P) \ R" by (rule classical) iprover lemma thin_refl: "\x = x; PROP W\ \ PROP W" . ML \ structure Hypsubst = Hypsubst ( val dest_eq = HOLogic.dest_eq val dest_Trueprop = HOLogic.dest_Trueprop val dest_imp = HOLogic.dest_imp val eq_reflection = @{thm eq_reflection} val rev_eq_reflection = @{thm meta_eq_to_obj_eq} val imp_intr = @{thm impI} val rev_mp = @{thm rev_mp} val subst = @{thm subst} val sym = @{thm sym} val thin_refl = @{thm thin_refl}; ); open Hypsubst; structure Classical = Classical ( val imp_elim = @{thm imp_elim} val not_elim = @{thm notE} val swap = @{thm swap} val classical = @{thm classical} val sizef = Drule.size_of_thm val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] ); structure Basic_Classical: BASIC_CLASSICAL = Classical; open Basic_Classical; \ setup \ (*prevent substitution on bool*) let fun non_bool_eq (\<^const_name>\HOL.eq\, Type (_, [T, _])) = T <> \<^typ>\bool\ | non_bool_eq _ = false; fun hyp_subst_tac' ctxt = SUBGOAL (fn (goal, i) => if Term.exists_Const non_bool_eq goal then Hypsubst.hyp_subst_tac ctxt i else no_tac); in Context_Rules.addSWrapper (fn ctxt => fn tac => hyp_subst_tac' ctxt ORELSE' tac) end \ declare iffI [intro!] and notI [intro!] and impI [intro!] and disjCI [intro!] and conjI [intro!] and TrueI [intro!] and refl [intro!] declare iffCE [elim!] and FalseE [elim!] and impCE [elim!] and disjE [elim!] and conjE [elim!] declare ex_ex1I [intro!] and allI [intro!] and exI [intro] declare exE [elim!] allE [elim] ML \val HOL_cs = claset_of \<^context>\ lemma contrapos_np: "\ Q \ (\ P \ Q) \ P" by (erule swap) declare ex_ex1I [rule del, intro! 2] and ex1I [intro] declare ext [intro] lemmas [intro?] = ext and [elim?] = ex1_implies_ex text \Better than \ex1E\ for classical reasoner: needs no quantifier duplication!\ lemma alt_ex1E [elim!]: assumes major: "\!x. P x" and minor: "\x. \P x; \y y'. P y \ P y' \ y = y'\ \ R" shows R proof (rule ex1E [OF major minor]) show "\y y'. P y \ P y' \ y = y'" if "P x" and \
: "\y. P y \ y = x" for x using \P x\ \
\
by fast qed assumption text \And again using Uniq\ lemma alt_ex1E': assumes "\!x. P x" "\x. \P x; \\<^sub>\\<^sub>1x. P x\ \ R" shows R using assms unfolding Uniq_def by fast lemma ex1_iff_ex_Uniq: "(\!x. P x) \ (\x. P x) \ (\\<^sub>\\<^sub>1x. P x)" unfolding Uniq_def by fast ML \ structure Blast = Blast ( structure Classical = Classical val Trueprop_const = dest_Const \<^Const>\Trueprop\ val equality_name = \<^const_name>\HOL.eq\ val not_name = \<^const_name>\Not\ val notE = @{thm notE} val ccontr = @{thm ccontr} val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac ); val blast_tac = Blast.blast_tac; \ subsubsection \THE: definite description operator\ lemma the_equality [intro]: assumes "P a" and "\x. P x \ x = a" shows "(THE x. P x) = a" by (blast intro: assms trans [OF arg_cong [where f=The] the_eq_trivial]) lemma theI: assumes "P a" and "\x. P x \ x = a" shows "P (THE x. P x)" by (iprover intro: assms the_equality [THEN ssubst]) lemma theI': "\!x. P x \ P (THE x. P x)" by (blast intro: theI) text \Easier to apply than \theI\: only one occurrence of \P\.\ lemma theI2: assumes "P a" "\x. P x \ x = a" "\x. P x \ Q x" shows "Q (THE x. P x)" by (iprover intro: assms theI) lemma the1I2: assumes "\!x. P x" "\x. P x \ Q x" shows "Q (THE x. P x)" by (iprover intro: assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)] elim: allE impE) lemma the1_equality [elim?]: "\\!x. P x; P a\ \ (THE x. P x) = a" by blast lemma the1_equality': "\\\<^sub>\\<^sub>1x. P x; P a\ \ (THE x. P x) = a" unfolding Uniq_def by blast lemma the_sym_eq_trivial: "(THE y. x = y) = x" by blast subsubsection \Simplifier\ lemma eta_contract_eq: "(\s. f s) = f" .. lemma subst_all: \(\x. x = a \ PROP P x) \ PROP P a\ \(\x. a = x \ PROP P x) \ PROP P a\ proof - show \(\x. x = a \ PROP P x) \ PROP P a\ proof (rule equal_intr_rule) assume *: \\x. x = a \ PROP P x\ show \PROP P a\ by (rule *) (rule refl) next fix x assume \PROP P a\ and \x = a\ from \x = a\ have \x \ a\ by (rule eq_reflection) with \PROP P a\ show \PROP P x\ by simp qed show \(\x. a = x \ PROP P x) \ PROP P a\ proof (rule equal_intr_rule) assume *: \\x. a = x \ PROP P x\ show \PROP P a\ by (rule *) (rule refl) next fix x assume \PROP P a\ and \a = x\ from \a = x\ have \a \ x\ by (rule eq_reflection) with \PROP P a\ show \PROP P x\ by simp qed qed lemma simp_thms: shows not_not: "(\ \ P) = P" and Not_eq_iff: "((\ P) = (\ Q)) = (P = Q)" and "(P \ Q) = (P = (\ Q))" "(P \ \P) = True" "(\ P \ P) = True" "(x = x) = True" and not_True_eq_False [code]: "(\ True) = False" and not_False_eq_True [code]: "(\ False) = True" and "(\ P) \ P" "P \ (\ P)" "(True = P) = P" and eq_True: "(P = True) = P" and "(False = P) = (\ P)" and eq_False: "(P = False) = (\ P)" and "(True \ P) = P" "(False \ P) = True" "(P \ True) = True" "(P \ P) = True" "(P \ False) = (\ P)" "(P \ \ P) = (\ P)" "(P \ True) = P" "(True \ P) = P" "(P \ False) = False" "(False \ P) = False" "(P \ P) = P" "(P \ (P \ Q)) = (P \ Q)" "(P \ \ P) = False" "(\ P \ P) = False" "(P \ True) = True" "(True \ P) = True" "(P \ False) = P" "(False \ P) = P" "(P \ P) = P" "(P \ (P \ Q)) = (P \ Q)" and "(\x. P) = P" "(\x. P) = P" "\x. x = t" "\x. t = x" and "\P. (\x. x = t \ P x) = P t" "\P. (\x. t = x \ P x) = P t" "\P. (\x. x = t \ P x) = P t" "\P. (\x. t = x \ P x) = P t" "(\x. x \ t) = False" "(\x. t \ x) = False" by (blast, blast, blast, blast, blast, iprover+) lemma disj_absorb: "A \ A \ A" by blast lemma disj_left_absorb: "A \ (A \ B) \ A \ B" by blast lemma conj_absorb: "A \ A \ A" by blast lemma conj_left_absorb: "A \ (A \ B) \ A \ B" by blast lemma eq_ac: shows eq_commute: "a = b \ b = a" and iff_left_commute: "(P \ (Q \ R)) \ (Q \ (P \ R))" and iff_assoc: "((P \ Q) \ R) \ (P \ (Q \ R))" by (iprover, blast+) lemma neq_commute: "a \ b \ b \ a" by iprover lemma conj_comms: shows conj_commute: "P \ Q \ Q \ P" and conj_left_commute: "P \ (Q \ R) \ Q \ (P \ R)" by iprover+ lemma conj_assoc: "(P \ Q) \ R \ P \ (Q \ R)" by iprover lemmas conj_ac = conj_commute conj_left_commute conj_assoc lemma disj_comms: shows disj_commute: "P \ Q \ Q \ P" and disj_left_commute: "P \ (Q \ R) \ Q \ (P \ R)" by iprover+ lemma disj_assoc: "(P \ Q) \ R \ P \ (Q \ R)" by iprover lemmas disj_ac = disj_commute disj_left_commute disj_assoc lemma conj_disj_distribL: "P \ (Q \ R) \ P \ Q \ P \ R" by iprover lemma conj_disj_distribR: "(P \ Q) \ R \ P \ R \ Q \ R" by iprover lemma disj_conj_distribL: "P \ (Q \ R) \ (P \ Q) \ (P \ R)" by iprover lemma disj_conj_distribR: "(P \ Q) \ R \ (P \ R) \ (Q \ R)" by iprover lemma imp_conjR: "(P \ (Q \ R)) = ((P \ Q) \ (P \ R))" by iprover lemma imp_conjL: "((P \ Q) \ R) = (P \ (Q \ R))" by iprover lemma imp_disjL: "((P \ Q) \ R) = ((P \ R) \ (Q \ R))" by iprover text \These two are specialized, but \imp_disj_not1\ is useful in \Auth/Yahalom\.\ lemma imp_disj_not1: "(P \ Q \ R) \ (\ Q \ P \ R)" by blast lemma imp_disj_not2: "(P \ Q \ R) \ (\ R \ P \ Q)" by blast lemma imp_disj1: "((P \ Q) \ R) \ (P \ Q \ R)" by blast lemma imp_disj2: "(Q \ (P \ R)) \ (P \ Q \ R)" by blast lemma imp_cong: "(P = P') \ (P' \ (Q = Q')) \ ((P \ Q) \ (P' \ Q'))" by iprover lemma de_Morgan_disj: "\ (P \ Q) \ \ P \ \ Q" by iprover lemma de_Morgan_conj: "\ (P \ Q) \ \ P \ \ Q" by blast lemma not_imp: "\ (P \ Q) \ P \ \ Q" by blast lemma not_iff: "P \ Q \ (P \ \ Q)" by blast lemma disj_not1: "\ P \ Q \ (P \ Q)" by blast lemma disj_not2: "P \ \ Q \ (Q \ P)" by blast \ \changes orientation :-(\ lemma imp_conv_disj: "(P \ Q) \ (\ P) \ Q" by blast lemma disj_imp: "P \ Q \ \ P \ Q" by blast lemma iff_conv_conj_imp: "(P \ Q) \ (P \ Q) \ (Q \ P)" by iprover lemma cases_simp: "(P \ Q) \ (\ P \ Q) \ Q" \ \Avoids duplication of subgoals after \if_split\, when the true and false\ \ \cases boil down to the same thing.\ by blast lemma not_all: "\ (\x. P x) \ (\x. \ P x)" by blast lemma imp_all: "((\x. P x) \ Q) \ (\x. P x \ Q)" by blast lemma not_ex: "\ (\x. P x) \ (\x. \ P x)" by iprover lemma imp_ex: "((\x. P x) \ Q) \ (\x. P x \ Q)" by iprover lemma all_not_ex: "(\x. P x) \ \ (\x. \ P x)" by blast declare All_def [no_atp] lemma ex_disj_distrib: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma all_conj_distrib: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover text \ \<^medskip> The \\\ congruence rule: not included by default! May slow rewrite proofs down by as much as 50\%\ lemma conj_cong: "P = P' \ (P' \ Q = Q') \ (P \ Q) = (P' \ Q')" by iprover lemma rev_conj_cong: "Q = Q' \ (Q' \ P = P') \ (P \ Q) = (P' \ Q')" by iprover text \The \|\ congruence rule: not included by default!\ lemma disj_cong: "P = P' \ (\ P' \ Q = Q') \ (P \ Q) = (P' \ Q')" by blast text \\<^medskip> if-then-else rules\ lemma if_True [code]: "(if True then x else y) = x" unfolding If_def by blast lemma if_False [code]: "(if False then x else y) = y" unfolding If_def by blast lemma if_P: "P \ (if P then x else y) = x" unfolding If_def by blast lemma if_not_P: "\ P \ (if P then x else y) = y" unfolding If_def by blast lemma if_split: "P (if Q then x else y) = ((Q \ P x) \ (\ Q \ P y))" proof (rule case_split [of Q]) show ?thesis if Q using that by (simplesubst if_P) blast+ show ?thesis if "\ Q" using that by (simplesubst if_not_P) blast+ qed lemma if_split_asm: "P (if Q then x else y) = (\ ((Q \ \ P x) \ (\ Q \ \ P y)))" by (simplesubst if_split) blast lemmas if_splits [no_atp] = if_split if_split_asm lemma if_cancel: "(if c then x else x) = x" by (simplesubst if_split) blast lemma if_eq_cancel: "(if x = y then y else x) = x" by (simplesubst if_split) blast lemma if_bool_eq_conj: "(if P then Q else R) = ((P \ Q) \ (\ P \ R))" \ \This form is useful for expanding \if\s on the RIGHT of the \\\ symbol.\ by (rule if_split) lemma if_bool_eq_disj: "(if P then Q else R) = ((P \ Q) \ (\ P \ R))" \ \And this form is useful for expanding \if\s on the LEFT.\ by (simplesubst if_split) blast lemma Eq_TrueI: "P \ P \ True" unfolding atomize_eq by iprover lemma Eq_FalseI: "\ P \ P \ False" unfolding atomize_eq by iprover text \\<^medskip> let rules for simproc\ lemma Let_folded: "f x \ g x \ Let x f \ Let x g" by (unfold Let_def) lemma Let_unfold: "f x \ g \ Let x f \ g" by (unfold Let_def) text \ The following copy of the implication operator is useful for fine-tuning congruence rules. It instructs the simplifier to simplify its premise. \ definition simp_implies :: "prop \ prop \ prop" (infixr "=simp=>" 1) where "simp_implies \ (\)" lemma simp_impliesI: assumes PQ: "(PROP P \ PROP Q)" shows "PROP P =simp=> PROP Q" unfolding simp_implies_def by (iprover intro: PQ) lemma simp_impliesE: assumes PQ: "PROP P =simp=> PROP Q" and P: "PROP P" and QR: "PROP Q \ PROP R" shows "PROP R" by (iprover intro: QR P PQ [unfolded simp_implies_def]) lemma simp_implies_cong: assumes PP' :"PROP P \ PROP P'" and P'QQ': "PROP P' \ (PROP Q \ PROP Q')" shows "(PROP P =simp=> PROP Q) \ (PROP P' =simp=> PROP Q')" unfolding simp_implies_def proof (rule equal_intr_rule) assume PQ: "PROP P \ PROP Q" and P': "PROP P'" from PP' [symmetric] and P' have "PROP P" by (rule equal_elim_rule1) then have "PROP Q" by (rule PQ) with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1) next assume P'Q': "PROP P' \ PROP Q'" and P: "PROP P" from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) then have "PROP Q'" by (rule P'Q') with P'QQ' [OF P', symmetric] show "PROP Q" by (rule equal_elim_rule1) qed lemma uncurry: assumes "P \ Q \ R" shows "P \ Q \ R" using assms by blast lemma iff_allI: assumes "\x. P x = Q x" shows "(\x. P x) = (\x. Q x)" using assms by blast lemma iff_exI: assumes "\x. P x = Q x" shows "(\x. P x) = (\x. Q x)" using assms by blast lemma all_comm: "(\x y. P x y) = (\y x. P x y)" by blast lemma ex_comm: "(\x y. P x y) = (\y x. P x y)" by blast ML_file \Tools/simpdata.ML\ ML \open Simpdata\ setup \ map_theory_simpset (put_simpset HOL_basic_ss) #> Simplifier.method_setup Splitter.split_modifiers \ simproc_setup defined_Ex ("\x. P x") = \K Quantifier1.rearrange_Ex\ simproc_setup defined_All ("\x. P x") = \K Quantifier1.rearrange_All\ simproc_setup defined_all("\x. PROP P x") = \K Quantifier1.rearrange_all\ text \Simproc for proving \(y = x) \ False\ from premise \\ (x = y)\:\ simproc_setup neq ("x = y") = \fn _ => let val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI}; fun is_neq eq lhs rhs thm = (case Thm.prop_of thm of _ $ (Not $ (eq' $ l' $ r')) => Not = HOLogic.Not andalso eq' = eq andalso r' aconv lhs andalso l' aconv rhs | _ => false); fun proc ss ct = (case Thm.term_of ct of eq $ lhs $ rhs => (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of ss) of SOME thm => SOME (thm RS neq_to_EQ_False) | NONE => NONE) | _ => NONE); in proc end \ simproc_setup let_simp ("Let x f") = \ let fun count_loose (Bound i) k = if i >= k then 1 else 0 | count_loose (s $ t) k = count_loose s k + count_loose t k | count_loose (Abs (_, _, t)) k = count_loose t (k + 1) | count_loose _ _ = 0; fun is_trivial_let (Const (\<^const_name>\Let\, _) $ x $ t) = (case t of Abs (_, _, t') => count_loose t' 0 <= 1 | _ => true); in fn _ => fn ctxt => fn ct => if is_trivial_let (Thm.term_of ct) then SOME @{thm Let_def} (*no or one ocurrence of bound variable*) else let (*Norbert Schirmer's case*) val t = Thm.term_of ct; val (t', ctxt') = yield_singleton (Variable.import_terms false) t ctxt; in Option.map (hd o Variable.export ctxt' ctxt o single) (case t' of Const (\<^const_name>\Let\,_) $ x $ f => (* x and f are already in normal form *) if is_Free x orelse is_Bound x orelse is_Const x then SOME @{thm Let_def} else let val n = case f of (Abs (x, _, _)) => x | _ => "x"; val cx = Thm.cterm_of ctxt x; val xT = Thm.typ_of_cterm cx; val cf = Thm.cterm_of ctxt f; val fx_g = Simplifier.rewrite ctxt (Thm.apply cf cx); val (_ $ _ $ g) = Thm.prop_of fx_g; val g' = abstract_over (x, g); val abs_g'= Abs (n, xT, g'); in if g aconv g' then let val rl = infer_instantiate ctxt [(("f", 0), cf), (("x", 0), cx)] @{thm Let_unfold}; in SOME (rl OF [fx_g]) end else if (Envir.beta_eta_contract f) aconv (Envir.beta_eta_contract abs_g') then NONE (*avoid identity conversion*) else let val g'x = abs_g' $ x; val g_g'x = Thm.symmetric (Thm.beta_conversion false (Thm.cterm_of ctxt g'x)); val rl = @{thm Let_folded} |> infer_instantiate ctxt [(("f", 0), Thm.cterm_of ctxt f), (("x", 0), cx), (("g", 0), Thm.cterm_of ctxt abs_g')]; in SOME (rl OF [Thm.transitive fx_g g_g'x]) end end | _ => NONE) end end \ lemma True_implies_equals: "(True \ PROP P) \ PROP P" proof assume "True \ PROP P" from this [OF TrueI] show "PROP P" . next assume "PROP P" then show "PROP P" . qed lemma implies_True_equals: "(PROP P \ True) \ Trueprop True" by standard (intro TrueI) lemma False_implies_equals: "(False \ P) \ Trueprop True" by standard simp_all (* It seems that making this a simp rule is slower than using the simproc below *) lemma implies_False_swap: "(False \ PROP P \ PROP Q) \ (PROP P \ False \ PROP Q)" by (rule swap_prems_eq) ML \ fun eliminate_false_implies ct = let val (prems, concl) = Logic.strip_horn (Thm.term_of ct) fun go n = if n > 1 then Conv.rewr_conv @{thm Pure.swap_prems_eq} then_conv Conv.arg_conv (go (n - 1)) then_conv Conv.rewr_conv @{thm HOL.implies_True_equals} else Conv.rewr_conv @{thm HOL.False_implies_equals} in case concl of Const (@{const_name HOL.Trueprop}, _) $ _ => SOME (go (length prems) ct) | _ => NONE end \ simproc_setup eliminate_false_implies ("False \ PROP P") = \K (K eliminate_false_implies)\ lemma ex_simps: "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" \ \Miniscoping: pushing in existential quantifiers.\ by (iprover | blast)+ lemma all_simps: "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" \ \Miniscoping: pushing in universal quantifiers.\ by (iprover | blast)+ lemmas [simp] = triv_forall_equality \ \prunes params\ True_implies_equals implies_True_equals \ \prune \True\ in asms\ False_implies_equals \ \prune \False\ in asms\ if_True if_False if_cancel if_eq_cancel imp_disjL \ \In general it seems wrong to add distributive laws by default: they might cause exponential blow-up. But \imp_disjL\ has been in for a while and cannot be removed without affecting existing proofs. Moreover, rewriting by \(P \ Q \ R) = ((P \ R) \ (Q \ R))\ might be justified on the grounds that it allows simplification of \R\ in the two cases.\ conj_assoc disj_assoc de_Morgan_conj de_Morgan_disj imp_disj1 imp_disj2 not_imp disj_not1 not_all not_ex cases_simp the_eq_trivial the_sym_eq_trivial ex_simps all_simps simp_thms subst_all lemmas [cong] = imp_cong simp_implies_cong lemmas [split] = if_split ML \val HOL_ss = simpset_of \<^context>\ text \Simplifies \x\ assuming \c\ and \y\ assuming \\ c\.\ lemma if_cong: assumes "b = c" and "c \ x = u" and "\ c \ y = v" shows "(if b then x else y) = (if c then u else v)" using assms by simp text \Prevents simplification of \x\ and \y\: faster and allows the execution of functional programs.\ lemma if_weak_cong [cong]: assumes "b = c" shows "(if b then x else y) = (if c then x else y)" using assms by (rule arg_cong) text \Prevents simplification of t: much faster\ lemma let_weak_cong: assumes "a = b" shows "(let x = a in t x) = (let x = b in t x)" using assms by (rule arg_cong) text \To tidy up the result of a simproc. Only the RHS will be simplified.\ lemma eq_cong2: assumes "u = u'" shows "(t \ u) \ (t \ u')" using assms by simp lemma if_distrib: "f (if c then x else y) = (if c then f x else f y)" by simp lemma if_distribR: "(if b then f else g) x = (if b then f x else g x)" by simp lemma all_if_distrib: "(\x. if x = a then P x else Q x) \ P a \ (\x. x\a \ Q x)" by auto lemma ex_if_distrib: "(\x. if x = a then P x else Q x) \ P a \ (\x. x\a \ Q x)" by auto lemma if_if_eq_conj: "(if P then if Q then x else y else y) = (if P \ Q then x else y)" by simp text \As a simplification rule, it replaces all function equalities by first-order equalities.\ lemma fun_eq_iff: "f = g \ (\x. f x = g x)" by auto subsubsection \Generic cases and induction\ text \Rule projections:\ ML \ structure Project_Rule = Project_Rule ( val conjunct1 = @{thm conjunct1} val conjunct2 = @{thm conjunct2} val mp = @{thm mp} ); \ context begin qualified definition "induct_forall P \ \x. P x" qualified definition "induct_implies A B \ A \ B" qualified definition "induct_equal x y \ x = y" qualified definition "induct_conj A B \ A \ B" qualified definition "induct_true \ True" qualified definition "induct_false \ False" lemma induct_forall_eq: "(\x. P x) \ Trueprop (induct_forall (\x. P x))" by (unfold atomize_all induct_forall_def) lemma induct_implies_eq: "(A \ B) \ Trueprop (induct_implies A B)" by (unfold atomize_imp induct_implies_def) lemma induct_equal_eq: "(x \ y) \ Trueprop (induct_equal x y)" by (unfold atomize_eq induct_equal_def) lemma induct_conj_eq: "(A &&& B) \ Trueprop (induct_conj A B)" by (unfold atomize_conj induct_conj_def) lemmas induct_atomize' = induct_forall_eq induct_implies_eq induct_conj_eq lemmas induct_atomize = induct_atomize' induct_equal_eq lemmas induct_rulify' [symmetric] = induct_atomize' lemmas induct_rulify [symmetric] = induct_atomize lemmas induct_rulify_fallback = induct_forall_def induct_implies_def induct_equal_def induct_conj_def induct_true_def induct_false_def lemma induct_forall_conj: "induct_forall (\x. induct_conj (A x) (B x)) = induct_conj (induct_forall A) (induct_forall B)" by (unfold induct_forall_def induct_conj_def) iprover lemma induct_implies_conj: "induct_implies C (induct_conj A B) = induct_conj (induct_implies C A) (induct_implies C B)" by (unfold induct_implies_def induct_conj_def) iprover lemma induct_conj_curry: "(induct_conj A B \ PROP C) \ (A \ B \ PROP C)" proof assume r: "induct_conj A B \ PROP C" assume ab: A B show "PROP C" by (rule r) (simp add: induct_conj_def ab) next assume r: "A \ B \ PROP C" assume ab: "induct_conj A B" show "PROP C" by (rule r) (simp_all add: ab [unfolded induct_conj_def]) qed lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry lemma induct_trueI: "induct_true" by (simp add: induct_true_def) text \Method setup.\ ML_file \~~/src/Tools/induct.ML\ ML \ structure Induct = Induct ( val cases_default = @{thm case_split} val atomize = @{thms induct_atomize} val rulify = @{thms induct_rulify'} val rulify_fallback = @{thms induct_rulify_fallback} val equal_def = @{thm induct_equal_def} fun dest_def (Const (\<^const_name>\induct_equal\, _) $ t $ u) = SOME (t, u) | dest_def _ = NONE fun trivial_tac ctxt = match_tac ctxt @{thms induct_trueI} ) \ ML_file \~~/src/Tools/induction.ML\ declaration \ fn _ => Induct.map_simpset (fn ss => ss addsimprocs [Simplifier.make_simproc \<^context> "swap_induct_false" {lhss = [\<^term>\induct_false \ PROP P \ PROP Q\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of _ $ (P as _ $ \<^Const_>\induct_false\) $ (_ $ Q $ _) => if P <> Q then SOME Drule.swap_prems_eq else NONE | _ => NONE)}, Simplifier.make_simproc \<^context> "induct_equal_conj_curry" {lhss = [\<^term>\induct_conj P Q \ PROP R\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of _ $ (_ $ P) $ _ => let fun is_conj \<^Const_>\induct_conj for P Q\ = is_conj P andalso is_conj Q | is_conj \<^Const_>\induct_equal _ for _ _\ = true | is_conj \<^Const_>\induct_true\ = true | is_conj \<^Const_>\induct_false\ = true | is_conj _ = false in if is_conj P then SOME @{thm induct_conj_curry} else NONE end | _ => NONE)}] |> Simplifier.set_mksimps (fn ctxt => Simpdata.mksimps Simpdata.mksimps_pairs ctxt #> map (rewrite_rule ctxt (map Thm.symmetric @{thms induct_rulify_fallback})))) \ text \Pre-simplification of induction and cases rules\ lemma [induct_simp]: "(\x. induct_equal x t \ PROP P x) \ PROP P t" unfolding induct_equal_def proof assume r: "\x. x = t \ PROP P x" show "PROP P t" by (rule r [OF refl]) next fix x assume "PROP P t" "x = t" then show "PROP P x" by simp qed lemma [induct_simp]: "(\x. induct_equal t x \ PROP P x) \ PROP P t" unfolding induct_equal_def proof assume r: "\x. t = x \ PROP P x" show "PROP P t" by (rule r [OF refl]) next fix x assume "PROP P t" "t = x" then show "PROP P x" by simp qed lemma [induct_simp]: "(induct_false \ P) \ Trueprop induct_true" unfolding induct_false_def induct_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "(induct_true \ PROP P) \ PROP P" unfolding induct_true_def proof assume "True \ PROP P" then show "PROP P" using TrueI . next assume "PROP P" then show "PROP P" . qed lemma [induct_simp]: "(PROP P \ induct_true) \ Trueprop induct_true" unfolding induct_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "(\x::'a::{}. induct_true) \ Trueprop induct_true" unfolding induct_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "induct_implies induct_true P \ P" by (simp add: induct_implies_def induct_true_def) lemma [induct_simp]: "x = x \ True" by (rule simp_thms) end ML_file \~~/src/Tools/induct_tacs.ML\ subsubsection \Coherent logic\ ML_file \~~/src/Tools/coherent.ML\ ML \ structure Coherent = Coherent ( val atomize_elimL = @{thm atomize_elimL}; val atomize_exL = @{thm atomize_exL}; val atomize_conjL = @{thm atomize_conjL}; val atomize_disjL = @{thm atomize_disjL}; val operator_names = [\<^const_name>\HOL.disj\, \<^const_name>\HOL.conj\, \<^const_name>\Ex\]; ); \ subsubsection \Reorienting equalities\ ML \ signature REORIENT_PROC = sig val add : (term -> bool) -> theory -> theory val proc : morphism -> Proof.context -> cterm -> thm option end; structure Reorient_Proc : REORIENT_PROC = struct structure Data = Theory_Data ( type T = ((term -> bool) * stamp) list; val empty = []; fun merge data : T = Library.merge (eq_snd (op =)) data; ); fun add m = Data.map (cons (m, stamp ())); fun matches thy t = exists (fn (m, _) => m t) (Data.get thy); val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; fun proc phi ctxt ct = let val thy = Proof_Context.theory_of ctxt; in case Thm.term_of ct of (_ $ t $ u) => if matches thy u then NONE else SOME meta_reorient | _ => NONE end; end; \ subsection \Other simple lemmas and lemma duplicates\ lemma eq_iff_swap: "(x = y \ P) \ (y = x \ P)" by blast lemma all_cong1: "(\x. P x = P' x) \ (\x. P x) = (\x. P' x)" by auto lemma ex_cong1: "(\x. P x = P' x) \ (\x. P x) = (\x. P' x)" by auto lemma all_cong: "(\x. Q x \ P x = P' x) \ (\x. Q x \ P x) = (\x. Q x \ P' x)" by auto lemma ex_cong: "(\x. Q x \ P x = P' x) \ (\x. Q x \ P x) = (\x. Q x \ P' x)" by auto lemma ex1_eq [iff]: "\!x. x = t" "\!x. t = x" by blast+ lemma choice_eq: "(\x. \!y. P x y) = (\!f. \x. P x (f x))" (is "?lhs = ?rhs") proof (intro iffI allI) assume L: ?lhs then have \
: "\x. P x (THE y. P x y)" by (best intro: theI') show ?rhs by (rule ex1I) (use L \
in \fast+\) next fix x assume R: ?rhs then obtain f where f: "\x. P x (f x)" and f1: "\y. (\x. P x (y x)) \ y = f" by (blast elim: ex1E) show "\!y. P x y" proof (rule ex1I) show "P x (f x)" using f by blast show "y = f x" if "P x y" for y proof - have "P z (if z = x then y else f z)" for z using f that by (auto split: if_split) with f1 [of "\z. if z = x then y else f z"] f show ?thesis by (auto simp add: split: if_split_asm dest: fun_cong) qed qed qed lemmas eq_sym_conv = eq_commute lemma nnf_simps: "(\ (P \ Q)) = (\ P \ \ Q)" "(\ (P \ Q)) = (\ P \ \ Q)" "(P \ Q) = (\ P \ Q)" "(P = Q) = ((P \ Q) \ (\ P \ \ Q))" "(\ (P = Q)) = ((P \ \ Q) \ (\ P \ Q))" "(\ \ P) = P" by blast+ subsection \Basic ML bindings\ ML \ val FalseE = @{thm FalseE} val Let_def = @{thm Let_def} val TrueI = @{thm TrueI} val allE = @{thm allE} val allI = @{thm allI} val all_dupE = @{thm all_dupE} val arg_cong = @{thm arg_cong} val box_equals = @{thm box_equals} val ccontr = @{thm ccontr} val classical = @{thm classical} val conjE = @{thm conjE} val conjI = @{thm conjI} val conjunct1 = @{thm conjunct1} val conjunct2 = @{thm conjunct2} val disjCI = @{thm disjCI} val disjE = @{thm disjE} val disjI1 = @{thm disjI1} val disjI2 = @{thm disjI2} val eq_reflection = @{thm eq_reflection} val ex1E = @{thm ex1E} val ex1I = @{thm ex1I} val ex1_implies_ex = @{thm ex1_implies_ex} val exE = @{thm exE} val exI = @{thm exI} val excluded_middle = @{thm excluded_middle} val ext = @{thm ext} val fun_cong = @{thm fun_cong} val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val iffI = @{thm iffI} val impE = @{thm impE} val impI = @{thm impI} val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq} val mp = @{thm mp} val notE = @{thm notE} val notI = @{thm notI} val not_all = @{thm not_all} val not_ex = @{thm not_ex} val not_iff = @{thm not_iff} val not_not = @{thm not_not} val not_sym = @{thm not_sym} val refl = @{thm refl} val rev_mp = @{thm rev_mp} val spec = @{thm spec} val ssubst = @{thm ssubst} val subst = @{thm subst} val sym = @{thm sym} val trans = @{thm trans} \ locale cnf begin lemma clause2raw_notE: "\P; \P\ \ False" by auto lemma clause2raw_not_disj: "\\ P; \ Q\ \ \ (P \ Q)" by auto lemma clause2raw_not_not: "P \ \\ P" by auto lemma iff_refl: "(P::bool) = P" by auto lemma iff_trans: "[| (P::bool) = Q; Q = R |] ==> P = R" by auto lemma conj_cong: "[| P = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto lemma disj_cong: "[| P = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto lemma make_nnf_imp: "[| (\P) = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto lemma make_nnf_iff: "[| P = P'; (\P) = NP; Q = Q'; (\Q) = NQ |] ==> (P = Q) = ((P' \ NQ) \ (NP \ Q'))" by auto lemma make_nnf_not_false: "(\False) = True" by auto lemma make_nnf_not_true: "(\True) = False" by auto lemma make_nnf_not_conj: "[| (\P) = P'; (\Q) = Q' |] ==> (\(P \ Q)) = (P' \ Q')" by auto lemma make_nnf_not_disj: "[| (\P) = P'; (\Q) = Q' |] ==> (\(P \ Q)) = (P' \ Q')" by auto lemma make_nnf_not_imp: "[| P = P'; (\Q) = Q' |] ==> (\(P \ Q)) = (P' \ Q')" by auto lemma make_nnf_not_iff: "[| P = P'; (\P) = NP; Q = Q'; (\Q) = NQ |] ==> (\(P = Q)) = ((P' \ Q') \ (NP \ NQ))" by auto lemma make_nnf_not_not: "P = P' ==> (\\P) = P'" by auto lemma simp_TF_conj_True_l: "[| P = True; Q = Q' |] ==> (P \ Q) = Q'" by auto lemma simp_TF_conj_True_r: "[| P = P'; Q = True |] ==> (P \ Q) = P'" by auto lemma simp_TF_conj_False_l: "P = False ==> (P \ Q) = False" by auto lemma simp_TF_conj_False_r: "Q = False ==> (P \ Q) = False" by auto lemma simp_TF_disj_True_l: "P = True ==> (P \ Q) = True" by auto lemma simp_TF_disj_True_r: "Q = True ==> (P \ Q) = True" by auto lemma simp_TF_disj_False_l: "[| P = False; Q = Q' |] ==> (P \ Q) = Q'" by auto lemma simp_TF_disj_False_r: "[| P = P'; Q = False |] ==> (P \ Q) = P'" by auto lemma make_cnf_disj_conj_l: "[| (P \ R) = PR; (Q \ R) = QR |] ==> ((P \ Q) \ R) = (PR \ QR)" by auto lemma make_cnf_disj_conj_r: "[| (P \ Q) = PQ; (P \ R) = PR |] ==> (P \ (Q \ R)) = (PQ \ PR)" by auto lemma make_cnfx_disj_ex_l: "((\(x::bool). P x) \ Q) = (\x. P x \ Q)" by auto lemma make_cnfx_disj_ex_r: "(P \ (\(x::bool). Q x)) = (\x. P \ Q x)" by auto lemma make_cnfx_newlit: "(P \ Q) = (\x. (P \ x) \ (Q \ \x))" by auto lemma make_cnfx_ex_cong: "(\(x::bool). P x = Q x) \ (\x. P x) = (\x. Q x)" by auto lemma weakening_thm: "[| P; Q |] ==> Q" by auto lemma cnftac_eq_imp: "[| P = Q; P |] ==> Q" by auto end ML_file \Tools/cnf.ML\ section \\NO_MATCH\ simproc\ text \ The simplification procedure can be used to avoid simplification of terms of a certain form. \ definition NO_MATCH :: "'a \ 'b \ bool" where "NO_MATCH pat val \ True" lemma NO_MATCH_cong[cong]: "NO_MATCH pat val = NO_MATCH pat val" by (rule refl) declare [[coercion_args NO_MATCH - -]] simproc_setup NO_MATCH ("NO_MATCH pat val") = \fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt val dest_binop = Term.dest_comb #> apfst (Term.dest_comb #> snd) val m = Pattern.matches thy (dest_binop (Thm.term_of ct)) in if m then NONE else SOME @{thm NO_MATCH_def} end \ text \ This setup ensures that a rewrite rule of the form \<^term>\NO_MATCH pat val \ t\ is only applied, if the pattern \pat\ does not match the value \val\. \ text\ Tagging a premise of a simp rule with ASSUMPTION forces the simplifier not to simplify the argument and to solve it by an assumption. \ definition ASSUMPTION :: "bool \ bool" where "ASSUMPTION A \ A" lemma ASSUMPTION_cong[cong]: "ASSUMPTION A = ASSUMPTION A" by (rule refl) lemma ASSUMPTION_I: "A \ ASSUMPTION A" by (simp add: ASSUMPTION_def) lemma ASSUMPTION_D: "ASSUMPTION A \ A" by (simp add: ASSUMPTION_def) setup \ let val asm_sol = mk_solver "ASSUMPTION" (fn ctxt => resolve_tac ctxt [@{thm ASSUMPTION_I}] THEN' resolve_tac ctxt (Simplifier.prems_of ctxt)) in map_theory_simpset (fn ctxt => Simplifier.addSolver (ctxt,asm_sol)) end \ subsection \Code generator setup\ subsubsection \Generic code generator preprocessor setup\ lemma conj_left_cong: "P \ Q \ P \ R \ Q \ R" by (fact arg_cong) lemma disj_left_cong: "P \ Q \ P \ R \ Q \ R" by (fact arg_cong) setup \ Code_Preproc.map_pre (put_simpset HOL_basic_ss) #> Code_Preproc.map_post (put_simpset HOL_basic_ss) #> Code_Simp.map_ss (put_simpset HOL_basic_ss #> Simplifier.add_cong @{thm conj_left_cong} #> Simplifier.add_cong @{thm disj_left_cong}) \ subsubsection \Equality\ class equal = fixes equal :: "'a \ 'a \ bool" assumes equal_eq: "equal x y \ x = y" begin lemma equal: "equal = (=)" by (rule ext equal_eq)+ lemma equal_refl: "equal x x \ True" - unfolding equal by rule+ + unfolding equal by (rule iffI TrueI refl)+ lemma eq_equal: "(=) \ equal" by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq) end declare eq_equal [symmetric, code_post] declare eq_equal [code] setup \ Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs [Simplifier.make_simproc \<^context> "equal" {lhss = [\<^term>\HOL.eq\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (_, Type (\<^type_name>\fun\, [Type _, _])) => SOME @{thm eq_equal} | _ => NONE)}]) \ subsubsection \Generic code generator foundation\ text \Datatype \<^typ>\bool\\ code_datatype True False lemma [code]: shows "False \ P \ False" and "True \ P \ P" and "P \ False \ False" and "P \ True \ P" by simp_all lemma [code]: shows "False \ P \ P" and "True \ P \ True" and "P \ False \ P" and "P \ True \ True" by simp_all lemma [code]: shows "(False \ P) \ True" and "(True \ P) \ P" and "(P \ False) \ \ P" and "(P \ True) \ True" by simp_all text \More about \<^typ>\prop\\ lemma [code nbe]: shows "(True \ PROP Q) \ PROP Q" and "(PROP Q \ True) \ Trueprop True" and "(P \ R) \ Trueprop (P \ R)" by (auto intro!: equal_intr_rule) lemma Trueprop_code [code]: "Trueprop True \ Code_Generator.holds" by (auto intro!: equal_intr_rule holds) declare Trueprop_code [symmetric, code_post] text \Equality\ declare simp_thms(6) [code nbe] instantiation itself :: (type) equal begin definition equal_itself :: "'a itself \ 'a itself \ bool" where "equal_itself x y \ x = y" instance by standard (fact equal_itself_def) end lemma equal_itself_code [code]: "equal TYPE('a) TYPE('a) \ True" by (simp add: equal) setup \Sign.add_const_constraint (\<^const_name>\equal\, SOME \<^typ>\'a::type \ 'a \ bool\)\ lemma equal_alias_cert: "OFCLASS('a, equal_class) \ (((=) :: 'a \ 'a \ bool) \ equal)" (is "?ofclass \ ?equal") proof assume "PROP ?ofclass" show "PROP ?equal" by (tactic \ALLGOALS (resolve_tac \<^context> [Thm.unconstrainT @{thm eq_equal}])\) (fact \PROP ?ofclass\) next assume "PROP ?equal" show "PROP ?ofclass" proof qed (simp add: \PROP ?equal\) qed setup \Sign.add_const_constraint (\<^const_name>\equal\, SOME \<^typ>\'a::equal \ 'a \ bool\)\ setup \Nbe.add_const_alias @{thm equal_alias_cert}\ text \Cases\ lemma Let_case_cert: assumes "CASE \ (\x. Let x f)" shows "CASE x \ f x" using assms by simp_all setup \ Code.declare_case_global @{thm Let_case_cert} #> Code.declare_undefined_global \<^const_name>\undefined\ \ declare [[code abort: undefined]] subsubsection \Generic code generator target languages\ text \type \<^typ>\bool\\ code_printing type_constructor bool \ (SML) "bool" and (OCaml) "bool" and (Haskell) "Bool" and (Scala) "Boolean" | constant True \ (SML) "true" and (OCaml) "true" and (Haskell) "True" and (Scala) "true" | constant False \ (SML) "false" and (OCaml) "false" and (Haskell) "False" and (Scala) "false" code_reserved SML bool true false code_reserved OCaml bool code_reserved Scala Boolean code_printing constant Not \ (SML) "not" and (OCaml) "not" and (Haskell) "not" and (Scala) "'! _" | constant HOL.conj \ (SML) infixl 1 "andalso" and (OCaml) infixl 3 "&&" and (Haskell) infixr 3 "&&" and (Scala) infixl 3 "&&" | constant HOL.disj \ (SML) infixl 0 "orelse" and (OCaml) infixl 2 "||" and (Haskell) infixl 2 "||" and (Scala) infixl 1 "||" | constant HOL.implies \ (SML) "!(if (_)/ then (_)/ else true)" and (OCaml) "!(if (_)/ then (_)/ else true)" and (Haskell) "!(if (_)/ then (_)/ else True)" and (Scala) "!(if ((_))/ (_)/ else true)" | constant If \ (SML) "!(if (_)/ then (_)/ else (_))" and (OCaml) "!(if (_)/ then (_)/ else (_))" and (Haskell) "!(if (_)/ then (_)/ else (_))" and (Scala) "!(if ((_))/ (_)/ else (_))" code_reserved SML not code_reserved OCaml not code_identifier code_module Pure \ (SML) HOL and (OCaml) HOL and (Haskell) HOL and (Scala) HOL text \Using built-in Haskell equality.\ code_printing type_class equal \ (Haskell) "Eq" | constant HOL.equal \ (Haskell) infix 4 "==" | constant HOL.eq \ (Haskell) infix 4 "==" text \\undefined\\ code_printing constant undefined \ (SML) "!(raise/ Fail/ \"undefined\")" and (OCaml) "failwith/ \"undefined\"" and (Haskell) "error/ \"undefined\"" and (Scala) "!sys.error(\"undefined\")" subsubsection \Evaluation and normalization by evaluation\ method_setup eval = \ let fun eval_tac ctxt = let val conv = Code_Runtime.dynamic_holds_conv in CONVERSION (Conv.params_conv ~1 (Conv.concl_conv ~1 o conv) ctxt) THEN' resolve_tac ctxt [TrueI] end in Scan.succeed (SIMPLE_METHOD' o eval_tac) end \ "solve goal by evaluation" method_setup normalization = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o (CONVERSION (Nbe.dynamic_conv ctxt) THEN_ALL_NEW (TRY o resolve_tac ctxt [TrueI])))) \ "solve goal by normalization" subsection \Counterexample Search Units\ subsubsection \Quickcheck\ quickcheck_params [size = 5, iterations = 50] subsubsection \Nitpick setup\ named_theorems nitpick_unfold "alternative definitions of constants as needed by Nitpick" and nitpick_simp "equational specification of constants as needed by Nitpick" and nitpick_psimp "partial equational specification of constants as needed by Nitpick" and nitpick_choice_spec "choice specification of constants as needed by Nitpick" declare if_bool_eq_conj [nitpick_unfold, no_atp] and if_bool_eq_disj [no_atp] subsection \Preprocessing for the predicate compiler\ named_theorems code_pred_def "alternative definitions of constants for the Predicate Compiler" and code_pred_inline "inlining definitions for the Predicate Compiler" and code_pred_simp "simplification rules for the optimisations in the Predicate Compiler" subsection \Legacy tactics and ML bindings\ ML \ (* combination of (spec RS spec RS ...(j times) ... spec RS mp) *) local fun wrong_prem (Const (\<^const_name>\All\, _) $ Abs (_, _, t)) = wrong_prem t | wrong_prem (Bound _) = true | wrong_prem _ = false; val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of); fun smp i = funpow i (fn m => filter_right ([spec] RL m)) [mp]; in fun smp_tac ctxt j = EVERY' [dresolve_tac ctxt (smp j), assume_tac ctxt]; end; local val nnf_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms nnf_simps}); in fun nnf_conv ctxt = Simplifier.rewrite (put_simpset nnf_ss ctxt); end \ hide_const (open) eq equal end diff --git a/src/HOL/Int.thy b/src/HOL/Int.thy --- a/src/HOL/Int.thy +++ b/src/HOL/Int.thy @@ -1,2252 +1,2272 @@ (* Title: HOL/Int.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) section \The Integers as Equivalence Classes over Pairs of Natural Numbers\ theory Int imports Quotient Groups_Big Fun_Def begin subsection \Definition of integers as a quotient type\ definition intrel :: "(nat \ nat) \ (nat \ nat) \ bool" where "intrel = (\(x, y) (u, v). x + v = u + y)" lemma intrel_iff [simp]: "intrel (x, y) (u, v) \ x + v = u + y" by (simp add: intrel_def) quotient_type int = "nat \ nat" / "intrel" morphisms Rep_Integ Abs_Integ proof (rule equivpI) show "reflp intrel" by (auto simp: reflp_def) show "symp intrel" by (auto simp: symp_def) show "transp intrel" by (auto simp: transp_def) qed lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: "(\x y. z = Abs_Integ (x, y) \ P) \ P" by (induct z) auto subsection \Integers form a commutative ring\ instantiation int :: comm_ring_1 begin lift_definition zero_int :: "int" is "(0, 0)" . lift_definition one_int :: "int" is "(1, 0)" . lift_definition plus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + u, y + v)" by clarsimp lift_definition uminus_int :: "int \ int" is "\(x, y). (y, x)" by clarsimp lift_definition minus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + v, y + u)" by clarsimp lift_definition times_int :: "int \ int \ int" is "\(x, y) (u, v). (x*u + y*v, x*v + y*u)" -proof (clarsimp) +proof (unfold intrel_def, clarify) fix s t u v w x y z :: nat assume "s + v = u + t" and "w + z = y + x" then have "(s + v) * w + (u + t) * x + u * (w + z) + v * (y + x) = (u + t) * w + (s + v) * x + u * (y + x) + v * (w + z)" by simp then show "(s * w + t * x) + (u * z + v * y) = (u * y + v * z) + (s * x + t * w)" by (simp add: algebra_simps) qed instance by standard (transfer; clarsimp simp: algebra_simps)+ end abbreviation int :: "nat \ int" where "int \ of_nat" lemma int_def: "int n = Abs_Integ (n, 0)" by (induct n) (simp add: zero_int.abs_eq, simp add: one_int.abs_eq plus_int.abs_eq) lemma int_transfer [transfer_rule]: includes lifting_syntax shows "rel_fun (=) pcr_int (\n. (n, 0)) int" by (simp add: rel_fun_def int.pcr_cr_eq cr_int_def int_def) lemma int_diff_cases: obtains (diff) m n where "z = int m - int n" by transfer clarsimp subsection \Integers are totally ordered\ instantiation int :: linorder begin lift_definition less_eq_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v \ u + y" by auto lift_definition less_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v < u + y" by auto instance by standard (transfer, force)+ end instantiation int :: distrib_lattice begin definition "(inf :: int \ int \ int) = min" definition "(sup :: int \ int \ int) = max" instance by standard (auto simp add: inf_int_def sup_int_def max_min_distrib2) end subsection \Ordering properties of arithmetic operations\ instance int :: ordered_cancel_ab_semigroup_add proof fix i j k :: int show "i \ j \ k + i \ k + j" by transfer clarsimp qed text \Strict Monotonicity of Multiplication.\ text \Strict, in 1st argument; proof is by induction on \k > 0\.\ lemma zmult_zless_mono2_lemma: "i < j \ 0 < k \ int k * i < int k * j" for i j :: int proof (induct k) case 0 then show ?case by simp next case (Suc k) then show ?case by (cases "k = 0") (simp_all add: distrib_right add_strict_mono) qed lemma zero_le_imp_eq_int: assumes "k \ (0::int)" shows "\n. k = int n" proof - have "b \ a \ \n::nat. a = n + b" for a b - by (rule_tac x="a - b" in exI) simp + using exI[of _ "a - b"] by simp with assms show ?thesis by transfer auto qed lemma zero_less_imp_eq_int: assumes "k > (0::int)" shows "\n>0. k = int n" proof - have "b < a \ \n::nat. n>0 \ a = n + b" for a b - by (rule_tac x="a - b" in exI) simp + using exI[of _ "a - b"] by simp with assms show ?thesis by transfer auto qed lemma zmult_zless_mono2: "i < j \ 0 < k \ k * i < k * j" for i j k :: int by (drule zero_less_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma) text \The integers form an ordered integral domain.\ instantiation int :: linordered_idom begin definition zabs_def: "\i::int\ = (if i < 0 then - i else i)" definition zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)" instance proof fix i j k :: int show "i < j \ 0 < k \ k * i < k * j" by (rule zmult_zless_mono2) show "\i\ = (if i < 0 then -i else i)" by (simp only: zabs_def) show "sgn (i::int) = (if i=0 then 0 else if 0 w + 1 \ z" for w z :: int by transfer clarsimp lemma zless_iff_Suc_zadd: "w < z \ (\n. z = w + int (Suc n))" for w z :: int proof - have "\a b c d. a + d < c + b \ \n. c + b = Suc (a + n + d)" - by (rule_tac x="c+b - Suc(a+d)" in exI) arith + proof - + fix a b c d :: nat + assume "a + d < c + b" + then have "c + b = Suc (a + (c + b - Suc (a + d)) + d) " + by arith + then show "\n. c + b = Suc (a + n + d)" + by (rule exI) + qed then show ?thesis by transfer auto qed lemma zabs_less_one_iff [simp]: "\z\ < 1 \ z = 0" (is "?lhs \ ?rhs") for z :: int proof assume ?rhs then show ?lhs by simp next assume ?lhs with zless_imp_add1_zle [of "\z\" 1] have "\z\ + 1 \ 1" by simp then have "\z\ \ 0" by simp then show ?rhs by simp qed subsection \Embedding of the Integers into any \ring_1\: \of_int\\ context ring_1 begin lift_definition of_int :: "int \ 'a" is "\(i, j). of_nat i - of_nat j" by (clarsimp simp add: diff_eq_eq eq_diff_eq diff_add_eq of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_0 [simp]: "of_int 0 = 0" by transfer simp lemma of_int_1 [simp]: "of_int 1 = 1" by transfer simp lemma of_int_add [simp]: "of_int (w + z) = of_int w + of_int z" by transfer (clarsimp simp add: algebra_simps) lemma of_int_minus [simp]: "of_int (- z) = - (of_int z)" by (transfer fixing: uminus) clarsimp lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z" using of_int_add [of w "- z"] by simp lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" by (transfer fixing: times) (clarsimp simp add: algebra_simps) lemma mult_of_int_commute: "of_int x * y = y * of_int x" by (transfer fixing: times) (auto simp: algebra_simps mult_of_nat_commute) text \Collapse nested embeddings.\ lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n" by (induct n) auto lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k" by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric]) lemma of_int_neg_numeral [code_post]: "of_int (- numeral k) = - numeral k" by simp lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n" by (induct n) simp_all lemma of_int_of_bool [simp]: "of_int (of_bool P) = of_bool P" by auto end context ring_char_0 begin lemma of_int_eq_iff [simp]: "of_int w = of_int z \ w = z" by transfer (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) text \Special cases where either operand is zero.\ lemma of_int_eq_0_iff [simp]: "of_int z = 0 \ z = 0" using of_int_eq_iff [of z 0] by simp lemma of_int_0_eq_iff [simp]: "0 = of_int z \ z = 0" using of_int_eq_iff [of 0 z] by simp lemma of_int_eq_1_iff [iff]: "of_int z = 1 \ z = 1" using of_int_eq_iff [of z 1] by simp lemma numeral_power_eq_of_int_cancel_iff [simp]: "numeral x ^ n = of_int y \ numeral x ^ n = y" using of_int_eq_iff[of "numeral x ^ n" y, unfolded of_int_numeral of_int_power] . lemma of_int_eq_numeral_power_cancel_iff [simp]: "of_int y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_int_cancel_iff [of x n y] by (metis (mono_tags)) lemma neg_numeral_power_eq_of_int_cancel_iff [simp]: "(- numeral x) ^ n = of_int y \ (- numeral x) ^ n = y" using of_int_eq_iff[of "(- numeral x) ^ n" y] by simp lemma of_int_eq_neg_numeral_power_cancel_iff [simp]: "of_int y = (- numeral x) ^ n \ y = (- numeral x) ^ n" using neg_numeral_power_eq_of_int_cancel_iff[of x n y] by (metis (mono_tags)) lemma of_int_eq_of_int_power_cancel_iff[simp]: "(of_int b) ^ w = of_int x \ b ^ w = x" by (metis of_int_power of_int_eq_iff) lemma of_int_power_eq_of_int_cancel_iff[simp]: "of_int x = (of_int b) ^ w \ x = b ^ w" by (metis of_int_eq_of_int_power_cancel_iff) end context linordered_idom begin text \Every \linordered_idom\ has characteristic zero.\ subclass ring_char_0 .. lemma of_int_le_iff [simp]: "of_int w \ of_int z \ w \ z" by (transfer fixing: less_eq) (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_less_iff [simp]: "of_int w < of_int z \ w < z" by (simp add: less_le order_less_le) lemma of_int_0_le_iff [simp]: "0 \ of_int z \ 0 \ z" using of_int_le_iff [of 0 z] by simp lemma of_int_le_0_iff [simp]: "of_int z \ 0 \ z \ 0" using of_int_le_iff [of z 0] by simp lemma of_int_0_less_iff [simp]: "0 < of_int z \ 0 < z" using of_int_less_iff [of 0 z] by simp lemma of_int_less_0_iff [simp]: "of_int z < 0 \ z < 0" using of_int_less_iff [of z 0] by simp lemma of_int_1_le_iff [simp]: "1 \ of_int z \ 1 \ z" using of_int_le_iff [of 1 z] by simp lemma of_int_le_1_iff [simp]: "of_int z \ 1 \ z \ 1" using of_int_le_iff [of z 1] by simp lemma of_int_1_less_iff [simp]: "1 < of_int z \ 1 < z" using of_int_less_iff [of 1 z] by simp lemma of_int_less_1_iff [simp]: "of_int z < 1 \ z < 1" using of_int_less_iff [of z 1] by simp lemma of_int_pos: "z > 0 \ of_int z > 0" by simp lemma of_int_nonneg: "z \ 0 \ of_int z \ 0" by simp lemma of_int_abs [simp]: "of_int \x\ = \of_int x\" by (auto simp add: abs_if) lemma of_int_lessD: assumes "\of_int n\ < x" shows "n = 0 \ x > 1" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 < x" using assms by (rule le_less_trans) then show ?thesis .. qed lemma of_int_leD: assumes "\of_int n\ \ x" shows "n = 0 \ 1 \ x" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 \ x" using assms by (rule order_trans) then show ?thesis .. qed lemma numeral_power_le_of_int_cancel_iff [simp]: "numeral x ^ n \ of_int a \ numeral x ^ n \ a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_le_iff) lemma of_int_le_numeral_power_cancel_iff [simp]: "of_int a \ numeral x ^ n \ a \ numeral x ^ n" by (metis (mono_tags) local.numeral_power_eq_of_int_cancel_iff of_int_le_iff) lemma numeral_power_less_of_int_cancel_iff [simp]: "numeral x ^ n < of_int a \ numeral x ^ n < a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma of_int_less_numeral_power_cancel_iff [simp]: "of_int a < numeral x ^ n \ a < numeral x ^ n" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma neg_numeral_power_le_of_int_cancel_iff [simp]: "(- numeral x) ^ n \ of_int a \ (- numeral x) ^ n \ a" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma of_int_le_neg_numeral_power_cancel_iff [simp]: "of_int a \ (- numeral x) ^ n \ a \ (- numeral x) ^ n" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma neg_numeral_power_less_of_int_cancel_iff [simp]: "(- numeral x) ^ n < of_int a \ (- numeral x) ^ n < a" using of_int_less_iff[of "(- numeral x) ^ n" a] by simp lemma of_int_less_neg_numeral_power_cancel_iff [simp]: "of_int a < (- numeral x) ^ n \ a < (- numeral x::int) ^ n" using of_int_less_iff[of a "(- numeral x) ^ n"] by simp lemma of_int_le_of_int_power_cancel_iff[simp]: "(of_int b) ^ w \ of_int x \ b ^ w \ x" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_power_le_of_int_cancel_iff[simp]: "of_int x \ (of_int b) ^ w\ x \ b ^ w" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_less_of_int_power_cancel_iff[simp]: "(of_int b) ^ w < of_int x \ b ^ w < x" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_power_less_of_int_cancel_iff[simp]: "of_int x < (of_int b) ^ w\ x < b ^ w" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_max: "of_int (max x y) = max (of_int x) (of_int y)" by (auto simp: max_def) lemma of_int_min: "of_int (min x y) = min (of_int x) (of_int y)" by (auto simp: min_def) end context division_ring begin lemmas mult_inverse_of_int_commute = mult_commute_imp_mult_inverse_commute[OF mult_of_int_commute] end text \Comparisons involving \<^term>\of_int\.\ lemma of_int_eq_numeral_iff [iff]: "of_int z = (numeral n :: 'a::ring_char_0) \ z = numeral n" using of_int_eq_iff by fastforce lemma of_int_le_numeral_iff [simp]: "of_int z \ (numeral n :: 'a::linordered_idom) \ z \ numeral n" using of_int_le_iff [of z "numeral n"] by simp lemma of_int_numeral_le_iff [simp]: "(numeral n :: 'a::linordered_idom) \ of_int z \ numeral n \ z" using of_int_le_iff [of "numeral n"] by simp lemma of_int_less_numeral_iff [simp]: "of_int z < (numeral n :: 'a::linordered_idom) \ z < numeral n" using of_int_less_iff [of z "numeral n"] by simp lemma of_int_numeral_less_iff [simp]: "(numeral n :: 'a::linordered_idom) < of_int z \ numeral n < z" using of_int_less_iff [of "numeral n" z] by simp lemma of_nat_less_of_int_iff: "(of_nat n::'a::linordered_idom) < of_int x \ int n < x" by (metis of_int_of_nat_eq of_int_less_iff) lemma of_int_eq_id [simp]: "of_int = id" proof show "of_int z = id z" for z by (cases z rule: int_diff_cases) simp qed instance int :: no_top proof - show "\x::int. \y. x < y" - by (rule_tac x="x + 1" in exI) simp + fix x::int + have "x < x + 1" + by simp + then show "\y. x < y" + by (rule exI) qed instance int :: no_bot proof - show "\x::int. \y. y < x" - by (rule_tac x="x - 1" in exI) simp + fix x::int + have "x - 1< x" + by simp + then show "\y. y < x" + by (rule exI) qed subsection \Magnitude of an Integer, as a Natural Number: \nat\\ lift_definition nat :: "int \ nat" is "\(x, y). x - y" by auto lemma nat_int [simp]: "nat (int n) = n" by transfer simp lemma int_nat_eq [simp]: "int (nat z) = (if 0 \ z then z else 0)" by transfer clarsimp lemma nat_0_le: "0 \ z \ int (nat z) = z" by simp lemma nat_le_0 [simp]: "z \ 0 \ nat z = 0" by transfer clarsimp lemma nat_le_eq_zle: "0 < w \ 0 \ z \ nat w \ nat z \ w \ z" by transfer (clarsimp, arith) text \An alternative condition is \<^term>\0 \ w\.\ lemma nat_mono_iff: "0 < z \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma nat_less_eq_zless: "0 \ w \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma zless_nat_conj [simp]: "nat w < nat z \ 0 < z \ w < z" by transfer (clarsimp, arith) lemma nonneg_int_cases: assumes "0 \ k" obtains n where "k = int n" proof - from assms have "k = int (nat k)" by simp then show thesis by (rule that) qed lemma pos_int_cases: assumes "0 < k" obtains n where "k = int n" and "n > 0" proof - from assms have "0 \ k" by simp then obtain n where "k = int n" by (rule nonneg_int_cases) moreover have "n > 0" using \k = int n\ assms by simp ultimately show thesis by (rule that) qed lemma nonpos_int_cases: assumes "k \ 0" obtains n where "k = - int n" proof - from assms have "- k \ 0" by simp then obtain n where "- k = int n" by (rule nonneg_int_cases) then have "k = - int n" by simp then show thesis by (rule that) qed lemma neg_int_cases: assumes "k < 0" obtains n where "k = - int n" and "n > 0" proof - from assms have "- k > 0" by simp then obtain n where "- k = int n" and "- k > 0" by (blast elim: pos_int_cases) then have "k = - int n" and "n > 0" by simp_all then show thesis by (rule that) qed lemma nat_eq_iff: "nat w = m \ (if 0 \ w then w = int m else m = 0)" by transfer (clarsimp simp add: le_imp_diff_is_add) lemma nat_eq_iff2: "m = nat w \ (if 0 \ w then w = int m else m = 0)" using nat_eq_iff [of w m] by auto lemma nat_0 [simp]: "nat 0 = 0" by (simp add: nat_eq_iff) lemma nat_1 [simp]: "nat 1 = Suc 0" by (simp add: nat_eq_iff) lemma nat_numeral [simp]: "nat (numeral k) = numeral k" by (simp add: nat_eq_iff) lemma nat_neg_numeral [simp]: "nat (- numeral k) = 0" by simp lemma nat_2: "nat 2 = Suc (Suc 0)" by simp lemma nat_less_iff: "0 \ w \ nat w < m \ w < of_nat m" by transfer (clarsimp, arith) lemma nat_le_iff: "nat x \ n \ x \ int n" by transfer (clarsimp simp add: le_diff_conv) lemma nat_mono: "x \ y \ nat x \ nat y" by transfer auto lemma nat_0_iff[simp]: "nat i = 0 \ i \ 0" for i :: int by transfer clarsimp lemma int_eq_iff: "of_nat m = z \ m = nat z \ 0 \ z" by (auto simp add: nat_eq_iff2) lemma zero_less_nat_eq [simp]: "0 < nat z \ 0 < z" using zless_nat_conj [of 0] by auto lemma nat_add_distrib: "0 \ z \ 0 \ z' \ nat (z + z') = nat z + nat z'" by transfer clarsimp lemma nat_diff_distrib': "0 \ x \ 0 \ y \ nat (x - y) = nat x - nat y" by transfer clarsimp lemma nat_diff_distrib: "0 \ z' \ z' \ z \ nat (z - z') = nat z - nat z'" by (rule nat_diff_distrib') auto lemma nat_zminus_int [simp]: "nat (- int n) = 0" by transfer simp lemma le_nat_iff: "k \ 0 \ n \ nat k \ int n \ k" by transfer auto lemma zless_nat_eq_int_zless: "m < nat z \ int m < z" by transfer (clarsimp simp add: less_diff_conv) lemma (in ring_1) of_nat_nat [simp]: "0 \ z \ of_nat (nat z) = of_int z" by transfer (clarsimp simp add: of_nat_diff) lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')" by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral) lemma nat_abs_triangle_ineq: "nat \k + l\ \ nat \k\ + nat \l\" by (simp add: nat_add_distrib [symmetric] nat_le_eq_zle abs_triangle_ineq) lemma nat_of_bool [simp]: "nat (of_bool P) = of_bool P" by auto lemma split_nat [arith_split]: "P (nat i) \ ((\n. i = int n \ P n) \ (i < 0 \ P 0))" (is "?P = (?L \ ?R)") for i :: int proof (cases "i < 0") case True then show ?thesis by auto next case False have "?P = ?L" proof assume ?P then show ?L using False by auto next assume ?L moreover from False have "int (nat i) = i" by (simp add: not_less) ultimately show ?P by simp qed with False show ?thesis by simp qed lemma all_nat: "(\x. P x) \ (\x\0. P (nat x))" by (auto split: split_nat) lemma ex_nat: "(\x. P x) \ (\x. 0 \ x \ P (nat x))" proof assume "\x. P x" then obtain x where "P x" .. then have "int x \ 0 \ P (nat (int x))" by simp then show "\x\0. P (nat x)" .. next assume "\x\0. P (nat x)" then show "\x. P x" by auto qed text \For termination proofs:\ lemma measure_function_int[measure_function]: "is_measure (nat \ abs)" .. subsection \Lemmas about the Function \<^term>\of_nat\ and Orderings\ lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)" by (simp add: order_less_le del: of_nat_Suc) lemma negative_zless [iff]: "- (int (Suc n)) < int m" by (rule negative_zless_0 [THEN order_less_le_trans], simp) lemma negative_zle_0: "- int n \ 0" by (simp add: minus_le_iff) lemma negative_zle [iff]: "- int n \ int m" by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) lemma not_zle_0_negative [simp]: "\ 0 \ - int (Suc n)" by (subst le_minus_iff) (simp del: of_nat_Suc) lemma int_zle_neg: "int n \ - int m \ n = 0 \ m = 0" by transfer simp lemma not_int_zless_negative [simp]: "\ int n < - int m" by (simp add: linorder_not_less) lemma negative_eq_positive [simp]: "- int n = of_nat m \ n = 0 \ m = 0" by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) lemma zle_iff_zadd: "w \ z \ (\n. z = w + int n)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by auto next assume ?lhs then have "0 \ z - w" by simp then obtain n where "z - w = int n" using zero_le_imp_eq_int [of "z - w"] by blast then have "z = w + int n" by simp then show ?rhs .. qed lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z" by simp text \ This version is proved for all ordered rings, not just integers! It is proved here because attribute \arith_split\ is not available in theory \Rings\. But is it really better than just rewriting with \abs_if\? \ lemma abs_split [arith_split, no_atp]: "P \a\ \ (0 \ a \ P a) \ (a < 0 \ P (- a))" for a :: "'a::linordered_idom" by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) lemma negD: assumes "x < 0" shows "\n. x = - (int (Suc n))" proof - have "\a b. a < b \ \n. Suc (a + n) = b" - by (rule_tac x="b - Suc a" in exI) arith + proof - + fix a b:: nat + assume "a < b" + then have "Suc (a + (b - Suc a)) = b" + by arith + then show "\n. Suc (a + n) = b" + by (rule exI) + qed with assms show ?thesis by transfer auto qed subsection \Cases and induction\ text \ Now we replace the case analysis rule by a more conventional one: whether an integer is negative or not. \ text \This version is symmetric in the two subgoals.\ lemma int_cases2 [case_names nonneg nonpos, cases type: int]: "(\n. z = int n \ P) \ (\n. z = - (int n) \ P) \ P" by (cases "z < 0") (auto simp add: linorder_not_less dest!: negD nat_0_le [THEN sym]) text \This is the default, with a negative case.\ lemma int_cases [case_names nonneg neg, cases type: int]: assumes pos: "\n. z = int n \ P" and neg: "\n. z = - (int (Suc n)) \ P" shows P proof (cases "z < 0") case True with neg show ?thesis by (blast dest!: negD) next case False with pos show ?thesis by (force simp add: linorder_not_less dest: nat_0_le [THEN sym]) qed lemma int_cases3 [case_names zero pos neg]: fixes k :: int assumes "k = 0 \ P" and "\n. k = int n \ n > 0 \ P" and "\n. k = - int n \ n > 0 \ P" shows "P" proof (cases k "0::int" rule: linorder_cases) case equal with assms(1) show P by simp next case greater then have *: "nat k > 0" by simp moreover from * have "k = int (nat k)" by auto ultimately show P using assms(2) by blast next case less then have *: "nat (- k) > 0" by simp moreover from * have "k = - int (nat (- k))" by auto ultimately show P using assms(3) by blast qed lemma int_of_nat_induct [case_names nonneg neg, induct type: int]: "(\n. P (int n)) \ (\n. P (- (int (Suc n)))) \ P z" by (cases z) auto lemma sgn_mult_dvd_iff [simp]: "sgn r * l dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int by (cases r rule: int_cases3) auto lemma mult_sgn_dvd_iff [simp]: "l * sgn r dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int using sgn_mult_dvd_iff [of r l k] by (simp add: ac_simps) lemma dvd_sgn_mult_iff [simp]: "l dvd sgn r * k \ l dvd k \ r = 0" for k l r :: int by (cases r rule: int_cases3) simp_all lemma dvd_mult_sgn_iff [simp]: "l dvd k * sgn r \ l dvd k \ r = 0" for k l r :: int using dvd_sgn_mult_iff [of l r k] by (simp add: ac_simps) lemma int_sgnE: fixes k :: int obtains n and l where "k = sgn l * int n" proof - have "k = sgn k * int (nat \k\)" by (simp add: sgn_mult_abs) then show ?thesis .. qed subsubsection \Binary comparisons\ text \Preliminaries\ lemma le_imp_0_less: fixes z :: int assumes le: "0 \ z" shows "0 < 1 + z" proof - have "0 \ z" by fact also have "\ < z + 1" by (rule less_add_one) also have "\ = 1 + z" by (simp add: ac_simps) finally show "0 < 1 + z" . qed lemma odd_less_0_iff: "1 + z + z < 0 \ z < 0" for z :: int proof (cases z) case (nonneg n) then show ?thesis by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le]) next case (neg n) then show ?thesis by (simp del: of_nat_Suc of_nat_add of_nat_1 add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric]) qed subsubsection \Comparisons, for Ordered Rings\ lemma odd_nonzero: "1 + z + z \ 0" for z :: int proof (cases z) case (nonneg n) have le: "0 \ z + z" by (simp add: nonneg add_increasing) then show ?thesis using le_imp_0_less [OF le] by (auto simp: ac_simps) next case (neg n) show ?thesis proof assume eq: "1 + z + z = 0" have "0 < 1 + (int n + int n)" by (simp add: le_imp_0_less add_increasing) also have "\ = - (1 + z + z)" by (simp add: neg add.assoc [symmetric]) also have "\ = 0" by (simp add: eq) finally have "0<0" .. then show False by blast qed qed subsection \The Set of Integers\ context ring_1 begin definition Ints :: "'a set" ("\") where "\ = range of_int" lemma Ints_of_int [simp]: "of_int z \ \" by (simp add: Ints_def) lemma Ints_of_nat [simp]: "of_nat n \ \" using Ints_of_int [of "of_nat n"] by simp lemma Ints_0 [simp]: "0 \ \" using Ints_of_int [of "0"] by simp lemma Ints_1 [simp]: "1 \ \" using Ints_of_int [of "1"] by simp lemma Ints_numeral [simp]: "numeral n \ \" by (subst of_nat_numeral [symmetric], rule Ints_of_nat) lemma Ints_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (force simp add: Ints_def simp flip: of_int_add intro: range_eqI) lemma Ints_minus [simp]: "a \ \ \ -a \ \" by (force simp add: Ints_def simp flip: of_int_minus intro: range_eqI) lemma minus_in_Ints_iff: "-x \ \ \ x \ \" using Ints_minus[of x] Ints_minus[of "-x"] by auto lemma Ints_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (force simp add: Ints_def simp flip: of_int_diff intro: range_eqI) lemma Ints_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (force simp add: Ints_def simp flip: of_int_mult intro: range_eqI) lemma Ints_power [simp]: "a \ \ \ a ^ n \ \" by (induct n) simp_all lemma Ints_cases [cases set: Ints]: assumes "q \ \" obtains (of_int) z where "q = of_int z" unfolding Ints_def proof - from \q \ \\ have "q \ range of_int" unfolding Ints_def . then obtain z where "q = of_int z" .. then show thesis .. qed lemma Ints_induct [case_names of_int, induct set: Ints]: "q \ \ \ (\z. P (of_int z)) \ P q" by (rule Ints_cases) auto lemma Nats_subset_Ints: "\ \ \" unfolding Nats_def Ints_def by (rule subsetI, elim imageE, hypsubst, subst of_int_of_nat_eq[symmetric], rule imageI) simp_all lemma Nats_altdef1: "\ = {of_int n |n. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {of_int n |n. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp next fix x :: 'a assume "x \ \" then obtain n where "x = of_nat n" by (auto elim!: Nats_cases) then have "x = of_int (int n)" by simp also have "int n \ 0" by simp then have "of_int (int n) \ {of_int n |n. n \ 0}" by blast finally show "x \ {of_int n |n. n \ 0}" . qed end lemma Ints_sum [intro]: "(\x. x \ A \ f x \ \) \ sum f A \ \" by (induction A rule: infinite_finite_induct) auto lemma Ints_prod [intro]: "(\x. x \ A \ f x \ \) \ prod f A \ \" by (induction A rule: infinite_finite_induct) auto lemma (in linordered_idom) Ints_abs [simp]: shows "a \ \ \ abs a \ \" by (auto simp: abs_if) lemma (in linordered_idom) Nats_altdef2: "\ = {n \ \. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {n \ \. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp qed (auto elim!: Nats_cases) lemma (in idom_divide) of_int_divide_in_Ints: "of_int a div of_int b \ \" if "b dvd a" proof - from that obtain c where "a = b * c" .. then show ?thesis by (cases "of_int b = 0") simp_all qed text \The premise involving \<^term>\Ints\ prevents \<^term>\a = 1/2\.\ lemma Ints_double_eq_0_iff: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "a + a = 0 \ a = 0" (is "?lhs \ ?rhs") proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume ?rhs then show ?lhs by simp next assume ?lhs with a have "of_int (z + z) = (of_int 0 :: 'a)" by simp then have "z + z = 0" by (simp only: of_int_eq_iff) then have "z = 0" by (simp only: double_zero) with a show ?rhs by simp qed qed lemma Ints_odd_nonzero: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "1 + a + a \ 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume "1 + a + a = 0" with a have "of_int (1 + z + z) = (of_int 0 :: 'a)" by simp then have "1 + z + z = 0" by (simp only: of_int_eq_iff) with odd_nonzero show False by blast qed qed lemma Nats_numeral [simp]: "numeral w \ \" using of_nat_in_Nats [of "numeral w"] by simp lemma Ints_odd_less_0: fixes a :: "'a::linordered_idom" assumes in_Ints: "a \ \" shows "1 + a + a < 0 \ a < 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. with a have "1 + a + a < 0 \ of_int (1 + z + z) < (of_int 0 :: 'a)" by simp also have "\ \ z < 0" by (simp only: of_int_less_iff odd_less_0_iff) also have "\ \ a < 0" by (simp add: a) finally show ?thesis . qed subsection \\<^term>\sum\ and \<^term>\prod\\ context semiring_1 begin lemma of_nat_sum [simp]: "of_nat (sum f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context ring_1 begin lemma of_int_sum [simp]: "of_int (sum f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_semiring_1 begin lemma of_nat_prod [simp]: "of_nat (prod f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_ring_1 begin lemma of_int_prod [simp]: "of_int (prod f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end subsection \Setting up simplification procedures\ ML_file \Tools/int_arith.ML\ declaration \K ( Lin_Arith.add_discrete_type \<^type_name>\Int.int\ #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle} #> Lin_Arith.add_inj_thms @{thms of_nat_le_iff [THEN iffD2] of_nat_eq_iff [THEN iffD2]} #> Lin_Arith.add_inj_const (\<^const_name>\of_nat\, \<^typ>\nat \ int\) #> Lin_Arith.add_simps @{thms of_int_0 of_int_1 of_int_add of_int_mult of_int_numeral of_int_neg_numeral nat_0 nat_1 diff_nat_numeral nat_numeral neg_less_iff_less True_implies_equals distrib_left [where a = "numeral v" for v] distrib_left [where a = "- numeral v" for v] div_by_1 div_0 times_divide_eq_right times_divide_eq_left minus_divide_left [THEN sym] minus_divide_right [THEN sym] add_divide_distrib diff_divide_distrib of_int_minus of_int_diff of_int_of_nat_eq} #> Lin_Arith.add_simprocs [Int_Arith.zero_one_idom_simproc] )\ simproc_setup fast_arith ("(m::'a::linordered_idom) < n" | "(m::'a::linordered_idom) \ n" | "(m::'a::linordered_idom) = n") = \K Lin_Arith.simproc\ subsection\More Inequality Reasoning\ lemma zless_add1_eq: "w < z + 1 \ w < z \ w = z" for w z :: int by arith lemma add1_zle_eq: "w + 1 \ z \ w < z" for w z :: int by arith lemma zle_diff1_eq [simp]: "w \ z - 1 \ w < z" for w z :: int by arith lemma zle_add1_eq_le [simp]: "w < z + 1 \ w \ z" for w z :: int by arith lemma int_one_le_iff_zero_less: "1 \ z \ 0 < z" for z :: int by arith lemma Ints_nonzero_abs_ge1: fixes x:: "'a :: linordered_idom" assumes "x \ Ints" "x \ 0" shows "1 \ abs x" proof (rule Ints_cases [OF \x \ Ints\]) fix z::int assume "x = of_int z" - with \x \ 0\ + with \x \ 0\ show "1 \ \x\" - apply (auto simp add: abs_if) + apply (auto simp: abs_if) by (metis diff_0 of_int_1 of_int_le_iff of_int_minus zle_diff1_eq) qed lemma Ints_nonzero_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; abs x < 1\ \ x = 0" using Ints_nonzero_abs_ge1 [of x] by auto lemma Ints_eq_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; y \ Ints\ \ x = y \ abs (x-y) < 1" using eq_iff_diff_eq_0 by (fastforce intro: Ints_nonzero_abs_less1) subsection \The functions \<^term>\nat\ and \<^term>\int\\ text \Simplify the term \<^term>\w + - z\.\ lemma one_less_nat_eq [simp]: "Suc 0 < nat z \ 1 < z" using zless_nat_conj [of 1 z] by auto lemma int_eq_iff_numeral [simp]: "int m = numeral v \ m = numeral v" by (simp add: int_eq_iff) lemma nat_abs_int_diff: "nat \int a - int b\ = (if a \ b then b - a else a - b)" by auto lemma nat_int_add: "nat (int a + int b) = a + b" by auto context ring_1 begin lemma of_int_of_nat [nitpick_simp]: "of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))" proof (cases "k < 0") case True then have "0 \ - k" by simp then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat) with True show ?thesis by simp next case False then show ?thesis by (simp add: not_less) qed end lemma transfer_rule_of_int: includes lifting_syntax fixes R :: "'a::ring_1 \ 'b::ring_1 \ bool" assumes [transfer_rule]: "R 0 0" "R 1 1" "(R ===> R ===> R) (+) (+)" "(R ===> R) uminus uminus" shows "((=) ===> R) of_int of_int" proof - note assms note transfer_rule_of_nat [transfer_rule] have [transfer_rule]: "((=) ===> R) of_nat of_nat" by transfer_prover show ?thesis by (unfold of_int_of_nat [abs_def]) transfer_prover qed lemma nat_mult_distrib: fixes z z' :: int assumes "0 \ z" shows "nat (z * z') = nat z * nat z'" proof (cases "0 \ z'") case False with assms have "z * z' \ 0" by (simp add: not_le mult_le_0_iff) then have "nat (z * z') = 0" by simp moreover from False have "nat z' = 0" by simp ultimately show ?thesis by simp next case True with assms have ge_0: "z * z' \ 0" by (simp add: zero_le_mult_iff) show ?thesis by (rule injD [of "of_nat :: nat \ int", OF inj_of_nat]) (simp only: of_nat_mult of_nat_nat [OF True] of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp) qed lemma nat_mult_distrib_neg: assumes "z \ (0::int)" shows "nat (z * z') = nat (- z) * nat (- z')" (is "?L = ?R") proof - have "?L = nat (- z * - z')" using assms by auto also have "... = ?R" by (rule nat_mult_distrib) (use assms in auto) finally show ?thesis . qed lemma nat_abs_mult_distrib: "nat \w * z\ = nat \w\ * nat \z\" by (cases "z = 0 \ w = 0") (auto simp add: abs_if nat_mult_distrib [symmetric] nat_mult_distrib_neg [symmetric] mult_less_0_iff) lemma int_in_range_abs [simp]: "int n \ range abs" proof (rule range_eqI) show "int n = \int n\" by simp qed lemma range_abs_Nats [simp]: "range abs = (\ :: int set)" proof - have "\k\ \ \" for k :: int by (cases k) simp_all moreover have "k \ range abs" if "k \ \" for k :: int using that by induct simp ultimately show ?thesis by blast qed lemma Suc_nat_eq_nat_zadd1: "0 \ z \ Suc (nat z) = nat (1 + z)" for z :: int by (rule sym) (simp add: nat_eq_iff) lemma diff_nat_eq_if: "nat z - nat z' = (if z' < 0 then nat z else let d = z - z' in if d < 0 then 0 else nat d)" by (simp add: Let_def nat_diff_distrib [symmetric]) lemma nat_numeral_diff_1 [simp]: "numeral v - (1::nat) = nat (numeral v - 1)" using diff_nat_numeral [of v Num.One] by simp subsection \Induction principles for int\ text \Well-founded segments of the integers.\ definition int_ge_less_than :: "int \ (int \ int) set" where "int_ge_less_than d = {(z', z). d \ z' \ z' < z}" lemma wf_int_ge_less_than: "wf (int_ge_less_than d)" proof - have "int_ge_less_than d \ measure (\z. nat (z - d))" by (auto simp add: int_ge_less_than_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed text \ This variant looks odd, but is typical of the relations suggested by RankFinder.\ definition int_ge_less_than2 :: "int \ (int \ int) set" where "int_ge_less_than2 d = {(z',z). d \ z \ z' < z}" lemma wf_int_ge_less_than2: "wf (int_ge_less_than2 d)" proof - have "int_ge_less_than2 d \ measure (\z. nat (1 + z - d))" by (auto simp add: int_ge_less_than2_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed (* `set:int': dummy construction *) theorem int_ge_induct [case_names base step, induct set: int]: fixes i :: int assumes ge: "k \ i" and base: "P k" and step: "\i. k \ i \ P i \ P (i + 1)" shows "P i" proof - have "\i::int. n = nat (i - k) \ k \ i \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat ((i - 1) - k)" by arith moreover have k: "k \ i - 1" using Suc.prems by arith ultimately have "P (i - 1)" by (rule Suc.hyps) from step [OF k this] show ?case by simp qed with ge show ?thesis by fast qed (* `set:int': dummy construction *) theorem int_gr_induct [case_names base step, induct set: int]: fixes i k :: int assumes "k < i" "P (k + 1)" "\i. k < i \ P i \ P (i + 1)" shows "P i" proof - have "k+1 \ i" using assms by auto then show ?thesis by (induction i rule: int_ge_induct) (auto simp: assms) qed theorem int_le_induct [consumes 1, case_names base step]: fixes i k :: int assumes le: "i \ k" and base: "P k" and step: "\i. i \ k \ P i \ P (i - 1)" shows "P i" proof - have "\i::int. n = nat(k-i) \ i \ k \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat (k - (i + 1))" by arith moreover have k: "i + 1 \ k" using Suc.prems by arith ultimately have "P (i + 1)" by (rule Suc.hyps) from step[OF k this] show ?case by simp qed with le show ?thesis by fast qed theorem int_less_induct [consumes 1, case_names base step]: fixes i k :: int assumes "i < k" "P (k - 1)" "\i. i < k \ P i \ P (i - 1)" shows "P i" proof - have "i \ k-1" using assms by auto then show ?thesis by (induction i rule: int_le_induct) (auto simp: assms) qed theorem int_induct [case_names base step1 step2]: fixes k :: int assumes base: "P k" and step1: "\i. k \ i \ P i \ P (i + 1)" and step2: "\i. k \ i \ P i \ P (i - 1)" shows "P i" proof - have "i \ k \ i \ k" by arith then show ?thesis proof assume "i \ k" then show ?thesis using base by (rule int_ge_induct) (fact step1) next assume "i \ k" then show ?thesis using base by (rule int_le_induct) (fact step2) qed qed subsection \Intermediate value theorems\ lemma nat_ivt_aux: "\\if (Suc i) - f i\ \ 1; f 0 \ k; k \ f n\ \ \i \ n. f i = k" for m n :: nat and k :: int proof (induct n) case (Suc n) show ?case proof (cases "k = f (Suc n)") case False with Suc have "k \ f n" by auto with Suc show ?thesis by (auto simp add: abs_if split: if_split_asm intro: le_SucI) qed (use Suc in auto) qed auto lemma nat_intermed_int_val: fixes m n :: nat and k :: int assumes "\i. m \ i \ i < n \ \f (Suc i) - f i\ \ 1" "m \ n" "f m \ k" "k \ f n" shows "\i. m \ i \ i \ n \ f i = k" proof - obtain i where "i \ n - m" "k = f (m + i)" using nat_ivt_aux [of "n - m" "f \ plus m" k] assms by auto with assms show ?thesis - by (rule_tac x = "m + i" in exI) auto + using exI[of _ "m + i"] by auto qed lemma nat0_intermed_int_val: "\i\n. f i = k" if "\if (i + 1) - f i\ \ 1" "f 0 \ k" "k \ f n" for n :: nat and k :: int using nat_intermed_int_val [of 0 n f k] that by auto subsection \Products and 1, by T. M. Rasmussen\ lemma abs_zmult_eq_1: fixes m n :: int assumes mn: "\m * n\ = 1" shows "\m\ = 1" proof - from mn have 0: "m \ 0" "n \ 0" by auto have "\ 2 \ \m\" proof assume "2 \ \m\" then have "2 * \n\ \ \m\ * \n\" by (simp add: mult_mono 0) also have "\ = \m * n\" by (simp add: abs_mult) also from mn have "\ = 1" by simp finally have "2 * \n\ \ 1" . with 0 show "False" by arith qed with 0 show ?thesis by auto qed lemma pos_zmult_eq_1_iff_lemma: "m * n = 1 \ m = 1 \ m = - 1" for m n :: int using abs_zmult_eq_1 [of m n] by arith lemma pos_zmult_eq_1_iff: fixes m n :: int assumes "0 < m" shows "m * n = 1 \ m = 1 \ n = 1" proof - from assms have "m * n = 1 \ m = 1" by (auto dest: pos_zmult_eq_1_iff_lemma) then show ?thesis by (auto dest: pos_zmult_eq_1_iff_lemma) qed lemma zmult_eq_1_iff: "m * n = 1 \ (m = 1 \ n = 1) \ (m = - 1 \ n = - 1)" (is "?L = ?R") for m n :: int proof assume L: ?L show ?R using pos_zmult_eq_1_iff_lemma [OF L] L by force qed auto lemma infinite_UNIV_int [simp]: "\ finite (UNIV::int set)" proof assume "finite (UNIV::int set)" moreover have "inj (\i::int. 2 * i)" by (rule injI) simp ultimately have "surj (\i::int. 2 * i)" by (rule finite_UNIV_inj_surj) then obtain i :: int where "1 = 2 * i" by (rule surjE) then show False by (simp add: pos_zmult_eq_1_iff) qed subsection \The divides relation\ lemma zdvd_antisym_nonneg: "0 \ m \ 0 \ n \ m dvd n \ n dvd m \ m = n" for m n :: int by (auto simp add: dvd_def mult.assoc zero_le_mult_iff zmult_eq_1_iff) lemma zdvd_antisym_abs: fixes a b :: int assumes "a dvd b" and "b dvd a" shows "\a\ = \b\" proof (cases "a = 0") case True with assms show ?thesis by simp next case False from \a dvd b\ obtain k where k: "b = a * k" unfolding dvd_def by blast from \b dvd a\ obtain k' where k': "a = b * k'" unfolding dvd_def by blast from k k' have "a = a * k * k'" by simp with mult_cancel_left1[where c="a" and b="k*k'"] have kk': "k * k' = 1" using \a \ 0\ by (simp add: mult.assoc) then have "k = 1 \ k' = 1 \ k = -1 \ k' = -1" by (simp add: zmult_eq_1_iff) with k k' show ?thesis by auto qed lemma zdvd_zdiffD: "k dvd m - n \ k dvd n \ k dvd m" for k m n :: int using dvd_add_right_iff [of k "- n" m] by simp lemma zdvd_reduce: "k dvd n + k * m \ k dvd n" for k m n :: int using dvd_add_times_triv_right_iff [of k n m] by (simp add: ac_simps) lemma dvd_imp_le_int: fixes d i :: int assumes "i \ 0" and "d dvd i" shows "\d\ \ \i\" proof - from \d dvd i\ obtain k where "i = d * k" .. with \i \ 0\ have "k \ 0" by auto then have "1 \ \k\" and "0 \ \d\" by auto then have "\d\ * 1 \ \d\ * \k\" by (rule mult_left_mono) with \i = d * k\ show ?thesis by (simp add: abs_mult) qed lemma zdvd_not_zless: fixes m n :: int assumes "0 < m" and "m < n" shows "\ n dvd m" proof from assms have "0 < n" by auto assume "n dvd m" then obtain k where k: "m = n * k" .. with \0 < m\ have "0 < n * k" by auto with \0 < n\ have "0 < k" by (simp add: zero_less_mult_iff) with k \0 < n\ \m < n\ have "n * k < n * 1" by simp with \0 < n\ \0 < k\ show False unfolding mult_less_cancel_left by auto qed lemma zdvd_mult_cancel: fixes k m n :: int assumes d: "k * m dvd k * n" and "k \ 0" shows "m dvd n" proof - from d obtain h where h: "k * n = k * m * h" unfolding dvd_def by blast have "n = m * h" proof (rule ccontr) assume "\ ?thesis" with \k \ 0\ have "k * n \ k * (m * h)" by simp with h show False by (simp add: mult.assoc) qed then show ?thesis by simp qed lemma int_dvd_int_iff [simp]: "int m dvd int n \ m dvd n" proof - have "m dvd n" if "int n = int m * k" for k proof (cases k) case (nonneg q) with that have "n = m * q" by (simp del: of_nat_mult add: of_nat_mult [symmetric]) then show ?thesis .. next case (neg q) with that have "int n = int m * (- int (Suc q))" by simp also have "\ = - (int m * int (Suc q))" by (simp only: mult_minus_right) also have "\ = - int (m * Suc q)" by (simp only: of_nat_mult [symmetric]) finally have "- int (m * Suc q) = int n" .. then show ?thesis by (simp only: negative_eq_positive) auto qed then show ?thesis by (auto simp add: dvd_def) qed lemma dvd_nat_abs_iff [simp]: "n dvd nat \k\ \ int n dvd k" proof - have "n dvd nat \k\ \ int n dvd int (nat \k\)" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma nat_abs_dvd_iff [simp]: "nat \k\ dvd n \ k dvd int n" proof - have "nat \k\ dvd n \ int (nat \k\) dvd int n" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma zdvd1_eq [simp]: "x dvd 1 \ \x\ = 1" (is "?lhs \ ?rhs") for x :: int proof assume ?lhs then have "nat \x\ dvd nat \1\" by (simp only: nat_abs_dvd_iff) simp then have "nat \x\ = 1" by simp then show ?rhs by (cases "x < 0") simp_all next assume ?rhs then have "x = 1 \ x = - 1" by auto then show ?lhs by (auto intro: dvdI) qed lemma zdvd_mult_cancel1: fixes m :: int assumes mp: "m \ 0" shows "m * n dvd m \ \n\ = 1" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (cases "n > 0") (auto simp add: minus_equation_iff) next assume ?lhs then have "m * n dvd m * 1" by simp from zdvd_mult_cancel[OF this mp] show ?rhs by (simp only: zdvd1_eq) qed lemma nat_dvd_iff: "nat z dvd m \ (if 0 \ z then z dvd int m else m = 0)" using nat_abs_dvd_iff [of z m] by (cases "z \ 0") auto lemma eq_nat_nat_iff: "0 \ z \ 0 \ z' \ nat z = nat z' \ z = z'" by (auto elim: nonneg_int_cases) lemma nat_power_eq: "0 \ z \ nat (z ^ n) = nat z ^ n" by (induct n) (simp_all add: nat_mult_distrib) lemma numeral_power_eq_nat_cancel_iff [simp]: "numeral x ^ n = nat y \ numeral x ^ n = y" using nat_eq_iff2 by auto lemma nat_eq_numeral_power_cancel_iff [simp]: "nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_nat_cancel_iff[of x n y] by (metis (mono_tags)) lemma numeral_power_le_nat_cancel_iff [simp]: "numeral x ^ n \ nat a \ numeral x ^ n \ a" using nat_le_eq_zle[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_le_numeral_power_cancel_iff [simp]: "nat a \ numeral x ^ n \ a \ numeral x ^ n" by (simp add: nat_le_iff) lemma numeral_power_less_nat_cancel_iff [simp]: "numeral x ^ n < nat a \ numeral x ^ n < a" using nat_less_eq_zless[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_less_numeral_power_cancel_iff [simp]: "nat a < numeral x ^ n \ a < numeral x ^ n" using nat_less_eq_zless[of a "numeral x ^ n"] by (cases "a < 0") (auto simp: nat_power_eq less_le_trans[where y=0]) lemma zdvd_imp_le: "z \ n" if "z dvd n" "0 < n" for n z :: int proof (cases n) case (nonneg n) show ?thesis by (cases z) (use nonneg dvd_imp_le that in auto) qed (use that in auto) lemma zdvd_period: fixes a d :: int assumes "a dvd d" shows "a dvd (x + t) \ a dvd ((x + c * d) + t)" (is "?lhs \ ?rhs") proof - from assms have "a dvd (x + t) \ a dvd ((x + t) + c * d)" by (simp add: dvd_add_left_iff) then show ?thesis by (simp add: ac_simps) qed subsection \Powers with integer exponents\ text \ The following allows writing powers with an integer exponent. While the type signature is very generic, most theorems will assume that the underlying type is a division ring or a field. The notation `powi' is inspired by the `powr' notation for real/complex exponentiation. \ definition power_int :: "'a :: {inverse, power} \ int \ 'a" (infixr "powi" 80) where "power_int x n = (if n \ 0 then x ^ nat n else inverse x ^ (nat (-n)))" lemma power_int_0_right [simp]: "power_int x 0 = 1" and power_int_1_right [simp]: "power_int (y :: 'a :: {power, inverse, monoid_mult}) 1 = y" and power_int_minus1_right [simp]: "power_int (y :: 'a :: {power, inverse, monoid_mult}) (-1) = inverse y" by (simp_all add: power_int_def) lemma power_int_of_nat [simp]: "power_int x (int n) = x ^ n" by (simp add: power_int_def) lemma power_int_numeral [simp]: "power_int x (numeral n) = x ^ numeral n" by (simp add: power_int_def) lemma int_cases4 [case_names nonneg neg]: fixes m :: int obtains n where "m = int n" | n where "n > 0" "m = -int n" proof (cases "m \ 0") case True thus ?thesis using that(1)[of "nat m"] by auto next case False thus ?thesis using that(2)[of "nat (-m)"] by auto qed context assumes "SORT_CONSTRAINT('a::division_ring)" begin lemma power_int_minus: "power_int (x::'a) (-n) = inverse (power_int x n)" by (auto simp: power_int_def power_inverse) lemma power_int_eq_0_iff [simp]: "power_int (x::'a) n = 0 \ x = 0 \ n \ 0" by (auto simp: power_int_def) lemma power_int_0_left_If: "power_int (0 :: 'a) m = (if m = 0 then 1 else 0)" by (auto simp: power_int_def) lemma power_int_0_left [simp]: "m \ 0 \ power_int (0 :: 'a) m = 0" by (simp add: power_int_0_left_If) lemma power_int_1_left [simp]: "power_int 1 n = (1 :: 'a :: division_ring)" by (auto simp: power_int_def) lemma power_diff_conv_inverse: "x \ 0 \ m \ n \ (x :: 'a) ^ (n - m) = x ^ n * inverse x ^ m" by (simp add: field_simps flip: power_add) lemma power_mult_inverse_distrib: "x ^ m * inverse (x :: 'a) = inverse x * x ^ m" proof (cases "x = 0") case [simp]: False show ?thesis proof (cases m) case (Suc m') have "x ^ Suc m' * inverse x = x ^ m'" by (subst power_Suc2) (auto simp: mult.assoc) also have "\ = inverse x * x ^ Suc m'" by (subst power_Suc) (auto simp: mult.assoc [symmetric]) finally show ?thesis using Suc by simp qed auto qed auto lemma power_mult_power_inverse_commute: "x ^ m * inverse (x :: 'a) ^ n = inverse x ^ n * x ^ m" proof (induction n) case (Suc n) have "x ^ m * inverse x ^ Suc n = (x ^ m * inverse x ^ n) * inverse x" by (simp only: power_Suc2 mult.assoc) also have "x ^ m * inverse x ^ n = inverse x ^ n * x ^ m" by (rule Suc) also have "\ * inverse x = (inverse x ^ n * inverse x) * x ^ m" by (simp add: mult.assoc power_mult_inverse_distrib) also have "\ = inverse x ^ (Suc n) * x ^ m" by (simp only: power_Suc2) finally show ?case . qed auto lemma power_int_add: assumes "x \ 0 \ m + n \ 0" shows "power_int (x::'a) (m + n) = power_int x m * power_int x n" proof (cases "x = 0") case True thus ?thesis using assms by (auto simp: power_int_0_left_If) next case [simp]: False show ?thesis proof (cases m n rule: int_cases4[case_product int_cases4]) case (nonneg_nonneg a b) thus ?thesis by (auto simp: power_int_def nat_add_distrib power_add) next case (nonneg_neg a b) thus ?thesis by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse power_mult_power_inverse_commute) next case (neg_nonneg a b) thus ?thesis by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse power_mult_power_inverse_commute) next case (neg_neg a b) thus ?thesis by (auto simp: power_int_def nat_add_distrib add.commute simp flip: power_add) qed qed lemma power_int_add_1: assumes "x \ 0 \ m \ -1" shows "power_int (x::'a) (m + 1) = power_int x m * x" using assms by (subst power_int_add) auto lemma power_int_add_1': assumes "x \ 0 \ m \ -1" shows "power_int (x::'a) (m + 1) = x * power_int x m" using assms by (subst add.commute, subst power_int_add) auto lemma power_int_commutes: "power_int (x :: 'a) n * x = x * power_int x n" by (cases "x = 0") (auto simp flip: power_int_add_1 power_int_add_1') lemma power_int_inverse [field_simps, field_split_simps, divide_simps]: "power_int (inverse (x :: 'a)) n = inverse (power_int x n)" by (auto simp: power_int_def power_inverse) lemma power_int_mult: "power_int (x :: 'a) (m * n) = power_int (power_int x m) n" by (auto simp: power_int_def zero_le_mult_iff simp flip: power_mult power_inverse nat_mult_distrib) end context assumes "SORT_CONSTRAINT('a::field)" begin lemma power_int_diff: assumes "x \ 0 \ m \ n" shows "power_int (x::'a) (m - n) = power_int x m / power_int x n" using power_int_add[of x m "-n"] assms by (auto simp: field_simps power_int_minus) lemma power_int_minus_mult: "x \ 0 \ n \ 0 \ power_int (x :: 'a) (n - 1) * x = power_int x n" by (auto simp flip: power_int_add_1) lemma power_int_mult_distrib: "power_int (x * y :: 'a) m = power_int x m * power_int y m" by (auto simp: power_int_def power_mult_distrib) lemmas power_int_mult_distrib_numeral1 = power_int_mult_distrib [where x = "numeral w" for w, simp] lemmas power_int_mult_distrib_numeral2 = power_int_mult_distrib [where y = "numeral w" for w, simp] lemma power_int_divide_distrib: "power_int (x / y :: 'a) m = power_int x m / power_int y m" using power_int_mult_distrib[of x "inverse y" m] unfolding power_int_inverse by (simp add: field_simps) end lemma power_int_add_numeral [simp]: "power_int x (numeral m) * power_int x (numeral n) = power_int x (numeral (m + n))" for x :: "'a :: division_ring" by (simp add: power_int_add [symmetric]) lemma power_int_add_numeral2 [simp]: "power_int x (numeral m) * (power_int x (numeral n) * b) = power_int x (numeral (m + n)) * b" for x :: "'a :: division_ring" by (simp add: mult.assoc [symmetric]) lemma power_int_mult_numeral [simp]: "power_int (power_int x (numeral m)) (numeral n) = power_int x (numeral (m * n))" for x :: "'a :: division_ring" by (simp only: numeral_mult power_int_mult) lemma power_int_not_zero: "(x :: 'a :: division_ring) \ 0 \ n = 0 \ power_int x n \ 0" by (subst power_int_eq_0_iff) auto lemma power_int_one_over [field_simps, field_split_simps, divide_simps]: "power_int (1 / x :: 'a :: division_ring) n = 1 / power_int x n" using power_int_inverse[of x] by (simp add: divide_inverse) context assumes "SORT_CONSTRAINT('a :: linordered_field)" begin lemma power_int_numeral_neg_numeral [simp]: "power_int (numeral m) (-numeral n) = (inverse (numeral (Num.pow m n)) :: 'a)" by (simp add: power_int_minus) lemma zero_less_power_int [simp]: "0 < (x :: 'a) \ 0 < power_int x n" by (auto simp: power_int_def) lemma zero_le_power_int [simp]: "0 \ (x :: 'a) \ 0 \ power_int x n" by (auto simp: power_int_def) lemma power_int_mono: "(x :: 'a) \ y \ n \ 0 \ 0 \ x \ power_int x n \ power_int y n" by (cases n rule: int_cases4) (auto intro: power_mono) lemma one_le_power_int [simp]: "1 \ (x :: 'a) \ n \ 0 \ 1 \ power_int x n" using power_int_mono [of 1 x n] by simp lemma power_int_le_one: "0 \ (x :: 'a) \ n \ 0 \ x \ 1 \ power_int x n \ 1" using power_int_mono [of x 1 n] by simp lemma power_int_le_imp_le_exp: assumes gt1: "1 < (x :: 'a :: linordered_field)" assumes "power_int x m \ power_int x n" "n \ 0" shows "m \ n" proof (cases "m < 0") case True with \n \ 0\ show ?thesis by simp next case False with assms have "x ^ nat m \ x ^ nat n" by (simp add: power_int_def) from gt1 and this show ?thesis using False \n \ 0\ by auto qed lemma power_int_le_imp_less_exp: assumes gt1: "1 < (x :: 'a :: linordered_field)" assumes "power_int x m < power_int x n" "n \ 0" shows "m < n" proof (cases "m < 0") case True with \n \ 0\ show ?thesis by simp next case False with assms have "x ^ nat m < x ^ nat n" by (simp add: power_int_def) from gt1 and this show ?thesis using False \n \ 0\ by auto qed lemma power_int_strict_mono: "(a :: 'a :: linordered_field) < b \ 0 \ a \ 0 < n \ power_int a n < power_int b n" by (auto simp: power_int_def intro!: power_strict_mono) lemma power_int_mono_iff [simp]: fixes a b :: "'a :: linordered_field" shows "\a \ 0; b \ 0; n > 0\ \ power_int a n \ power_int b n \ a \ b" by (auto simp: power_int_def intro!: power_strict_mono) lemma power_int_strict_increasing: fixes a :: "'a :: linordered_field" assumes "n < N" "1 < a" shows "power_int a N > power_int a n" proof - have *: "a ^ nat (N - n) > a ^ 0" using assms by (intro power_strict_increasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ > power_int a n * 1" using assms * by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_increasing: fixes a :: "'a :: linordered_field" assumes "n \ N" "a \ 1" shows "power_int a N \ power_int a n" proof - have *: "a ^ nat (N - n) \ a ^ 0" using assms by (intro power_increasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ \ power_int a n * 1" using assms * by (intro mult_left_mono) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_strict_decreasing: fixes a :: "'a :: linordered_field" assumes "n < N" "0 < a" "a < 1" shows "power_int a N < power_int a n" proof - have *: "a ^ nat (N - n) < a ^ 0" using assms by (intro power_strict_decreasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ < power_int a n * 1" using assms * by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_decreasing: fixes a :: "'a :: linordered_field" assumes "n \ N" "0 \ a" "a \ 1" "a \ 0 \ N \ 0 \ n = 0" shows "power_int a N \ power_int a n" proof (cases "a = 0") case False have *: "a ^ nat (N - n) \ a ^ 0" using assms by (intro power_decreasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms False by (simp flip: power_int_add) also have "\ \ power_int a n * 1" using assms * by (intro mult_left_mono) (auto simp: power_int_def) finally show ?thesis by simp qed (use assms in \auto simp: power_int_0_left_If\) lemma one_less_power_int: "1 < (a :: 'a) \ 0 < n \ 1 < power_int a n" using power_int_strict_increasing[of 0 n a] by simp lemma power_int_abs: "\power_int a n :: 'a\ = power_int \a\ n" by (auto simp: power_int_def power_abs) lemma power_int_sgn [simp]: "sgn (power_int a n :: 'a) = power_int (sgn a) n" by (auto simp: power_int_def) lemma abs_power_int_minus [simp]: "\power_int (- a) n :: 'a\ = \power_int a n\" by (simp add: power_int_abs) lemma power_int_strict_antimono: assumes "(a :: 'a :: linordered_field) < b" "0 < a" "n < 0" shows "power_int a n > power_int b n" proof - have "inverse (power_int a (-n)) > inverse (power_int b (-n))" using assms by (intro less_imp_inverse_less power_int_strict_mono zero_less_power_int) auto thus ?thesis by (simp add: power_int_minus) qed lemma power_int_antimono: assumes "(a :: 'a :: linordered_field) \ b" "0 < a" "n < 0" shows "power_int a n \ power_int b n" using power_int_strict_antimono[of a b n] assms by (cases "a = b") auto end subsection \Finiteness of intervals\ lemma finite_interval_int1 [iff]: "finite {i :: int. a \ i \ i \ b}" proof (cases "a \ b") case True then show ?thesis proof (induct b rule: int_ge_induct) case base have "{i. a \ i \ i \ a} = {a}" by auto then show ?case by simp next case (step b) then have "{i. a \ i \ i \ b + 1} = {i. a \ i \ i \ b} \ {b + 1}" by auto with step show ?case by simp qed next case False then show ?thesis by (metis (lifting, no_types) Collect_empty_eq finite.emptyI order_trans) qed lemma finite_interval_int2 [iff]: "finite {i :: int. a \ i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int3 [iff]: "finite {i :: int. a < i \ i \ b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int4 [iff]: "finite {i :: int. a < i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto subsection \Configuration of the code generator\ text \Constructors\ definition Pos :: "num \ int" where [simp, code_abbrev]: "Pos = numeral" definition Neg :: "num \ int" where [simp, code_abbrev]: "Neg n = - (Pos n)" code_datatype "0::int" Pos Neg text \Auxiliary operations.\ definition dup :: "int \ int" where [simp]: "dup k = k + k" lemma dup_code [code]: "dup 0 = 0" "dup (Pos n) = Pos (Num.Bit0 n)" "dup (Neg n) = Neg (Num.Bit0 n)" by (simp_all add: numeral_Bit0) definition sub :: "num \ num \ int" where [simp]: "sub m n = numeral m - numeral n" lemma sub_code [code]: "sub Num.One Num.One = 0" "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" by (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM) lemma sub_BitM_One_eq: \(Num.sub (Num.BitM n) num.One) = 2 * (Num.sub n Num.One :: int)\ by (cases n) simp_all text \Implementations.\ lemma one_int_code [code]: "1 = Pos Num.One" by simp lemma plus_int_code [code]: "k + 0 = k" "0 + l = l" "Pos m + Pos n = Pos (m + n)" "Pos m + Neg n = sub m n" "Neg m + Pos n = sub n m" "Neg m + Neg n = Neg (m + n)" for k l :: int by simp_all lemma uminus_int_code [code]: "uminus 0 = (0::int)" "uminus (Pos m) = Neg m" "uminus (Neg m) = Pos m" by simp_all lemma minus_int_code [code]: "k - 0 = k" "0 - l = uminus l" "Pos m - Pos n = sub m n" "Pos m - Neg n = Pos (m + n)" "Neg m - Pos n = Neg (m + n)" "Neg m - Neg n = sub n m" for k l :: int by simp_all lemma times_int_code [code]: "k * 0 = 0" "0 * l = 0" "Pos m * Pos n = Pos (m * n)" "Pos m * Neg n = Neg (m * n)" "Neg m * Pos n = Neg (m * n)" "Neg m * Neg n = Pos (m * n)" for k l :: int by simp_all instantiation int :: equal begin definition "HOL.equal k l \ k = (l::int)" instance by standard (rule equal_int_def) end lemma equal_int_code [code]: "HOL.equal 0 (0::int) \ True" "HOL.equal 0 (Pos l) \ False" "HOL.equal 0 (Neg l) \ False" "HOL.equal (Pos k) 0 \ False" "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" "HOL.equal (Pos k) (Neg l) \ False" "HOL.equal (Neg k) 0 \ False" "HOL.equal (Neg k) (Pos l) \ False" "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" by (auto simp add: equal) lemma equal_int_refl [code nbe]: "HOL.equal k k \ True" for k :: int by (fact equal_refl) lemma less_eq_int_code [code]: "0 \ (0::int) \ True" "0 \ Pos l \ True" "0 \ Neg l \ False" "Pos k \ 0 \ False" "Pos k \ Pos l \ k \ l" "Pos k \ Neg l \ False" "Neg k \ 0 \ True" "Neg k \ Pos l \ True" "Neg k \ Neg l \ l \ k" by simp_all lemma less_int_code [code]: "0 < (0::int) \ False" "0 < Pos l \ True" "0 < Neg l \ False" "Pos k < 0 \ False" "Pos k < Pos l \ k < l" "Pos k < Neg l \ False" "Neg k < 0 \ True" "Neg k < Pos l \ True" "Neg k < Neg l \ l < k" by simp_all lemma nat_code [code]: "nat (Int.Neg k) = 0" "nat 0 = 0" "nat (Int.Pos k) = nat_of_num k" by (simp_all add: nat_of_num_numeral) lemma (in ring_1) of_int_code [code]: "of_int (Int.Neg k) = - numeral k" "of_int 0 = 0" "of_int (Int.Pos k) = numeral k" by simp_all text \Serializer setup.\ code_identifier code_module Int \ (SML) Arith and (OCaml) Arith and (Haskell) Arith quickcheck_params [default_type = int] hide_const (open) Pos Neg sub dup text \De-register \int\ as a quotient type:\ lifting_update int.lifting lifting_forget int.lifting subsection \Duplicates\ lemmas int_sum = of_nat_sum [where 'a=int] lemmas int_prod = of_nat_prod [where 'a=int] lemmas zle_int = of_nat_le_iff [where 'a=int] lemmas int_int_eq = of_nat_eq_iff [where 'a=int] lemmas nonneg_eq_int = nonneg_int_cases lemmas double_eq_0_iff = double_zero lemmas int_distrib = distrib_right [of z1 z2 w] distrib_left [of w z1 z2] left_diff_distrib [of z1 z2 w] right_diff_distrib [of w z1 z2] for z1 z2 w :: int end diff --git a/src/HOL/Lattices_Big.thy b/src/HOL/Lattices_Big.thy --- a/src/HOL/Lattices_Big.thy +++ b/src/HOL/Lattices_Big.thy @@ -1,1071 +1,1077 @@ (* Title: HOL/Lattices_Big.thy Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel with contributions by Jeremy Avigad *) section \Big infimum (minimum) and supremum (maximum) over finite (non-empty) sets\ theory Lattices_Big imports Groups_Big Option begin subsection \Generic lattice operations over a set\ subsubsection \Without neutral element\ locale semilattice_set = semilattice begin interpretation comp_fun_idem f by standard (simp_all add: fun_eq_iff left_commute) definition F :: "'a set \ 'a" where eq_fold': "F A = the (Finite_Set.fold (\x y. Some (case y of None \ x | Some z \ f x z)) None A)" lemma eq_fold: assumes "finite A" shows "F (insert x A) = Finite_Set.fold f x A" proof (rule sym) let ?f = "\x y. Some (case y of None \ x | Some z \ f x z)" interpret comp_fun_idem "?f" by standard (simp_all add: fun_eq_iff commute left_commute split: option.split) from assms show "Finite_Set.fold f x A = F (insert x A)" proof induct case empty then show ?case by (simp add: eq_fold') next case (insert y B) then show ?case by (simp add: insert_commute [of x] eq_fold') qed qed lemma singleton [simp]: "F {x} = x" by (simp add: eq_fold) lemma insert_not_elem: assumes "finite A" and "x \ A" and "A \ {}" shows "F (insert x A) = x \<^bold>* F A" proof - from \A \ {}\ obtain b where "b \ A" by blast then obtain B where *: "A = insert b B" "b \ B" by (blast dest: mk_disjoint_insert) with \finite A\ and \x \ A\ have "finite (insert x B)" and "b \ insert x B" by auto then have "F (insert b (insert x B)) = x \<^bold>* F (insert b B)" by (simp add: eq_fold) then show ?thesis by (simp add: * insert_commute) qed lemma in_idem: assumes "finite A" and "x \ A" shows "x \<^bold>* F A = F A" proof - from assms have "A \ {}" by auto with \finite A\ show ?thesis using \x \ A\ by (induct A rule: finite_ne_induct) (auto simp add: ac_simps insert_not_elem) qed lemma insert [simp]: assumes "finite A" and "A \ {}" shows "F (insert x A) = x \<^bold>* F A" using assms by (cases "x \ A") (simp_all add: insert_absorb in_idem insert_not_elem) lemma union: assumes "finite A" "A \ {}" and "finite B" "B \ {}" shows "F (A \ B) = F A \<^bold>* F B" using assms by (induct A rule: finite_ne_induct) (simp_all add: ac_simps) lemma remove: assumes "finite A" and "x \ A" shows "F A = (if A - {x} = {} then x else x \<^bold>* F (A - {x}))" proof - from assms obtain B where "A = insert x B" and "x \ B" by (blast dest: mk_disjoint_insert) with assms show ?thesis by simp qed lemma insert_remove: assumes "finite A" shows "F (insert x A) = (if A - {x} = {} then x else x \<^bold>* F (A - {x}))" using assms by (cases "x \ A") (simp_all add: insert_absorb remove) lemma subset: assumes "finite A" "B \ {}" and "B \ A" shows "F B \<^bold>* F A = F A" proof - from assms have "A \ {}" and "finite B" by (auto dest: finite_subset) with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) qed lemma closed: assumes "finite A" "A \ {}" and elem: "\x y. x \<^bold>* y \ {x, y}" shows "F A \ A" using \finite A\ \A \ {}\ proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case insert with elem show ?case by force qed lemma hom_commute: assumes hom: "\x y. h (x \<^bold>* y) = h x \<^bold>* h y" and N: "finite N" "N \ {}" shows "h (F N) = F (h ` N)" using N proof (induct rule: finite_ne_induct) case singleton thus ?case by simp next case (insert n N) then have "h (F (insert n N)) = h (n \<^bold>* F N)" by simp also have "\ = h n \<^bold>* h (F N)" by (rule hom) also have "h (F N) = F (h ` N)" by (rule insert) also have "h n \<^bold>* \ = F (insert (h n) (h ` N))" using insert by simp also have "insert (h n) (h ` N) = h ` insert n N" by simp finally show ?case . qed lemma infinite: "\ finite A \ F A = the None" unfolding eq_fold' by (cases "finite (UNIV::'a set)") (auto intro: finite_subset fold_infinite) end locale semilattice_order_set = binary?: semilattice_order + semilattice_set begin lemma bounded_iff: assumes "finite A" and "A \ {}" shows "x \<^bold>\ F A \ (\a\A. x \<^bold>\ a)" using assms by (induct rule: finite_ne_induct) simp_all lemma boundedI: assumes "finite A" assumes "A \ {}" assumes "\a. a \ A \ x \<^bold>\ a" shows "x \<^bold>\ F A" using assms by (simp add: bounded_iff) lemma boundedE: assumes "finite A" and "A \ {}" and "x \<^bold>\ F A" obtains "\a. a \ A \ x \<^bold>\ a" using assms by (simp add: bounded_iff) lemma coboundedI: assumes "finite A" and "a \ A" shows "F A \<^bold>\ a" proof - from assms have "A \ {}" by auto from \finite A\ \A \ {}\ \a \ A\ show ?thesis proof (induct rule: finite_ne_induct) case singleton thus ?case by (simp add: refl) next case (insert x B) from insert have "a = x \ a \ B" by simp then show ?case using insert by (auto intro: coboundedI2) qed qed lemma subset_imp: assumes "A \ B" and "A \ {}" and "finite B" shows "F B \<^bold>\ F A" proof (cases "A = B") case True then show ?thesis by (simp add: refl) next case False have B: "B = A \ (B - A)" using \A \ B\ by blast then have "F B = F (A \ (B - A))" by simp also have "\ = F A \<^bold>* F (B - A)" using False assms by (subst union) (auto intro: finite_subset) also have "\ \<^bold>\ F A" by simp finally show ?thesis . qed end subsubsection \With neutral element\ locale semilattice_neutr_set = semilattice_neutr begin interpretation comp_fun_idem f by standard (simp_all add: fun_eq_iff left_commute) definition F :: "'a set \ 'a" where eq_fold: "F A = Finite_Set.fold f \<^bold>1 A" lemma infinite [simp]: "\ finite A \ F A = \<^bold>1" by (simp add: eq_fold) lemma empty [simp]: "F {} = \<^bold>1" by (simp add: eq_fold) lemma insert [simp]: assumes "finite A" shows "F (insert x A) = x \<^bold>* F A" using assms by (simp add: eq_fold) lemma in_idem: assumes "finite A" and "x \ A" shows "x \<^bold>* F A = F A" proof - from assms have "A \ {}" by auto with \finite A\ show ?thesis using \x \ A\ by (induct A rule: finite_ne_induct) (auto simp add: ac_simps) qed lemma union: assumes "finite A" and "finite B" shows "F (A \ B) = F A \<^bold>* F B" using assms by (induct A) (simp_all add: ac_simps) lemma remove: assumes "finite A" and "x \ A" shows "F A = x \<^bold>* F (A - {x})" proof - from assms obtain B where "A = insert x B" and "x \ B" by (blast dest: mk_disjoint_insert) with assms show ?thesis by simp qed lemma insert_remove: assumes "finite A" shows "F (insert x A) = x \<^bold>* F (A - {x})" using assms by (cases "x \ A") (simp_all add: insert_absorb remove) lemma subset: assumes "finite A" and "B \ A" shows "F B \<^bold>* F A = F A" proof - from assms have "finite B" by (auto dest: finite_subset) with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) qed lemma closed: assumes "finite A" "A \ {}" and elem: "\x y. x \<^bold>* y \ {x, y}" shows "F A \ A" using \finite A\ \A \ {}\ proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case insert with elem show ?case by force qed end locale semilattice_order_neutr_set = binary?: semilattice_neutr_order + semilattice_neutr_set begin lemma bounded_iff: assumes "finite A" shows "x \<^bold>\ F A \ (\a\A. x \<^bold>\ a)" using assms by (induct A) simp_all lemma boundedI: assumes "finite A" assumes "\a. a \ A \ x \<^bold>\ a" shows "x \<^bold>\ F A" using assms by (simp add: bounded_iff) lemma boundedE: assumes "finite A" and "x \<^bold>\ F A" obtains "\a. a \ A \ x \<^bold>\ a" using assms by (simp add: bounded_iff) lemma coboundedI: assumes "finite A" and "a \ A" shows "F A \<^bold>\ a" proof - from assms have "A \ {}" by auto from \finite A\ \A \ {}\ \a \ A\ show ?thesis proof (induct rule: finite_ne_induct) case singleton thus ?case by (simp add: refl) next case (insert x B) from insert have "a = x \ a \ B" by simp then show ?case using insert by (auto intro: coboundedI2) qed qed lemma subset_imp: assumes "A \ B" and "finite B" shows "F B \<^bold>\ F A" proof (cases "A = B") case True then show ?thesis by (simp add: refl) next case False have B: "B = A \ (B - A)" using \A \ B\ by blast then have "F B = F (A \ (B - A))" by simp also have "\ = F A \<^bold>* F (B - A)" using False assms by (subst union) (auto intro: finite_subset) also have "\ \<^bold>\ F A" by simp finally show ?thesis . qed end subsection \Lattice operations on finite sets\ context semilattice_inf begin sublocale Inf_fin: semilattice_order_set inf less_eq less defines Inf_fin ("\\<^sub>f\<^sub>i\<^sub>n _" [900] 900) = Inf_fin.F .. end context semilattice_sup begin sublocale Sup_fin: semilattice_order_set sup greater_eq greater defines Sup_fin ("\\<^sub>f\<^sub>i\<^sub>n _" [900] 900) = Sup_fin.F .. end subsection \Infimum and Supremum over non-empty sets\ context lattice begin lemma Inf_fin_le_Sup_fin [simp]: assumes "finite A" and "A \ {}" shows "\\<^sub>f\<^sub>i\<^sub>nA \ \\<^sub>f\<^sub>i\<^sub>nA" proof - from \A \ {}\ obtain a where "a \ A" by blast with \finite A\ have "\\<^sub>f\<^sub>i\<^sub>nA \ a" by (rule Inf_fin.coboundedI) moreover from \finite A\ \a \ A\ have "a \ \\<^sub>f\<^sub>i\<^sub>nA" by (rule Sup_fin.coboundedI) ultimately show ?thesis by (rule order_trans) qed lemma sup_Inf_absorb [simp]: "finite A \ a \ A \ \\<^sub>f\<^sub>i\<^sub>nA \ a = a" by (rule sup_absorb2) (rule Inf_fin.coboundedI) lemma inf_Sup_absorb [simp]: "finite A \ a \ A \ a \ \\<^sub>f\<^sub>i\<^sub>nA = a" by (rule inf_absorb1) (rule Sup_fin.coboundedI) end context distrib_lattice begin lemma sup_Inf1_distrib: assumes "finite A" and "A \ {}" shows "sup x (\\<^sub>f\<^sub>i\<^sub>nA) = \\<^sub>f\<^sub>i\<^sub>n{sup x a|a. a \ A}" using assms by (simp add: image_def Inf_fin.hom_commute [where h="sup x", OF sup_inf_distrib1]) (rule arg_cong [where f="Inf_fin"], blast) lemma sup_Inf2_distrib: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" shows "sup (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB) = \\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \ A \ b \ B}" using A proof (induct rule: finite_ne_induct) case singleton then show ?case by (simp add: sup_Inf1_distrib [OF B]) next case (insert x A) have finB: "finite {sup x b |b. b \ B}" by (rule finite_surj [where f = "sup x", OF B(1)], auto) have finAB: "finite {sup a b |a b. a \ A \ b \ B}" proof - have "{sup a b |a b. a \ A \ b \ B} = (\a\A. \b\B. {sup a b})" by blast thus ?thesis by(simp add: insert(1) B(1)) qed have ne: "{sup a b |a b. a \ A \ b \ B} \ {}" using insert B by blast have "sup (\\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\\<^sub>f\<^sub>i\<^sub>nB) = sup (inf x (\\<^sub>f\<^sub>i\<^sub>nA)) (\\<^sub>f\<^sub>i\<^sub>nB)" using insert by simp also have "\ = inf (sup x (\\<^sub>f\<^sub>i\<^sub>nB)) (sup (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB))" by(rule sup_inf_distrib2) also have "\ = inf (\\<^sub>f\<^sub>i\<^sub>n{sup x b|b. b \ B}) (\\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \ A \ b \ B})" using insert by(simp add:sup_Inf1_distrib[OF B]) also have "\ = \\<^sub>f\<^sub>i\<^sub>n({sup x b |b. b \ B} \ {sup a b |a b. a \ A \ b \ B})" (is "_ = \\<^sub>f\<^sub>i\<^sub>n?M") using B insert by (simp add: Inf_fin.union [OF finB _ finAB ne]) also have "?M = {sup a b |a b. a \ insert x A \ b \ B}" by blast finally show ?case . qed lemma inf_Sup1_distrib: assumes "finite A" and "A \ {}" shows "inf x (\\<^sub>f\<^sub>i\<^sub>nA) = \\<^sub>f\<^sub>i\<^sub>n{inf x a|a. a \ A}" using assms by (simp add: image_def Sup_fin.hom_commute [where h="inf x", OF inf_sup_distrib1]) (rule arg_cong [where f="Sup_fin"], blast) lemma inf_Sup2_distrib: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" shows "inf (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB) = \\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \ A \ b \ B}" using A proof (induct rule: finite_ne_induct) case singleton thus ?case by(simp add: inf_Sup1_distrib [OF B]) next case (insert x A) have finB: "finite {inf x b |b. b \ B}" by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto) have finAB: "finite {inf a b |a b. a \ A \ b \ B}" proof - have "{inf a b |a b. a \ A \ b \ B} = (\a\A. \b\B. {inf a b})" by blast thus ?thesis by(simp add: insert(1) B(1)) qed have ne: "{inf a b |a b. a \ A \ b \ B} \ {}" using insert B by blast have "inf (\\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\\<^sub>f\<^sub>i\<^sub>nB) = inf (sup x (\\<^sub>f\<^sub>i\<^sub>nA)) (\\<^sub>f\<^sub>i\<^sub>nB)" using insert by simp also have "\ = sup (inf x (\\<^sub>f\<^sub>i\<^sub>nB)) (inf (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB))" by(rule inf_sup_distrib2) also have "\ = sup (\\<^sub>f\<^sub>i\<^sub>n{inf x b|b. b \ B}) (\\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \ A \ b \ B})" using insert by(simp add:inf_Sup1_distrib[OF B]) also have "\ = \\<^sub>f\<^sub>i\<^sub>n({inf x b |b. b \ B} \ {inf a b |a b. a \ A \ b \ B})" (is "_ = \\<^sub>f\<^sub>i\<^sub>n?M") using B insert by (simp add: Sup_fin.union [OF finB _ finAB ne]) also have "?M = {inf a b |a b. a \ insert x A \ b \ B}" by blast finally show ?case . qed end context complete_lattice begin lemma Inf_fin_Inf: assumes "finite A" and "A \ {}" shows "\\<^sub>f\<^sub>i\<^sub>nA = \A" proof - from assms obtain b B where "A = insert b B" and "finite B" by auto then show ?thesis by (simp add: Inf_fin.eq_fold inf_Inf_fold_inf inf.commute [of b]) qed lemma Sup_fin_Sup: assumes "finite A" and "A \ {}" shows "\\<^sub>f\<^sub>i\<^sub>nA = \A" proof - from assms obtain b B where "A = insert b B" and "finite B" by auto then show ?thesis by (simp add: Sup_fin.eq_fold sup_Sup_fold_sup sup.commute [of b]) qed end subsection \Minimum and Maximum over non-empty sets\ context linorder begin sublocale Min: semilattice_order_set min less_eq less + Max: semilattice_order_set max greater_eq greater defines Min = Min.F and Max = Max.F .. end syntax "_MIN1" :: "pttrns \ 'b \ 'b" ("(3MIN _./ _)" [0, 10] 10) "_MIN" :: "pttrn \ 'a set \ 'b \ 'b" ("(3MIN _\_./ _)" [0, 0, 10] 10) "_MAX1" :: "pttrns \ 'b \ 'b" ("(3MAX _./ _)" [0, 10] 10) "_MAX" :: "pttrn \ 'a set \ 'b \ 'b" ("(3MAX _\_./ _)" [0, 0, 10] 10) translations "MIN x y. f" \ "MIN x. MIN y. f" "MIN x. f" \ "CONST Min (CONST range (\x. f))" "MIN x\A. f" \ "CONST Min ((\x. f) ` A)" "MAX x y. f" \ "MAX x. MAX y. f" "MAX x. f" \ "CONST Max (CONST range (\x. f))" "MAX x\A. f" \ "CONST Max ((\x. f) ` A)" text \An aside: \<^const>\Min\/\<^const>\Max\ on linear orders as special case of \<^const>\Inf_fin\/\<^const>\Sup_fin\\ lemma Inf_fin_Min: "Inf_fin = (Min :: 'a::{semilattice_inf, linorder} set \ 'a)" by (simp add: Inf_fin_def Min_def inf_min) lemma Sup_fin_Max: "Sup_fin = (Max :: 'a::{semilattice_sup, linorder} set \ 'a)" by (simp add: Sup_fin_def Max_def sup_max) context linorder begin lemma dual_min: "ord.min greater_eq = max" by (auto simp add: ord.min_def max_def fun_eq_iff) lemma dual_max: "ord.max greater_eq = min" by (auto simp add: ord.max_def min_def fun_eq_iff) lemma dual_Min: "linorder.Min greater_eq = Max" proof - interpret dual: linorder greater_eq greater by (fact dual_linorder) show ?thesis by (simp add: dual.Min_def dual_min Max_def) qed lemma dual_Max: "linorder.Max greater_eq = Min" proof - interpret dual: linorder greater_eq greater by (fact dual_linorder) show ?thesis by (simp add: dual.Max_def dual_max Min_def) qed lemmas Min_singleton = Min.singleton lemmas Max_singleton = Max.singleton lemmas Min_insert = Min.insert lemmas Max_insert = Max.insert lemmas Min_Un = Min.union lemmas Max_Un = Max.union lemmas hom_Min_commute = Min.hom_commute lemmas hom_Max_commute = Max.hom_commute lemma Min_in [simp]: assumes "finite A" and "A \ {}" shows "Min A \ A" using assms by (auto simp add: min_def Min.closed) lemma Max_in [simp]: assumes "finite A" and "A \ {}" shows "Max A \ A" using assms by (auto simp add: max_def Max.closed) lemma Min_insert2: assumes "finite A" and min: "\b. b \ A \ a \ b" shows "Min (insert a A) = a" proof (cases "A = {}") case True then show ?thesis by simp next case False with \finite A\ have "Min (insert a A) = min a (Min A)" by simp moreover from \finite A\ \A \ {}\ min have "a \ Min A" by simp ultimately show ?thesis by (simp add: min.absorb1) qed lemma Max_insert2: assumes "finite A" and max: "\b. b \ A \ b \ a" shows "Max (insert a A) = a" proof (cases "A = {}") case True then show ?thesis by simp next case False with \finite A\ have "Max (insert a A) = max a (Max A)" by simp moreover from \finite A\ \A \ {}\ max have "Max A \ a" by simp ultimately show ?thesis by (simp add: max.absorb1) qed lemma Max_const[simp]: "\ finite A; A \ {} \ \ Max ((\_. c) ` A) = c" using Max_in image_is_empty by blast lemma Min_const[simp]: "\ finite A; A \ {} \ \ Min ((\_. c) ` A) = c" using Min_in image_is_empty by blast lemma Min_le [simp]: assumes "finite A" and "x \ A" shows "Min A \ x" using assms by (fact Min.coboundedI) lemma Max_ge [simp]: assumes "finite A" and "x \ A" shows "x \ Max A" using assms by (fact Max.coboundedI) lemma Min_eqI: assumes "finite A" assumes "\y. y \ A \ y \ x" and "x \ A" shows "Min A = x" proof (rule order.antisym) from \x \ A\ have "A \ {}" by auto with assms show "Min A \ x" by simp next from assms show "x \ Min A" by simp qed lemma Max_eqI: assumes "finite A" assumes "\y. y \ A \ y \ x" and "x \ A" shows "Max A = x" proof (rule order.antisym) from \x \ A\ have "A \ {}" by auto with assms show "Max A \ x" by simp next from assms show "x \ Max A" by simp qed lemma eq_Min_iff: "\ finite A; A \ {} \ \ m = Min A \ m \ A \ (\a \ A. m \ a)" by (meson Min_in Min_le order.antisym) lemma Min_eq_iff: "\ finite A; A \ {} \ \ Min A = m \ m \ A \ (\a \ A. m \ a)" by (meson Min_in Min_le order.antisym) lemma eq_Max_iff: "\ finite A; A \ {} \ \ m = Max A \ m \ A \ (\a \ A. a \ m)" by (meson Max_in Max_ge order.antisym) lemma Max_eq_iff: "\ finite A; A \ {} \ \ Max A = m \ m \ A \ (\a \ A. a \ m)" by (meson Max_in Max_ge order.antisym) context fixes A :: "'a set" assumes fin_nonempty: "finite A" "A \ {}" begin lemma Min_ge_iff [simp]: "x \ Min A \ (\a\A. x \ a)" using fin_nonempty by (fact Min.bounded_iff) lemma Max_le_iff [simp]: "Max A \ x \ (\a\A. a \ x)" using fin_nonempty by (fact Max.bounded_iff) lemma Min_gr_iff [simp]: "x < Min A \ (\a\A. x < a)" using fin_nonempty by (induct rule: finite_ne_induct) simp_all lemma Max_less_iff [simp]: "Max A < x \ (\a\A. a < x)" using fin_nonempty by (induct rule: finite_ne_induct) simp_all lemma Min_le_iff: "Min A \ x \ (\a\A. a \ x)" using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj) lemma Max_ge_iff: "x \ Max A \ (\a\A. x \ a)" using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj) lemma Min_less_iff: "Min A < x \ (\a\A. a < x)" using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj) lemma Max_gr_iff: "x < Max A \ (\a\A. x < a)" using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj) end lemma Max_eq_if: assumes "finite A" "finite B" "\a\A. \b\B. a \ b" "\b\B. \a\A. b \ a" shows "Max A = Max B" proof cases assume "A = {}" thus ?thesis using assms by simp next assume "A \ {}" thus ?thesis using assms by(blast intro: order.antisym Max_in Max_ge_iff[THEN iffD2]) qed lemma Min_antimono: assumes "M \ N" and "M \ {}" and "finite N" shows "Min N \ Min M" using assms by (fact Min.subset_imp) lemma Max_mono: assumes "M \ N" and "M \ {}" and "finite N" shows "Max M \ Max N" using assms by (fact Max.subset_imp) lemma mono_Min_commute: assumes "mono f" assumes "finite A" and "A \ {}" shows "f (Min A) = Min (f ` A)" proof (rule linorder_class.Min_eqI [symmetric]) from \finite A\ show "finite (f ` A)" by simp from assms show "f (Min A) \ f ` A" by simp fix x assume "x \ f ` A" then obtain y where "y \ A" and "x = f y" .. with assms have "Min A \ y" by auto with \mono f\ have "f (Min A) \ f y" by (rule monoE) with \x = f y\ show "f (Min A) \ x" by simp qed lemma mono_Max_commute: assumes "mono f" assumes "finite A" and "A \ {}" shows "f (Max A) = Max (f ` A)" proof (rule linorder_class.Max_eqI [symmetric]) from \finite A\ show "finite (f ` A)" by simp from assms show "f (Max A) \ f ` A" by simp fix x assume "x \ f ` A" then obtain y where "y \ A" and "x = f y" .. with assms have "y \ Max A" by auto with \mono f\ have "f y \ f (Max A)" by (rule monoE) with \x = f y\ show "x \ f (Max A)" by simp qed lemma finite_linorder_max_induct [consumes 1, case_names empty insert]: assumes fin: "finite A" and empty: "P {}" and insert: "\b A. finite A \ \a\A. a < b \ P A \ P (insert b A)" shows "P A" using fin empty insert proof (induct rule: finite_psubset_induct) case (psubset A) have IH: "\B. \B < A; P {}; (\A b. \finite A; \a\A. a \ P (insert b A))\ \ P B" by fact have fin: "finite A" by fact have empty: "P {}" by fact have step: "\b A. \finite A; \a\A. a < b; P A\ \ P (insert b A)" by fact show "P A" proof (cases "A = {}") assume "A = {}" then show "P A" using \P {}\ by simp next let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B" have "finite ?B" using \finite A\ by simp assume "A \ {}" with \finite A\ have "Max A \ A" by auto then have A: "?A = A" using insert_Diff_single insert_absorb by auto then have "P ?B" using \P {}\ step IH [of ?B] by blast moreover have "\a\?B. a < Max A" using Max_ge [OF \finite A\] by fastforce ultimately show "P A" using A insert_Diff_single step [OF \finite ?B\] by fastforce qed qed lemma finite_linorder_min_induct [consumes 1, case_names empty insert]: "\finite A; P {}; \b A. \finite A; \a\A. b < a; P A\ \ P (insert b A)\ \ P A" by (rule linorder.finite_linorder_max_induct [OF dual_linorder]) lemma finite_ranking_induct[consumes 1, case_names empty insert]: fixes f :: "'b \ 'a" assumes "finite S" assumes "P {}" assumes "\x S. finite S \ (\y. y \ S \ f y \ f x) \ P S \ P (insert x S)" shows "P S" using \finite S\ proof (induction rule: finite_psubset_induct) case (psubset A) { assume "A \ {}" hence "f ` A \ {}" and "finite (f ` A)" using psubset finite_image_iff by simp+ then obtain a where "f a = Max (f ` A)" and "a \ A" by (metis Max_in[of "f ` A"] imageE) then have "P (A - {a})" using psubset member_remove by blast moreover have "\y. y \ A \ f y \ f a" using \f a = Max (f ` A)\ \finite (f ` A)\ by simp ultimately have ?case by (metis \a \ A\ DiffD1 insert_Diff assms(3) finite_Diff psubset.hyps) } thus ?case using assms(2) by blast qed lemma Least_Min: assumes "finite {a. P a}" and "\a. P a" shows "(LEAST a. P a) = Min {a. P a}" proof - { fix A :: "'a set" assume A: "finite A" "A \ {}" have "(LEAST a. a \ A) = Min A" using A proof (induct A rule: finite_ne_induct) case singleton show ?case by (rule Least_equality) simp_all next case (insert a A) have "(LEAST b. b = a \ b \ A) = min a (LEAST a. a \ A)" by (auto intro!: Least_equality simp add: min_def not_le Min_le_iff insert.hyps dest!: less_imp_le) with insert show ?case by simp qed } from this [of "{a. P a}"] assms show ?thesis by simp qed lemma infinite_growing: assumes "X \ {}" assumes *: "\x. x \ X \ \y\X. y > x" shows "\ finite X" proof assume "finite X" with \X \ {}\ have "Max X \ X" "\x\X. x \ Max X" by auto with *[of "Max X"] show False by auto qed end lemma sum_le_card_Max: "finite A \ sum f A \ card A * Max (f ` A)" using sum_bounded_above[of A f "Max (f ` A)"] by simp lemma card_Min_le_sum: "finite A \ card A * Min (f ` A) \ sum f A" using sum_bounded_below[of A "Min (f ` A)" f] by simp context linordered_ab_semigroup_add begin lemma Min_add_commute: fixes k assumes "finite S" and "S \ {}" shows "Min ((\x. f x + k) ` S) = Min(f ` S) + k" proof - have m: "\x y. min x y + k = min (x+k) (y+k)" by (simp add: min_def order.antisym add_right_mono) have "(\x. f x + k) ` S = (\y. y + k) ` (f ` S)" by auto also have "Min \ = Min (f ` S) + k" using assms hom_Min_commute [of "\y. y+k" "f ` S", OF m, symmetric] by simp finally show ?thesis by simp qed lemma Max_add_commute: fixes k assumes "finite S" and "S \ {}" shows "Max ((\x. f x + k) ` S) = Max(f ` S) + k" proof - have m: "\x y. max x y + k = max (x+k) (y+k)" by (simp add: max_def order.antisym add_right_mono) have "(\x. f x + k) ` S = (\y. y + k) ` (f ` S)" by auto also have "Max \ = Max (f ` S) + k" using assms hom_Max_commute [of "\y. y+k" "f ` S", OF m, symmetric] by simp finally show ?thesis by simp qed end context linordered_ab_group_add begin lemma minus_Max_eq_Min [simp]: "finite S \ S \ {} \ - Max S = Min (uminus ` S)" by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min) lemma minus_Min_eq_Max [simp]: "finite S \ S \ {} \ - Min S = Max (uminus ` S)" by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max) end context complete_linorder begin lemma Min_Inf: assumes "finite A" and "A \ {}" shows "Min A = Inf A" proof - from assms obtain b B where "A = insert b B" and "finite B" by auto then show ?thesis by (simp add: Min.eq_fold complete_linorder_inf_min [symmetric] inf_Inf_fold_inf inf.commute [of b]) qed lemma Max_Sup: assumes "finite A" and "A \ {}" shows "Max A = Sup A" proof - from assms obtain b B where "A = insert b B" and "finite B" by auto then show ?thesis by (simp add: Max.eq_fold complete_linorder_sup_max [symmetric] sup_Sup_fold_sup sup.commute [of b]) qed end lemma disjnt_ge_max: \<^marker>\contributor \Lars Hupel\\ \disjnt X Y\ if \finite Y\ \\x. x \ X \ x > Max Y\ using that by (auto simp add: disjnt_def) (use Max_less_iff in blast) subsection \Arg Min\ context ord begin definition is_arg_min :: "('b \ 'a) \ ('b \ bool) \ 'b \ bool" where "is_arg_min f P x = (P x \ \(\y. P y \ f y < f x))" definition arg_min :: "('b \ 'a) \ ('b \ bool) \ 'b" where "arg_min f P = (SOME x. is_arg_min f P x)" definition arg_min_on :: "('b \ 'a) \ 'b set \ 'b" where "arg_min_on f S = arg_min f (\x. x \ S)" end syntax "_arg_min" :: "('b \ 'a) \ pttrn \ bool \ 'b" ("(3ARG'_MIN _ _./ _)" [1000, 0, 10] 10) translations "ARG_MIN f x. P" \ "CONST arg_min f (\x. P)" lemma is_arg_min_linorder: fixes f :: "'a \ 'b :: linorder" shows "is_arg_min f P x = (P x \ (\y. P y \ f x \ f y))" by(auto simp add: is_arg_min_def) lemma is_arg_min_antimono: fixes f :: "'a \ ('b::order)" shows "\ is_arg_min f P x; f y \ f x; P y \ \ is_arg_min f P y" by (simp add: order.order_iff_strict is_arg_min_def) lemma arg_minI: "\ P x; \y. P y \ \ f y < f x; \x. \ P x; \y. P y \ \ f y < f x \ \ Q x \ \ Q (arg_min f P)" apply (simp add: arg_min_def is_arg_min_def) apply (rule someI2_ex) apply blast apply blast done lemma arg_min_equality: "\ P k; \x. P x \ f k \ f x \ \ f (arg_min f P) = f k" for f :: "_ \ 'a::order" apply (rule arg_minI) apply assumption apply (simp add: less_le_not_le) by (metis le_less) lemma wf_linord_ex_has_least: "\ wf r; \x y. (x, y) \ r\<^sup>+ \ (y, x) \ r\<^sup>*; P k \ \ \x. P x \ (\y. P y \ (m x, m y) \ r\<^sup>*)" apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]]) apply (drule_tac x = "m ` Collect P" in spec) by force lemma ex_has_least_nat: "P k \ \x. P x \ (\y. P y \ m x \ m y)" for m :: "'a \ nat" apply (simp only: pred_nat_trancl_eq_le [symmetric]) apply (rule wf_pred_nat [THEN wf_linord_ex_has_least]) apply (simp add: less_eq linorder_not_le pred_nat_trancl_eq_le) by assumption lemma arg_min_nat_lemma: "P k \ P(arg_min m P) \ (\y. P y \ m (arg_min m P) \ m y)" for m :: "'a \ nat" apply (simp add: arg_min_def is_arg_min_linorder) apply (rule someI_ex) apply (erule ex_has_least_nat) done lemmas arg_min_natI = arg_min_nat_lemma [THEN conjunct1] lemma is_arg_min_arg_min_nat: fixes m :: "'a \ nat" shows "P x \ is_arg_min m P (arg_min m P)" by (metis arg_min_nat_lemma is_arg_min_linorder) lemma arg_min_nat_le: "P x \ m (arg_min m P) \ m x" for m :: "'a \ nat" by (rule arg_min_nat_lemma [THEN conjunct2, THEN spec, THEN mp]) lemma ex_min_if_finite: "\ finite S; S \ {} \ \ \m\S. \(\x\S. x < (m::'a::order))" by(induction rule: finite.induct) (auto intro: order.strict_trans) lemma ex_is_arg_min_if_finite: fixes f :: "'a \ 'b :: order" shows "\ finite S; S \ {} \ \ \x. is_arg_min f (\x. x \ S) x" unfolding is_arg_min_def using ex_min_if_finite[of "f ` S"] by auto lemma arg_min_SOME_Min: "finite S \ arg_min_on f S = (SOME y. y \ S \ f y = Min(f ` S))" unfolding arg_min_on_def arg_min_def is_arg_min_linorder apply(rule arg_cong[where f = Eps]) apply (auto simp: fun_eq_iff intro: Min_eqI[symmetric]) done lemma arg_min_if_finite: fixes f :: "'a \ 'b :: order" assumes "finite S" "S \ {}" shows "arg_min_on f S \ S" and "\(\x\S. f x < f (arg_min_on f S))" using ex_is_arg_min_if_finite[OF assms, of f] unfolding arg_min_on_def arg_min_def is_arg_min_def by(auto dest!: someI_ex) lemma arg_min_least: fixes f :: "'a \ 'b :: linorder" shows "\ finite S; S \ {}; y \ S \ \ f(arg_min_on f S) \ f y" by(simp add: arg_min_SOME_Min inv_into_def2[symmetric] f_inv_into_f) lemma arg_min_inj_eq: fixes f :: "'a \ 'b :: order" shows "\ inj_on f {x. P x}; P a; \y. P y \ f a \ f y \ \ arg_min f P = a" apply(simp add: arg_min_def is_arg_min_def) apply(rule someI2[of _ a]) apply (simp add: less_le_not_le) by (metis inj_on_eq_iff less_le mem_Collect_eq) subsection \Arg Max\ context ord begin definition is_arg_max :: "('b \ 'a) \ ('b \ bool) \ 'b \ bool" where "is_arg_max f P x = (P x \ \(\y. P y \ f y > f x))" definition arg_max :: "('b \ 'a) \ ('b \ bool) \ 'b" where "arg_max f P = (SOME x. is_arg_max f P x)" definition arg_max_on :: "('b \ 'a) \ 'b set \ 'b" where "arg_max_on f S = arg_max f (\x. x \ S)" end syntax "_arg_max" :: "('b \ 'a) \ pttrn \ bool \ 'a" ("(3ARG'_MAX _ _./ _)" [1000, 0, 10] 10) translations "ARG_MAX f x. P" \ "CONST arg_max f (\x. P)" lemma is_arg_max_linorder: fixes f :: "'a \ 'b :: linorder" shows "is_arg_max f P x = (P x \ (\y. P y \ f x \ f y))" by(auto simp add: is_arg_max_def) lemma arg_maxI: "P x \ (\y. P y \ \ f y > f x) \ (\x. P x \ \y. P y \ \ f y > f x \ Q x) \ Q (arg_max f P)" apply (simp add: arg_max_def is_arg_max_def) apply (rule someI2_ex) apply blast apply blast done lemma arg_max_equality: "\ P k; \x. P x \ f x \ f k \ \ f (arg_max f P) = f k" for f :: "_ \ 'a::order" apply (rule arg_maxI [where f = f]) apply assumption apply (simp add: less_le_not_le) by (metis le_less) lemma ex_has_greatest_nat_lemma: "P k \ \x. P x \ (\y. P y \ \ f y \ f x) \ \y. P y \ \ f y < f k + n" for f :: "'a \ nat" -by (induct n) (force simp: le_Suc_eq)+ + by (induct n) (force simp: le_Suc_eq)+ lemma ex_has_greatest_nat: - "P k \ \y. P y \ f y < b \ \x. P x \ (\y. P y \ f y \ f x)" - for f :: "'a \ nat" -apply (rule ccontr) -apply (cut_tac P = P and n = "b - f k" in ex_has_greatest_nat_lemma) - apply (subgoal_tac [3] "f k \ b") - apply auto -done + assumes "P k" + and "\y. P y \ (f:: 'a \ nat) y < b" +shows "\x. P x \ (\y. P y \ f y \ f x)" +proof (rule ccontr) + assume "\x. P x \ (\y. P y \ f y \ f x)" + then have "\x. P x \ (\y. P y \ \ f y \ f x)" + by auto + then have "\y. P y \ \ f y < f k + (b - f k)" + using assms ex_has_greatest_nat_lemma[of P k f "b - f k"] + by blast + then show "False" + using assms by auto +qed lemma arg_max_nat_lemma: "\ P k; \y. P y \ f y < b \ \ P (arg_max f P) \ (\y. P y \ f y \ f (arg_max f P))" for f :: "'a \ nat" apply (simp add: arg_max_def is_arg_max_linorder) apply (rule someI_ex) apply (erule (1) ex_has_greatest_nat) done lemmas arg_max_natI = arg_max_nat_lemma [THEN conjunct1] lemma arg_max_nat_le: "P x \ \y. P y \ f y < b \ f x \ f (arg_max f P)" for f :: "'a \ nat" by (blast dest: arg_max_nat_lemma [THEN conjunct2, THEN spec, of P]) end diff --git a/src/HOL/Nat.thy b/src/HOL/Nat.thy --- a/src/HOL/Nat.thy +++ b/src/HOL/Nat.thy @@ -1,2560 +1,2572 @@ (* Title: HOL/Nat.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Natural numbers\ theory Nat imports Inductive Typedef Fun Rings begin subsection \Type \ind\\ typedecl ind axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \ ind" \ \The axiom of infinity in 2 parts:\ where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \ x = y" and Suc_Rep_not_Zero_Rep: "Suc_Rep x \ Zero_Rep" subsection \Type nat\ text \Type definition\ inductive Nat :: "ind \ bool" where Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ Nat (Suc_Rep i)" typedef nat = "{n. Nat n}" morphisms Rep_Nat Abs_Nat using Nat.Zero_RepI by auto lemma Nat_Rep_Nat: "Nat (Rep_Nat n)" using Rep_Nat by simp lemma Nat_Abs_Nat_inverse: "Nat n \ Rep_Nat (Abs_Nat n) = n" using Abs_Nat_inverse by simp lemma Nat_Abs_Nat_inject: "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" using Abs_Nat_inject by simp instantiation nat :: zero begin definition Zero_nat_def: "0 = Abs_Nat Zero_Rep" instance .. end definition Suc :: "nat \ nat" where "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" lemma Suc_not_Zero: "Suc m \ 0" by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat) lemma Zero_not_Suc: "0 \ Suc m" by (rule not_sym) (rule Suc_not_Zero) lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \ x = y" by (rule iffI, rule Suc_Rep_inject) simp_all lemma nat_induct0: assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" proof - have "P (Abs_Nat (Rep_Nat n))" using assms unfolding Zero_nat_def Suc_def by (iprover intro: Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst]) then show ?thesis by (simp add: Rep_Nat_inverse) qed free_constructors case_nat for "0 :: nat" | Suc pred where "pred (0 :: nat) = (0 :: nat)" - apply atomize_elim - apply (rename_tac n, induct_tac n rule: nat_induct0, auto) - apply (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) - apply (simp only: Suc_not_Zero) - done +proof atomize_elim + fix n + show "n = 0 \ (\m. n = Suc m)" + by (induction n rule: nat_induct0) auto +next + fix n m + show "(Suc n = Suc m) = (n = m)" + by (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) +next + fix n + show "0 \ Suc n" + by (simp add: Suc_not_Zero) +qed + \ \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "0 :: nat" Suc by (erule nat_induct0) auto setup \Sign.parent_path\ \ \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "nat"\ declare old.nat.inject[iff del] and old.nat.distinct(1)[simp del, induct_simp del] lemmas induct = old.nat.induct lemmas inducts = old.nat.inducts lemmas rec = old.nat.rec lemmas simps = nat.inject nat.distinct nat.case nat.rec setup \Sign.parent_path\ abbreviation rec_nat :: "'a \ (nat \ 'a \ 'a) \ nat \ 'a" where "rec_nat \ old.rec_nat" declare nat.sel[code del] hide_const (open) Nat.pred \ \hide everything related to the selector\ hide_fact nat.case_eq_if nat.collapse nat.expand nat.sel nat.exhaust_sel nat.split_sel nat.split_sel_asm lemma nat_exhaust [case_names 0 Suc, cases type: nat]: "(y = 0 \ P) \ (\nat. y = Suc nat \ P) \ P" \ \for backward compatibility -- names of variables differ\ by (rule old.nat.exhaust) lemma nat_induct [case_names 0 Suc, induct type: nat]: fixes n assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" \ \for backward compatibility -- names of variables differ\ using assms by (rule nat.induct) hide_fact nat_exhaust nat_induct0 ML \ val nat_basic_lfp_sugar = let val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}} end; \ setup \ let fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; in BNF_LFP_Rec_Sugar.register_lfp_rec_extension {nested_simps = [], special_endgame_tac = K (K (K (K no_tac))), is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of, rewrite_nested_rec_call = NONE} end \ text \Injectiveness and distinctness lemmas\ lemma inj_Suc [simp]: "inj_on Suc N" by (simp add: inj_on_def) lemma bij_betw_Suc [simp]: "bij_betw Suc M N \ Suc ` M = N" by (simp add: bij_betw_def) lemma Suc_neq_Zero: "Suc m = 0 \ R" by (rule notE) (rule Suc_not_Zero) lemma Zero_neq_Suc: "0 = Suc m \ R" by (rule Suc_neq_Zero) (erule sym) lemma Suc_inject: "Suc x = Suc y \ x = y" by (rule inj_Suc [THEN injD]) lemma n_not_Suc_n: "n \ Suc n" by (induct n) simp_all lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" and "\x y. P x y \ P (Suc x) (Suc y)" shows "P m n" proof (induct n arbitrary: m) case 0 show ?case by (rule assms(1)) next case (Suc n) show ?case proof (induct m) case 0 show ?case by (rule assms(2)) next case (Suc m) from \P m n\ show ?case by (rule assms(3)) qed qed subsection \Arithmetic operators\ instantiation nat :: comm_monoid_diff begin primrec plus_nat where add_0: "0 + n = (n::nat)" | add_Suc: "Suc m + n = Suc (m + n)" lemma add_0_right [simp]: "m + 0 = m" for m :: nat by (induct m) simp_all lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)" by (induct m) simp_all declare add_0 [code] lemma add_Suc_shift [code]: "Suc m + n = m + Suc n" by simp primrec minus_nat where diff_0 [code]: "m - 0 = (m::nat)" | diff_Suc: "m - Suc n = (case m - n of 0 \ 0 | Suc k \ k)" declare diff_Suc [simp del] lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat by (induct n) (simp_all add: diff_Suc) lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n" by (induct n) (simp_all add: diff_Suc) instance proof fix n m q :: nat show "(n + m) + q = n + (m + q)" by (induct n) simp_all show "n + m = m + n" by (induct n) simp_all show "m + n - m = n" by (induct m) simp_all show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc) show "0 + n = n" by simp show "0 - n = 0" by simp qed end hide_fact (open) add_0 add_0_right diff_0 instantiation nat :: comm_semiring_1_cancel begin definition One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0::nat)" | mult_Suc: "Suc m * n = n + (m * n)" lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat by (induct m) (simp_all add: add.assoc) instance proof fix k n m q :: nat show "0 \ (1::nat)" by simp show "1 * n = n" by simp show "n * m = m * n" by (induct n) simp_all show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib) show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib) show "k * (m - n) = (k * m) - (k * n)" by (induct m n rule: diff_induct) simp_all qed end subsubsection \Addition\ text \Reasoning about \m + 0 = 0\, etc.\ lemma add_is_0 [iff]: "m + n = 0 \ m = 0 \ n = 0" for m n :: nat by (cases m) simp_all lemma add_is_1: "m + n = Suc 0 \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (cases m) simp_all lemma one_is_add: "Suc 0 = m + n \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (rule trans, rule eq_commute, rule add_is_1) lemma add_eq_self_zero: "m + n = m \ n = 0" for m n :: nat by (induct m) simp_all lemma plus_1_eq_Suc: "plus 1 = Suc" by (simp add: fun_eq_iff) lemma Suc_eq_plus1: "Suc n = n + 1" by simp lemma Suc_eq_plus1_left: "Suc n = 1 + n" by simp subsubsection \Difference\ lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_add) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" by simp subsubsection \Multiplication\ lemma mult_is_0 [simp]: "m * n = 0 \ m = 0 \ n = 0" for m n :: nat by (induct m) auto lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \ m = Suc 0 \ n = Suc 0" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (induct n) auto qed lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \ m = Suc 0 \ n = Suc 0" by (simp add: eq_commute flip: mult_eq_1_iff) lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" and nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" for m n :: nat by auto lemma mult_cancel1 [simp]: "k * m = k * n \ m = n \ k = 0" for k m n :: nat proof - have "k \ 0 \ k * m = k * n \ m = n" proof (induct n arbitrary: m) case 0 then show "m = 0" by simp next case (Suc n) then show "m = Suc n" by (cases m) (simp_all add: eq_commute [of 0]) qed then show ?thesis by auto qed lemma mult_cancel2 [simp]: "m * k = n * k \ m = n \ k = 0" for k m n :: nat by (simp add: mult.commute) lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \ m = n" by (subst mult_cancel1) simp subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ instantiation nat :: linorder begin primrec less_eq_nat where "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] lemma le0 [iff]: "0 \ n" for n :: nat by (simp add: less_eq_nat.simps) lemma [code]: "0 \ n \ True" for n :: nat by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" lemma Suc_le_mono [iff]: "Suc n \ Suc m \ n \ m" by (simp add: less_eq_nat.simps(2)) lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. lemma le_0_eq [iff]: "n \ 0 \ n = 0" for n :: nat by (induct n) (simp_all add: less_eq_nat.simps(2)) lemma not_less0 [iff]: "\ n < 0" for n :: nat by (simp add: less_eq_Suc_le) lemma less_nat_zero_code [code]: "n < 0 \ False" for n :: nat by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" by (simp add: less_eq_Suc_le) lemma less_Suc_eq_le [code]: "m < Suc n \ m \ n" by (simp add: less_eq_Suc_le) lemma Suc_less_eq2: "Suc n < m \ (\m'. m = Suc m' \ n < m')" by (cases m) auto lemma le_SucI: "m \ n \ m \ Suc n" by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits) lemma Suc_leD: "Suc m \ n \ m \ n" by (cases n) (auto intro: le_SucI) lemma less_SucI: "m < n \ m < Suc n" by (simp add: less_eq_Suc_le) (erule Suc_leD) lemma Suc_lessD: "Suc m < n \ m < n" by (simp add: less_eq_Suc_le) (erule Suc_leD) instance proof fix n m q :: nat show "n < m \ n \ m \ \ m \ n" proof (induct n arbitrary: m) case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le) next case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le) qed show "n \ n" by (induct n) simp_all then show "n = m" if "n \ m" and "m \ n" using that by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) show "n \ q" if "n \ m" and "m \ q" using that proof (induct n arbitrary: m q) case 0 show ?case by simp next case (Suc n) then show ?case by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits) qed show "n \ m \ m \ n" by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) qed end instantiation nat :: order_bot begin definition bot_nat :: nat where "bot_nat = 0" instance by standard (simp add: bot_nat_def) end instance nat :: no_top by standard (auto intro: less_Suc_eq_le [THEN iffD2]) subsubsection \Introduction properties\ lemma lessI [iff]: "n < Suc n" by (simp add: less_Suc_eq_le) lemma zero_less_Suc [iff]: "0 < Suc n" by (simp add: less_Suc_eq_le) subsubsection \Elimination properties\ lemma less_not_refl: "\ n < n" for n :: nat by (rule order_less_irrefl) lemma less_not_refl2: "n < m \ m \ n" for m n :: nat by (rule not_sym) (rule less_imp_neq) lemma less_not_refl3: "s < t \ s \ t" for s t :: nat by (rule less_imp_neq) lemma less_irrefl_nat: "n < n \ R" for n :: nat by (rule notE, rule less_not_refl) lemma less_zeroE: "n < 0 \ R" for n :: nat by (rule notE) (rule not_less0) lemma less_Suc_eq: "m < Suc n \ m < n \ m = n" unfolding less_Suc_eq_le le_less .. lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)" by (simp add: less_Suc_eq) lemma less_one [iff]: "n < 1 \ n = 0" for n :: nat unfolding One_nat_def by (rule less_Suc0) lemma Suc_mono: "m < n \ Suc m < Suc n" by simp text \"Less than" is antisymmetric, sort of.\ lemma less_antisym: "\ n < m \ n < Suc m \ m = n" unfolding not_less less_Suc_eq_le by (rule antisym) lemma nat_neq_iff: "m \ n \ m < n \ n < m" for m n :: nat by (rule linorder_neq_iff) subsubsection \Inductive (?) properties\ lemma Suc_lessI: "m < n \ Suc m \ n \ Suc m < n" unfolding less_eq_Suc_le [of m] le_less by simp lemma lessE: assumes major: "i < k" and 1: "k = Suc i \ P" and 2: "\j. i < j \ k = Suc j \ P" shows P proof - from major have "\j. i \ j \ k = Suc j" unfolding less_eq_Suc_le by (induct k) simp_all then have "(\j. i < j \ k = Suc j) \ k = Suc i" by (auto simp add: less_le) with 1 2 show P by auto qed lemma less_SucE: assumes major: "m < Suc n" and less: "m < n \ P" and eq: "m = n \ P" shows P proof (rule major [THEN lessE]) show "Suc n = Suc m \ P" using eq by blast show "\j. \m < j; Suc n = Suc j\ \ P" by (blast intro: less) qed lemma Suc_lessE: assumes major: "Suc i < k" and minor: "\j. i < j \ k = Suc j \ P" shows P proof (rule major [THEN lessE]) show "k = Suc (Suc i) \ P" using lessI minor by iprover show "\j. \Suc i < j; k = Suc j\ \ P" using Suc_lessD minor by iprover qed lemma Suc_less_SucD: "Suc m < Suc n \ m < n" by simp lemma less_trans_Suc: assumes le: "i < j" shows "j < k \ Suc i < k" proof (induct k) case 0 then show ?case by simp next case (Suc k) with le show ?case by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) lemma not_less_eq_eq: "\ m \ n \ Suc n \ m" by (simp only: not_le Suc_le_eq) text \Properties of "less than or equal".\ lemma le_imp_less_Suc: "m \ n \ m < Suc n" by (simp only: less_Suc_eq_le) lemma Suc_n_not_le_n: "\ Suc n \ n" by (simp add: not_le less_Suc_eq_le) lemma le_Suc_eq: "m \ Suc n \ m \ n \ m = Suc n" by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq) lemma le_SucE: "m \ Suc n \ (m \ n \ R) \ (m = Suc n \ R) \ R" by (drule le_Suc_eq [THEN iffD1], iprover+) lemma Suc_leI: "m < n \ Suc m \ n" by (simp only: Suc_le_eq) text \Stronger version of \Suc_leD\.\ lemma Suc_le_lessD: "Suc m \ n \ m < n" by (simp only: Suc_le_eq) lemma less_imp_le_nat: "m < n \ m \ n" for m n :: nat unfolding less_eq_Suc_le by (rule Suc_leD) text \For instance, \(Suc m < Suc n) = (Suc m \ n) = (m < n)\\ lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq text \Equivalence of \m \ n\ and \m < n \ m = n\\ lemma less_or_eq_imp_le: "m < n \ m = n \ m \ n" for m n :: nat unfolding le_less . lemma le_eq_less_or_eq: "m \ n \ m < n \ m = n" for m n :: nat by (rule le_less) text \Useful with \blast\.\ lemma eq_imp_le: "m = n \ m \ n" for m n :: nat by auto lemma le_refl: "n \ n" for n :: nat by simp lemma le_trans: "i \ j \ j \ k \ i \ k" for i j k :: nat by (rule order_trans) lemma le_antisym: "m \ n \ n \ m \ m = n" for m n :: nat by (rule antisym) lemma nat_less_le: "m < n \ m \ n \ m \ n" for m n :: nat by (rule less_le) lemma le_neq_implies_less: "m \ n \ m \ n \ m < n" for m n :: nat unfolding less_le .. lemma nat_le_linear: "m \ n \ n \ m" for m n :: nat by (rule linear) lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat] lemma le_less_Suc_eq: "m \ n \ n < Suc m \ n = m" unfolding less_Suc_eq_le by auto lemma not_less_less_Suc_eq: "\ n < m \ n < Suc m \ n = m" unfolding not_less by (rule le_less_Suc_eq) lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq lemma not0_implies_Suc: "n \ 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr0_implies_Suc: "n > 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr_implies_not0: "m < n \ n \ 0" for m n :: nat by (cases n) simp_all lemma neq0_conv[iff]: "n \ 0 \ 0 < n" for n :: nat by (cases n) simp_all text \This theorem is useful with \blast\\ lemma gr0I: "(n = 0 \ False) \ 0 < n" for n :: nat by (rule neq0_conv[THEN iffD1]) iprover lemma gr0_conv_Suc: "0 < n \ (\m. n = Suc m)" by (fast intro: not0_implies_Suc) lemma not_gr0 [iff]: "\ 0 < n \ n = 0" for n :: nat using neq0_conv by blast lemma Suc_le_D: "Suc n \ m' \ \m. m' = Suc m" by (induct m') simp_all text \Useful in certain inductive arguments\ lemma less_Suc_eq_0_disj: "m < Suc n \ m = 0 \ (\j. m = Suc j \ j < n)" by (cases m) simp_all lemma All_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma All_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) lemma Ex_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma Ex_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) text \@{term mono} (non-strict) doesn't imply increasing, as the function could be constant\ lemma strict_mono_imp_increasing: fixes n::nat assumes "strict_mono f" shows "f n \ n" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case unfolding not_less_eq_eq [symmetric] using Suc_n_not_le_n assms order_trans strict_mono_less_eq by blast qed subsubsection \Monotonicity of Addition\ lemma Suc_pred [simp]: "n > 0 \ Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) lemma Suc_diff_1 [simp]: "0 < n \ Suc (n - 1) = n" unfolding One_nat_def by (rule Suc_pred) lemma nat_add_left_cancel_le [simp]: "k + m \ k + n \ m \ n" for k m n :: nat by (induct k) simp_all lemma nat_add_left_cancel_less [simp]: "k + m < k + n \ m < n" for k m n :: nat by (induct k) simp_all lemma add_gr_0 [iff]: "m + n > 0 \ m > 0 \ n > 0" for m n :: nat by (auto dest: gr0_implies_Suc) text \strict, in 1st argument\ lemma add_less_mono1: "i < j \ i + k < j + k" for i j k :: nat by (induct k) simp_all text \strict, in both arguments\ lemma add_less_mono: fixes i j k l :: nat assumes "i < j" "k < l" shows "i + k < j + l" proof - have "i + k < j + k" by (simp add: add_less_mono1 assms) also have "... < j + l" using \i < j\ by (induction j) (auto simp: assms) finally show ?thesis . qed lemma less_imp_Suc_add: "m < n \ \k. n = Suc (m + k)" proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: order_le_less) (blast elim!: less_SucE intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric]) qed lemma le_Suc_ex: "k \ l \ (\n. l = k + n)" for k l :: nat by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add) lemma less_natE: assumes \m < n\ obtains q where \n = Suc (m + q)\ using assms by (auto dest: less_imp_Suc_add intro: that) text \strict, in 1st argument; proof is by induction on \k > 0\\ lemma mult_less_mono2: fixes i j :: nat assumes "i < j" and "0 < k" shows "k * i < k * j" using \0 < k\ proof (induct k) case 0 then show ?case by simp next case (Suc k) with \i < j\ show ?case by (cases k) (simp_all add: add_less_mono) qed text \Addition is the inverse of subtraction: if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma nat_le_iff_add: "m \ n \ (\k. n = m + k)" for m n :: nat using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex) text \The naturals form an ordered \semidom\ and a \dioid\.\ instance nat :: linordered_semidom proof fix m n q :: nat show "0 < (1::nat)" by simp show "m \ n \ q + m \ q + n" by simp show "m < n \ 0 < q \ q * m < q * n" by (simp add: mult_less_mono2) show "m \ 0 \ n \ 0 \ m * n \ 0" by simp show "n \ m \ (m - n) + n = m" by (simp add: add_diff_inverse_nat add.commute linorder_not_less) qed instance nat :: dioid by standard (rule nat_le_iff_add) declare le0[simp del] \ \This is now @{thm zero_le}\ declare le_0_eq[simp del] \ \This is now @{thm le_zero_eq}\ declare not_less0[simp del] \ \This is now @{thm not_less_zero}\ declare not_gr0[simp del] \ \This is now @{thm not_gr_zero}\ instance nat :: ordered_cancel_comm_monoid_add .. instance nat :: ordered_cancel_comm_monoid_diff .. subsubsection \\<^term>\min\ and \<^term>\max\\ global_interpretation bot_nat_0: ordering_top \(\)\ \(>)\ \0::nat\ by standard simp global_interpretation max_nat: semilattice_neutr_order max \0::nat\ \(\)\ \(>)\ by standard (simp add: max_def) lemma mono_Suc: "mono Suc" by (rule monoI) simp lemma min_0L [simp]: "min 0 n = 0" for n :: nat by (rule min_absorb1) simp lemma min_0R [simp]: "min n 0 = 0" for n :: nat by (rule min_absorb2) simp lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)" by (simp add: mono_Suc min_of_mono) lemma min_Suc1: "min (Suc n) m = (case m of 0 \ 0 | Suc m' \ Suc(min n m'))" by (simp split: nat.split) lemma min_Suc2: "min m (Suc n) = (case m of 0 \ 0 | Suc m' \ Suc(min m' n))" by (simp split: nat.split) lemma max_0L [simp]: "max 0 n = n" for n :: nat by (fact max_nat.left_neutral) lemma max_0R [simp]: "max n 0 = n" for n :: nat by (fact max_nat.right_neutral) lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)" by (simp add: mono_Suc max_of_mono) lemma max_Suc1: "max (Suc n) m = (case m of 0 \ Suc n | Suc m' \ Suc (max n m'))" by (simp split: nat.split) lemma max_Suc2: "max m (Suc n) = (case m of 0 \ Suc n | Suc m' \ Suc (max m' n))" by (simp split: nat.split) lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat by (simp add: max_def) lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat by (simp add: max_def) lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ instance nat :: wellorder proof fix P and n :: nat assume step: "(\m. m < n \ P m) \ P n" for n :: nat have "\q. q \ n \ P q" proof (induct n) case (0 n) have "P 0" by (rule step) auto with 0 show ?case by auto next case (Suc m n) then have "n \ m \ n = Suc m" by (simp add: le_Suc_eq) then show ?case proof assume "n \ m" then show "P n" by (rule Suc(1)) next assume n: "n = Suc m" show "P n" by (rule step) (rule Suc(1), simp add: n le_simps) qed qed then show "P n" by auto qed lemma Least_eq_0[simp]: "P 0 \ Least P = 0" for P :: "nat \ bool" by (rule Least_equality[OF _ le0]) lemma Least_Suc: assumes "P n" "\ P 0" shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))" proof (cases n) case (Suc m) show ?thesis proof (rule antisym) show "(LEAST x. P x) \ Suc (LEAST x. P (Suc x))" using assms Suc by (force intro: LeastI Least_le) have \
: "P (LEAST x. P x)" by (blast intro: LeastI assms) show "Suc (LEAST m. P (Suc m)) \ (LEAST n. P n)" proof (cases "(LEAST n. P n)") case 0 then show ?thesis using \
by (simp add: assms) next case Suc with \
show ?thesis by (auto simp: Least_le) qed qed qed (use assms in auto) lemma Least_Suc2: "P n \ Q m \ \ P 0 \ \k. P (Suc k) = Q k \ Least P = Suc (Least Q)" by (erule (1) Least_Suc [THEN ssubst]) simp lemma ex_least_nat_le: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\k\n. (\i P i) \ P k" proof (cases n) case (Suc m) with assms show ?thesis by (blast intro: Least_le LeastI_ex dest: not_less_Least) qed (use assms in auto) lemma ex_least_nat_less: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\ki\k. \ P i) \ P (Suc k)" proof (cases n) case (Suc m) then obtain k where k: "k \ n" "\i P i" "P k" using ex_least_nat_le [OF assms] by blast show ?thesis by (cases k) (use assms k less_eq_Suc_le in auto) qed (use assms in auto) lemma nat_less_induct: fixes P :: "nat \ bool" assumes "\n. \m. m < n \ P m \ P n" shows "P n" using assms less_induct by blast lemma measure_induct_rule [case_names less]: fixes f :: "'a \ 'b::wellorder" assumes step: "\x. (\y. f y < f x \ P y) \ P x" shows "P a" by (induct m \ "f a" arbitrary: a rule: less_induct) (auto intro: step) text \old style induction rules:\ lemma measure_induct: fixes f :: "'a \ 'b::wellorder" shows "(\x. \y. f y < f x \ P y \ P x) \ P a" by (rule measure_induct_rule [of f P a]) iprover lemma full_nat_induct: assumes step: "\n. (\m. Suc m \ n \ P m) \ P n" shows "P n" by (rule less_induct) (auto intro: step simp:le_simps) text\An induction rule for establishing binary relations\ lemma less_Suc_induct [consumes 1]: assumes less: "i < j" and step: "\i. P i (Suc i)" and trans: "\i j k. i < j \ j < k \ P i j \ P j k \ P i k" shows "P i j" proof - from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add) have "P i (Suc (i + k))" proof (induct k) case 0 show ?case by (simp add: step) next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp then have "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed then show "P i j" by (simp add: j) qed text \ The method of infinite descent, frequently used in number theory. Provided by Roelof Oosterhuis. \P n\ is true for all natural numbers if \<^item> case ``0'': given \n = 0\ prove \P n\ \<^item> case ``smaller'': given \n > 0\ and \\ P n\ prove there exists a smaller natural number \m\ such that \\ P m\. \ lemma infinite_descent: "(\n. \ P n \ \m P m) \ P n" for P :: "nat \ bool" \ \compact version without explicit base case\ by (induct n rule: less_induct) auto lemma infinite_descent0 [case_names 0 smaller]: fixes P :: "nat \ bool" assumes "P 0" and "\n. n > 0 \ \ P n \ \m. m < n \ \ P m" shows "P n" proof (rule infinite_descent) - show "\n. \ P n \ \m P m" - using assms by (case_tac "n > 0") auto + fix n + show "\ P n \ \m P m" + using assms by (cases "n > 0") auto qed text \ Infinite descent using a mapping to \nat\: \P x\ is true for all \x \ D\ if there exists a \V \ D \ nat\ and \<^item> case ``0'': given \V x = 0\ prove \P x\ \<^item> ``smaller'': given \V x > 0\ and \\ P x\ prove there exists a \y \ D\ such that \V y < V x\ and \\ P y\. \ corollary infinite_descent0_measure [case_names 0 smaller]: fixes V :: "'a \ nat" assumes 1: "\x. V x = 0 \ P x" and 2: "\x. V x > 0 \ \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent0) case 0 with 1 show "P x" by auto next case (smaller n) then obtain x where *: "V x = n " and "V x > 0 \ \ P x" by auto with 2 obtain y where "V y < V x \ \ P y" by auto with * obtain m where "m = V y \ m < n \ \ P y" by auto then show ?case by auto qed ultimately show "P x" by auto qed text \Again, without explicit base case:\ lemma infinite_descent_measure: fixes V :: "'a \ nat" assumes "\x. \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - from assms obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" - proof (induct n rule: infinite_descent, auto) - show "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x + proof - + have "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x using assms and that by auto + then show "\x. V x = n \ P x" + by (induct n rule: infinite_descent, auto) qed ultimately show "P x" by auto qed text \A (clumsy) way of lifting \<\ monotonicity to \\\ monotonicity\ lemma less_mono_imp_le_mono: fixes f :: "nat \ nat" and i j :: nat assumes "\i j::nat. i < j \ f i < f j" and "i \ j" shows "f i \ f j" using assms by (auto simp add: order_le_less) text \non-strict, in 1st argument\ lemma add_le_mono1: "i \ j \ i + k \ j + k" for i j k :: nat by (rule add_right_mono) text \non-strict, in both arguments\ lemma add_le_mono: "i \ j \ k \ l \ i + k \ j + l" for i j k l :: nat by (rule add_mono) lemma le_add2: "n \ m + n" for m n :: nat by simp lemma le_add1: "n \ n + m" for m n :: nat by simp lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) lemma less_add_Suc2: "i < Suc (m + i)" by (rule le_less_trans, rule le_add2, rule lessI) lemma less_iff_Suc_add: "m < n \ (\k. n = Suc (m + k))" by (iprover intro!: less_add_Suc1 less_imp_Suc_add) lemma trans_le_add1: "i \ j \ i \ j + m" for i j m :: nat by (rule le_trans, assumption, rule le_add1) lemma trans_le_add2: "i \ j \ i \ m + j" for i j m :: nat by (rule le_trans, assumption, rule le_add2) lemma trans_less_add1: "i < j \ i < j + m" for i j m :: nat by (rule less_le_trans, assumption, rule le_add1) lemma trans_less_add2: "i < j \ i < m + j" for i j m :: nat by (rule less_le_trans, assumption, rule le_add2) lemma add_lessD1: "i + j < k \ i < k" for i j k :: nat by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1) lemma not_add_less1 [iff]: "\ i + j < i" for i j :: nat by simp lemma not_add_less2 [iff]: "\ j + i < i" for i j :: nat by simp lemma add_leD1: "m + k \ n \ m \ n" for k m n :: nat by (rule order_trans [of _ "m + k"]) (simp_all add: le_add1) lemma add_leD2: "m + k \ n \ k \ n" for k m n :: nat by (force simp add: add.commute dest: add_leD1) lemma add_leE: "m + k \ n \ (m \ n \ k \ n \ R) \ R" for k m n :: nat by (blast dest: add_leD1 add_leD2) text \needs \\k\ for \ac_simps\ to work\ lemma less_add_eq_less: "\k. k < l \ m + l = k + n \ m < n" for l m n :: nat by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps) subsubsection \More results about difference\ lemma Suc_diff_le: "n \ m \ Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all lemma diff_less_Suc: "m - n < Suc m" by (induct m n rule: diff_induct) (auto simp: less_Suc_eq) lemma diff_le_self [simp]: "m - n \ m" for m n :: nat by (induct m n rule: diff_induct) (simp_all add: le_SucI) lemma less_imp_diff_less: "j < k \ j - n < k" for j k n :: nat by (rule le_less_trans, rule diff_le_self) lemma diff_Suc_less [simp]: "0 < n \ n - Suc i < n" by (cases n) (auto simp add: le_simps) lemma diff_add_assoc: "k \ j \ (i + j) - k = i + (j - k)" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc) lemma add_diff_assoc [simp]: "k \ j \ i + (j - k) = i + j - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc) lemma diff_add_assoc2: "k \ j \ (j + i) - k = (j - k) + i" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc2) lemma add_diff_assoc2 [simp]: "k \ j \ j - k + i = j + i - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc2) lemma le_imp_diff_is_add: "i \ j \ (j - i = k) = (j = k + i)" for i j k :: nat by auto lemma diff_is_0_eq [simp]: "m - n = 0 \ m \ n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma diff_is_0_eq' [simp]: "m \ n \ m - n = 0" for m n :: nat by (rule iffD2, rule diff_is_0_eq) lemma zero_less_diff [simp]: "0 < n - m \ m < n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma less_imp_add_positive: assumes "i < j" shows "\k::nat. 0 < k \ i + k = j" proof from assms show "0 < j - i \ i + (j - i) = j" by (simp add: order_less_imp_le) qed text \a nice rewrite for bounded subtraction\ lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat by (simp add: max_def not_le order_less_imp_le) lemma nat_diff_split: "P (a - b) \ (a < b \ P 0) \ (\d. a = b + d \ P d)" for a b :: nat \ \elimination of \-\ on \nat\\ by (cases "a < b") (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym]) lemma nat_diff_split_asm: "P (a - b) \ \ (a < b \ \ P 0 \ (\d. a = b + d \ \ P d))" for a b :: nat \ \elimination of \-\ on \nat\ in assumptions\ by (auto split: nat_diff_split) lemma Suc_pred': "0 < n \ n = Suc(n - 1)" by simp lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat by (cases m) simp_all lemma Suc_diff_eq_diff_pred: "0 < n \ Suc m - n = m - (n - 1)" by (cases n) simp_all lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (cases m) simp_all lemma Let_Suc [simp]: "Let (Suc n) f \ f (Suc n)" by (fact Let_def) subsubsection \Monotonicity of multiplication\ lemma mult_le_mono1: "i \ j \ i * k \ j * k" for i j k :: nat by (simp add: mult_right_mono) lemma mult_le_mono2: "i \ j \ k * i \ k * j" for i j k :: nat by (simp add: mult_left_mono) text \\\\ monotonicity, BOTH arguments\ lemma mult_le_mono: "i \ j \ k \ l \ i * k \ j * l" for i j k l :: nat by (simp add: mult_mono) lemma mult_less_mono1: "i < j \ 0 < k \ i * k < j * k" for i j k :: nat by (simp add: mult_strict_right_mono) text \Differs from the standard \zero_less_mult_iff\ in that there are no negative numbers.\ lemma nat_0_less_mult_iff [simp]: "0 < m * n \ 0 < m \ 0 < n" for m n :: nat proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma one_le_mult_iff [simp]: "Suc 0 \ m * n \ Suc 0 \ m \ Suc 0 \ n" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma mult_less_cancel2 [simp]: "m * k < n * k \ 0 < k \ m < n" for k m n :: nat proof (intro iffI conjI) assume m: "m * k < n * k" then show "0 < k" by (cases k) auto show "m < n" proof (cases k) case 0 then show ?thesis using m by auto next case (Suc k') then show ?thesis using m by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1) qed next assume "0 < k \ m < n" then show "m * k < n * k" by (blast intro: mult_less_mono1) qed lemma mult_less_cancel1 [simp]: "k * m < k * n \ 0 < k \ m < n" for k m n :: nat by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "k * m \ k * n \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma mult_le_cancel2 [simp]: "m * k \ n * k \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \ m < n" by (subst mult_less_cancel1) simp lemma Suc_mult_le_cancel1: "Suc k * m \ Suc k * n \ m \ n" by (subst mult_le_cancel1) simp lemma le_square: "m \ m * m" for m :: nat by (cases m) (auto intro: le_add1) lemma le_cube: "m \ m * (m * m)" for m :: nat by (cases m) (auto intro: le_add1) text \Lemma for \gcd\\ lemma mult_eq_self_implies_10: fixes m n :: nat assumes "m = m * n" shows "n = 1 \ m = 0" proof (rule disjCI) assume "m \ 0" show "n = 1" proof (cases n "1::nat" rule: linorder_cases) case greater show ?thesis using assms mult_less_mono2 [OF greater, of m] \m \ 0\ by auto qed (use assms \m \ 0\ in auto) qed lemma mono_times_nat: fixes n :: nat assumes "n > 0" shows "mono (times n)" proof fix m q :: nat assume "m \ q" with assms show "n * m \ n * q" by simp qed text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin definition "(inf :: nat \ nat \ nat) = min" definition "(sup :: nat \ nat \ nat) = max" instance by intro_classes (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def intro: order_less_imp_le antisym elim!: order_trans order_less_trans) end subsection \Natural operation of natural numbers on functions\ text \ We use the same logical constant for the power operations on functions and relations, in order to share the same syntax. \ consts compow :: "nat \ 'a \ 'a" abbreviation compower :: "'a \ nat \ 'a" (infixr "^^" 80) where "f ^^ n \ compow n f" notation (latex output) compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \\f ^^ n = f \ \ \ f\, the \n\-fold composition of \f\\ overloading funpow \ "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f \ funpow n f" end lemma funpow_0 [simp]: "(f ^^ 0) x = x" by simp lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \ f" proof (induct n) case 0 then show ?case by simp next fix n assume "f ^^ Suc n = f ^^ n \ f" then show "f ^^ Suc (Suc n) = f ^^ Suc n \ f" by (simp add: o_assoc) qed lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right text \For code generation.\ context begin qualified definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where funpow_code_def [code_abbrev]: "funpow = compow" lemma [code]: "funpow (Suc n) f = f \ funpow n f" "funpow 0 f = id" by (simp_all add: funpow_code_def) end lemma funpow_add: "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \ 'a" by (induct n) (simp_all add: funpow_add) lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp also have "\ = (f ^^ n \ f ^^ 1) x" by (simp only: funpow_add) also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \ 'a" by (induct n) simp_all lemma Suc_funpow[simp]: "Suc ^^ n = ((+) n)" by (induct n) simp_all lemma id_funpow[simp]: "id ^^ n = id" by (induct n) simp_all lemma funpow_mono: "mono f \ A \ B \ (f ^^ n) A \ (f ^^ n) B" for f :: "'a \ ('a::order)" by (induct n arbitrary: A B) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def) lemma funpow_mono2: assumes "mono f" and "i \ j" and "x \ y" and "x \ f x" shows "(f ^^ i) x \ (f ^^ j) y" using assms(2,3) proof (induct j arbitrary: y) case 0 then show ?case by simp next case (Suc j) show ?case proof(cases "i = Suc j") case True with assms(1) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right monoD funpow_mono) next case False with assms(1,4) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right le_eq_less_or_eq less_Suc_eq_le) (simp add: Suc.hyps monoD order_subst1) qed qed lemma inj_fn[simp]: fixes f::"'a \ 'a" assumes "inj f" shows "inj (f^^n)" proof (induction n) case Suc thus ?case using inj_compose[OF assms Suc.IH] by (simp del: comp_apply) qed simp lemma surj_fn[simp]: fixes f::"'a \ 'a" assumes "surj f" shows "surj (f^^n)" proof (induction n) case Suc thus ?case by (simp add: comp_surj[OF Suc.IH assms] del: comp_apply) qed simp lemma bij_fn[simp]: fixes f::"'a \ 'a" assumes "bij f" shows "bij (f^^n)" by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]]) lemma bij_betw_funpow: \<^marker>\contributor \Lars Noschinski\\ assumes "bij_betw f S S" shows "bij_betw (f ^^ n) S S" proof (induct n) case 0 then show ?case by (auto simp: id_def[symmetric]) next case (Suc n) then show ?case unfolding funpow.simps using assms by (rule bij_betw_trans) qed subsection \Kleene iteration\ lemma Kleene_iter_lpfp: fixes f :: "'a::order_bot \ 'a" assumes "mono f" and "f p \ p" shows "(f ^^ k) bot \ p" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma lfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) bot = (f ^^ k) bot" shows "lfp f = (f ^^ k) bot" proof (rule antisym) show "lfp f \ (f ^^ k) bot" proof (rule lfp_lowerbound) show "f ((f ^^ k) bot) \ (f ^^ k) bot" using assms(2) by simp qed show "(f ^^ k) bot \ lfp f" using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp qed lemma mono_pow: "mono f \ mono (f ^^ n)" for f :: "'a \ 'a::complete_lattice" by (induct n) (auto simp: mono_def) lemma lfp_funpow: assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f" proof (rule antisym) show "lfp f \ lfp (f ^^ Suc n)" proof (rule lfp_lowerbound) have "f (lfp (f ^^ Suc n)) = lfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: lfp_rolling f mono_pow comp_def) then show "f (lfp (f ^^ Suc n)) \ lfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (lfp f) = lfp f" for n by (induct n) (auto intro: f lfp_fixpoint) then show "lfp (f ^^ Suc n) \ lfp f" by (intro lfp_lowerbound) (simp del: funpow.simps) qed lemma gfp_funpow: assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f" proof (rule antisym) show "gfp f \ gfp (f ^^ Suc n)" proof (rule gfp_upperbound) have "f (gfp (f ^^ Suc n)) = gfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: gfp_rolling f mono_pow comp_def) then show "f (gfp (f ^^ Suc n)) \ gfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (gfp f) = gfp f" for n by (induct n) (auto intro: f gfp_fixpoint) then show "gfp (f ^^ Suc n) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed lemma Kleene_iter_gpfp: fixes f :: "'a::order_top \ 'a" assumes "mono f" and "p \ f p" shows "p \ (f ^^ k) top" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma gfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) top = (f ^^ k) top" shows "gfp f = (f ^^ k) top" (is "?lhs = ?rhs") proof (rule antisym) have "?rhs \ f ?rhs" using assms(2) by simp then show "?rhs \ ?lhs" by (rule gfp_upperbound) show "?lhs \ ?rhs" using Kleene_iter_gpfp[OF assms(1)] gfp_unfold[OF assms(1)] by simp qed subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin definition of_nat :: "nat \ 'a" where "of_nat n = (plus 1 ^^ n) 0" lemma of_nat_simps [simp]: shows of_nat_0: "of_nat 0 = 0" and of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m" by (simp_all add: of_nat_def) lemma of_nat_1 [simp]: "of_nat 1 = 1" by (simp add: of_nat_def) lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n" by (induct m) (simp_all add: ac_simps) lemma of_nat_mult [simp]: "of_nat (m * n) = of_nat m * of_nat n" by (induct m) (simp_all add: ac_simps distrib_right) lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x" by (induct x) (simp_all add: algebra_simps) primrec of_nat_aux :: "('a \ 'a) \ nat \ 'a \ 'a" where "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \ \tail recursive\ lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "\i. of_nat_aux (\i. i + 1) n (i + 1) = of_nat_aux (\i. i + 1) n i + 1" by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp with Suc show ?case by (simp add: add.commute) qed lemma of_nat_of_bool [simp]: "of_nat (of_bool P) = of_bool P" by auto end declare of_nat_code [code] context semiring_1_cancel begin lemma of_nat_diff: \of_nat (m - n) = of_nat m - of_nat n\ if \n \ m\ proof - from that obtain q where \m = n + q\ by (blast dest: le_Suc_ex) then show ?thesis by simp qed end text \Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.\ class semiring_char_0 = semiring_1 + assumes inj_of_nat: "inj of_nat" begin lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \ m = n" by (auto intro: inj_of_nat injD) text \Special cases where either operand is zero\ lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \ 0 = n" by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0]) lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \ m = 0" by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0]) lemma of_nat_1_eq_iff [simp]: "1 = of_nat n \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_eq_1_iff [simp]: "of_nat n = 1 \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_neq_0 [simp]: "of_nat (Suc n) \ 0" unfolding of_nat_eq_0_iff by simp lemma of_nat_0_neq [simp]: "0 \ of_nat (Suc n)" unfolding of_nat_0_eq_iff by simp end class ring_char_0 = ring_1 + semiring_char_0 context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" by (induct n) simp_all lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" by (auto simp: le_iff_add intro!: add_increasing2) lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" proof(induct m n rule: diff_induct) case (1 m) then show ?case by auto next case (2 n) then show ?case by (simp add: add_pos_nonneg) next case (3 m n) then show ?case by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) lemma less_imp_of_nat_less: "m < n \ of_nat m < of_nat n" by simp lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: order.eq_iff) text \Special cases where either operand is zero\ lemma of_nat_le_0_iff [simp]: "of_nat m \ 0 \ m = 0" by (rule of_nat_le_iff [of _ 0, simplified]) lemma of_nat_0_less_iff [simp]: "0 < of_nat n \ 0 < n" by (rule of_nat_less_iff [of 0, simplified]) end context linordered_nonzero_semiring begin lemma of_nat_max: "of_nat (max x y) = max (of_nat x) (of_nat y)" by (auto simp: max_def ord_class.max_def) lemma of_nat_min: "of_nat (min x y) = min (of_nat x) (of_nat y)" by (auto simp: min_def ord_class.min_def) end context linordered_semidom begin subclass linordered_nonzero_semiring .. subclass semiring_char_0 .. end context linordered_idom begin lemma abs_of_nat [simp]: "\of_nat n\ = of_nat n" by (simp add: abs_if) lemma sgn_of_nat [simp]: "sgn (of_nat n) = of_bool (n > 0)" by simp end lemma of_nat_id [simp]: "of_nat n = n" by (induct n) simp_all lemma of_nat_eq_id [simp]: "of_nat = id" by (auto simp add: fun_eq_iff) subsection \The set of natural numbers\ context semiring_1 begin definition Nats :: "'a set" ("\") where "\ = range of_nat" lemma of_nat_in_Nats [simp]: "of_nat n \ \" by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ \" using of_nat_0 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_1 [simp]: "1 \ \" using of_nat_1 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" unfolding Nats_def using of_nat_add [symmetric] by (blast intro: range_eqI) lemma Nats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" unfolding Nats_def using of_nat_mult [symmetric] by (blast intro: range_eqI) lemma Nats_cases [cases set: Nats]: assumes "x \ \" obtains (of_nat) n where "x = of_nat n" unfolding Nats_def proof - from \x \ \\ have "x \ range of_nat" unfolding Nats_def . then obtain n where "x = of_nat n" .. then show thesis .. qed lemma Nats_induct [case_names of_nat, induct set: Nats]: "x \ \ \ (\n. P (of_nat n)) \ P x" by (rule Nats_cases) auto end lemma Nats_diff [simp]: fixes a:: "'a::linordered_idom" assumes "a \ \" "b \ \" "b \ a" shows "a - b \ \" proof - obtain i where i: "a = of_nat i" using Nats_cases assms by blast obtain j where j: "b = of_nat j" using Nats_cases assms by blast have "j \ i" using \b \ a\ i j of_nat_le_iff by blast then have *: "of_nat i - of_nat j = (of_nat (i-j) :: 'a)" by (simp add: of_nat_diff) then show ?thesis by (simp add: * i j) qed subsection \Further arithmetic facts concerning the natural numbers\ lemma subst_equals: assumes "t = s" and "u = t" shows "u = s" using assms(2,1) by (rule trans) locale nat_arith begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma suc1: "A == k + a \ Suc A \ k + Suc a" by (simp only: add_Suc_right) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/nat_arith.ML\ simproc_setup nateq_cancel_sums ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") = \fn phi => try o Nat_Arith.cancel_eq_conv\ simproc_setup natless_cancel_sums ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") = \fn phi => try o Nat_Arith.cancel_less_conv\ simproc_setup natle_cancel_sums ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "Suc m \ n" | "m \ Suc n") = \fn phi => try o Nat_Arith.cancel_le_conv\ simproc_setup natdiff_cancel_sums ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") = \fn phi => try o Nat_Arith.cancel_diff_conv\ context order begin lemma lift_Suc_mono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_antimono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_mono_less: assumes mono: "\n. f n < f (Suc n)" and "n < n'" shows "f n < f n'" using \n < n'\ by (induct n n' rule: less_Suc_induct) (auto intro: mono) lemma lift_Suc_mono_less_iff: "(\n. f n < f (Suc n)) \ f n < f m \ n < m" by (blast intro: less_asym' lift_Suc_mono_less [of f] dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1]) end lemma mono_iff_le_Suc: "mono f \ (\n. f n \ f (Suc n))" unfolding mono_def by (auto intro: lift_Suc_mono_le [of f]) lemma antimono_iff_le_Suc: "antimono f \ (\n. f (Suc n) \ f n)" unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f]) lemma mono_nat_linear_lb: fixes f :: "nat \ nat" assumes "\m n. m < n \ f m < f n" shows "f m + k \ f (m + k)" proof (induct k) case 0 then show ?case by simp next case (Suc k) then have "Suc (f m + k) \ Suc (f (m + k))" by simp also from assms [of "m + k" "Suc (m + k)"] have "Suc (f (m + k)) \ f (Suc (m + k))" by (simp add: Suc_le_eq) finally show ?case by simp qed text \Subtraction laws, mostly by Clemens Ballarin\ lemma diff_less_mono: fixes a b c :: nat assumes "a < b" and "c \ a" shows "a - c < b - c" proof - from assms obtain d e where "b = c + (d + e)" and "a = c + e" and "d > 0" by (auto dest!: le_Suc_ex less_imp_Suc_add simp add: ac_simps) then show ?thesis by simp qed lemma less_diff_conv: "i < j - k \ i + k < j" for i j k :: nat by (cases "k \ j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex) lemma less_diff_conv2: "k \ j \ j - k < i \ j < i + k" for j k i :: nat by (auto dest: le_Suc_ex) lemma le_diff_conv: "j - k \ i \ j \ i + k" for j k i :: nat by (cases "k \ j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex) lemma diff_diff_cancel [simp]: "i \ n \ n - (n - i) = i" for i n :: nat by (auto dest: le_Suc_ex) lemma diff_less [simp]: "0 < n \ 0 < m \ m - n < m" for i n :: nat by (auto dest: less_imp_Suc_add) text \Simplification of relational expressions involving subtraction\ lemma diff_diff_eq: "k \ m \ k \ n \ m - k - (n - k) = m - n" for m n k :: nat by (auto dest!: le_Suc_ex) hide_fact (open) diff_diff_eq lemma eq_diff_iff: "k \ m \ k \ n \ m - k = n - k \ m = n" for m n k :: nat by (auto dest: le_Suc_ex) lemma less_diff_iff: "k \ m \ k \ n \ m - k < n - k \ m < n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff: "k \ m \ k \ n \ m - k \ n - k \ m \ n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff': "a \ c \ b \ c \ c - a \ c - b \ b \ a" for a b c :: nat by (force dest: le_Suc_ex) text \(Anti)Monotonicity of subtraction -- by Stephan Merz\ lemma diff_le_mono: "m \ n \ m - l \ n - l" for m n l :: nat by (auto dest: less_imp_le less_imp_Suc_add split: nat_diff_split) lemma diff_le_mono2: "m \ n \ l - n \ l - m" for m n l :: nat by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split: nat_diff_split) lemma diff_less_mono2: "m < n \ m < l \ l - n < l - m" for m n l :: nat by (auto dest: less_imp_Suc_add split: nat_diff_split) lemma diffs0_imp_equal: "m - n = 0 \ n - m = 0 \ m = n" for m n :: nat by (simp split: nat_diff_split) lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat by (cases m n rule: le_cases) (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono) lemma inj_on_diff_nat: fixes k :: nat assumes "\n. n \ N \ k \ n" shows "inj_on (\n. n - k) N" proof (rule inj_onI) fix x y assume a: "x \ N" "y \ N" "x - k = y - k" with assms have "x - k + k = y - k + k" by auto with a assms show "x = y" by (auto simp add: eq_diff_iff) qed text \Rewriting to pull differences out\ lemma diff_diff_right [simp]: "k \ j \ i - (j - k) = i + k - j" for i j k :: nat by (fact diff_diff_right) lemma diff_Suc_diff_eq1 [simp]: assumes "k \ j" shows "i - Suc (j - k) = i + k - Suc j" proof - from assms have *: "Suc (j - k) = Suc j - k" by (simp add: Suc_diff_le) from assms have "k \ Suc j" by (rule order_trans) simp with diff_diff_right [of k "Suc j" i] * show ?thesis by simp qed lemma diff_Suc_diff_eq2 [simp]: assumes "k \ j" shows "Suc (j - k) - i = Suc j - (k + i)" proof - from assms obtain n where "j = k + n" by (auto dest: le_Suc_ex) moreover have "Suc n - i = (k + Suc n) - (k + i)" using add_diff_cancel_left [of k "Suc n" i] by simp ultimately show ?thesis by simp qed lemma Suc_diff_Suc: assumes "n < m" shows "Suc (m - Suc n) = m - n" proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp qed lemma one_less_mult: "Suc 0 < n \ Suc 0 < m \ Suc 0 < m * n" using less_1_mult [of n m] by (simp add: ac_simps) lemma n_less_m_mult_n: "0 < n \ Suc 0 < m \ n < m * n" using mult_strict_right_mono [of 1 m n] by simp lemma n_less_n_mult_m: "0 < n \ Suc 0 < m \ n < n * m" using mult_strict_left_mono [of 1 m n] by simp text \Induction starting beyond zero\ lemma nat_induct_at_least [consumes 1, case_names base Suc]: "P n" if "n \ m" "P m" "\n. n \ m \ P n \ P (Suc n)" proof - define q where "q = n - m" with \n \ m\ have "n = m + q" by simp moreover have "P (m + q)" by (induction q) (use that in simp_all) ultimately show "P n" by simp qed lemma nat_induct_non_zero [consumes 1, case_names 1 Suc]: "P n" if "n > 0" "P 1" "\n. n > 0 \ P n \ P (Suc n)" proof - from \n > 0\ have "n \ 1" by (cases n) simp_all moreover note \P 1\ moreover have "\n. n \ 1 \ P n \ P (Suc n)" using \\n. n > 0 \ P n \ P (Suc n)\ by (simp add: Suc_le_eq) ultimately show "P n" by (rule nat_induct_at_least) qed text \Specialized induction principles that work "backwards":\ lemma inc_induct [consumes 1, case_names base step]: assumes less: "i \ j" and base: "P j" and step: "\n. i \ n \ n < j \ P (Suc n) \ P n" shows "P i" using less step proof (induct "j - i" arbitrary: i) case (0 i) then have "i = j" by simp with base show ?case by simp next case (Suc d n) from Suc.hyps have "n \ j" by auto with Suc have "n < j" by (simp add: less_le) from \Suc d = j - n\ have "d + 1 = j - n" by simp then have "d + 1 - 1 = j - n - 1" by simp then have "d = j - n - 1" by simp then have "d = j - (n + 1)" by (simp add: diff_diff_eq) then have "d = j - Suc n" by simp moreover from \n < j\ have "Suc n \ j" by (simp add: Suc_le_eq) ultimately have "P (Suc n)" proof (rule Suc.hyps) fix q assume "Suc n \ q" then have "n \ q" by (simp add: Suc_le_eq less_imp_le) moreover assume "q < j" moreover assume "P (Suc q)" ultimately show "P q" by (rule Suc.prems) qed with order_refl \n < j\ show "P n" by (rule Suc.prems) qed lemma strict_inc_induct [consumes 1, case_names base step]: assumes less: "i < j" and base: "\i. j = Suc i \ P i" and step: "\i. i < j \ P (Suc i) \ P i" shows "P i" using less proof (induct "j - i - 1" arbitrary: i) case (0 i) from \i < j\ obtain n where "j = i + n" and "n > 0" by (auto dest!: less_imp_Suc_add) with 0 have "j = Suc i" by (auto intro: order_antisym simp add: Suc_le_eq) with base show ?case by simp next case (Suc d i) from \Suc d = j - i - 1\ have *: "Suc d = j - Suc i" by (simp add: diff_diff_add) then have "Suc d - 1 = j - Suc i - 1" by simp then have "d = j - Suc i - 1" by simp moreover from * have "j - Suc i \ 0" by auto then have "Suc i < j" by (simp add: not_le) ultimately have "P (Suc i)" by (rule Suc.hyps) with \i < j\ show "P i" by (rule step) qed lemma zero_induct_lemma: "P k \ (\n. P (Suc n) \ P n) \ P (k - i)" using inc_induct[of "k - i" k P, simplified] by blast lemma zero_induct: "P k \ (\n. P (Suc n) \ P n) \ P 0" using inc_induct[of 0 k P] by blast text \Further induction rule similar to @{thm inc_induct}.\ lemma dec_induct [consumes 1, case_names base step]: "i \ j \ P i \ (\n. i \ n \ n < j \ P n \ P (Suc n)) \ P j" proof (induct j arbitrary: i) case 0 then show ?case by simp next case (Suc j) from Suc.prems consider "i \ j" | "i = Suc j" by (auto simp add: le_Suc_eq) then show ?case proof cases case 1 moreover have "j < Suc j" by simp moreover have "P j" using \i \ j\ \P i\ proof (rule Suc.hyps) fix q assume "i \ q" moreover assume "q < j" then have "q < Suc j" by (simp add: less_Suc_eq) moreover assume "P q" ultimately show "P (Suc q)" by (rule Suc.prems) qed ultimately show "P (Suc j)" by (rule Suc.prems) next case 2 with \P i\ show "P (Suc j)" by simp qed qed lemma transitive_stepwise_le: assumes "m \ n" "\x. R x x" "\x y z. R x y \ R y z \ R x z" and "\n. R n (Suc n)" shows "R m n" using \m \ n\ by (induction rule: dec_induct) (use assms in blast)+ subsubsection \Greatest operator\ lemma ex_has_greatest_nat: "P (k::nat) \ \y. P y \ y \ b \ \x. P x \ (\y. P y \ y \ x)" proof (induction "b-k" arbitrary: b k rule: less_induct) case less show ?case proof cases assume "\n>k. P n" then obtain n where "n>k" "P n" by blast have "n \ b" using \P n\ less.prems(2) by auto hence "b-n < b-k" by(rule diff_less_mono2[OF \k less_le_trans[OF \k]]) from less.hyps[OF this \P n\ less.prems(2)] show ?thesis . next assume "\ (\n>k. P n)" hence "\y. P y \ y \ k" by (auto simp: not_less) thus ?thesis using less.prems(1) by auto qed qed lemma fixes k::nat assumes "P k" and minor: "\y. P y \ y \ b" shows GreatestI_nat: "P (Greatest P)" and Greatest_le_nat: "k \ Greatest P" proof - obtain x where "P x" "\y. P y \ y \ x" using assms ex_has_greatest_nat by blast with \P k\ show "P (Greatest P)" "k \ Greatest P" using GreatestI2_order by blast+ qed lemma GreatestI_ex_nat: "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" by (blast intro: GreatestI_nat) subsection \Monotonicity of \funpow\\ lemma funpow_increasing: "m \ n \ mono f \ (f ^^ n) \ \ (f ^^ m) \" for f :: "'a::{lattice,order_top} \ 'a" by (induct rule: inc_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma funpow_decreasing: "m \ n \ mono f \ (f ^^ m) \ \ (f ^^ n) \" for f :: "'a::{lattice,order_bot} \ 'a" by (induct rule: dec_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma mono_funpow: "mono Q \ mono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_bot} \ 'a" by (auto intro!: funpow_decreasing simp: mono_def) lemma antimono_funpow: "mono Q \ antimono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_top} \ 'a" by (auto intro!: funpow_increasing simp: antimono_def) subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def) lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \ m = Suc 0" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \ m = 1" for m :: nat by (simp add: dvd_def) lemma dvd_antisym: "m dvd n \ n dvd m \ m = n" for m n :: nat unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) lemma dvd_diff_nat [simp]: "k dvd m \ k dvd n \ k dvd (m - n)" for k m n :: nat unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric]) lemma dvd_diffD: fixes k m n :: nat assumes "k dvd m - n" "k dvd n" "n \ m" shows "k dvd m" proof - have "k dvd n + (m - n)" using assms by (blast intro: dvd_add) with assms show ?thesis by simp qed lemma dvd_diffD1: "k dvd m - n \ k dvd m \ n \ m \ k dvd n" for k m n :: nat by (drule_tac m = m in dvd_diff_nat) auto lemma dvd_mult_cancel: fixes m n k :: nat assumes "k * m dvd k * n" and "0 < k" shows "m dvd n" proof - from assms(1) obtain q where "k * n = (k * m) * q" .. then have "k * n = k * (m * q)" by (simp add: ac_simps) with \0 < k\ have "n = m * q" by (auto simp add: mult_left_cancel) then show ?thesis .. qed lemma dvd_mult_cancel1: fixes m n :: nat assumes "0 < m" shows "m * n dvd m \ n = 1" proof assume "m * n dvd m" then have "m * n dvd m * 1" by simp then have "n dvd 1" by (iprover intro: assms dvd_mult_cancel) then show "n = 1" by auto qed auto lemma dvd_mult_cancel2: "0 < m \ n * m dvd m \ n = 1" for m n :: nat using dvd_mult_cancel1 [of m n] by (simp add: ac_simps) lemma dvd_imp_le: "k dvd n \ 0 < n \ k \ n" for k n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma nat_dvd_not_less: "0 < m \ m < n \ \ n dvd m" for m n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma less_eq_dvd_minus: fixes m n :: nat assumes "m \ n" shows "m dvd n \ m dvd n - m" proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. then show ?thesis by (simp add: add.commute [of m]) qed lemma dvd_minus_self: "m dvd n - m \ n < m \ m dvd n" for m n :: nat by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le) lemma dvd_minus_add: fixes m n q r :: nat assumes "q \ n" "q \ r * m" shows "m dvd n - q \ m dvd n + (r * m - q)" proof - have "m dvd n - q \ m dvd r * m + (n - q)" using dvd_add_times_triv_left_iff [of m r] by simp also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed subsection \Aliasses\ lemma nat_mult_1: "1 * n = n" for n :: nat by (fact mult_1_left) lemma nat_mult_1_right: "n * 1 = n" for n :: nat by (fact mult_1_right) lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat by (fact left_diff_distrib') lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat by (fact right_diff_distrib') (*Used in AUTO2 and Groups.le_diff_conv2 (with variables renamed) doesn't work for some reason*) lemma le_diff_conv2: "k \ j \ (i \ j - k) = (i + k \ j)" for i j k :: nat by (fact le_diff_conv2) lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat by (fact diff_cancel) lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat by (fact diff_diff_add) lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat by (fact add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = 0" for m n :: nat by (fact diff_add_zero) lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat by (fact distrib_left) lemmas nat_distrib = add_mult_distrib distrib_left diff_mult_distrib diff_mult_distrib2 subsection \Size of a datatype value\ class size = fixes size :: "'a \ nat" \ \see further theory \Wellfounded\\ instantiation nat :: size begin definition size_nat where [simp, code]: "size (n::nat) = n" instance .. end lemmas size_nat = size_nat_def lemma size_neq_size_imp_neq: "size x \ size y \ x \ y" by (erule contrapos_nn) (rule arg_cong) subsection \Code module namespace\ code_identifier code_module Nat \ (SML) Arith and (OCaml) Arith and (Haskell) Arith hide_const (open) of_nat_aux end diff --git a/src/HOL/Num.thy b/src/HOL/Num.thy --- a/src/HOL/Num.thy +++ b/src/HOL/Num.thy @@ -1,1530 +1,1553 @@ (* Title: HOL/Num.thy Author: Florian Haftmann Author: Brian Huffman *) section \Binary Numerals\ theory Num imports BNF_Least_Fixpoint Transfer begin subsection \The \num\ type\ datatype num = One | Bit0 num | Bit1 num text \Increment function for type \<^typ>\num\\ primrec inc :: "num \ num" where "inc One = Bit0 One" | "inc (Bit0 x) = Bit1 x" | "inc (Bit1 x) = Bit0 (inc x)" text \Converting between type \<^typ>\num\ and type \<^typ>\nat\\ primrec nat_of_num :: "num \ nat" where "nat_of_num One = Suc 0" | "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" primrec num_of_nat :: "nat \ num" where "num_of_nat 0 = One" | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" lemma nat_of_num_pos: "0 < nat_of_num x" by (induct x) simp_all lemma nat_of_num_neq_0: " nat_of_num x \ 0" by (induct x) simp_all lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" by (induct x) simp_all lemma num_of_nat_double: "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" by (induct n) simp_all text \Type \<^typ>\num\ is isomorphic to the strictly positive natural numbers.\ lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" by (induct n) (simp_all add: nat_of_num_inc) lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" apply safe apply (drule arg_cong [where f=num_of_nat]) apply (simp add: nat_of_num_inverse) done lemma num_induct [case_names One inc]: fixes P :: "num \ bool" assumes One: "P One" and inc: "\x. P x \ P (inc x)" shows "P x" proof - obtain n where n: "Suc n = nat_of_num x" by (cases "nat_of_num x") (simp_all add: nat_of_num_neq_0) have "P (num_of_nat (Suc n))" proof (induct n) case 0 from One show ?case by simp next case (Suc n) then have "P (inc (num_of_nat (Suc n)))" by (rule inc) then show "P (num_of_nat (Suc (Suc n)))" by simp qed with n show "P x" by (simp add: nat_of_num_inverse) qed text \ From now on, there are two possible models for \<^typ>\num\: as positive naturals (rule \num_induct\) and as digit representation (rules \num.induct\, \num.cases\). \ subsection \Numeral operations\ instantiation num :: "{plus,times,linorder}" begin definition [code del]: "m + n = num_of_nat (nat_of_num m + nat_of_num n)" definition [code del]: "m * n = num_of_nat (nat_of_num m * nat_of_num n)" definition [code del]: "m \ n \ nat_of_num m \ nat_of_num n" definition [code del]: "m < n \ nat_of_num m < nat_of_num n" instance by standard (auto simp add: less_num_def less_eq_num_def num_eq_iff) end lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" unfolding plus_num_def by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" unfolding times_num_def by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) lemma add_num_simps [simp, code]: "One + One = Bit0 One" "One + Bit0 n = Bit1 n" "One + Bit1 n = Bit0 (n + One)" "Bit0 m + One = Bit1 m" "Bit0 m + Bit0 n = Bit0 (m + n)" "Bit0 m + Bit1 n = Bit1 (m + n)" "Bit1 m + One = Bit0 (m + One)" "Bit1 m + Bit0 n = Bit1 (m + n)" "Bit1 m + Bit1 n = Bit0 (m + n + One)" by (simp_all add: num_eq_iff nat_of_num_add) lemma mult_num_simps [simp, code]: "m * One = m" "One * n = n" "Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))" "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult distrib_right distrib_left) lemma eq_num_simps: "One = One \ True" "One = Bit0 n \ False" "One = Bit1 n \ False" "Bit0 m = One \ False" "Bit1 m = One \ False" "Bit0 m = Bit0 n \ m = n" "Bit0 m = Bit1 n \ False" "Bit1 m = Bit0 n \ False" "Bit1 m = Bit1 n \ m = n" by simp_all lemma le_num_simps [simp, code]: "One \ n \ True" "Bit0 m \ One \ False" "Bit1 m \ One \ False" "Bit0 m \ Bit0 n \ m \ n" "Bit0 m \ Bit1 n \ m \ n" "Bit1 m \ Bit1 n \ m \ n" "Bit1 m \ Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma less_num_simps [simp, code]: "m < One \ False" "One < Bit0 n \ True" "One < Bit1 n \ True" "Bit0 m < Bit0 n \ m < n" "Bit0 m < Bit1 n \ m \ n" "Bit1 m < Bit1 n \ m < n" "Bit1 m < Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma le_num_One_iff: "x \ num.One \ x = num.One" by (simp add: antisym_conv) text \Rules using \One\ and \inc\ as constructors.\ lemma add_One: "x + One = inc x" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma add_One_commute: "One + n = n + One" by (induct n) simp_all lemma add_inc: "x + inc y = inc (x + y)" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma mult_inc: "x * inc y = x * y + x" by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) text \The \<^const>\num_of_nat\ conversion.\ lemma num_of_nat_One: "n \ 1 \ num_of_nat n = One" by (cases n) simp_all lemma num_of_nat_plus_distrib: "0 < m \ 0 < n \ num_of_nat (m + n) = num_of_nat m + num_of_nat n" by (induct n) (auto simp add: add_One add_One_commute add_inc) text \A double-and-decrement function.\ primrec BitM :: "num \ num" where "BitM One = One" | "BitM (Bit0 n) = Bit1 (BitM n)" | "BitM (Bit1 n) = Bit1 (Bit0 n)" lemma BitM_plus_one: "BitM n + One = Bit0 n" by (induct n) simp_all lemma one_plus_BitM: "One + BitM n = Bit0 n" unfolding add_One_commute BitM_plus_one .. lemma BitM_inc_eq: \Num.BitM (Num.inc n) = Num.Bit1 n\ by (induction n) simp_all lemma inc_BitM_eq: \Num.inc (Num.BitM n) = num.Bit0 n\ by (simp add: BitM_plus_one[symmetric] add_One) text \Squaring and exponentiation.\ primrec sqr :: "num \ num" where "sqr One = One" | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" primrec pow :: "num \ num \ num" where "pow x One = x" | "pow x (Bit0 y) = sqr (pow x y)" | "pow x (Bit1 y) = sqr (pow x y) * x" lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" by (induct x) (simp_all add: algebra_simps nat_of_num_add) lemma sqr_conv_mult: "sqr x = x * x" by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) lemma num_double [simp]: "num.Bit0 num.One * n = num.Bit0 n" by (simp add: num_eq_iff nat_of_num_mult) subsection \Binary numerals\ text \ We embed binary representations into a generic algebraic structure using \numeral\. \ class numeral = one + semigroup_add begin primrec numeral :: "num \ 'a" where numeral_One: "numeral One = 1" | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" lemma numeral_code [code]: "numeral One = 1" "numeral (Bit0 n) = (let m = numeral n in m + m)" "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" by (simp_all add: Let_def) lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) next case Bit1 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) qed lemma numeral_inc: "numeral (inc x) = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by simp next case (Bit1 x) have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" by (simp only: one_plus_numeral_commute) with Bit1 show ?case by (simp add: add.assoc) qed declare numeral.simps [simp del] abbreviation "Numeral1 \ numeral One" declare numeral_One [code_post] end text \Numeral syntax.\ syntax "_Numeral" :: "num_const \ 'a" ("_") ML_file \Tools/numeral.ML\ parse_translation \ let fun numeral_tr [(c as Const (\<^syntax_const>\_constrain\, _)) $ t $ u] = c $ numeral_tr [t] $ u | numeral_tr [Const (num, _)] = (Numeral.mk_number_syntax o #value o Lexicon.read_num) num | numeral_tr ts = raise TERM ("numeral_tr", ts); in [(\<^syntax_const>\_Numeral\, K numeral_tr)] end \ typed_print_translation \ let fun num_tr' ctxt T [n] = let val k = Numeral.dest_num_syntax n; val t' = Syntax.const \<^syntax_const>\_Numeral\ $ Syntax.free (string_of_int k); in (case T of Type (\<^type_name>\fun\, [_, T']) => if Printer.type_emphasis ctxt T' then Syntax.const \<^syntax_const>\_constrain\ $ t' $ Syntax_Phases.term_of_typ ctxt T' else t' | _ => if T = dummyT then t' else raise Match) end; in [(\<^const_syntax>\numeral\, num_tr')] end \ subsection \Class-specific numeral rules\ text \\<^const>\numeral\ is a morphism.\ subsubsection \Structures with addition: class \numeral\\ context numeral begin lemma numeral_add: "numeral (m + n) = numeral m + numeral n" by (induct n rule: num_induct) (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" by (rule numeral_add [symmetric]) lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" using numeral_add [of n One] by (simp add: numeral_One) lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" using numeral_add [of One n] by (simp add: numeral_One) lemma one_add_one: "1 + 1 = 2" using numeral_add [of One One] by (simp add: numeral_One) lemmas add_numeral_special = numeral_plus_one one_plus_numeral one_add_one end subsubsection \Structures with negation: class \neg_numeral\\ class neg_numeral = numeral + group_add begin lemma uminus_numeral_One: "- Numeral1 = - 1" by (simp add: numeral_One) text \Numerals form an abelian subgroup.\ inductive is_num :: "'a \ bool" where "is_num 1" | "is_num x \ is_num (- x)" | "is_num x \ is_num y \ is_num (x + y)" lemma is_num_numeral: "is_num (numeral k)" by (induct k) (simp_all add: numeral.simps is_num.intros) lemma is_num_add_commute: "is_num x \ is_num y \ x + y = y + x" - apply (induct x rule: is_num.induct) - apply (induct y rule: is_num.induct) - apply simp - apply (rule_tac a=x in add_left_imp_eq) - apply (rule_tac a=x in add_right_imp_eq) - apply (simp add: add.assoc) - apply (simp add: add.assoc [symmetric]) - apply (simp add: add.assoc) - apply (rule_tac a=x in add_left_imp_eq) - apply (rule_tac a=x in add_right_imp_eq) - apply (simp add: add.assoc) - apply (simp add: add.assoc) - apply (simp add: add.assoc [symmetric]) - done +proof(induction x rule: is_num.induct) + case 1 + then show ?case + proof (induction y rule: is_num.induct) + case 1 + then show ?case by simp + next + case (2 y) + then have "y + (1 + - y) + y = y + (- y + 1) + y" + by (simp add: add.assoc) + then have "y + (1 + - y) = y + (- y + 1)" + by simp + then show ?case + by (rule add_left_imp_eq[of y]) + next + case (3 x y) + then have "1 + (x + y) = x + 1 + y" + by (simp add: add.assoc [symmetric]) + then show ?case using 3 + by (simp add: add.assoc) + qed +next + case (2 x) + then have "x + (- x + y) + x = x + (y + - x) + x" + by (simp add: add.assoc) + then have "x + (- x + y) = x + (y + - x)" + by simp + then show ?case + by (rule add_left_imp_eq[of x]) +next + case (3 x z) + moreover have "x + (y + z) = (x + y) + z" + by (simp add: add.assoc[symmetric]) + ultimately show ?case + by (simp add: add.assoc) +qed lemma is_num_add_left_commute: "is_num x \ is_num y \ x + (y + z) = y + (x + z)" by (simp only: add.assoc [symmetric] is_num_add_commute) lemmas is_num_normalize = add.assoc is_num_add_commute is_num_add_left_commute is_num.intros is_num_numeral minus_add definition dbl :: "'a \ 'a" where "dbl x = x + x" definition dbl_inc :: "'a \ 'a" where "dbl_inc x = x + x + 1" definition dbl_dec :: "'a \ 'a" where "dbl_dec x = x + x - 1" definition sub :: "num \ num \ 'a" where "sub k l = numeral k - numeral l" lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) lemma sub_inc_One_eq: \Num.sub (Num.inc n) num.One = numeral n\ by (simp_all add: sub_def diff_eq_eq numeral_inc numeral.numeral_One) lemma dbl_simps [simp]: "dbl (- numeral k) = - dbl (numeral k)" "dbl 0 = 0" "dbl 1 = 2" "dbl (- 1) = - 2" "dbl (numeral k) = numeral (Bit0 k)" by (simp_all add: dbl_def numeral.simps minus_add) lemma dbl_inc_simps [simp]: "dbl_inc (- numeral k) = - dbl_dec (numeral k)" "dbl_inc 0 = 1" "dbl_inc 1 = 3" "dbl_inc (- 1) = - 1" "dbl_inc (numeral k) = numeral (Bit1 k)" by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff) lemma dbl_dec_simps [simp]: "dbl_dec (- numeral k) = - dbl_inc (numeral k)" "dbl_dec 0 = - 1" "dbl_dec 1 = 1" "dbl_dec (- 1) = - 3" "dbl_dec (numeral k) = numeral (BitM k)" by (simp_all add: dbl_dec_def dbl_inc_def numeral.simps numeral_BitM is_num_normalize) lemma sub_num_simps [simp]: "sub One One = 0" "sub One (Bit0 l) = - numeral (BitM l)" "sub One (Bit1 l) = - numeral (Bit0 l)" "sub (Bit0 k) One = numeral (BitM k)" "sub (Bit1 k) One = numeral (Bit0 k)" "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def numeral.simps numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_simps: "numeral m + - numeral n = sub m n" "- numeral m + numeral n = sub n m" "- numeral m + - numeral n = - (numeral m + numeral n)" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_special: "1 + - numeral m = sub One m" "- numeral m + 1 = sub One m" "numeral m + - 1 = sub m One" "- 1 + numeral n = sub n One" "- 1 + - numeral n = - numeral (inc n)" "- numeral m + - 1 = - numeral (inc m)" "1 + - 1 = 0" "- 1 + 1 = 0" "- 1 + - 1 = - 2" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_simps: "numeral m - numeral n = sub m n" "numeral m - - numeral n = numeral (m + n)" "- numeral m - numeral n = - numeral (m + n)" "- numeral m - - numeral n = sub n m" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_special: "1 - numeral n = sub One n" "numeral m - 1 = sub m One" "1 - - numeral n = numeral (One + n)" "- numeral m - 1 = - numeral (m + One)" "- 1 - numeral n = - numeral (inc n)" "numeral m - - 1 = numeral (inc m)" "- 1 - - numeral n = sub n One" "- numeral m - - 1 = sub One m" "1 - 1 = 0" "- 1 - 1 = - 2" "1 - - 1 = 2" "- 1 - - 1 = 0" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) end subsubsection \Structures with multiplication: class \semiring_numeral\\ class semiring_numeral = semiring + monoid_mult begin subclass numeral .. lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" by (induct n rule: num_induct) (simp_all add: numeral_One mult_inc numeral_inc numeral_add distrib_left) lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" by (rule numeral_mult [symmetric]) lemma mult_2: "2 * z = z + z" by (simp add: one_add_one [symmetric] distrib_right) lemma mult_2_right: "z * 2 = z + z" by (simp add: one_add_one [symmetric] distrib_left) lemma left_add_twice: "a + (a + b) = 2 * a + b" by (simp add: mult_2 ac_simps) end subsubsection \Structures with a zero: class \semiring_1\\ context semiring_1 begin subclass semiring_numeral .. lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) end lemma nat_of_num_numeral [code_abbrev]: "nat_of_num = numeral" proof fix n have "numeral n = nat_of_num n" by (induct n) (simp_all add: numeral.simps) then show "nat_of_num n = numeral n" by simp qed lemma nat_of_num_code [code]: "nat_of_num One = 1" "nat_of_num (Bit0 n) = (let m = nat_of_num n in m + m)" "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))" by (simp_all add: Let_def) subsubsection \Equality: class \semiring_char_0\\ context semiring_char_0 begin lemma numeral_eq_iff: "numeral m = numeral n \ m = n" by (simp only: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] of_nat_eq_iff num_eq_iff) lemma numeral_eq_one_iff: "numeral n = 1 \ n = One" by (rule numeral_eq_iff [of n One, unfolded numeral_One]) lemma one_eq_numeral_iff: "1 = numeral n \ One = n" by (rule numeral_eq_iff [of One n, unfolded numeral_One]) lemma numeral_neq_zero: "numeral n \ 0" by (simp add: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] nat_of_num_pos) lemma zero_neq_numeral: "0 \ numeral n" unfolding eq_commute [of 0] by (rule numeral_neq_zero) lemmas eq_numeral_simps [simp] = numeral_eq_iff numeral_eq_one_iff one_eq_numeral_iff numeral_neq_zero zero_neq_numeral end subsubsection \Comparisons: class \linordered_nonzero_semiring\\ context linordered_nonzero_semiring begin lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" proof - have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" by (simp only: less_eq_num_def nat_of_num_numeral of_nat_le_iff) then show ?thesis by simp qed lemma one_le_numeral: "1 \ numeral n" using numeral_le_iff [of num.One n] by (simp add: numeral_One) lemma numeral_le_one_iff: "numeral n \ 1 \ n \ num.One" using numeral_le_iff [of n num.One] by (simp add: numeral_One) lemma numeral_less_iff: "numeral m < numeral n \ m < n" proof - have "of_nat (numeral m) < of_nat (numeral n) \ m < n" unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. then show ?thesis by simp qed lemma not_numeral_less_one: "\ numeral n < 1" using numeral_less_iff [of n num.One] by (simp add: numeral_One) lemma one_less_numeral_iff: "1 < numeral n \ num.One < n" using numeral_less_iff [of num.One n] by (simp add: numeral_One) lemma zero_le_numeral: "0 \ numeral n" using dual_order.trans one_le_numeral zero_le_one by blast lemma zero_less_numeral: "0 < numeral n" using less_linear not_numeral_less_one order.strict_trans zero_less_one by blast lemma not_numeral_le_zero: "\ numeral n \ 0" by (simp add: not_le zero_less_numeral) lemma not_numeral_less_zero: "\ numeral n < 0" by (simp add: not_less zero_le_numeral) lemmas le_numeral_extra = zero_le_one not_one_le_zero order_refl [of 0] order_refl [of 1] lemmas less_numeral_extra = zero_less_one not_one_less_zero less_irrefl [of 0] less_irrefl [of 1] lemmas le_numeral_simps [simp] = numeral_le_iff one_le_numeral numeral_le_one_iff zero_le_numeral not_numeral_le_zero lemmas less_numeral_simps [simp] = numeral_less_iff one_less_numeral_iff not_numeral_less_one zero_less_numeral not_numeral_less_zero lemma min_0_1 [simp]: fixes min' :: "'a \ 'a \ 'a" defines "min' \ min" shows "min' 0 1 = 0" "min' 1 0 = 0" "min' 0 (numeral x) = 0" "min' (numeral x) 0 = 0" "min' 1 (numeral x) = 1" "min' (numeral x) 1 = 1" by (simp_all add: min'_def min_def le_num_One_iff) lemma max_0_1 [simp]: fixes max' :: "'a \ 'a \ 'a" defines "max' \ max" shows "max' 0 1 = 1" "max' 1 0 = 1" "max' 0 (numeral x) = numeral x" "max' (numeral x) 0 = numeral x" "max' 1 (numeral x) = numeral x" "max' (numeral x) 1 = numeral x" by (simp_all add: max'_def max_def le_num_One_iff) end text \Unfold \min\ and \max\ on numerals.\ lemmas max_number_of [simp] = max_def [of "numeral u" "numeral v"] max_def [of "numeral u" "- numeral v"] max_def [of "- numeral u" "numeral v"] max_def [of "- numeral u" "- numeral v"] for u v lemmas min_number_of [simp] = min_def [of "numeral u" "numeral v"] min_def [of "numeral u" "- numeral v"] min_def [of "- numeral u" "numeral v"] min_def [of "- numeral u" "- numeral v"] for u v subsubsection \Multiplication and negation: class \ring_1\\ context ring_1 begin subclass neg_numeral .. lemma mult_neg_numeral_simps: "- numeral m * - numeral n = numeral (m * n)" "- numeral m * numeral n = - numeral (m * n)" "numeral m * - numeral n = - numeral (m * n)" by (simp_all only: mult_minus_left mult_minus_right minus_minus numeral_mult) lemma mult_minus1 [simp]: "- 1 * z = - z" by (simp add: numeral.simps) lemma mult_minus1_right [simp]: "z * - 1 = - z" by (simp add: numeral.simps) lemma minus_sub_one_diff_one [simp]: \- sub m One - 1 = - numeral m\ proof - have \sub m One + 1 = numeral m\ by (simp flip: eq_diff_eq add: diff_numeral_special) then have \- (sub m One + 1) = - numeral m\ by simp then show ?thesis by simp qed end subsubsection \Equality using \iszero\ for rings with non-zero characteristic\ context ring_1 begin definition iszero :: "'a \ bool" where "iszero z \ z = 0" lemma iszero_0 [simp]: "iszero 0" by (simp add: iszero_def) lemma not_iszero_1 [simp]: "\ iszero 1" by (simp add: iszero_def) lemma not_iszero_Numeral1: "\ iszero Numeral1" by (simp add: numeral_One) lemma not_iszero_neg_1 [simp]: "\ iszero (- 1)" by (simp add: iszero_def) lemma not_iszero_neg_Numeral1: "\ iszero (- Numeral1)" by (simp add: numeral_One) lemma iszero_neg_numeral [simp]: "iszero (- numeral w) \ iszero (numeral w)" unfolding iszero_def by (rule neg_equal_0_iff_equal) lemma eq_iff_iszero_diff: "x = y \ iszero (x - y)" unfolding iszero_def by (rule eq_iff_diff_eq_0) text \ The \eq_numeral_iff_iszero\ lemmas are not declared \[simp]\ by default, because for rings of characteristic zero, better simp rules are possible. For a type like integers mod \n\, type-instantiated versions of these rules should be added to the simplifier, along with a type-specific rule for deciding propositions of the form \iszero (numeral w)\. bh: Maybe it would not be so bad to just declare these as simp rules anyway? I should test whether these rules take precedence over the \ring_char_0\ rules in the simplifier. \ lemma eq_numeral_iff_iszero: "numeral x = numeral y \ iszero (sub x y)" "numeral x = - numeral y \ iszero (numeral (x + y))" "- numeral x = numeral y \ iszero (numeral (x + y))" "- numeral x = - numeral y \ iszero (sub y x)" "numeral x = 1 \ iszero (sub x One)" "1 = numeral y \ iszero (sub One y)" "- numeral x = 1 \ iszero (numeral (x + One))" "1 = - numeral y \ iszero (numeral (One + y))" "numeral x = 0 \ iszero (numeral x)" "0 = numeral y \ iszero (numeral y)" "- numeral x = 0 \ iszero (numeral x)" "0 = - numeral y \ iszero (numeral y)" unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special by simp_all end subsubsection \Equality and negation: class \ring_char_0\\ context ring_char_0 begin lemma not_iszero_numeral [simp]: "\ iszero (numeral w)" by (simp add: iszero_def) lemma neg_numeral_eq_iff: "- numeral m = - numeral n \ m = n" by simp lemma numeral_neq_neg_numeral: "numeral m \ - numeral n" by (simp add: eq_neg_iff_add_eq_0 numeral_plus_numeral) lemma neg_numeral_neq_numeral: "- numeral m \ numeral n" by (rule numeral_neq_neg_numeral [symmetric]) lemma zero_neq_neg_numeral: "0 \ - numeral n" by simp lemma neg_numeral_neq_zero: "- numeral n \ 0" by simp lemma one_neq_neg_numeral: "1 \ - numeral n" using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) lemma neg_numeral_neq_one: "- numeral n \ 1" using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) lemma neg_one_neq_numeral: "- 1 \ numeral n" using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One) lemma numeral_neq_neg_one: "numeral n \ - 1" using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One) lemma neg_one_eq_numeral_iff: "- 1 = - numeral n \ n = One" using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One) lemma numeral_eq_neg_one_iff: "- numeral n = - 1 \ n = One" using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One) lemma neg_one_neq_zero: "- 1 \ 0" by simp lemma zero_neq_neg_one: "0 \ - 1" by simp lemma neg_one_neq_one: "- 1 \ 1" using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemma one_neq_neg_one: "1 \ - 1" using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemmas eq_neg_numeral_simps [simp] = neg_numeral_eq_iff numeral_neq_neg_numeral neg_numeral_neq_numeral one_neq_neg_numeral neg_numeral_neq_one zero_neq_neg_numeral neg_numeral_neq_zero neg_one_neq_numeral numeral_neq_neg_one neg_one_eq_numeral_iff numeral_eq_neg_one_iff neg_one_neq_zero zero_neq_neg_one neg_one_neq_one one_neq_neg_one end subsubsection \Structures with negation and order: class \linordered_idom\\ context linordered_idom begin subclass ring_char_0 .. lemma neg_numeral_le_iff: "- numeral m \ - numeral n \ n \ m" by (simp only: neg_le_iff_le numeral_le_iff) lemma neg_numeral_less_iff: "- numeral m < - numeral n \ n < m" by (simp only: neg_less_iff_less numeral_less_iff) lemma neg_numeral_less_zero: "- numeral n < 0" by (simp only: neg_less_0_iff_less zero_less_numeral) lemma neg_numeral_le_zero: "- numeral n \ 0" by (simp only: neg_le_0_iff_le zero_le_numeral) lemma not_zero_less_neg_numeral: "\ 0 < - numeral n" by (simp only: not_less neg_numeral_le_zero) lemma not_zero_le_neg_numeral: "\ 0 \ - numeral n" by (simp only: not_le neg_numeral_less_zero) lemma neg_numeral_less_numeral: "- numeral m < numeral n" using neg_numeral_less_zero zero_less_numeral by (rule less_trans) lemma neg_numeral_le_numeral: "- numeral m \ numeral n" by (simp only: less_imp_le neg_numeral_less_numeral) lemma not_numeral_less_neg_numeral: "\ numeral m < - numeral n" by (simp only: not_less neg_numeral_le_numeral) lemma not_numeral_le_neg_numeral: "\ numeral m \ - numeral n" by (simp only: not_le neg_numeral_less_numeral) lemma neg_numeral_less_one: "- numeral m < 1" by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) lemma neg_numeral_le_one: "- numeral m \ 1" by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) lemma not_one_less_neg_numeral: "\ 1 < - numeral m" by (simp only: not_less neg_numeral_le_one) lemma not_one_le_neg_numeral: "\ 1 \ - numeral m" by (simp only: not_le neg_numeral_less_one) lemma not_numeral_less_neg_one: "\ numeral m < - 1" using not_numeral_less_neg_numeral [of m One] by (simp add: numeral_One) lemma not_numeral_le_neg_one: "\ numeral m \ - 1" using not_numeral_le_neg_numeral [of m One] by (simp add: numeral_One) lemma neg_one_less_numeral: "- 1 < numeral m" using neg_numeral_less_numeral [of One m] by (simp add: numeral_One) lemma neg_one_le_numeral: "- 1 \ numeral m" using neg_numeral_le_numeral [of One m] by (simp add: numeral_One) lemma neg_numeral_less_neg_one_iff: "- numeral m < - 1 \ m \ One" by (cases m) simp_all lemma neg_numeral_le_neg_one: "- numeral m \ - 1" by simp lemma not_neg_one_less_neg_numeral: "\ - 1 < - numeral m" by simp lemma not_neg_one_le_neg_numeral_iff: "\ - 1 \ - numeral m \ m \ One" by (cases m) simp_all lemma sub_non_negative: "sub n m \ 0 \ n \ m" by (simp only: sub_def le_diff_eq) simp lemma sub_positive: "sub n m > 0 \ n > m" by (simp only: sub_def less_diff_eq) simp lemma sub_non_positive: "sub n m \ 0 \ n \ m" by (simp only: sub_def diff_le_eq) simp lemma sub_negative: "sub n m < 0 \ n < m" by (simp only: sub_def diff_less_eq) simp lemmas le_neg_numeral_simps [simp] = neg_numeral_le_iff neg_numeral_le_numeral not_numeral_le_neg_numeral neg_numeral_le_zero not_zero_le_neg_numeral neg_numeral_le_one not_one_le_neg_numeral neg_one_le_numeral not_numeral_le_neg_one neg_numeral_le_neg_one not_neg_one_le_neg_numeral_iff lemma le_minus_one_simps [simp]: "- 1 \ 0" "- 1 \ 1" "\ 0 \ - 1" "\ 1 \ - 1" by simp_all lemmas less_neg_numeral_simps [simp] = neg_numeral_less_iff neg_numeral_less_numeral not_numeral_less_neg_numeral neg_numeral_less_zero not_zero_less_neg_numeral neg_numeral_less_one not_one_less_neg_numeral neg_one_less_numeral not_numeral_less_neg_one neg_numeral_less_neg_one_iff not_neg_one_less_neg_numeral lemma less_minus_one_simps [simp]: "- 1 < 0" "- 1 < 1" "\ 0 < - 1" "\ 1 < - 1" by (simp_all add: less_le) lemma abs_numeral [simp]: "\numeral n\ = numeral n" by simp lemma abs_neg_numeral [simp]: "\- numeral n\ = numeral n" by (simp only: abs_minus_cancel abs_numeral) lemma abs_neg_one [simp]: "\- 1\ = 1" by simp end subsubsection \Natural numbers\ lemma numeral_num_of_nat: "numeral (num_of_nat n) = n" if "n > 0" using that nat_of_num_numeral num_of_nat_inverse by simp lemma Suc_1 [simp]: "Suc 1 = 2" unfolding Suc_eq_plus1 by (rule one_add_one) lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" unfolding Suc_eq_plus1 by (rule numeral_plus_one) definition pred_numeral :: "num \ nat" where "pred_numeral k = numeral k - 1" declare [[code drop: pred_numeral]] lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" by (simp add: pred_numeral_def) lemma eval_nat_numeral: "numeral One = Suc 0" "numeral (Bit0 n) = Suc (numeral (BitM n))" "numeral (Bit1 n) = Suc (numeral (Bit0 n))" by (simp_all add: numeral.simps BitM_plus_one) lemma pred_numeral_simps [simp]: "pred_numeral One = 0" "pred_numeral (Bit0 k) = numeral (BitM k)" "pred_numeral (Bit1 k) = numeral (Bit0 k)" by (simp_all only: pred_numeral_def eval_nat_numeral diff_Suc_Suc diff_0) lemma pred_numeral_inc [simp]: "pred_numeral (Num.inc k) = numeral k" by (simp only: pred_numeral_def numeral_inc diff_add_inverse2) lemma numeral_2_eq_2: "2 = Suc (Suc 0)" by (simp add: eval_nat_numeral) lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" by (simp only: numeral_One One_nat_def) lemma Suc_nat_number_of_add: "Suc (numeral v + n) = numeral (v + One) + n" by simp lemma numerals: "Numeral1 = (1::nat)" "2 = Suc (Suc 0)" by (rule numeral_One) (rule numeral_2_eq_2) lemmas numeral_nat = eval_nat_numeral BitM.simps One_nat_def text \Comparisons involving \<^term>\Suc\.\ lemma eq_numeral_Suc [simp]: "numeral k = Suc n \ pred_numeral k = n" by (simp add: numeral_eq_Suc) lemma Suc_eq_numeral [simp]: "Suc n = numeral k \ n = pred_numeral k" by (simp add: numeral_eq_Suc) lemma less_numeral_Suc [simp]: "numeral k < Suc n \ pred_numeral k < n" by (simp add: numeral_eq_Suc) lemma less_Suc_numeral [simp]: "Suc n < numeral k \ n < pred_numeral k" by (simp add: numeral_eq_Suc) lemma le_numeral_Suc [simp]: "numeral k \ Suc n \ pred_numeral k \ n" by (simp add: numeral_eq_Suc) lemma le_Suc_numeral [simp]: "Suc n \ numeral k \ n \ pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_Suc_numeral [simp]: "Suc n - numeral k = n - pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_numeral_Suc [simp]: "numeral k - Suc n = pred_numeral k - n" by (simp add: numeral_eq_Suc) lemma max_Suc_numeral [simp]: "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma max_numeral_Suc [simp]: "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" by (simp add: numeral_eq_Suc) lemma min_Suc_numeral [simp]: "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma min_numeral_Suc [simp]: "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" by (simp add: numeral_eq_Suc) text \For \<^term>\case_nat\ and \<^term>\rec_nat\.\ lemma case_nat_numeral [simp]: "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" by (simp add: numeral_eq_Suc) lemma case_nat_add_eq_if [simp]: "case_nat a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))" by (simp add: numeral_eq_Suc) lemma rec_nat_numeral [simp]: "rec_nat a f (numeral v) = (let pv = pred_numeral v in f pv (rec_nat a f pv))" by (simp add: numeral_eq_Suc Let_def) lemma rec_nat_add_eq_if [simp]: "rec_nat a f (numeral v + n) = (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" by (simp add: numeral_eq_Suc Let_def) text \Case analysis on \<^term>\n < 2\.\ lemma less_2_cases: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) lemma less_2_cases_iff: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) text \Removal of Small Numerals: 0, 1 and (in additive positions) 2.\ text \bh: Are these rules really a good idea? LCP: well, it already happens for 0 and 1!\ lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" by simp lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" by simp text \Can be used to eliminate long strings of Sucs, but not by default.\ lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" by simp lemmas nat_1_add_1 = one_add_one [where 'a=nat] (* legacy *) context semiring_numeral begin lemma numeral_add_unfold_funpow: \numeral k + a = ((+) 1 ^^ numeral k) a\ proof (rule sym, induction k arbitrary: a) case One then show ?case by (simp add: numeral_One Num.numeral_One) next case (Bit0 k) then show ?case by (simp add: numeral_Bit0 Num.numeral_Bit0 ac_simps funpow_add) next case (Bit1 k) then show ?case by (simp add: numeral_Bit1 Num.numeral_Bit1 ac_simps funpow_add) qed end context semiring_1 begin lemma numeral_unfold_funpow: \numeral k = ((+) 1 ^^ numeral k) 0\ using numeral_add_unfold_funpow [of k 0] by simp end context includes lifting_syntax begin lemma transfer_rule_numeral: \((=) ===> R) numeral numeral\ if [transfer_rule]: \R 0 0\ \R 1 1\ \(R ===> R ===> R) (+) (+)\ for R :: \'a::{semiring_numeral,monoid_add} \ 'b::{semiring_numeral,monoid_add} \ bool\ proof - have "((=) ===> R) (\k. ((+) 1 ^^ numeral k) 0) (\k. ((+) 1 ^^ numeral k) 0)" by transfer_prover moreover have \numeral = (\k. ((+) (1::'a) ^^ numeral k) 0)\ using numeral_add_unfold_funpow [where ?'a = 'a, of _ 0] by (simp add: fun_eq_iff) moreover have \numeral = (\k. ((+) (1::'b) ^^ numeral k) 0)\ using numeral_add_unfold_funpow [where ?'a = 'b, of _ 0] by (simp add: fun_eq_iff) ultimately show ?thesis by simp qed end subsection \Particular lemmas concerning \<^term>\2\\ context linordered_field begin subclass field_char_0 .. lemma half_gt_zero_iff: "0 < a / 2 \ 0 < a" by (auto simp add: field_simps) lemma half_gt_zero [simp]: "0 < a \ 0 < a / 2" by (simp add: half_gt_zero_iff) end subsection \Numeral equations as default simplification rules\ declare (in numeral) numeral_One [simp] declare (in numeral) numeral_plus_numeral [simp] declare (in numeral) add_numeral_special [simp] declare (in neg_numeral) add_neg_numeral_simps [simp] declare (in neg_numeral) add_neg_numeral_special [simp] declare (in neg_numeral) diff_numeral_simps [simp] declare (in neg_numeral) diff_numeral_special [simp] declare (in semiring_numeral) numeral_times_numeral [simp] declare (in ring_1) mult_neg_numeral_simps [simp] subsubsection \Special Simplification for Constants\ text \These distributive laws move literals inside sums and differences.\ lemmas distrib_right_numeral [simp] = distrib_right [of _ _ "numeral v"] for v lemmas distrib_left_numeral [simp] = distrib_left [of "numeral v"] for v lemmas left_diff_distrib_numeral [simp] = left_diff_distrib [of _ _ "numeral v"] for v lemmas right_diff_distrib_numeral [simp] = right_diff_distrib [of "numeral v"] for v text \These are actually for fields, like real\ lemmas zero_less_divide_iff_numeral [simp, no_atp] = zero_less_divide_iff [of "numeral w"] for w lemmas divide_less_0_iff_numeral [simp, no_atp] = divide_less_0_iff [of "numeral w"] for w lemmas zero_le_divide_iff_numeral [simp, no_atp] = zero_le_divide_iff [of "numeral w"] for w lemmas divide_le_0_iff_numeral [simp, no_atp] = divide_le_0_iff [of "numeral w"] for w text \Replaces \inverse #nn\ by \1/#nn\. It looks strange, but then other simprocs simplify the quotient.\ lemmas inverse_eq_divide_numeral [simp] = inverse_eq_divide [of "numeral w"] for w lemmas inverse_eq_divide_neg_numeral [simp] = inverse_eq_divide [of "- numeral w"] for w text \These laws simplify inequalities, moving unary minus from a term into the literal.\ lemmas equation_minus_iff_numeral [no_atp] = equation_minus_iff [of "numeral v"] for v lemmas minus_equation_iff_numeral [no_atp] = minus_equation_iff [of _ "numeral v"] for v lemmas le_minus_iff_numeral [no_atp] = le_minus_iff [of "numeral v"] for v lemmas minus_le_iff_numeral [no_atp] = minus_le_iff [of _ "numeral v"] for v lemmas less_minus_iff_numeral [no_atp] = less_minus_iff [of "numeral v"] for v lemmas minus_less_iff_numeral [no_atp] = minus_less_iff [of _ "numeral v"] for v (* FIXME maybe simproc *) text \Cancellation of constant factors in comparisons (\<\ and \\\)\ lemmas mult_less_cancel_left_numeral [simp, no_atp] = mult_less_cancel_left [of "numeral v"] for v lemmas mult_less_cancel_right_numeral [simp, no_atp] = mult_less_cancel_right [of _ "numeral v"] for v lemmas mult_le_cancel_left_numeral [simp, no_atp] = mult_le_cancel_left [of "numeral v"] for v lemmas mult_le_cancel_right_numeral [simp, no_atp] = mult_le_cancel_right [of _ "numeral v"] for v text \Multiplying out constant divisors in comparisons (\<\, \\\ and \=\)\ named_theorems divide_const_simps "simplification rules to simplify comparisons involving constant divisors" lemmas le_divide_eq_numeral1 [simp,divide_const_simps] = pos_le_divide_eq [of "numeral w", OF zero_less_numeral] neg_le_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_le_eq_numeral1 [simp,divide_const_simps] = pos_divide_le_eq [of "numeral w", OF zero_less_numeral] neg_divide_le_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas less_divide_eq_numeral1 [simp,divide_const_simps] = pos_less_divide_eq [of "numeral w", OF zero_less_numeral] neg_less_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_less_eq_numeral1 [simp,divide_const_simps] = pos_divide_less_eq [of "numeral w", OF zero_less_numeral] neg_divide_less_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas eq_divide_eq_numeral1 [simp,divide_const_simps] = eq_divide_eq [of _ _ "numeral w"] eq_divide_eq [of _ _ "- numeral w"] for w lemmas divide_eq_eq_numeral1 [simp,divide_const_simps] = divide_eq_eq [of _ "numeral w"] divide_eq_eq [of _ "- numeral w"] for w subsubsection \Optional Simplification Rules Involving Constants\ text \Simplify quotients that are compared with a literal constant.\ lemmas le_divide_eq_numeral [divide_const_simps] = le_divide_eq [of "numeral w"] le_divide_eq [of "- numeral w"] for w lemmas divide_le_eq_numeral [divide_const_simps] = divide_le_eq [of _ _ "numeral w"] divide_le_eq [of _ _ "- numeral w"] for w lemmas less_divide_eq_numeral [divide_const_simps] = less_divide_eq [of "numeral w"] less_divide_eq [of "- numeral w"] for w lemmas divide_less_eq_numeral [divide_const_simps] = divide_less_eq [of _ _ "numeral w"] divide_less_eq [of _ _ "- numeral w"] for w lemmas eq_divide_eq_numeral [divide_const_simps] = eq_divide_eq [of "numeral w"] eq_divide_eq [of "- numeral w"] for w lemmas divide_eq_eq_numeral [divide_const_simps] = divide_eq_eq [of _ _ "numeral w"] divide_eq_eq [of _ _ "- numeral w"] for w text \Not good as automatic simprules because they cause case splits.\ lemmas [divide_const_simps] = le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1 subsection \Setting up simprocs\ lemma mult_numeral_1: "Numeral1 * a = a" for a :: "'a::semiring_numeral" by simp lemma mult_numeral_1_right: "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp lemma divide_numeral_1: "a / Numeral1 = a" for a :: "'a::field" by simp lemma inverse_numeral_1: "inverse Numeral1 = (Numeral1::'a::division_ring)" by simp text \ Theorem lists for the cancellation simprocs. The use of a binary numeral for 1 reduces the number of special cases. \ lemma mult_1s_semiring_numeral: "Numeral1 * a = a" "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp_all lemma mult_1s_ring_1: "- Numeral1 * b = - b" "b * - Numeral1 = - b" for b :: "'a::ring_1" by simp_all lemmas mult_1s = mult_1s_semiring_numeral mult_1s_ring_1 setup \ Reorient_Proc.add (fn Const (\<^const_name>\numeral\, _) $ _ => true | Const (\<^const_name>\uminus\, _) $ (Const (\<^const_name>\numeral\, _) $ _) => true | _ => false) \ simproc_setup reorient_numeral ("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc subsubsection \Simplification of arithmetic operations on integer constants\ lemmas arith_special = (* already declared simp above *) add_numeral_special add_neg_numeral_special diff_numeral_special lemmas arith_extra_simps = (* rules already in simpset *) numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right minus_zero diff_numeral_simps diff_0 diff_0_right numeral_times_numeral mult_neg_numeral_simps mult_zero_left mult_zero_right abs_numeral abs_neg_numeral text \ For making a minimal simpset, one must include these default simprules. Also include \simp_thms\. \ lemmas arith_simps = add_num_simps mult_num_simps sub_num_simps BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps abs_zero abs_one arith_extra_simps lemmas more_arith_simps = neg_le_iff_le minus_zero left_minus right_minus mult_1_left mult_1_right mult_minus_left mult_minus_right minus_add_distrib minus_minus mult.assoc lemmas of_nat_simps = of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult text \Simplification of relational operations.\ lemmas eq_numeral_extra = zero_neq_one one_neq_zero lemmas rel_simps = le_num_simps less_num_simps eq_num_simps le_numeral_simps le_neg_numeral_simps le_minus_one_simps le_numeral_extra less_numeral_simps less_neg_numeral_simps less_minus_one_simps less_numeral_extra eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. declaration \ let fun number_of ctxt T n = if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\numeral\)) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; in K ( Lin_Arith.set_number_of number_of #> Lin_Arith.add_simps @{thms arith_simps more_arith_simps rel_simps pred_numeral_simps arith_special numeral_One of_nat_simps uminus_numeral_One Suc_numeral Let_numeral Let_neg_numeral Let_0 Let_1 le_Suc_numeral le_numeral_Suc less_Suc_numeral less_numeral_Suc Suc_eq_numeral eq_numeral_Suc mult_Suc mult_Suc_right of_nat_numeral}) end \ subsubsection \Simplification of arithmetic when nested to the right\ lemma add_numeral_left [simp]: "numeral v + (numeral w + z) = (numeral(v + w) + z)" by (simp_all add: add.assoc [symmetric]) lemma add_neg_numeral_left [simp]: "numeral v + (- numeral w + y) = (sub v w + y)" "- numeral v + (numeral w + y) = (sub w v + y)" "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)" by (simp_all add: add.assoc [symmetric]) lemma mult_numeral_left_semiring_numeral: "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" by (simp add: mult.assoc [symmetric]) lemma mult_numeral_left_ring_1: "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'a::ring_1)" by (simp_all add: mult.assoc [symmetric]) lemmas mult_numeral_left [simp] = mult_numeral_left_semiring_numeral mult_numeral_left_ring_1 hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec subsection \Code module namespace\ code_identifier code_module Num \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Printing of evaluated natural numbers as numerals\ lemma [code_post]: "Suc 0 = 1" "Suc 1 = 2" "Suc (numeral n) = numeral (Num.inc n)" by (simp_all add: numeral_inc) lemmas [code_post] = Num.inc.simps subsection \More on auxiliary conversion\ context semiring_1 begin lemma numeral_num_of_nat_unfold: \numeral (num_of_nat n) = (if n = 0 then 1 else of_nat n)\ by (induction n) (simp_all add: numeral_inc ac_simps) lemma num_of_nat_numeral_eq [simp]: \num_of_nat (numeral q) = q\ proof (induction q) case One then show ?case by simp next case (Bit0 q) - then show ?case - apply (simp only: Num.numeral_Bit0 Num.numeral_add) - apply (subst num_of_nat_double) - apply simp_all - done + then have "num_of_nat (numeral (num.Bit0 q)) = num_of_nat (numeral q + numeral q)" + by (simp only: Num.numeral_Bit0 Num.numeral_add) + also have "\ = num.Bit0 (num_of_nat (numeral q))" + by (rule num_of_nat_double) simp + finally show ?case + using Bit0.IH by simp next case (Bit1 q) - then show ?case - apply (simp only: Num.numeral_Bit1 Num.numeral_add) - apply (subst num_of_nat_plus_distrib) - apply simp - apply simp - apply (subst num_of_nat_double) - apply simp_all - done + then have "num_of_nat (numeral (num.Bit1 q)) = num_of_nat (numeral q + numeral q + 1)" + by (simp only: Num.numeral_Bit1 Num.numeral_add) + also have "\ = num_of_nat (numeral q + numeral q) + num_of_nat 1" + by (rule num_of_nat_plus_distrib) auto + also have "\ = num.Bit0 (num_of_nat (numeral q)) + num_of_nat 1" + by (subst num_of_nat_double) auto + finally show ?case + using Bit1.IH by simp qed end end diff --git a/src/HOL/Order_Relation.thy b/src/HOL/Order_Relation.thy --- a/src/HOL/Order_Relation.thy +++ b/src/HOL/Order_Relation.thy @@ -1,585 +1,613 @@ (* Title: HOL/Order_Relation.thy Author: Tobias Nipkow Author: Andrei Popescu, TU Muenchen *) section \Orders as Relations\ theory Order_Relation imports Wfrec begin subsection \Orders on a set\ definition "preorder_on A r \ refl_on A r \ trans r" definition "partial_order_on A r \ preorder_on A r \ antisym r" definition "linear_order_on A r \ partial_order_on A r \ total_on A r" definition "strict_linear_order_on A r \ trans r \ irrefl r \ total_on A r" definition "well_order_on A r \ linear_order_on A r \ wf(r - Id)" lemmas order_on_defs = preorder_on_def partial_order_on_def linear_order_on_def strict_linear_order_on_def well_order_on_def lemma partial_order_onD: assumes "partial_order_on A r" shows "refl_on A r" and "trans r" and "antisym r" using assms unfolding partial_order_on_def preorder_on_def by auto lemma preorder_on_empty[simp]: "preorder_on {} {}" by (simp add: preorder_on_def trans_def) lemma partial_order_on_empty[simp]: "partial_order_on {} {}" by (simp add: partial_order_on_def) lemma lnear_order_on_empty[simp]: "linear_order_on {} {}" by (simp add: linear_order_on_def) lemma well_order_on_empty[simp]: "well_order_on {} {}" by (simp add: well_order_on_def) lemma preorder_on_converse[simp]: "preorder_on A (r\) = preorder_on A r" by (simp add: preorder_on_def) lemma partial_order_on_converse[simp]: "partial_order_on A (r\) = partial_order_on A r" by (simp add: partial_order_on_def) lemma linear_order_on_converse[simp]: "linear_order_on A (r\) = linear_order_on A r" by (simp add: linear_order_on_def) lemma partial_order_on_acyclic: "partial_order_on A r \ acyclic (r - Id)" by (simp add: acyclic_irrefl partial_order_on_def preorder_on_def trans_diff_Id) lemma partial_order_on_well_order_on: "finite r \ partial_order_on A r \ wf (r - Id)" by (simp add: finite_acyclic_wf partial_order_on_acyclic) lemma strict_linear_order_on_diff_Id: "linear_order_on A r \ strict_linear_order_on A (r - Id)" by (simp add: order_on_defs trans_diff_Id) lemma linear_order_on_singleton [simp]: "linear_order_on {x} {(x, x)}" by (simp add: order_on_defs) lemma linear_order_on_acyclic: assumes "linear_order_on A r" shows "acyclic (r - Id)" using strict_linear_order_on_diff_Id[OF assms] by (auto simp add: acyclic_irrefl strict_linear_order_on_def) lemma linear_order_on_well_order_on: assumes "finite r" shows "linear_order_on A r \ well_order_on A r" unfolding well_order_on_def using assms finite_acyclic_wf[OF _ linear_order_on_acyclic, of r] by blast subsection \Orders on the field\ abbreviation "Refl r \ refl_on (Field r) r" abbreviation "Preorder r \ preorder_on (Field r) r" abbreviation "Partial_order r \ partial_order_on (Field r) r" abbreviation "Total r \ total_on (Field r) r" abbreviation "Linear_order r \ linear_order_on (Field r) r" abbreviation "Well_order r \ well_order_on (Field r) r" lemma subset_Image_Image_iff: "Preorder r \ A \ Field r \ B \ Field r \ r `` A \ r `` B \ (\a\A.\b\B. (b, a) \ r)" apply (simp add: preorder_on_def refl_on_def Image_def subset_eq) apply (simp only: trans_def) apply fast done lemma subset_Image1_Image1_iff: "Preorder r \ a \ Field r \ b \ Field r \ r `` {a} \ r `` {b} \ (b, a) \ r" by (simp add: subset_Image_Image_iff) lemma Refl_antisym_eq_Image1_Image1_iff: assumes "Refl r" and as: "antisym r" and abf: "a \ Field r" "b \ Field r" shows "r `` {a} = r `` {b} \ a = b" (is "?lhs \ ?rhs") proof assume ?lhs then have *: "\x. (a, x) \ r \ (b, x) \ r" by (simp add: set_eq_iff) have "(a, a) \ r" "(b, b) \ r" using \Refl r\ abf by (simp_all add: refl_on_def) then have "(a, b) \ r" "(b, a) \ r" using *[of a] *[of b] by simp_all then show ?rhs using \antisym r\[unfolded antisym_def] by blast next assume ?rhs then show ?lhs by fast qed lemma Partial_order_eq_Image1_Image1_iff: "Partial_order r \ a \ Field r \ b \ Field r \ r `` {a} = r `` {b} \ a = b" by (auto simp: order_on_defs Refl_antisym_eq_Image1_Image1_iff) lemma Total_Id_Field: assumes "Total r" and not_Id: "\ r \ Id" shows "Field r = Field (r - Id)" - using mono_Field[of "r - Id" r] Diff_subset[of r Id] -proof auto - fix a assume *: "a \ Field r" - from not_Id have "r \ {}" by fast - with not_Id obtain b and c where "b \ c \ (b,c) \ r" by auto - then have "b \ c \ {b, c} \ Field r" by (auto simp: Field_def) - with * obtain d where "d \ Field r" "d \ a" by auto - with * \Total r\ have "(a, d) \ r \ (d, a) \ r" by (simp add: total_on_def) - with \d \ a\ show "a \ Field (r - Id)" unfolding Field_def by blast +proof - + have "Field r \ Field (r - Id)" + proof (rule subsetI) + fix a assume *: "a \ Field r" + from not_Id have "r \ {}" by fast + with not_Id obtain b and c where "b \ c \ (b,c) \ r" by auto + then have "b \ c \ {b, c} \ Field r" by (auto simp: Field_def) + with * obtain d where "d \ Field r" "d \ a" by auto + with * \Total r\ have "(a, d) \ r \ (d, a) \ r" by (simp add: total_on_def) + with \d \ a\ show "a \ Field (r - Id)" unfolding Field_def by blast + qed + then show ?thesis + using mono_Field[of "r - Id" r] Diff_subset[of r Id] by auto qed subsection\Relations given by a predicate and the field\ definition relation_of :: "('a \ 'a \ bool) \ 'a set \ ('a \ 'a) set" where "relation_of P A \ { (a, b) \ A \ A. P a b }" lemma Field_relation_of: assumes "refl_on A (relation_of P A)" shows "Field (relation_of P A) = A" using assms unfolding refl_on_def Field_def by auto lemma partial_order_on_relation_ofI: assumes refl: "\a. a \ A \ P a a" and trans: "\a b c. \ a \ A; b \ A; c \ A \ \ P a b \ P b c \ P a c" and antisym: "\a b. \ a \ A; b \ A \ \ P a b \ P b a \ a = b" shows "partial_order_on A (relation_of P A)" proof - from refl have "refl_on A (relation_of P A)" unfolding refl_on_def relation_of_def by auto moreover have "trans (relation_of P A)" and "antisym (relation_of P A)" unfolding relation_of_def by (auto intro: transI dest: trans, auto intro: antisymI dest: antisym) ultimately show ?thesis unfolding partial_order_on_def preorder_on_def by simp qed lemma Partial_order_relation_ofI: assumes "partial_order_on A (relation_of P A)" shows "Partial_order (relation_of P A)" using Field_relation_of assms partial_order_on_def preorder_on_def by fastforce subsection \Orders on a type\ abbreviation "strict_linear_order \ strict_linear_order_on UNIV" abbreviation "linear_order \ linear_order_on UNIV" abbreviation "well_order \ well_order_on UNIV" subsection \Order-like relations\ text \ In this subsection, we develop basic concepts and results pertaining to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or total relations. We also further define upper and lower bounds operators. \ subsubsection \Auxiliaries\ lemma refl_on_domain: "refl_on A r \ (a, b) \ r \ a \ A \ b \ A" by (auto simp add: refl_on_def) corollary well_order_on_domain: "well_order_on A r \ (a, b) \ r \ a \ A \ b \ A" by (auto simp add: refl_on_domain order_on_defs) lemma well_order_on_Field: "well_order_on A r \ A = Field r" by (auto simp add: refl_on_def Field_def order_on_defs) lemma well_order_on_Well_order: "well_order_on A r \ A = Field r \ Well_order r" using well_order_on_Field [of A] by auto lemma Total_subset_Id: assumes "Total r" and "r \ Id" shows "r = {} \ (\a. r = {(a, a)})" proof - have "\a. r = {(a, a)}" if "r \ {}" proof - from that obtain a b where ab: "(a, b) \ r" by fast with \r \ Id\ have "a = b" by blast with ab have aa: "(a, a) \ r" by simp have "a = c \ a = d" if "(c, d) \ r" for c d proof - from that have "{a, c, d} \ Field r" using ab unfolding Field_def by blast then have "((a, c) \ r \ (c, a) \ r \ a = c) \ ((a, d) \ r \ (d, a) \ r \ a = d)" using \Total r\ unfolding total_on_def by blast with \r \ Id\ show ?thesis by blast qed then have "r \ {(a, a)}" by auto with aa show ?thesis by blast qed then show ?thesis by blast qed lemma Linear_order_in_diff_Id: assumes "Linear_order r" and "a \ Field r" and "b \ Field r" shows "(a, b) \ r \ (b, a) \ r - Id" using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force subsubsection \The upper and lower bounds operators\ text \ Here we define upper (``above") and lower (``below") bounds operators. We think of \r\ as a \<^emph>\non-strict\ relation. The suffix \S\ at the names of some operators indicates that the bounds are strict -- e.g., \underS a\ is the set of all strict lower bounds of \a\ (w.r.t. \r\). Capitalization of the first letter in the name reminds that the operator acts on sets, rather than on individual elements. \ definition under :: "'a rel \ 'a \ 'a set" where "under r a \ {b. (b, a) \ r}" definition underS :: "'a rel \ 'a \ 'a set" where "underS r a \ {b. b \ a \ (b, a) \ r}" definition Under :: "'a rel \ 'a set \ 'a set" where "Under r A \ {b \ Field r. \a \ A. (b, a) \ r}" definition UnderS :: "'a rel \ 'a set \ 'a set" where "UnderS r A \ {b \ Field r. \a \ A. b \ a \ (b, a) \ r}" definition above :: "'a rel \ 'a \ 'a set" where "above r a \ {b. (a, b) \ r}" definition aboveS :: "'a rel \ 'a \ 'a set" where "aboveS r a \ {b. b \ a \ (a, b) \ r}" definition Above :: "'a rel \ 'a set \ 'a set" where "Above r A \ {b \ Field r. \a \ A. (a, b) \ r}" definition AboveS :: "'a rel \ 'a set \ 'a set" where "AboveS r A \ {b \ Field r. \a \ A. b \ a \ (a, b) \ r}" definition ofilter :: "'a rel \ 'a set \ bool" where "ofilter r A \ A \ Field r \ (\a \ A. under r a \ A)" text \ Note: In the definitions of \Above[S]\ and \Under[S]\, we bounded comprehension by \Field r\ in order to properly cover the case of \A\ being empty. \ lemma underS_subset_under: "underS r a \ under r a" by (auto simp add: underS_def under_def) lemma underS_notIn: "a \ underS r a" by (simp add: underS_def) lemma Refl_under_in: "Refl r \ a \ Field r \ a \ under r a" by (simp add: refl_on_def under_def) lemma AboveS_disjoint: "A \ (AboveS r A) = {}" by (auto simp add: AboveS_def) lemma in_AboveS_underS: "a \ Field r \ a \ AboveS r (underS r a)" by (auto simp add: AboveS_def underS_def) lemma Refl_under_underS: "Refl r \ a \ Field r \ under r a = underS r a \ {a}" unfolding under_def underS_def using refl_on_def[of _ r] by fastforce lemma underS_empty: "a \ Field r \ underS r a = {}" by (auto simp: Field_def underS_def) lemma under_Field: "under r a \ Field r" by (auto simp: under_def Field_def) lemma underS_Field: "underS r a \ Field r" by (auto simp: underS_def Field_def) lemma underS_Field2: "a \ Field r \ underS r a \ Field r" using underS_notIn underS_Field by fast lemma underS_Field3: "Field r \ {} \ underS r a \ Field r" by (cases "a \ Field r") (auto simp: underS_Field2 underS_empty) lemma AboveS_Field: "AboveS r A \ Field r" by (auto simp: AboveS_def Field_def) lemma under_incr: assumes "trans r" and "(a, b) \ r" shows "under r a \ under r b" unfolding under_def -proof auto +proof safe fix x assume "(x, a) \ r" with assms trans_def[of r] show "(x, b) \ r" by blast qed lemma underS_incr: assumes "trans r" and "antisym r" and ab: "(a, b) \ r" shows "underS r a \ underS r b" unfolding underS_def -proof auto +proof safe assume *: "b \ a" and **: "(b, a) \ r" with \antisym r\ antisym_def[of r] ab show False by blast next fix x assume "x \ a" "(x, a) \ r" with ab \trans r\ trans_def[of r] show "(x, b) \ r" by blast qed lemma underS_incl_iff: assumes LO: "Linear_order r" and INa: "a \ Field r" and INb: "b \ Field r" shows "underS r a \ underS r b \ (a, b) \ r" (is "?lhs \ ?rhs") proof assume ?rhs with \Linear_order r\ show ?lhs by (simp add: order_on_defs underS_incr) next assume *: ?lhs have "(a, b) \ r" if "a = b" using assms that by (simp add: order_on_defs refl_on_def) moreover have False if "a \ b" "(b, a) \ r" proof - from that have "b \ underS r a" unfolding underS_def by blast with * have "b \ underS r b" by blast then show ?thesis by (simp add: underS_notIn) qed ultimately show "(a,b) \ r" using assms order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast qed lemma finite_Partial_order_induct[consumes 3, case_names step]: assumes "Partial_order r" and "x \ Field r" and "finite r" and step: "\x. x \ Field r \ (\y. y \ aboveS r x \ P y) \ P x" shows "P x" using assms(2) proof (induct rule: wf_induct[of "r\ - Id"]) case 1 from assms(1,3) show "wf (r\ - Id)" using partial_order_on_well_order_on partial_order_on_converse by blast next case prems: (2 x) show ?case by (rule step) (use prems in \auto simp: aboveS_def intro: FieldI2\) qed lemma finite_Linear_order_induct[consumes 3, case_names step]: assumes "Linear_order r" and "x \ Field r" and "finite r" and step: "\x. x \ Field r \ (\y. y \ aboveS r x \ P y) \ P x" shows "P x" using assms(2) proof (induct rule: wf_induct[of "r\ - Id"]) case 1 from assms(1,3) show "wf (r\ - Id)" using linear_order_on_well_order_on linear_order_on_converse unfolding well_order_on_def by blast next case prems: (2 x) show ?case by (rule step) (use prems in \auto simp: aboveS_def intro: FieldI2\) qed subsection \Variations on Well-Founded Relations\ text \ This subsection contains some variations of the results from \<^theory>\HOL.Wellfounded\: \<^item> means for slightly more direct definitions by well-founded recursion; \<^item> variations of well-founded induction; \<^item> means for proving a linear order to be a well-order. \ subsubsection \Characterizations of well-foundedness\ text \ A transitive relation is well-founded iff it is ``locally'' well-founded, i.e., iff its restriction to the lower bounds of of any element is well-founded. \ lemma trans_wf_iff: assumes "trans r" shows "wf r \ (\a. wf (r \ (r\``{a} \ r\``{a})))" proof - define R where "R a = r \ (r\``{a} \ r\``{a})" for a have "wf (R a)" if "wf r" for a using that R_def wf_subset[of r "R a"] by auto moreover have "wf r" if *: "\a. wf(R a)" unfolding wf_def proof clarify fix phi a assume **: "\a. (\b. (b, a) \ r \ phi b) \ phi a" define chi where "chi b \ (b, a) \ r \ phi b" for b with * have "wf (R a)" by auto then have "(\b. (\c. (c, b) \ R a \ chi c) \ chi b) \ (\b. chi b)" unfolding wf_def by blast also have "\b. (\c. (c, b) \ R a \ chi c) \ chi b" - proof (auto simp add: chi_def R_def) + proof safe fix b - assume "(b, a) \ r" and "\c. (c, b) \ r \ (c, a) \ r \ phi c" - then have "\c. (c, b) \ r \ phi c" - using assms trans_def[of r] by blast - with ** show "phi b" by blast + assume "\c. (c, b) \ R a \ chi c" + moreover have "(b, a) \ r \ \c. (c, b) \ r \ (c, a) \ r \ phi c \ phi b" + proof - + assume "(b, a) \ r" and "\c. (c, b) \ r \ (c, a) \ r \ phi c" + then have "\c. (c, b) \ r \ phi c" + using assms trans_def[of r] by blast + with ** show "phi b" by blast + qed + ultimately show "chi b" + by (auto simp add: chi_def R_def) qed finally have "\b. chi b" . with ** chi_def show "phi a" by blast qed ultimately show ?thesis unfolding R_def by blast qed text\A transitive relation is well-founded if all initial segments are finite.\ corollary wf_finite_segments: assumes "irrefl r" and "trans r" and "\x. finite {y. (y, x) \ r}" - shows "wf (r)" -proof (clarsimp simp: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms) - fix a - have "trans (r \ ({x. (x, a) \ r} \ {x. (x, a) \ r}))" - using assms unfolding trans_def Field_def by blast - then show "acyclic (r \ {x. (x, a) \ r} \ {x. (x, a) \ r})" - using assms acyclic_def assms irrefl_def by fastforce + shows "wf r" +proof - + have "\a. acyclic (r \ {x. (x, a) \ r} \ {x. (x, a) \ r})" + proof - + fix a + have "trans (r \ ({x. (x, a) \ r} \ {x. (x, a) \ r}))" + using assms unfolding trans_def Field_def by blast + then show "acyclic (r \ {x. (x, a) \ r} \ {x. (x, a) \ r})" + using assms acyclic_def assms irrefl_def by fastforce + qed + then show ?thesis + by (clarsimp simp: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms) qed text \The next lemma is a variation of \wf_eq_minimal\ from Wellfounded, allowing one to assume the set included in the field.\ lemma wf_eq_minimal2: "wf r \ (\A. A \ Field r \ A \ {} \ (\a \ A. \a' \ A. (a', a) \ r))" proof- let ?phi = "\A. A \ {} \ (\a \ A. \a' \ A. (a',a) \ r)" have "wf r \ (\A. ?phi A)" - apply (auto simp: ex_in_conv [THEN sym]) - apply (erule wfE_min) - apply assumption - apply blast - apply (rule wfI_min) - apply fast - done + proof + assume "wf r" + show "\A. ?phi A" + proof clarify + fix A:: "'a set" + assume "A \ {}" + then obtain x where "x \ A" + by auto + show "\a\A. \a'\A. (a', a) \ r" + apply (rule wfE_min[of r x A]) + apply fact+ + by blast + qed + next + assume *: "\A. ?phi A" + then show "wf r" + apply (clarsimp simp: ex_in_conv [THEN sym]) + apply (rule wfI_min) + by fast + qed also have "(\A. ?phi A) \ (\B \ Field r. ?phi B)" proof assume "\A. ?phi A" then show "\B \ Field r. ?phi B" by simp next assume *: "\B \ Field r. ?phi B" show "\A. ?phi A" proof clarify fix A :: "'a set" assume **: "A \ {}" define B where "B = A \ Field r" show "\a \ A. \a' \ A. (a', a) \ r" proof (cases "B = {}") case True with ** obtain a where a: "a \ A" "a \ Field r" unfolding B_def by blast with a have "\a' \ A. (a',a) \ r" unfolding Field_def by blast with a show ?thesis by blast next case False have "B \ Field r" unfolding B_def by blast with False * obtain a where a: "a \ B" "\a' \ B. (a', a) \ r" by blast have "(a', a) \ r" if "a' \ A" for a' proof assume a'a: "(a', a) \ r" with that have "a' \ B" unfolding B_def Field_def by blast with a a'a show False by blast qed with a show ?thesis unfolding B_def by blast qed qed qed finally show ?thesis by blast qed subsubsection \Characterizations of well-foundedness\ text \ The next lemma and its corollary enable one to prove that a linear order is a well-order in a way which is more standard than via well-foundedness of the strict version of the relation. \ lemma Linear_order_wf_diff_Id: assumes "Linear_order r" shows "wf (r - Id) \ (\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a, a') \ r))" proof (cases "r \ Id") case True then have *: "r - Id = {}" by blast have "wf (r - Id)" by (simp add: *) moreover have "\a \ A. \a' \ A. (a, a') \ r" if *: "A \ Field r" and **: "A \ {}" for A proof - from \Linear_order r\ True obtain a where a: "r = {} \ r = {(a, a)}" unfolding order_on_defs using Total_subset_Id [of r] by blast with * ** have "A = {a} \ r = {(a, a)}" unfolding Field_def by blast with a show ?thesis by blast qed ultimately show ?thesis by blast next case False with \Linear_order r\ have Field: "Field r = Field (r - Id)" unfolding order_on_defs using Total_Id_Field [of r] by blast show ?thesis proof assume *: "wf (r - Id)" show "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a, a') \ r)" proof clarify fix A assume **: "A \ Field r" and ***: "A \ {}" then have "\a \ A. \a' \ A. (a',a) \ r - Id" using Field * unfolding wf_eq_minimal2 by simp moreover have "\a \ A. \a' \ A. (a, a') \ r \ (a', a) \ r - Id" using Linear_order_in_diff_Id [OF \Linear_order r\] ** by blast ultimately show "\a \ A. \a' \ A. (a, a') \ r" by blast qed next assume *: "\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a, a') \ r)" show "wf (r - Id)" unfolding wf_eq_minimal2 proof clarify fix A assume **: "A \ Field(r - Id)" and ***: "A \ {}" then have "\a \ A. \a' \ A. (a,a') \ r" using Field * by simp moreover have "\a \ A. \a' \ A. (a, a') \ r \ (a', a) \ r - Id" using Linear_order_in_diff_Id [OF \Linear_order r\] ** mono_Field[of "r - Id" r] by blast ultimately show "\a \ A. \a' \ A. (a',a) \ r - Id" by blast qed qed qed corollary Linear_order_Well_order_iff: "Linear_order r \ Well_order r \ (\A \ Field r. A \ {} \ (\a \ A. \a' \ A. (a, a') \ r))" unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast end diff --git a/src/HOL/Orderings.thy b/src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy +++ b/src/HOL/Orderings.thy @@ -1,1712 +1,1736 @@ (* Title: HOL/Orderings.thy Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson *) section \Abstract orderings\ theory Orderings imports HOL keywords "print_orders" :: diag begin ML_file \~~/src/Provers/order_procedure.ML\ ML_file \~~/src/Provers/order_tac.ML\ subsection \Abstract ordering\ locale partial_preordering = fixes less_eq :: \'a \ 'a \ bool\ (infix \\<^bold>\\ 50) assumes refl: \a \<^bold>\ a\ \ \not \iff\: makes problems due to multiple (dual) interpretations\ and trans: \a \<^bold>\ b \ b \<^bold>\ c \ a \<^bold>\ c\ locale preordering = partial_preordering + fixes less :: \'a \ 'a \ bool\ (infix \\<^bold><\ 50) assumes strict_iff_not: \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ begin lemma strict_implies_order: \a \<^bold>< b \ a \<^bold>\ b\ by (simp add: strict_iff_not) lemma irrefl: \ \not \iff\: makes problems due to multiple (dual) interpretations\ \\ a \<^bold>< a\ by (simp add: strict_iff_not) lemma asym: \a \<^bold>< b \ b \<^bold>< a \ False\ by (auto simp add: strict_iff_not) lemma strict_trans1: \a \<^bold>\ b \ b \<^bold>< c \ a \<^bold>< c\ by (auto simp add: strict_iff_not intro: trans) lemma strict_trans2: \a \<^bold>< b \ b \<^bold>\ c \ a \<^bold>< c\ by (auto simp add: strict_iff_not intro: trans) lemma strict_trans: \a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ by (auto intro: strict_trans1 strict_implies_order) end lemma preordering_strictI: \ \Alternative introduction rule with bias towards strict order\ fixes less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) assumes less_eq_less: \\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b\ assumes asym: \\a b. a \<^bold>< b \ \ b \<^bold>< a\ assumes irrefl: \\a. \ a \<^bold>< a\ assumes trans: \\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ shows \preordering (\<^bold>\) (\<^bold><)\ proof fix a b show \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ by (auto simp add: less_eq_less asym irrefl) next fix a show \a \<^bold>\ a\ by (auto simp add: less_eq_less) next fix a b c assume \a \<^bold>\ b\ and \b \<^bold>\ c\ then show \a \<^bold>\ c\ by (auto simp add: less_eq_less intro: trans) qed lemma preordering_dualI: fixes less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) assumes \preordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)\ shows \preordering (\<^bold>\) (\<^bold><)\ proof - from assms interpret preordering \\a b. b \<^bold>\ a\ \\a b. b \<^bold>< a\ . show ?thesis by standard (auto simp: strict_iff_not refl intro: trans) qed locale ordering = partial_preordering + fixes less :: \'a \ 'a \ bool\ (infix \\<^bold><\ 50) assumes strict_iff_order: \a \<^bold>< b \ a \<^bold>\ b \ a \ b\ assumes antisym: \a \<^bold>\ b \ b \<^bold>\ a \ a = b\ begin sublocale preordering \(\<^bold>\)\ \(\<^bold><)\ proof show \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ for a b by (auto simp add: strict_iff_order intro: antisym) qed lemma strict_implies_not_eq: \a \<^bold>< b \ a \ b\ by (simp add: strict_iff_order) lemma not_eq_order_implies_strict: \a \ b \ a \<^bold>\ b \ a \<^bold>< b\ by (simp add: strict_iff_order) lemma order_iff_strict: \a \<^bold>\ b \ a \<^bold>< b \ a = b\ by (auto simp add: strict_iff_order refl) lemma eq_iff: \a = b \ a \<^bold>\ b \ b \<^bold>\ a\ by (auto simp add: refl intro: antisym) end lemma ordering_strictI: \ \Alternative introduction rule with bias towards strict order\ fixes less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) assumes less_eq_less: \\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b\ assumes asym: \\a b. a \<^bold>< b \ \ b \<^bold>< a\ assumes irrefl: \\a. \ a \<^bold>< a\ assumes trans: \\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ shows \ordering (\<^bold>\) (\<^bold><)\ proof fix a b show \a \<^bold>< b \ a \<^bold>\ b \ a \ b\ by (auto simp add: less_eq_less asym irrefl) next fix a show \a \<^bold>\ a\ by (auto simp add: less_eq_less) next fix a b c assume \a \<^bold>\ b\ and \b \<^bold>\ c\ then show \a \<^bold>\ c\ by (auto simp add: less_eq_less intro: trans) next fix a b assume \a \<^bold>\ b\ and \b \<^bold>\ a\ then show \a = b\ by (auto simp add: less_eq_less asym) qed lemma ordering_dualI: fixes less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) assumes \ordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)\ shows \ordering (\<^bold>\) (\<^bold><)\ proof - from assms interpret ordering \\a b. b \<^bold>\ a\ \\a b. b \<^bold>< a\ . show ?thesis by standard (auto simp: strict_iff_order refl intro: antisym trans) qed locale ordering_top = ordering + fixes top :: \'a\ (\\<^bold>\\) assumes extremum [simp]: \a \<^bold>\ \<^bold>\\ begin lemma extremum_uniqueI: \\<^bold>\ \<^bold>\ a \ a = \<^bold>\\ by (rule antisym) auto lemma extremum_unique: \\<^bold>\ \<^bold>\ a \ a = \<^bold>\\ by (auto intro: antisym) lemma extremum_strict [simp]: \\ (\<^bold>\ \<^bold>< a)\ using extremum [of a] by (auto simp add: order_iff_strict intro: asym irrefl) lemma not_eq_extremum: \a \ \<^bold>\ \ a \<^bold>< \<^bold>\\ by (auto simp add: order_iff_strict intro: not_eq_order_implies_strict extremum) end subsection \Syntactic orders\ class ord = fixes less_eq :: "'a \ 'a \ bool" and less :: "'a \ 'a \ bool" begin notation less_eq ("'(\')") and less_eq ("(_/ \ _)" [51, 51] 50) and less ("'(<')") and less ("(_/ < _)" [51, 51] 50) abbreviation (input) greater_eq (infix "\" 50) where "x \ y \ y \ x" abbreviation (input) greater (infix ">" 50) where "x > y \ y < x" notation (ASCII) less_eq ("'(<=')") and less_eq ("(_/ <= _)" [51, 51] 50) notation (input) greater_eq (infix ">=" 50) end subsection \Quasi orders\ class preorder = ord + assumes less_le_not_le: "x < y \ x \ y \ \ (y \ x)" and order_refl [iff]: "x \ x" and order_trans: "x \ y \ y \ z \ x \ z" begin sublocale order: preordering less_eq less + dual_order: preordering greater_eq greater proof - interpret preordering less_eq less by standard (auto intro: order_trans simp add: less_le_not_le) show \preordering less_eq less\ by (fact preordering_axioms) then show \preordering greater_eq greater\ by (rule preordering_dualI) qed text \Reflexivity.\ lemma eq_refl: "x = y \ x \ y" \ \This form is useful with the classical reasoner.\ by (erule ssubst) (rule order_refl) lemma less_irrefl [iff]: "\ x < x" by (simp add: less_le_not_le) lemma less_imp_le: "x < y \ x \ y" by (simp add: less_le_not_le) text \Asymmetry.\ lemma less_not_sym: "x < y \ \ (y < x)" by (simp add: less_le_not_le) lemma less_asym: "x < y \ (\ P \ y < x) \ P" by (drule less_not_sym, erule contrapos_np) simp text \Transitivity.\ lemma less_trans: "x < y \ y < z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) lemma le_less_trans: "x \ y \ y < z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) lemma less_le_trans: "x < y \ y \ z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) text \Useful for simplification, but too risky to include by default.\ lemma less_imp_not_less: "x < y \ (\ y < x) \ True" by (blast elim: less_asym) lemma less_imp_triv: "x < y \ (y < x \ P) \ True" by (blast elim: less_asym) text \Transitivity rules for calculational reasoning\ lemma less_asym': "a < b \ b < a \ P" by (rule less_asym) text \Dual order\ lemma dual_preorder: \class.preorder (\) (>)\ by standard (auto simp add: less_le_not_le intro: order_trans) end lemma preordering_preorderI: \class.preorder (\<^bold>\) (\<^bold><)\ if \preordering (\<^bold>\) (\<^bold><)\ for less_eq (infix \\<^bold>\\ 50) and less (infix \\<^bold><\ 50) proof - from that interpret preordering \(\<^bold>\)\ \(\<^bold><)\ . show ?thesis by standard (auto simp add: strict_iff_not refl intro: trans) qed subsection \Partial orders\ class order = preorder + assumes order_antisym: "x \ y \ y \ x \ x = y" begin lemma less_le: "x < y \ x \ y \ x \ y" by (auto simp add: less_le_not_le intro: order_antisym) sublocale order: ordering less_eq less + dual_order: ordering greater_eq greater proof - interpret ordering less_eq less by standard (auto intro: order_antisym order_trans simp add: less_le) show "ordering less_eq less" by (fact ordering_axioms) then show "ordering greater_eq greater" by (rule ordering_dualI) qed print_theorems text \Reflexivity.\ lemma le_less: "x \ y \ x < y \ x = y" \ \NOT suitable for iff, since it can cause PROOF FAILED.\ by (fact order.order_iff_strict) lemma le_imp_less_or_eq: "x \ y \ x < y \ x = y" by (simp add: less_le) text \Useful for simplification, but too risky to include by default.\ lemma less_imp_not_eq: "x < y \ (x = y) \ False" by auto lemma less_imp_not_eq2: "x < y \ (y = x) \ False" by auto text \Transitivity rules for calculational reasoning\ lemma neq_le_trans: "a \ b \ a \ b \ a < b" by (fact order.not_eq_order_implies_strict) lemma le_neq_trans: "a \ b \ a \ b \ a < b" by (rule order.not_eq_order_implies_strict) text \Asymmetry.\ lemma order_eq_iff: "x = y \ x \ y \ y \ x" by (fact order.eq_iff) lemma antisym_conv: "y \ x \ x \ y \ x = y" by (simp add: order.eq_iff) lemma less_imp_neq: "x < y \ x \ y" by (fact order.strict_implies_not_eq) lemma antisym_conv1: "\ x < y \ x \ y \ x = y" by (simp add: local.le_less) lemma antisym_conv2: "x \ y \ \ x < y \ x = y" by (simp add: local.less_le) lemma leD: "y \ x \ \ x < y" by (auto simp: less_le order.antisym) text \Least value operator\ definition (in ord) Least :: "('a \ bool) \ 'a" (binder "LEAST " 10) where "Least P = (THE x. P x \ (\y. P y \ x \ y))" lemma Least_equality: assumes "P x" and "\y. P y \ x \ y" shows "Least P = x" unfolding Least_def by (rule the_equality) (blast intro: assms order.antisym)+ lemma LeastI2_order: assumes "P x" and "\y. P y \ x \ y" and "\x. P x \ \y. P y \ x \ y \ Q x" shows "Q (Least P)" unfolding Least_def by (rule theI2) (blast intro: assms order.antisym)+ lemma Least_ex1: assumes "\!x. P x \ (\y. P y \ x \ y)" shows Least1I: "P (Least P)" and Least1_le: "P z \ Least P \ z" using theI'[OF assms] unfolding Least_def by auto text \Greatest value operator\ definition Greatest :: "('a \ bool) \ 'a" (binder "GREATEST " 10) where "Greatest P = (THE x. P x \ (\y. P y \ x \ y))" lemma GreatestI2_order: "\ P x; \y. P y \ x \ y; \x. \ P x; \y. P y \ x \ y \ \ Q x \ \ Q (Greatest P)" unfolding Greatest_def by (rule theI2) (blast intro: order.antisym)+ lemma Greatest_equality: "\ P x; \y. P y \ x \ y \ \ Greatest P = x" unfolding Greatest_def by (rule the_equality) (blast intro: order.antisym)+ end lemma ordering_orderI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "ordering less_eq less" shows "class.order less_eq less" proof - from assms interpret ordering less_eq less . show ?thesis by standard (auto intro: antisym trans simp add: refl strict_iff_order) qed lemma order_strictI: fixes less (infix "\<^bold><" 50) and less_eq (infix "\<^bold>\" 50) assumes "\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b" assumes "\a b. a \<^bold>< b \ \ b \<^bold>< a" assumes "\a. \ a \<^bold>< a" assumes "\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c" shows "class.order less_eq less" by (rule ordering_orderI) (rule ordering_strictI, (fact assms)+) context order begin text \Dual order\ lemma dual_order: "class.order (\) (>)" using dual_order.ordering_axioms by (rule ordering_orderI) end subsection \Linear (total) orders\ class linorder = order + assumes linear: "x \ y \ y \ x" begin lemma less_linear: "x < y \ x = y \ y < x" unfolding less_le using less_le linear by blast lemma le_less_linear: "x \ y \ y < x" by (simp add: le_less less_linear) lemma le_cases [case_names le ge]: "(x \ y \ P) \ (y \ x \ P) \ P" using linear by blast lemma (in linorder) le_cases3: "\\x \ y; y \ z\ \ P; \y \ x; x \ z\ \ P; \x \ z; z \ y\ \ P; \z \ y; y \ x\ \ P; \y \ z; z \ x\ \ P; \z \ x; x \ y\ \ P\ \ P" by (blast intro: le_cases) lemma linorder_cases [case_names less equal greater]: "(x < y \ P) \ (x = y \ P) \ (y < x \ P) \ P" using less_linear by blast lemma linorder_wlog[case_names le sym]: "(\a b. a \ b \ P a b) \ (\a b. P b a \ P a b) \ P a b" by (cases rule: le_cases[of a b]) blast+ lemma not_less: "\ x < y \ y \ x" unfolding less_le using linear by (blast intro: order.antisym) lemma not_less_iff_gr_or_eq: "\(x < y) \ (x > y \ x = y)" by (auto simp add:not_less le_less) lemma not_le: "\ x \ y \ y < x" unfolding less_le using linear by (blast intro: order.antisym) lemma neq_iff: "x \ y \ x < y \ y < x" by (cut_tac x = x and y = y in less_linear, auto) lemma neqE: "x \ y \ (x < y \ R) \ (y < x \ R) \ R" by (simp add: neq_iff) blast lemma antisym_conv3: "\ y < x \ \ x < y \ x = y" by (blast intro: order.antisym dest: not_less [THEN iffD1]) lemma leI: "\ x < y \ y \ x" unfolding not_less . lemma not_le_imp_less: "\ y \ x \ x < y" unfolding not_le . lemma linorder_less_wlog[case_names less refl sym]: "\\a b. a < b \ P a b; \a. P a a; \a b. P b a \ P a b\ \ P a b" using antisym_conv3 by blast text \Dual order\ lemma dual_linorder: "class.linorder (\) (>)" by (rule class.linorder.intro, rule dual_order) (unfold_locales, rule linear) end text \Alternative introduction rule with bias towards strict order\ lemma linorder_strictI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "class.order less_eq less" assumes trichotomy: "\a b. a \<^bold>< b \ a = b \ b \<^bold>< a" shows "class.linorder less_eq less" proof - interpret order less_eq less by (fact \class.order less_eq less\) show ?thesis proof fix a b show "a \<^bold>\ b \ b \<^bold>\ a" using trichotomy by (auto simp add: le_less) qed qed subsection \Reasoning tools setup\ ML \ structure Logic_Signature : LOGIC_SIGNATURE = struct val mk_Trueprop = HOLogic.mk_Trueprop val dest_Trueprop = HOLogic.dest_Trueprop val Trueprop_conv = HOLogic.Trueprop_conv val Not = HOLogic.Not val conj = HOLogic.conj val disj = HOLogic.disj val notI = @{thm notI} val ccontr = @{thm ccontr} val conjI = @{thm conjI} val conjE = @{thm conjE} val disjE = @{thm disjE} val not_not_conv = Conv.rewr_conv @{thm eq_reflection[OF not_not]} val de_Morgan_conj_conv = Conv.rewr_conv @{thm eq_reflection[OF de_Morgan_conj]} val de_Morgan_disj_conv = Conv.rewr_conv @{thm eq_reflection[OF de_Morgan_disj]} val conj_disj_distribL_conv = Conv.rewr_conv @{thm eq_reflection[OF conj_disj_distribL]} val conj_disj_distribR_conv = Conv.rewr_conv @{thm eq_reflection[OF conj_disj_distribR]} end structure HOL_Base_Order_Tac = Base_Order_Tac( structure Logic_Sig = Logic_Signature; (* Exclude types with specialised solvers. *) val excluded_types = [HOLogic.natT, HOLogic.intT, HOLogic.realT] ) structure HOL_Order_Tac = Order_Tac(structure Base_Tac = HOL_Base_Order_Tac) fun print_orders ctxt0 = let val ctxt = Config.put show_sorts true ctxt0 val orders = HOL_Order_Tac.Data.get (Context.Proof ctxt) fun pretty_term t = Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t), Pretty.brk 1, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt (type_of t)), Pretty.brk 1] fun pretty_order ({kind = kind, ops = ops, ...}, _) = Pretty.block ([Pretty.str (@{make_string} kind), Pretty.str ":", Pretty.brk 1] @ map pretty_term ops) in Pretty.writeln (Pretty.big_list "order structures:" (map pretty_order orders)) end val _ = Outer_Syntax.command \<^command_keyword>\print_orders\ "print order structures available to transitivity reasoner" (Scan.succeed (Toplevel.keep (print_orders o Toplevel.context_of))) \ method_setup order = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (HOL_Order_Tac.tac [] ctxt)) \ "transitivity reasoner" text \Declarations to set up transitivity reasoner of partial and linear orders.\ context order begin lemma nless_le: "(\ a < b) \ (\ a \ b) \ a = b" using local.dual_order.order_iff_strict by blast local_setup \ HOL_Order_Tac.declare_order { ops = {eq = @{term \(=) :: 'a \ 'a \ bool\}, le = @{term \(\)\}, lt = @{term \(<)\}}, thms = {trans = @{thm order_trans}, refl = @{thm order_refl}, eqD1 = @{thm eq_refl}, eqD2 = @{thm eq_refl[OF sym]}, antisym = @{thm order_antisym}, contr = @{thm notE}}, conv_thms = {less_le = @{thm eq_reflection[OF less_le]}, nless_le = @{thm eq_reflection[OF nless_le]}} } \ end context linorder begin lemma nle_le: "(\ a \ b) \ b \ a \ b \ a" using not_le less_le by simp local_setup \ HOL_Order_Tac.declare_linorder { ops = {eq = @{term \(=) :: 'a \ 'a \ bool\}, le = @{term \(\)\}, lt = @{term \(<)\}}, thms = {trans = @{thm order_trans}, refl = @{thm order_refl}, eqD1 = @{thm eq_refl}, eqD2 = @{thm eq_refl[OF sym]}, antisym = @{thm order_antisym}, contr = @{thm notE}}, conv_thms = {less_le = @{thm eq_reflection[OF less_le]}, nless_le = @{thm eq_reflection[OF not_less]}, nle_le = @{thm eq_reflection[OF nle_le]}} } \ end setup \ map_theory_simpset (fn ctxt0 => ctxt0 addSolver mk_solver "Transitivity" (fn ctxt => HOL_Order_Tac.tac (Simplifier.prems_of ctxt) ctxt)) \ ML \ local fun prp t thm = Thm.prop_of thm = t; (* FIXME proper aconv!? *) in fun antisym_le_simproc ctxt ct = (case Thm.term_of ct of (le as Const (_, T)) $ r $ s => (let val prems = Simplifier.prems_of ctxt; val less = Const (\<^const_name>\less\, T); val t = HOLogic.mk_Trueprop(le $ s $ r); in (case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop(HOLogic.Not $ (less $ r $ s)) in (case find_first (prp t) prems of NONE => NONE | SOME thm => SOME(mk_meta_eq(thm RS @{thm antisym_conv1}))) end | SOME thm => SOME (mk_meta_eq (thm RS @{thm order_class.antisym_conv}))) end handle THM _ => NONE) | _ => NONE); fun antisym_less_simproc ctxt ct = (case Thm.term_of ct of NotC $ ((less as Const(_,T)) $ r $ s) => (let val prems = Simplifier.prems_of ctxt; val le = Const (\<^const_name>\less_eq\, T); val t = HOLogic.mk_Trueprop(le $ r $ s); in (case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop (NotC $ (less $ s $ r)) in (case find_first (prp t) prems of NONE => NONE | SOME thm => SOME (mk_meta_eq(thm RS @{thm linorder_class.antisym_conv3}))) end | SOME thm => SOME (mk_meta_eq (thm RS @{thm antisym_conv2}))) end handle THM _ => NONE) | _ => NONE); end; \ simproc_setup antisym_le ("(x::'a::order) \ y") = "K antisym_le_simproc" simproc_setup antisym_less ("\ (x::'a::linorder) < y") = "K antisym_less_simproc" subsection \Bounded quantifiers\ syntax (ASCII) "_All_less" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_All_greater" :: "[idt, 'a, bool] => bool" ("(3ALL _>_./ _)" [0, 0, 10] 10) "_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3EX _>_./ _)" [0, 0, 10] 10) "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3ALL _~=_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3EX _~=_./ _)" [0, 0, 10] 10) syntax "_All_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_greater" :: "[idt, 'a, bool] => bool" ("(3\_>_./ _)" [0, 0, 10] 10) "_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3\_>_./ _)" [0, 0, 10] 10) "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) syntax (input) "_All_less" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3! _~=_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3? _~=_./ _)" [0, 0, 10] 10) translations "\x "\x. x < y \ P" "\x "\x. x < y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x>y. P" \ "\x. x > y \ P" "\x>y. P" \ "\x. x > y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" print_translation \ let val All_binder = Mixfix.binder_name \<^const_syntax>\All\; val Ex_binder = Mixfix.binder_name \<^const_syntax>\Ex\; val impl = \<^const_syntax>\HOL.implies\; val conj = \<^const_syntax>\HOL.conj\; val less = \<^const_syntax>\less\; val less_eq = \<^const_syntax>\less_eq\; val trans = [((All_binder, impl, less), (\<^syntax_const>\_All_less\, \<^syntax_const>\_All_greater\)), ((All_binder, impl, less_eq), (\<^syntax_const>\_All_less_eq\, \<^syntax_const>\_All_greater_eq\)), ((Ex_binder, conj, less), (\<^syntax_const>\_Ex_less\, \<^syntax_const>\_Ex_greater\)), ((Ex_binder, conj, less_eq), (\<^syntax_const>\_Ex_less_eq\, \<^syntax_const>\_Ex_greater_eq\))]; fun matches_bound v t = (case t of Const (\<^syntax_const>\_bound\, _) $ Free (v', _) => v = v' | _ => false); fun contains_var v = Term.exists_subterm (fn Free (x, _) => x = v | _ => false); fun mk x c n P = Syntax.const c $ Syntax_Trans.mark_bound_body x $ n $ P; fun tr' q = (q, fn _ => (fn [Const (\<^syntax_const>\_bound\, _) $ Free (v, T), Const (c, _) $ (Const (d, _) $ t $ u) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME (l, g) => if matches_bound v t andalso not (contains_var v u) then mk (v, T) l u P else if matches_bound v u andalso not (contains_var v t) then mk (v, T) g t P else raise Match) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ subsection \Transitivity reasoning\ context ord begin lemma ord_le_eq_trans: "a \ b \ b = c \ a \ c" by (rule subst) lemma ord_eq_le_trans: "a = b \ b \ c \ a \ c" by (rule ssubst) lemma ord_less_eq_trans: "a < b \ b = c \ a < c" by (rule subst) lemma ord_eq_less_trans: "a = b \ b < c \ a < c" by (rule ssubst) end -lemma order_less_subst2: "(a::'a::order) < b ==> f b < (c::'c::order) ==> - (!!x y. x < y ==> f x < f y) ==> f a < c" +lemma order_less_subst2: "(a::'a::order) < b \ f b < (c::'c::order) \ + (!!x y. x < y \ f x < f y) \ f a < c" proof - - assume r: "!!x y. x < y ==> f x < f y" + assume r: "!!x y. x < y \ f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b < c" finally (less_trans) show ?thesis . qed -lemma order_less_subst1: "(a::'a::order) < f b ==> (b::'b::order) < c ==> - (!!x y. x < y ==> f x < f y) ==> a < f c" +lemma order_less_subst1: "(a::'a::order) < f b \ (b::'b::order) < c \ + (!!x y. x < y \ f x < f y) \ a < f c" proof - - assume r: "!!x y. x < y ==> f x < f y" + assume r: "!!x y. x < y \ f x < f y" assume "a < f b" also assume "b < c" hence "f b < f c" by (rule r) finally (less_trans) show ?thesis . qed -lemma order_le_less_subst2: "(a::'a::order) <= b ==> f b < (c::'c::order) ==> - (!!x y. x <= y ==> f x <= f y) ==> f a < c" +lemma order_le_less_subst2: "(a::'a::order) <= b \ f b < (c::'c::order) \ + (!!x y. x <= y \ f x <= f y) \ f a < c" proof - - assume r: "!!x y. x <= y ==> f x <= f y" + assume r: "!!x y. x <= y \ f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b < c" finally (le_less_trans) show ?thesis . qed -lemma order_le_less_subst1: "(a::'a::order) <= f b ==> (b::'b::order) < c ==> - (!!x y. x < y ==> f x < f y) ==> a < f c" +lemma order_le_less_subst1: "(a::'a::order) <= f b \ (b::'b::order) < c \ + (!!x y. x < y \ f x < f y) \ a < f c" proof - - assume r: "!!x y. x < y ==> f x < f y" + assume r: "!!x y. x < y \ f x < f y" assume "a <= f b" also assume "b < c" hence "f b < f c" by (rule r) finally (le_less_trans) show ?thesis . qed -lemma order_less_le_subst2: "(a::'a::order) < b ==> f b <= (c::'c::order) ==> - (!!x y. x < y ==> f x < f y) ==> f a < c" +lemma order_less_le_subst2: "(a::'a::order) < b \ f b <= (c::'c::order) \ + (!!x y. x < y \ f x < f y) \ f a < c" proof - - assume r: "!!x y. x < y ==> f x < f y" + assume r: "!!x y. x < y \ f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b <= c" finally (less_le_trans) show ?thesis . qed -lemma order_less_le_subst1: "(a::'a::order) < f b ==> (b::'b::order) <= c ==> - (!!x y. x <= y ==> f x <= f y) ==> a < f c" +lemma order_less_le_subst1: "(a::'a::order) < f b \ (b::'b::order) <= c \ + (!!x y. x <= y \ f x <= f y) \ a < f c" proof - - assume r: "!!x y. x <= y ==> f x <= f y" + assume r: "!!x y. x <= y \ f x <= f y" assume "a < f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (less_le_trans) show ?thesis . qed -lemma order_subst1: "(a::'a::order) <= f b ==> (b::'b::order) <= c ==> - (!!x y. x <= y ==> f x <= f y) ==> a <= f c" +lemma order_subst1: "(a::'a::order) <= f b \ (b::'b::order) <= c \ + (!!x y. x <= y \ f x <= f y) \ a <= f c" proof - - assume r: "!!x y. x <= y ==> f x <= f y" + assume r: "!!x y. x <= y \ f x <= f y" assume "a <= f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (order_trans) show ?thesis . qed -lemma order_subst2: "(a::'a::order) <= b ==> f b <= (c::'c::order) ==> - (!!x y. x <= y ==> f x <= f y) ==> f a <= c" +lemma order_subst2: "(a::'a::order) <= b \ f b <= (c::'c::order) \ + (!!x y. x <= y \ f x <= f y) \ f a <= c" proof - - assume r: "!!x y. x <= y ==> f x <= f y" + assume r: "!!x y. x <= y \ f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b <= c" finally (order_trans) show ?thesis . qed -lemma ord_le_eq_subst: "a <= b ==> f b = c ==> - (!!x y. x <= y ==> f x <= f y) ==> f a <= c" +lemma ord_le_eq_subst: "a <= b \ f b = c \ + (!!x y. x <= y \ f x <= f y) \ f a <= c" proof - - assume r: "!!x y. x <= y ==> f x <= f y" + assume r: "!!x y. x <= y \ f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b = c" finally (ord_le_eq_trans) show ?thesis . qed -lemma ord_eq_le_subst: "a = f b ==> b <= c ==> - (!!x y. x <= y ==> f x <= f y) ==> a <= f c" +lemma ord_eq_le_subst: "a = f b \ b <= c \ + (!!x y. x <= y \ f x <= f y) \ a <= f c" proof - - assume r: "!!x y. x <= y ==> f x <= f y" + assume r: "!!x y. x <= y \ f x <= f y" assume "a = f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (ord_eq_le_trans) show ?thesis . qed -lemma ord_less_eq_subst: "a < b ==> f b = c ==> - (!!x y. x < y ==> f x < f y) ==> f a < c" +lemma ord_less_eq_subst: "a < b \ f b = c \ + (!!x y. x < y \ f x < f y) \ f a < c" proof - - assume r: "!!x y. x < y ==> f x < f y" + assume r: "!!x y. x < y \ f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b = c" finally (ord_less_eq_trans) show ?thesis . qed -lemma ord_eq_less_subst: "a = f b ==> b < c ==> - (!!x y. x < y ==> f x < f y) ==> a < f c" +lemma ord_eq_less_subst: "a = f b \ b < c \ + (!!x y. x < y \ f x < f y) \ a < f c" proof - - assume r: "!!x y. x < y ==> f x < f y" + assume r: "!!x y. x < y \ f x < f y" assume "a = f b" also assume "b < c" hence "f b < f c" by (rule r) finally (ord_eq_less_trans) show ?thesis . qed text \ Note that this list of rules is in reverse order of priorities. \ lemmas [trans] = order_less_subst2 order_less_subst1 order_le_less_subst2 order_le_less_subst1 order_less_le_subst2 order_less_le_subst1 order_subst2 order_subst1 ord_le_eq_subst ord_eq_le_subst ord_less_eq_subst ord_eq_less_subst forw_subst back_subst rev_mp mp lemmas (in order) [trans] = neq_le_trans le_neq_trans lemmas (in preorder) [trans] = less_trans less_asym' le_less_trans less_le_trans order_trans lemmas (in order) [trans] = order.antisym lemmas (in ord) [trans] = ord_le_eq_trans ord_eq_le_trans ord_less_eq_trans ord_eq_less_trans lemmas [trans] = trans lemmas order_trans_rules = order_less_subst2 order_less_subst1 order_le_less_subst2 order_le_less_subst1 order_less_le_subst2 order_less_le_subst1 order_subst2 order_subst1 ord_le_eq_subst ord_eq_le_subst ord_less_eq_subst ord_eq_less_subst forw_subst back_subst rev_mp mp neq_le_trans le_neq_trans less_trans less_asym' le_less_trans less_le_trans order_trans order.antisym ord_le_eq_trans ord_eq_le_trans ord_less_eq_trans ord_eq_less_trans trans text \These support proving chains of decreasing inequalities - a >= b >= c ... in Isar proofs.\ + a \ b \ c ... in Isar proofs.\ lemma xt1 [no_atp]: "a = b \ b > c \ a > c" "a > b \ b = c \ a > c" "a = b \ b \ c \ a \ c" "a \ b \ b = c \ a \ c" "(x::'a::order) \ y \ y \ x \ x = y" "(x::'a::order) \ y \ y \ z \ x \ z" "(x::'a::order) > y \ y \ z \ x > z" "(x::'a::order) \ y \ y > z \ x > z" "(a::'a::order) > b \ b > a \ P" "(x::'a::order) > y \ y > z \ x > z" "(a::'a::order) \ b \ a \ b \ a > b" "(a::'a::order) \ b \ a \ b \ a > b" "a = f b \ b > c \ (\x y. x > y \ f x > f y) \ a > f c" "a > b \ f b = c \ (\x y. x > y \ f x > f y) \ f a > c" "a = f b \ b \ c \ (\x y. x \ y \ f x \ f y) \ a \ f c" "a \ b \ f b = c \ (\x y. x \ y \ f x \ f y) \ f a \ c" by auto lemma xt2 [no_atp]: - "(a::'a::order) >= f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c" -by (subgoal_tac "f b >= f c", force, force) - -lemma xt3 [no_atp]: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> - (!!x y. x >= y ==> f x >= f y) ==> f a >= c" -by (subgoal_tac "f a >= f b", force, force) - -lemma xt4 [no_atp]: "(a::'a::order) > f b ==> (b::'b::order) >= c ==> - (!!x y. x >= y ==> f x >= f y) ==> a > f c" -by (subgoal_tac "f b >= f c", force, force) - -lemma xt5 [no_atp]: "(a::'a::order) > b ==> (f b::'b::order) >= c==> - (!!x y. x > y ==> f x > f y) ==> f a > c" -by (subgoal_tac "f a > f b", force, force) + assumes "(a::'a::order) \ f b" + and "b \ c" + and "\x y. x \ y \ f x \ f y" + shows "a \ f c" + using assms by force -lemma xt6 [no_atp]: "(a::'a::order) >= f b ==> b > c ==> - (!!x y. x > y ==> f x > f y) ==> a > f c" -by (subgoal_tac "f b > f c", force, force) +lemma xt3 [no_atp]: + assumes "(a::'a::order) \ b" + and "(f b::'b::order) \ c" + and "\x y. x \ y \ f x \ f y" + shows "f a \ c" + using assms by force -lemma xt7 [no_atp]: "(a::'a::order) >= b ==> (f b::'b::order) > c ==> - (!!x y. x >= y ==> f x >= f y) ==> f a > c" -by (subgoal_tac "f a >= f b", force, force) +lemma xt4 [no_atp]: + assumes "(a::'a::order) > f b" + and "(b::'b::order) \ c" + and "\x y. x \ y \ f x \ f y" + shows "a > f c" + using assms by force -lemma xt8 [no_atp]: "(a::'a::order) > f b ==> (b::'b::order) > c ==> - (!!x y. x > y ==> f x > f y) ==> a > f c" -by (subgoal_tac "f b > f c", force, force) +lemma xt5 [no_atp]: + assumes "(a::'a::order) > b" + and "(f b::'b::order) \ c" + and "\x y. x > y \ f x > f y" + shows "f a > c" + using assms by force -lemma xt9 [no_atp]: "(a::'a::order) > b ==> (f b::'b::order) > c ==> - (!!x y. x > y ==> f x > f y) ==> f a > c" -by (subgoal_tac "f a > f b", force, force) +lemma xt6 [no_atp]: + assumes "(a::'a::order) \ f b" + and "b > c" + and "\x y. x > y \ f x > f y" + shows "a > f c" + using assms by force + +lemma xt7 [no_atp]: + assumes "(a::'a::order) \ b" + and "(f b::'b::order) > c" + and "\x y. x \ y \ f x \ f y" + shows "f a > c" + using assms by force + +lemma xt8 [no_atp]: + assumes "(a::'a::order) > f b" + and "(b::'b::order) > c" + and "\x y. x > y \ f x > f y" + shows "a > f c" + using assms by force + +lemma xt9 [no_atp]: + assumes "(a::'a::order) > b" + and "(f b::'b::order) > c" + and "\x y. x > y \ f x > f y" + shows "f a > c" + using assms by force lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 (* - Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands + Since "a \ b" abbreviates "b \ a", the abbreviation "..." stands for the wrong thing in an Isar proof. The extra transitivity rules can be used as follows: lemma "(a::'a::order) > z" proof - - have "a >= b" (is "_ >= ?rhs") + have "a \ b" (is "_ \ ?rhs") sorry - also have "?rhs >= c" (is "_ >= ?rhs") + also have "?rhs \ c" (is "_ \ ?rhs") sorry also (xtrans) have "?rhs = d" (is "_ = ?rhs") sorry - also (xtrans) have "?rhs >= e" (is "_ >= ?rhs") + also (xtrans) have "?rhs \ e" (is "_ \ ?rhs") sorry also (xtrans) have "?rhs > f" (is "_ > ?rhs") sorry also (xtrans) have "?rhs > z" sorry finally (xtrans) show ?thesis . qed Alternatively, one can use "declare xtrans [trans]" and then leave out the "(xtrans)" above. *) subsection \Monotonicity\ context order begin definition mono :: "('a \ 'b::order) \ bool" where "mono f \ (\x y. x \ y \ f x \ f y)" lemma monoI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ mono f" unfolding mono_def by iprover lemma monoD [dest?]: fixes f :: "'a \ 'b::order" shows "mono f \ x \ y \ f x \ f y" unfolding mono_def by iprover lemma monoE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: mono_def) qed definition antimono :: "('a \ 'b::order) \ bool" where "antimono f \ (\x y. x \ y \ f x \ f y)" lemma antimonoI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ antimono f" unfolding antimono_def by iprover lemma antimonoD [dest?]: fixes f :: "'a \ 'b::order" shows "antimono f \ x \ y \ f x \ f y" unfolding antimono_def by iprover lemma antimonoE: fixes f :: "'a \ 'b::order" assumes "antimono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: antimono_def) qed definition strict_mono :: "('a \ 'b::order) \ bool" where "strict_mono f \ (\x y. x < y \ f x < f y)" lemma strict_monoI [intro?]: assumes "\x y. x < y \ f x < f y" shows "strict_mono f" using assms unfolding strict_mono_def by auto lemma strict_monoD [dest?]: "strict_mono f \ x < y \ f x < f y" unfolding strict_mono_def by auto lemma strict_mono_mono [dest?]: assumes "strict_mono f" shows "mono f" proof (rule monoI) fix x y assume "x \ y" show "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with \x \ y\ have "x < y" by simp with assms strict_monoD have "f x < f y" by auto then show ?thesis by simp qed qed end context linorder begin lemma mono_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x \ y" proof show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma mono_strict_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x < y" proof show "x < y" proof (rule ccontr) assume "\ x < y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma strict_mono_eq: assumes "strict_mono f" shows "f x = f y \ x = y" proof assume "f x = f y" show "x = y" proof (cases x y rule: linorder_cases) case less with assms strict_monoD have "f x < f y" by auto with \f x = f y\ show ?thesis by simp next case equal then show ?thesis . next case greater with assms strict_monoD have "f y < f x" by auto with \f x = f y\ show ?thesis by simp qed qed simp lemma strict_mono_less_eq: assumes "strict_mono f" shows "f x \ f y \ x \ y" proof assume "x \ y" with assms strict_mono_mono monoD show "f x \ f y" by auto next assume "f x \ f y" show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y < x" by simp with assms strict_monoD have "f y < f x" by auto with \f x \ f y\ show False by simp qed qed lemma strict_mono_less: assumes "strict_mono f" shows "f x < f y \ x < y" using assms by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) end subsection \min and max -- fundamental\ definition (in ord) min :: "'a \ 'a \ 'a" where "min a b = (if a \ b then a else b)" definition (in ord) max :: "'a \ 'a \ 'a" where "max a b = (if a \ b then b else a)" lemma min_absorb1: "x \ y \ min x y = x" by (simp add: min_def) lemma max_absorb2: "x \ y \ max x y = y" by (simp add: max_def) lemma min_absorb2: "(y::'a::order) \ x \ min x y = y" by (simp add:min_def) lemma max_absorb1: "(y::'a::order) \ x \ max x y = x" by (simp add: max_def) lemma max_min_same [simp]: fixes x y :: "'a :: linorder" shows "max x (min x y) = x" "max (min x y) x = x" "max (min x y) y = y" "max y (min x y) = y" by(auto simp add: max_def min_def) subsection \(Unique) top and bottom elements\ class bot = fixes bot :: 'a ("\") class order_bot = order + bot + assumes bot_least: "\ \ a" begin sublocale bot: ordering_top greater_eq greater bot by standard (fact bot_least) lemma le_bot: "a \ \ \ a = \" by (fact bot.extremum_uniqueI) lemma bot_unique: "a \ \ \ a = \" by (fact bot.extremum_unique) lemma not_less_bot: "\ a < \" by (fact bot.extremum_strict) lemma bot_less: "a \ \ \ \ < a" by (fact bot.not_eq_extremum) lemma max_bot[simp]: "max bot x = x" by(simp add: max_def bot_unique) lemma max_bot2[simp]: "max x bot = x" by(simp add: max_def bot_unique) lemma min_bot[simp]: "min bot x = bot" by(simp add: min_def bot_unique) lemma min_bot2[simp]: "min x bot = bot" by(simp add: min_def bot_unique) end class top = fixes top :: 'a ("\") class order_top = order + top + assumes top_greatest: "a \ \" begin sublocale top: ordering_top less_eq less top by standard (fact top_greatest) lemma top_le: "\ \ a \ a = \" by (fact top.extremum_uniqueI) lemma top_unique: "\ \ a \ a = \" by (fact top.extremum_unique) lemma not_top_less: "\ \ < a" by (fact top.extremum_strict) lemma less_top: "a \ \ \ a < \" by (fact top.not_eq_extremum) lemma max_top[simp]: "max top x = top" by(simp add: max_def top_unique) lemma max_top2[simp]: "max x top = top" by(simp add: max_def top_unique) lemma min_top[simp]: "min top x = x" by(simp add: min_def top_unique) lemma min_top2[simp]: "min x top = x" by(simp add: min_def top_unique) end subsection \Dense orders\ class dense_order = order + assumes dense: "x < y \ (\z. x < z \ z < y)" class dense_linorder = linorder + dense_order begin lemma dense_le: fixes y z :: 'a assumes "\x. x < y \ x \ z" shows "y \ z" proof (rule ccontr) assume "\ ?thesis" hence "z < y" by simp from dense[OF this] obtain x where "x < y" and "z < x" by safe moreover have "x \ z" using assms[OF \x < y\] . ultimately show False by auto qed lemma dense_le_bounded: fixes x y z :: 'a assumes "x < y" assumes *: "\w. \ x < w ; w < y \ \ w \ z" shows "y \ z" proof (rule dense_le) fix w assume "w < y" from dense[OF \x < y\] obtain u where "x < u" "u < y" by safe from linear[of u w] show "w \ z" proof (rule disjE) assume "u \ w" from less_le_trans[OF \x < u\ \u \ w\] \w < y\ show "w \ z" by (rule *) next assume "w \ u" from \w \ u\ *[OF \x < u\ \u < y\] show "w \ z" by (rule order_trans) qed qed lemma dense_ge: fixes y z :: 'a assumes "\x. z < x \ y \ x" shows "y \ z" proof (rule ccontr) assume "\ ?thesis" hence "z < y" by simp from dense[OF this] obtain x where "x < y" and "z < x" by safe moreover have "y \ x" using assms[OF \z < x\] . ultimately show False by auto qed lemma dense_ge_bounded: fixes x y z :: 'a assumes "z < x" assumes *: "\w. \ z < w ; w < x \ \ y \ w" shows "y \ z" proof (rule dense_ge) fix w assume "z < w" from dense[OF \z < x\] obtain u where "z < u" "u < x" by safe from linear[of u w] show "y \ w" proof (rule disjE) assume "w \ u" from \z < w\ le_less_trans[OF \w \ u\ \u < x\] show "y \ w" by (rule *) next assume "u \ w" from *[OF \z < u\ \u < x\] \u \ w\ show "y \ w" by (rule order_trans) qed qed end class no_top = order + assumes gt_ex: "\y. x < y" class no_bot = order + assumes lt_ex: "\y. y < x" class unbounded_dense_linorder = dense_linorder + no_top + no_bot subsection \Wellorders\ class wellorder = linorder + assumes less_induct [case_names less]: "(\x. (\y. y < x \ P y) \ P x) \ P a" begin lemma wellorder_Least_lemma: fixes k :: 'a assumes "P k" shows LeastI: "P (LEAST x. P x)" and Least_le: "(LEAST x. P x) \ k" proof - have "P (LEAST x. P x) \ (LEAST x. P x) \ k" using assms proof (induct k rule: less_induct) case (less x) then have "P x" by simp show ?case proof (rule classical) assume assm: "\ (P (LEAST a. P a) \ (LEAST a. P a) \ x)" have "\y. P y \ x \ y" proof (rule classical) fix y assume "P y" and "\ x \ y" with less have "P (LEAST a. P a)" and "(LEAST a. P a) \ y" by (auto simp add: not_le) with assm have "x < (LEAST a. P a)" and "(LEAST a. P a) \ y" by auto then show "x \ y" by auto qed with \P x\ have Least: "(LEAST a. P a) = x" by (rule Least_equality) with \P x\ show ?thesis by simp qed qed then show "P (LEAST x. P x)" and "(LEAST x. P x) \ k" by auto qed \ \The following 3 lemmas are due to Brian Huffman\ lemma LeastI_ex: "\x. P x \ P (Least P)" by (erule exE) (erule LeastI) lemma LeastI2: "P a \ (\x. P x \ Q x) \ Q (Least P)" by (blast intro: LeastI) lemma LeastI2_ex: "\a. P a \ (\x. P x \ Q x) \ Q (Least P)" by (blast intro: LeastI_ex) lemma LeastI2_wellorder: assumes "P a" and "\a. \ P a; \b. P b \ a \ b \ \ Q a" shows "Q (Least P)" proof (rule LeastI2_order) show "P (Least P)" using \P a\ by (rule LeastI) next fix y assume "P y" thus "Least P \ y" by (rule Least_le) next fix x assume "P x" "\y. P y \ x \ y" thus "Q x" by (rule assms(2)) qed lemma LeastI2_wellorder_ex: assumes "\x. P x" and "\a. \ P a; \b. P b \ a \ b \ \ Q a" shows "Q (Least P)" using assms by clarify (blast intro!: LeastI2_wellorder) lemma not_less_Least: "k < (LEAST x. P x) \ \ P k" apply (simp add: not_le [symmetric]) apply (erule contrapos_nn) apply (erule Least_le) done lemma exists_least_iff: "(\n. P n) \ (\n. P n \ (\m < n. \ P m))" (is "?lhs \ ?rhs") proof assume ?rhs thus ?lhs by blast next assume H: ?lhs then obtain n where n: "P n" by blast let ?x = "Least P" { fix m assume m: "m < ?x" from not_less_Least[OF m] have "\ P m" . } with LeastI_ex[OF H] show ?rhs by blast qed end subsection \Order on \<^typ>\bool\\ instantiation bool :: "{order_bot, order_top, linorder}" begin definition le_bool_def [simp]: "P \ Q \ P \ Q" definition [simp]: "(P::bool) < Q \ \ P \ Q" definition [simp]: "\ \ False" definition [simp]: "\ \ True" instance proof qed auto end lemma le_boolI: "(P \ Q) \ P \ Q" by simp lemma le_boolI': "P \ Q \ P \ Q" by simp lemma le_boolE: "P \ Q \ P \ (Q \ R) \ R" by simp lemma le_boolD: "P \ Q \ P \ Q" by simp lemma bot_boolE: "\ \ P" by simp lemma top_boolI: \ by simp lemma [code]: "False \ b \ True" "True \ b \ b" "False < b \ b" "True < b \ False" by simp_all subsection \Order on \<^typ>\_ \ _\\ instantiation "fun" :: (type, ord) ord begin definition le_fun_def: "f \ g \ (\x. f x \ g x)" definition "(f::'a \ 'b) < g \ f \ g \ \ (g \ f)" instance .. end instance "fun" :: (type, preorder) preorder proof qed (auto simp add: le_fun_def less_fun_def intro: order_trans order.antisym) instance "fun" :: (type, order) order proof qed (auto simp add: le_fun_def intro: order.antisym) instantiation "fun" :: (type, bot) bot begin definition "\ = (\x. \)" instance .. end instantiation "fun" :: (type, order_bot) order_bot begin lemma bot_apply [simp, code]: "\ x = \" by (simp add: bot_fun_def) instance proof qed (simp add: le_fun_def) end instantiation "fun" :: (type, top) top begin definition [no_atp]: "\ = (\x. \)" instance .. end instantiation "fun" :: (type, order_top) order_top begin lemma top_apply [simp, code]: "\ x = \" by (simp add: top_fun_def) instance proof qed (simp add: le_fun_def) end lemma le_funI: "(\x. f x \ g x) \ f \ g" unfolding le_fun_def by simp lemma le_funE: "f \ g \ (f x \ g x \ P) \ P" unfolding le_fun_def by simp lemma le_funD: "f \ g \ f x \ g x" by (rule le_funE) lemma mono_compose: "mono Q \ mono (\i x. Q i (f x))" unfolding mono_def le_fun_def by auto subsection \Order on unary and binary predicates\ lemma predicate1I: assumes PQ: "\x. P x \ Q x" shows "P \ Q" apply (rule le_funI) apply (rule le_boolI) apply (rule PQ) apply assumption done lemma predicate1D: "P \ Q \ P x \ Q x" apply (erule le_funE) apply (erule le_boolE) apply assumption+ done lemma rev_predicate1D: "P x \ P \ Q \ Q x" by (rule predicate1D) lemma predicate2I: assumes PQ: "\x y. P x y \ Q x y" shows "P \ Q" apply (rule le_funI)+ apply (rule le_boolI) apply (rule PQ) apply assumption done lemma predicate2D: "P \ Q \ P x y \ Q x y" apply (erule le_funE)+ apply (erule le_boolE) apply assumption+ done lemma rev_predicate2D: "P x y \ P \ Q \ Q x y" by (rule predicate2D) lemma bot1E [no_atp]: "\ x \ P" by (simp add: bot_fun_def) lemma bot2E: "\ x y \ P" by (simp add: bot_fun_def) lemma top1I: "\ x" by (simp add: top_fun_def) lemma top2I: "\ x y" by (simp add: top_fun_def) subsection \Name duplicates\ lemmas antisym = order.antisym lemmas eq_iff = order.eq_iff lemmas order_eq_refl = preorder_class.eq_refl lemmas order_less_irrefl = preorder_class.less_irrefl lemmas order_less_imp_le = preorder_class.less_imp_le lemmas order_less_not_sym = preorder_class.less_not_sym lemmas order_less_asym = preorder_class.less_asym lemmas order_less_trans = preorder_class.less_trans lemmas order_le_less_trans = preorder_class.le_less_trans lemmas order_less_le_trans = preorder_class.less_le_trans lemmas order_less_imp_not_less = preorder_class.less_imp_not_less lemmas order_less_imp_triv = preorder_class.less_imp_triv lemmas order_less_asym' = preorder_class.less_asym' lemmas order_less_le = order_class.less_le lemmas order_le_less = order_class.le_less lemmas order_le_imp_less_or_eq = order_class.le_imp_less_or_eq lemmas order_less_imp_not_eq = order_class.less_imp_not_eq lemmas order_less_imp_not_eq2 = order_class.less_imp_not_eq2 lemmas order_neq_le_trans = order_class.neq_le_trans lemmas order_le_neq_trans = order_class.le_neq_trans lemmas order_eq_iff = order_class.order.eq_iff lemmas order_antisym_conv = order_class.antisym_conv lemmas linorder_linear = linorder_class.linear lemmas linorder_less_linear = linorder_class.less_linear lemmas linorder_le_less_linear = linorder_class.le_less_linear lemmas linorder_le_cases = linorder_class.le_cases lemmas linorder_not_less = linorder_class.not_less lemmas linorder_not_le = linorder_class.not_le lemmas linorder_neq_iff = linorder_class.neq_iff lemmas linorder_neqE = linorder_class.neqE end diff --git a/src/HOL/Partial_Function.thy b/src/HOL/Partial_Function.thy --- a/src/HOL/Partial_Function.thy +++ b/src/HOL/Partial_Function.thy @@ -1,459 +1,466 @@ (* Title: HOL/Partial_Function.thy Author: Alexander Krauss, TU Muenchen *) section \Partial Function Definitions\ theory Partial_Function imports Complete_Partial_Order Option keywords "partial_function" :: thy_defn begin named_theorems partial_function_mono "monotonicity rules for partial function definitions" ML_file \Tools/Function/partial_function.ML\ lemma (in ccpo) in_chain_finite: assumes "Complete_Partial_Order.chain (\) A" "finite A" "A \ {}" shows "\A \ A" using assms(2,1,3) proof induction case empty thus ?case by simp next case (insert x A) note chain = \Complete_Partial_Order.chain (\) (insert x A)\ show ?case proof(cases "A = {}") case True thus ?thesis by simp next case False from chain have chain': "Complete_Partial_Order.chain (\) A" by(rule chain_subset) blast hence "\A \ A" using False by(rule insert.IH) show ?thesis proof(cases "x \ \A") case True have "\(insert x A) \ \A" using chain by(rule ccpo_Sup_least)(auto simp add: True intro: ccpo_Sup_upper[OF chain']) hence "\(insert x A) = \A" by(rule order.antisym)(blast intro: ccpo_Sup_upper[OF chain] ccpo_Sup_least[OF chain']) with \\A \ A\ show ?thesis by simp next case False with chainD[OF chain, of x "\A"] \\A \ A\ have "\(insert x A) = x" by(auto intro: order.antisym ccpo_Sup_least[OF chain] order_trans[OF ccpo_Sup_upper[OF chain']] ccpo_Sup_upper[OF chain]) thus ?thesis by simp qed qed qed lemma (in ccpo) admissible_chfin: "(\S. Complete_Partial_Order.chain (\) S \ finite S) \ ccpo.admissible Sup (\) P" using in_chain_finite by (blast intro: ccpo.admissibleI) subsection \Axiomatic setup\ text \This techical locale constains the requirements for function definitions with ccpo fixed points.\ definition "fun_ord ord f g \ (\x. ord (f x) (g x))" definition "fun_lub L A = (\x. L {y. \f\A. y = f x})" definition "img_ord f ord = (\x y. ord (f x) (f y))" definition "img_lub f g Lub = (\A. g (Lub (f ` A)))" lemma chain_fun: assumes A: "chain (fun_ord ord) A" shows "chain ord {y. \f\A. y = f a}" (is "chain ord ?C") proof (rule chainI) fix x y assume "x \ ?C" "y \ ?C" then obtain f g where fg: "f \ A" "g \ A" and [simp]: "x = f a" "y = g a" by blast from chainD[OF A fg] show "ord x y \ ord y x" unfolding fun_ord_def by auto qed lemma call_mono[partial_function_mono]: "monotone (fun_ord ord) ord (\f. f t)" by (rule monotoneI) (auto simp: fun_ord_def) lemma let_mono[partial_function_mono]: "(\x. monotone orda ordb (\f. b f x)) \ monotone orda ordb (\f. Let t (b f))" by (simp add: Let_def) lemma if_mono[partial_function_mono]: "monotone orda ordb F \ monotone orda ordb G \ monotone orda ordb (\f. if c then F f else G f)" unfolding monotone_def by simp definition "mk_less R = (\x y. R x y \ \ R y x)" locale partial_function_definitions = fixes leq :: "'a \ 'a \ bool" fixes lub :: "'a set \ 'a" assumes leq_refl: "leq x x" assumes leq_trans: "leq x y \ leq y z \ leq x z" assumes leq_antisym: "leq x y \ leq y x \ x = y" assumes lub_upper: "chain leq A \ x \ A \ leq x (lub A)" assumes lub_least: "chain leq A \ (\x. x \ A \ leq x z) \ leq (lub A) z" lemma partial_function_lift: assumes "partial_function_definitions ord lb" shows "partial_function_definitions (fun_ord ord) (fun_lub lb)" (is "partial_function_definitions ?ordf ?lubf") proof - interpret partial_function_definitions ord lb by fact show ?thesis proof fix x show "?ordf x x" unfolding fun_ord_def by (auto simp: leq_refl) next fix x y z assume "?ordf x y" "?ordf y z" thus "?ordf x z" unfolding fun_ord_def by (force dest: leq_trans) next fix x y assume "?ordf x y" "?ordf y x" thus "x = y" unfolding fun_ord_def by (force intro!: dest: leq_antisym) next fix A f assume f: "f \ A" and A: "chain ?ordf A" thus "?ordf f (?lubf A)" unfolding fun_lub_def fun_ord_def by (blast intro: lub_upper chain_fun[OF A] f) next fix A :: "('b \ 'a) set" and g :: "'b \ 'a" assume A: "chain ?ordf A" and g: "\f. f \ A \ ?ordf f g" show "?ordf (?lubf A) g" unfolding fun_lub_def fun_ord_def by (blast intro: lub_least chain_fun[OF A] dest: g[unfolded fun_ord_def]) qed qed lemma ccpo: assumes "partial_function_definitions ord lb" shows "class.ccpo lb ord (mk_less ord)" using assms unfolding partial_function_definitions_def mk_less_def by unfold_locales blast+ lemma partial_function_image: assumes "partial_function_definitions ord Lub" assumes inj: "\x y. f x = f y \ x = y" assumes inv: "\x. f (g x) = x" shows "partial_function_definitions (img_ord f ord) (img_lub f g Lub)" proof - let ?iord = "img_ord f ord" let ?ilub = "img_lub f g Lub" interpret partial_function_definitions ord Lub by fact show ?thesis proof fix A x assume "chain ?iord A" "x \ A" then have "chain ord (f ` A)" "f x \ f ` A" by (auto simp: img_ord_def intro: chainI dest: chainD) thus "?iord x (?ilub A)" unfolding inv img_lub_def img_ord_def by (rule lub_upper) next fix A x assume "chain ?iord A" and 1: "\z. z \ A \ ?iord z x" then have "chain ord (f ` A)" by (auto simp: img_ord_def intro: chainI dest: chainD) thus "?iord (?ilub A) x" unfolding inv img_lub_def img_ord_def by (rule lub_least) (auto dest: 1[unfolded img_ord_def]) qed (auto simp: img_ord_def intro: leq_refl dest: leq_trans leq_antisym inj) qed context partial_function_definitions begin abbreviation "le_fun \ fun_ord leq" abbreviation "lub_fun \ fun_lub lub" abbreviation "fixp_fun \ ccpo.fixp lub_fun le_fun" abbreviation "mono_body \ monotone le_fun leq" abbreviation "admissible \ ccpo.admissible lub_fun le_fun" text \Interpret manually, to avoid flooding everything with facts about orders\ lemma ccpo: "class.ccpo lub_fun le_fun (mk_less le_fun)" apply (rule ccpo) apply (rule partial_function_lift) apply (rule partial_function_definitions_axioms) done text \The crucial fixed-point theorem\ lemma mono_body_fixp: "(\x. mono_body (\f. F f x)) \ fixp_fun F = F (fixp_fun F)" by (rule ccpo.fixp_unfold[OF ccpo]) (auto simp: monotone_def fun_ord_def) text \Version with curry/uncurry combinators, to be used by package\ lemma fixp_rule_uc: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a" and C :: "('b \ 'a) \ 'c" assumes mono: "\x. mono_body (\f. U (F (C f)) x)" assumes eq: "f \ C (fixp_fun (\f. U (F (C f))))" assumes inverse: "\f. C (U f) = f" shows "f = F f" proof - have "f = C (fixp_fun (\f. U (F (C f))))" by (simp add: eq) also have "... = C (U (F (C (fixp_fun (\f. U (F (C f)))))))" by (subst mono_body_fixp[of "%f. U (F (C f))", OF mono]) (rule refl) also have "... = F (C (fixp_fun (\f. U (F (C f)))))" by (rule inverse) also have "... = F f" by (simp add: eq) finally show "f = F f" . qed text \Fixpoint induction rule\ lemma fixp_induct_uc: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a" and C :: "('b \ 'a) \ 'c" and P :: "('b \ 'a) \ bool" assumes mono: "\x. mono_body (\f. U (F (C f)) x)" and eq: "f \ C (fixp_fun (\f. U (F (C f))))" and inverse: "\f. U (C f) = f" and adm: "ccpo.admissible lub_fun le_fun P" and bot: "P (\_. lub {})" and step: "\f. P (U f) \ P (U (F f))" shows "P (U f)" unfolding eq inverse -apply (rule ccpo.fixp_induct[OF ccpo adm]) -apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2] -apply (rule_tac f5="C x" in step) -apply (simp add: inverse) -done +proof (rule ccpo.fixp_induct[OF ccpo adm]) + show "monotone le_fun le_fun (\f. U (F (C f)))" + using mono by (auto simp: monotone_def fun_ord_def) +next + show "P (lub_fun {})" + by (auto simp: bot fun_lub_def) +next + fix x + assume "P x" + then show "P (U (F (C x)))" + using step[of "C x"] by (simp add: inverse) +qed text \Rules for \<^term>\mono_body\:\ lemma const_mono[partial_function_mono]: "monotone ord leq (\f. c)" by (rule monotoneI) (rule leq_refl) end subsection \Flat interpretation: tailrec and option\ definition "flat_ord b x y \ x = b \ x = y" definition "flat_lub b A = (if A \ {b} then b else (THE x. x \ A - {b}))" lemma flat_interpretation: "partial_function_definitions (flat_ord b) (flat_lub b)" proof fix A x assume 1: "chain (flat_ord b) A" "x \ A" show "flat_ord b x (flat_lub b A)" proof cases assume "x = b" thus ?thesis by (simp add: flat_ord_def) next assume "x \ b" with 1 have "A - {b} = {x}" by (auto elim: chainE simp: flat_ord_def) then have "flat_lub b A = x" by (auto simp: flat_lub_def) thus ?thesis by (auto simp: flat_ord_def) qed next fix A z assume A: "chain (flat_ord b) A" and z: "\x. x \ A \ flat_ord b x z" show "flat_ord b (flat_lub b A) z" proof cases assume "A \ {b}" thus ?thesis by (auto simp: flat_lub_def flat_ord_def) next assume nb: "\ A \ {b}" then obtain y where y: "y \ A" "y \ b" by auto with A have "A - {b} = {y}" by (auto elim: chainE simp: flat_ord_def) with nb have "flat_lub b A = y" by (auto simp: flat_lub_def) with z y show ?thesis by auto qed qed (auto simp: flat_ord_def) lemma flat_ordI: "(x \ a \ x = y) \ flat_ord a x y" by(auto simp add: flat_ord_def) lemma flat_ord_antisym: "\ flat_ord a x y; flat_ord a y x \ \ x = y" by(auto simp add: flat_ord_def) lemma antisymp_flat_ord: "antisymp (flat_ord a)" by(rule antisympI)(auto dest: flat_ord_antisym) interpretation tailrec: partial_function_definitions "flat_ord undefined" "flat_lub undefined" rewrites "flat_lub undefined {} \ undefined" by (rule flat_interpretation)(simp add: flat_lub_def) interpretation option: partial_function_definitions "flat_ord None" "flat_lub None" rewrites "flat_lub None {} \ None" by (rule flat_interpretation)(simp add: flat_lub_def) abbreviation "tailrec_ord \ flat_ord undefined" abbreviation "mono_tailrec \ monotone (fun_ord tailrec_ord) tailrec_ord" lemma tailrec_admissible: "ccpo.admissible (fun_lub (flat_lub c)) (fun_ord (flat_ord c)) (\a. \x. a x \ c \ P x (a x))" proof(intro ccpo.admissibleI strip) fix A x assume chain: "Complete_Partial_Order.chain (fun_ord (flat_ord c)) A" and P [rule_format]: "\f\A. \x. f x \ c \ P x (f x)" and defined: "fun_lub (flat_lub c) A x \ c" from defined obtain f where f: "f \ A" "f x \ c" by(auto simp add: fun_lub_def flat_lub_def split: if_split_asm) hence "P x (f x)" by(rule P) moreover from chain f have "\f' \ A. f' x = c \ f' x = f x" by(auto 4 4 simp add: Complete_Partial_Order.chain_def flat_ord_def fun_ord_def) hence "fun_lub (flat_lub c) A x = f x" using f by(auto simp add: fun_lub_def flat_lub_def) ultimately show "P x (fun_lub (flat_lub c) A x)" by simp qed lemma fixp_induct_tailrec: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a" and C :: "('b \ 'a) \ 'c" and P :: "'b \ 'a \ bool" and x :: "'b" assumes mono: "\x. monotone (fun_ord (flat_ord c)) (flat_ord c) (\f. U (F (C f)) x)" assumes eq: "f \ C (ccpo.fixp (fun_lub (flat_lub c)) (fun_ord (flat_ord c)) (\f. U (F (C f))))" assumes inverse2: "\f. U (C f) = f" assumes step: "\f x y. (\x y. U f x = y \ y \ c \ P x y) \ U (F f) x = y \ y \ c \ P x y" assumes result: "U f x = y" assumes defined: "y \ c" shows "P x y" proof - have "\x y. U f x = y \ y \ c \ P x y" by(rule partial_function_definitions.fixp_induct_uc[OF flat_interpretation, of _ U F C, OF mono eq inverse2]) (auto intro: step tailrec_admissible simp add: fun_lub_def flat_lub_def) thus ?thesis using result defined by blast qed lemma admissible_image: assumes pfun: "partial_function_definitions le lub" assumes adm: "ccpo.admissible lub le (P \ g)" assumes inj: "\x y. f x = f y \ x = y" assumes inv: "\x. f (g x) = x" shows "ccpo.admissible (img_lub f g lub) (img_ord f le) P" proof (rule ccpo.admissibleI) fix A assume "chain (img_ord f le) A" then have ch': "chain le (f ` A)" by (auto simp: img_ord_def intro: chainI dest: chainD) assume "A \ {}" assume P_A: "\x\A. P x" have "(P \ g) (lub (f ` A))" using adm ch' proof (rule ccpo.admissibleD) fix x assume "x \ f ` A" with P_A show "(P \ g) x" by (auto simp: inj[OF inv]) qed(simp add: \A \ {}\) thus "P (img_lub f g lub A)" unfolding img_lub_def by simp qed lemma admissible_fun: assumes pfun: "partial_function_definitions le lub" assumes adm: "\x. ccpo.admissible lub le (Q x)" shows "ccpo.admissible (fun_lub lub) (fun_ord le) (\f. \x. Q x (f x))" proof (rule ccpo.admissibleI) fix A :: "('b \ 'a) set" assume Q: "\f\A. \x. Q x (f x)" assume ch: "chain (fun_ord le) A" assume "A \ {}" hence non_empty: "\a. {y. \f\A. y = f a} \ {}" by auto show "\x. Q x (fun_lub lub A x)" unfolding fun_lub_def by (rule allI, rule ccpo.admissibleD[OF adm chain_fun[OF ch] non_empty]) (auto simp: Q) qed abbreviation "option_ord \ flat_ord None" abbreviation "mono_option \ monotone (fun_ord option_ord) option_ord" lemma bind_mono[partial_function_mono]: assumes mf: "mono_option B" and mg: "\y. mono_option (\f. C y f)" shows "mono_option (\f. Option.bind (B f) (\y. C y f))" proof (rule monotoneI) fix f g :: "'a \ 'b option" assume fg: "fun_ord option_ord f g" with mf have "option_ord (B f) (B g)" by (rule monotoneD[of _ _ _ f g]) then have "option_ord (Option.bind (B f) (\y. C y f)) (Option.bind (B g) (\y. C y f))" unfolding flat_ord_def by auto also from mg have "\y'. option_ord (C y' f) (C y' g)" by (rule monotoneD) (rule fg) then have "option_ord (Option.bind (B g) (\y'. C y' f)) (Option.bind (B g) (\y'. C y' g))" unfolding flat_ord_def by (cases "B g") auto finally (option.leq_trans) show "option_ord (Option.bind (B f) (\y. C y f)) (Option.bind (B g) (\y'. C y' g))" . qed lemma flat_lub_in_chain: assumes ch: "chain (flat_ord b) A " assumes lub: "flat_lub b A = a" shows "a = b \ a \ A" proof (cases "A \ {b}") case True then have "flat_lub b A = b" unfolding flat_lub_def by simp with lub show ?thesis by simp next case False then obtain c where "c \ A" and "c \ b" by auto { fix z assume "z \ A" from chainD[OF ch \c \ A\ this] have "z = c \ z = b" unfolding flat_ord_def using \c \ b\ by auto } with False have "A - {b} = {c}" by auto with False have "flat_lub b A = c" by (auto simp: flat_lub_def) with \c \ A\ lub show ?thesis by simp qed lemma option_admissible: "option.admissible (%(f::'a \ 'b option). (\x y. f x = Some y \ P x y))" proof (rule ccpo.admissibleI) fix A :: "('a \ 'b option) set" assume ch: "chain option.le_fun A" and IH: "\f\A. \x y. f x = Some y \ P x y" from ch have ch': "\x. chain option_ord {y. \f\A. y = f x}" by (rule chain_fun) show "\x y. option.lub_fun A x = Some y \ P x y" proof (intro allI impI) fix x y assume "option.lub_fun A x = Some y" from flat_lub_in_chain[OF ch' this[unfolded fun_lub_def]] have "Some y \ {y. \f\A. y = f x}" by simp then have "\f\A. f x = Some y" by auto with IH show "P x y" by auto qed qed lemma fixp_induct_option: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a option" and C :: "('b \ 'a option) \ 'c" and P :: "'b \ 'a \ bool" assumes mono: "\x. mono_option (\f. U (F (C f)) x)" assumes eq: "f \ C (ccpo.fixp (fun_lub (flat_lub None)) (fun_ord option_ord) (\f. U (F (C f))))" assumes inverse2: "\f. U (C f) = f" assumes step: "\f x y. (\x y. U f x = Some y \ P x y) \ U (F f) x = Some y \ P x y" assumes defined: "U f x = Some y" shows "P x y" using step defined option.fixp_induct_uc[of U F C, OF mono eq inverse2 option_admissible] unfolding fun_lub_def flat_lub_def by(auto 9 2) declaration \Partial_Function.init "tailrec" \<^term>\tailrec.fixp_fun\ \<^term>\tailrec.mono_body\ @{thm tailrec.fixp_rule_uc} @{thm tailrec.fixp_induct_uc} (SOME @{thm fixp_induct_tailrec[where c = undefined]})\ declaration \Partial_Function.init "option" \<^term>\option.fixp_fun\ \<^term>\option.mono_body\ @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc} (SOME @{thm fixp_induct_option})\ hide_const (open) chain end diff --git a/src/HOL/Power.thy b/src/HOL/Power.thy --- a/src/HOL/Power.thy +++ b/src/HOL/Power.thy @@ -1,1025 +1,1016 @@ (* Title: HOL/Power.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1997 University of Cambridge *) section \Exponentiation\ theory Power imports Num begin subsection \Powers for Arbitrary Monoids\ class power = one + times begin primrec power :: "'a \ nat \ 'a" (infixr "^" 80) where power_0: "a ^ 0 = 1" | power_Suc: "a ^ Suc n = a * a ^ n" notation (latex output) power ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \Special syntax for squares.\ abbreviation power2 :: "'a \ 'a" ("(_\<^sup>2)" [1000] 999) where "x\<^sup>2 \ x ^ 2" end context includes lifting_syntax begin lemma power_transfer [transfer_rule]: \(R ===> (=) ===> R) (^) (^)\ if [transfer_rule]: \R 1 1\ \(R ===> R ===> R) (*) (*)\ for R :: \'a::power \ 'b::power \ bool\ by (simp only: power_def [abs_def]) transfer_prover end context monoid_mult begin subclass power . lemma power_one [simp]: "1 ^ n = 1" by (induct n) simp_all lemma power_one_right [simp]: "a ^ 1 = a" by simp lemma power_Suc0_right [simp]: "a ^ Suc 0 = a" by simp lemma power_commutes: "a ^ n * a = a * a ^ n" by (induct n) (simp_all add: mult.assoc) lemma power_Suc2: "a ^ Suc n = a ^ n * a" by (simp add: power_commutes) lemma power_add: "a ^ (m + n) = a ^ m * a ^ n" by (induct m) (simp_all add: algebra_simps) lemma power_mult: "a ^ (m * n) = (a ^ m) ^ n" by (induct n) (simp_all add: power_add) lemma power_even_eq: "a ^ (2 * n) = (a ^ n)\<^sup>2" by (subst mult.commute) (simp add: power_mult) lemma power_odd_eq: "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2" by (simp add: power_even_eq) lemma power_numeral_even: "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)" by (simp only: numeral_Bit0 power_add Let_def) lemma power_numeral_odd: "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" by (simp only: numeral_Bit1 One_nat_def add_Suc_right add_0_right power_Suc power_add Let_def mult.assoc) lemma power2_eq_square: "a\<^sup>2 = a * a" by (simp add: numeral_2_eq_2) lemma power3_eq_cube: "a ^ 3 = a * a * a" by (simp add: numeral_3_eq_3 mult.assoc) lemma power4_eq_xxxx: "x^4 = x * x * x * x" by (simp add: mult.assoc power_numeral_even) lemma funpow_times_power: "(times x ^^ f x) = times (x ^ f x)" proof (induct "f x" arbitrary: f) case 0 then show ?case by (simp add: fun_eq_iff) next case (Suc n) define g where "g x = f x - 1" for x with Suc have "n = g x" by simp with Suc have "times x ^^ g x = times (x ^ g x)" by simp moreover from Suc g_def have "f x = g x + 1" by simp ultimately show ?case by (simp add: power_add funpow_add fun_eq_iff mult.assoc) qed lemma power_commuting_commutes: assumes "x * y = y * x" shows "x ^ n * y = y * x ^n" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "x ^ Suc n * y = x ^ n * y * x" by (subst power_Suc2) (simp add: assms ac_simps) also have "\ = y * x ^ Suc n" by (simp only: Suc power_Suc2) (simp add: ac_simps) finally show ?case . qed lemma power_minus_mult: "0 < n \ a ^ (n - 1) * a = a ^ n" by (simp add: power_commutes split: nat_diff_split) lemma left_right_inverse_power: assumes "x * y = 1" shows "x ^ n * y ^ n = 1" proof (induct n) case (Suc n) moreover have "x ^ Suc n * y ^ Suc n = x^n * (x * y) * y^n" by (simp add: power_Suc2[symmetric] mult.assoc[symmetric]) ultimately show ?case by (simp add: assms) qed simp end context comm_monoid_mult begin lemma power_mult_distrib [algebra_simps, algebra_split_simps, field_simps, field_split_simps, divide_simps]: "(a * b) ^ n = (a ^ n) * (b ^ n)" by (induction n) (simp_all add: ac_simps) end text \Extract constant factors from powers.\ declare power_mult_distrib [where a = "numeral w" for w, simp] declare power_mult_distrib [where b = "numeral w" for w, simp] lemma power_add_numeral [simp]: "a^numeral m * a^numeral n = a^numeral (m + n)" for a :: "'a::monoid_mult" by (simp add: power_add [symmetric]) lemma power_add_numeral2 [simp]: "a^numeral m * (a^numeral n * b) = a^numeral (m + n) * b" for a :: "'a::monoid_mult" by (simp add: mult.assoc [symmetric]) lemma power_mult_numeral [simp]: "(a^numeral m)^numeral n = a^numeral (m * n)" for a :: "'a::monoid_mult" by (simp only: numeral_mult power_mult) context semiring_numeral begin lemma numeral_sqr: "numeral (Num.sqr k) = numeral k * numeral k" by (simp only: sqr_conv_mult numeral_mult) lemma numeral_pow: "numeral (Num.pow k l) = numeral k ^ numeral l" by (induct l) (simp_all only: numeral_class.numeral.simps pow.simps numeral_sqr numeral_mult power_add power_one_right) lemma power_numeral [simp]: "numeral k ^ numeral l = numeral (Num.pow k l)" by (rule numeral_pow [symmetric]) end context semiring_1 begin lemma of_nat_power [simp]: "of_nat (m ^ n) = of_nat m ^ n" by (induct n) simp_all lemma zero_power: "0 < n \ 0 ^ n = 0" by (cases n) simp_all lemma power_zero_numeral [simp]: "0 ^ numeral k = 0" by (simp add: numeral_eq_Suc) lemma zero_power2: "0\<^sup>2 = 0" (* delete? *) by (rule power_zero_numeral) lemma one_power2: "1\<^sup>2 = 1" (* delete? *) by (rule power_one) lemma power_0_Suc [simp]: "0 ^ Suc n = 0" by simp text \It looks plausible as a simprule, but its effect can be strange.\ lemma power_0_left: "0 ^ n = (if n = 0 then 1 else 0)" by (cases n) simp_all end context semiring_char_0 begin lemma numeral_power_eq_of_nat_cancel_iff [simp]: "numeral x ^ n = of_nat y \ numeral x ^ n = y" using of_nat_eq_iff by fastforce lemma real_of_nat_eq_numeral_power_cancel_iff [simp]: "of_nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_nat_cancel_iff [of x n y] by (metis (mono_tags)) lemma of_nat_eq_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w = of_nat x \ b ^ w = x" by (metis of_nat_power of_nat_eq_iff) lemma of_nat_power_eq_of_nat_cancel_iff[simp]: "of_nat x = (of_nat b) ^ w \ x = b ^ w" by (metis of_nat_eq_of_nat_power_cancel_iff) end context comm_semiring_1 begin text \The divides relation.\ lemma le_imp_power_dvd: assumes "m \ n" shows "a ^ m dvd a ^ n" proof from assms have "a ^ n = a ^ (m + (n - m))" by simp also have "\ = a ^ m * a ^ (n - m)" by (rule power_add) finally show "a ^ n = a ^ m * a ^ (n - m)" . qed lemma power_le_dvd: "a ^ n dvd b \ m \ n \ a ^ m dvd b" by (rule dvd_trans [OF le_imp_power_dvd]) lemma dvd_power_same: "x dvd y \ x ^ n dvd y ^ n" by (induct n) (auto simp add: mult_dvd_mono) lemma dvd_power_le: "x dvd y \ m \ n \ x ^ n dvd y ^ m" by (rule power_le_dvd [OF dvd_power_same]) lemma dvd_power [simp]: fixes n :: nat assumes "n > 0 \ x = 1" shows "x dvd (x ^ n)" using assms proof assume "0 < n" then have "x ^ n = x ^ Suc (n - 1)" by simp then show "x dvd (x ^ n)" by simp next assume "x = 1" then show "x dvd (x ^ n)" by simp qed end context semiring_1_no_zero_divisors begin subclass power . lemma power_eq_0_iff [simp]: "a ^ n = 0 \ a = 0 \ n > 0" by (induct n) auto lemma power_not_zero: "a \ 0 \ a ^ n \ 0" by (induct n) auto lemma zero_eq_power2 [simp]: "a\<^sup>2 = 0 \ a = 0" unfolding power2_eq_square by simp end context ring_1 begin lemma power_minus: "(- a) ^ n = (- 1) ^ n * a ^ n" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp del: power_Suc add: power_Suc2 mult.assoc) qed lemma power_minus': "NO_MATCH 1 x \ (-x) ^ n = (-1)^n * x ^ n" by (rule power_minus) lemma power_minus_Bit0: "(- x) ^ numeral (Num.Bit0 k) = x ^ numeral (Num.Bit0 k)" by (induct k, simp_all only: numeral_class.numeral.simps power_add power_one_right mult_minus_left mult_minus_right minus_minus) lemma power_minus_Bit1: "(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))" by (simp only: eval_nat_numeral(3) power_Suc power_minus_Bit0 mult_minus_left) lemma power2_minus [simp]: "(- a)\<^sup>2 = a\<^sup>2" by (fact power_minus_Bit0) lemma power_minus1_even [simp]: "(- 1) ^ (2*n) = 1" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp add: power_add power2_eq_square) qed lemma power_minus1_odd: "(- 1) ^ Suc (2*n) = -1" by simp lemma power_minus_even [simp]: "(-a) ^ (2*n) = a ^ (2*n)" by (simp add: power_minus [of a]) end context ring_1_no_zero_divisors begin lemma power2_eq_1_iff: "a\<^sup>2 = 1 \ a = 1 \ a = - 1" using square_eq_1_iff [of a] by (simp add: power2_eq_square) end context idom begin lemma power2_eq_iff: "x\<^sup>2 = y\<^sup>2 \ x = y \ x = - y" unfolding power2_eq_square by (rule square_eq_iff) end context semidom_divide begin lemma power_diff: "a ^ (m - n) = (a ^ m) div (a ^ n)" if "a \ 0" and "n \ m" proof - define q where "q = m - n" with \n \ m\ have "m = q + n" by simp with \a \ 0\ q_def show ?thesis by (simp add: power_add) qed end context algebraic_semidom begin lemma div_power: "b dvd a \ (a div b) ^ n = a ^ n div b ^ n" by (induct n) (simp_all add: div_mult_div_if_dvd dvd_power_same) lemma is_unit_power_iff: "is_unit (a ^ n) \ is_unit a \ n = 0" by (induct n) (auto simp add: is_unit_mult_iff) lemma dvd_power_iff: assumes "x \ 0" shows "x ^ m dvd x ^ n \ is_unit x \ m \ n" proof assume *: "x ^ m dvd x ^ n" { assume "m > n" note * also have "x ^ n = x ^ n * 1" by simp also from \m > n\ have "m = n + (m - n)" by simp also have "x ^ \ = x ^ n * x ^ (m - n)" by (rule power_add) finally have "x ^ (m - n) dvd 1" - by (subst (asm) dvd_times_left_cancel_iff) (insert assms, simp_all) + using assms by (subst (asm) dvd_times_left_cancel_iff) simp_all with \m > n\ have "is_unit x" by (simp add: is_unit_power_iff) } thus "is_unit x \ m \ n" by force qed (auto intro: unit_imp_dvd simp: is_unit_power_iff le_imp_power_dvd) end context normalization_semidom_multiplicative begin lemma normalize_power: "normalize (a ^ n) = normalize a ^ n" by (induct n) (simp_all add: normalize_mult) lemma unit_factor_power: "unit_factor (a ^ n) = unit_factor a ^ n" by (induct n) (simp_all add: unit_factor_mult) end context division_ring begin text \Perhaps these should be simprules.\ lemma power_inverse [field_simps, field_split_simps, divide_simps]: "inverse a ^ n = inverse (a ^ n)" proof (cases "a = 0") case True then show ?thesis by (simp add: power_0_left) next case False then have "inverse (a ^ n) = inverse a ^ n" by (induct n) (simp_all add: nonzero_inverse_mult_distrib power_commutes) then show ?thesis by simp qed lemma power_one_over [field_simps, field_split_simps, divide_simps]: "(1 / a) ^ n = 1 / a ^ n" using power_inverse [of a] by (simp add: divide_inverse) end context field begin lemma power_divide [field_simps, field_split_simps, divide_simps]: "(a / b) ^ n = a ^ n / b ^ n" by (induct n) simp_all end subsection \Exponentiation on ordered types\ context linordered_semidom begin lemma zero_less_power [simp]: "0 < a \ 0 < a ^ n" by (induct n) simp_all lemma zero_le_power [simp]: "0 \ a \ 0 \ a ^ n" by (induct n) simp_all lemma power_mono: "a \ b \ 0 \ a \ a ^ n \ b ^ n" by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" using power_mono [of 1 a n] by simp lemma power_le_one: "0 \ a \ a \ 1 \ a ^ n \ 1" using power_mono [of a 1 n] by simp lemma power_gt1_lemma: assumes gt1: "1 < a" shows "1 < a * a ^ n" proof - from gt1 have "0 \ a" by (fact order_trans [OF zero_le_one less_imp_le]) from gt1 have "1 * 1 < a * 1" by simp also from gt1 have "\ \ a * a ^ n" by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed lemma power_gt1: "1 < a \ 1 < a ^ Suc n" by (simp add: power_gt1_lemma) lemma one_less_power [simp]: "1 < a \ 0 < n \ 1 < a ^ n" by (cases n) (simp_all add: power_gt1_lemma) lemma power_le_imp_le_exp: assumes gt1: "1 < a" shows "a ^ m \ a ^ n \ m \ n" proof (induct m arbitrary: n) case 0 show ?case by simp next case (Suc m) show ?case proof (cases n) case 0 with Suc have "a * a ^ m \ 1" by simp with gt1 show ?thesis by (force simp only: power_gt1_lemma not_less [symmetric]) next case (Suc n) with Suc.prems Suc.hyps show ?thesis by (force dest: mult_left_le_imp_le simp add: less_trans [OF zero_less_one gt1]) qed qed lemma of_nat_zero_less_power_iff [simp]: "of_nat x ^ n > 0 \ x > 0 \ n = 0" by (induct n) auto text \Surely we can strengthen this? It holds for \0 too.\ lemma power_inject_exp [simp]: \a ^ m = a ^ n \ m = n\ if \1 < a\ using that by (force simp add: order_class.order.antisym power_le_imp_le_exp) text \ Can relax the first premise to \<^term>\0 in the case of the natural numbers. \ lemma power_less_imp_less_exp: "1 < a \ a ^ m < a ^ n \ m < n" by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] power_le_imp_le_exp) - -lemma power_strict_mono [rule_format]: "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" - by (induct n) (auto simp: mult_strict_mono le_less_trans [of 0 a b]) + +lemma power_strict_mono: "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" +proof (induct n) + case 0 + then show ?case by simp +next + case (Suc n) + then show ?case + by (cases "n = 0") (auto simp: mult_strict_mono le_less_trans [of 0 a b]) +qed lemma power_mono_iff [simp]: shows "\a \ 0; b \ 0; n>0\ \ a ^ n \ b ^ n \ a \ b" using power_mono [of a b] power_strict_mono [of b a] not_le by auto text\Lemma for \power_strict_decreasing\\ lemma power_Suc_less: "0 < a \ a < 1 \ a * a ^ n < a ^ n" by (induct n) (auto simp: mult_strict_left_mono) -lemma power_strict_decreasing [rule_format]: "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" -proof (induct N) +lemma power_strict_decreasing: "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" +proof (induction N) + case 0 + then show ?case by simp + next + case (Suc N) + then show ?case + using mult_strict_mono[of a 1 "a ^ N" "a ^ n"] + by (auto simp add: power_Suc_less less_Suc_eq) + qed + +text \Proof resembles that of \power_strict_decreasing\.\ +lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" +proof (induction N) case 0 then show ?case by simp next case (Suc N) then show ?case - apply (auto simp add: power_Suc_less less_Suc_eq) - apply (subgoal_tac "a * a^N < 1 * a^n") - apply simp - apply (rule mult_strict_mono) - apply auto - done -qed - -text \Proof resembles that of \power_strict_decreasing\.\ -lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" -proof (induct N) - case 0 - then show ?case by simp -next - case (Suc N) - then show ?case - apply (auto simp add: le_Suc_eq) - apply (subgoal_tac "a * a^N \ 1 * a^n") - apply simp - apply (rule mult_mono) - apply auto - done + using mult_mono[of a 1 "a^N" "a ^ n"] + by (auto simp add: le_Suc_eq) qed lemma power_decreasing_iff [simp]: "\0 < b; b < 1\ \ b ^ m \ b ^ n \ n \ m" using power_strict_decreasing [of m n b] by (auto intro: power_decreasing ccontr) lemma power_strict_decreasing_iff [simp]: "\0 < b; b < 1\ \ b ^ m < b ^ n \ n < m" using power_decreasing_iff [of b m n] unfolding le_less by (auto dest: power_strict_decreasing le_neq_implies_less) lemma power_Suc_less_one: "0 < a \ a < 1 \ a ^ Suc n < 1" using power_strict_decreasing [of 0 "Suc n" a] by simp text \Proof again resembles that of \power_strict_decreasing\.\ lemma power_increasing: "n \ N \ 1 \ a \ a ^ n \ a ^ N" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case - apply (auto simp add: le_Suc_eq) - apply (subgoal_tac "1 * a^n \ a * a^N") - apply simp - apply (rule mult_mono) - apply (auto simp add: order_trans [OF zero_le_one]) - done + using mult_mono[of 1 a "a ^ n" "a ^ N"] + by (auto simp add: le_Suc_eq order_trans [OF zero_le_one]) qed text \Lemma for \power_strict_increasing\.\ lemma power_less_power_Suc: "1 < a \ a ^ n < a * a ^ n" by (induct n) (auto simp: mult_strict_left_mono less_trans [OF zero_less_one]) lemma power_strict_increasing: "n < N \ 1 < a \ a ^ n < a ^ N" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case - apply (auto simp add: power_less_power_Suc less_Suc_eq) - apply (subgoal_tac "1 * a^n < a * a^N") - apply simp - apply (rule mult_strict_mono) - apply (auto simp add: less_trans [OF zero_less_one] less_imp_le) - done + using mult_strict_mono[of 1 a "a^n" "a^N"] + by (auto simp add: power_less_power_Suc less_Suc_eq less_trans [OF zero_less_one] less_imp_le) qed lemma power_increasing_iff [simp]: "1 < b \ b ^ x \ b ^ y \ x \ y" by (blast intro: power_le_imp_le_exp power_increasing less_imp_le) lemma power_strict_increasing_iff [simp]: "1 < b \ b ^ x < b ^ y \ x < y" by (blast intro: power_less_imp_less_exp power_strict_increasing) lemma power_le_imp_le_base: assumes le: "a ^ Suc n \ b ^ Suc n" and "0 \ b" shows "a \ b" proof (rule ccontr) assume "\ ?thesis" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" by (simp only: assms(2) power_strict_mono) with le show False by (simp add: linorder_not_less [symmetric]) qed lemma power_less_imp_less_base: assumes less: "a ^ n < b ^ n" assumes nonneg: "0 \ b" shows "a < b" proof (rule contrapos_pp [OF less]) assume "\ ?thesis" then have "b \ a" by (simp only: linorder_not_less) from this nonneg have "b ^ n \ a ^ n" by (rule power_mono) then show "\ a ^ n < b ^ n" by (simp only: linorder_not_less) qed lemma power_inject_base: "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" by (blast intro: power_le_imp_le_base order.antisym eq_refl sym) lemma power_eq_imp_eq_base: "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" by (cases n) (simp_all del: power_Suc, rule power_inject_base) lemma power_eq_iff_eq_base: "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" using power_eq_imp_eq_base [of a n b] by auto lemma power2_le_imp_le: "x\<^sup>2 \ y\<^sup>2 \ 0 \ y \ x \ y" unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) lemma power2_less_imp_less: "x\<^sup>2 < y\<^sup>2 \ 0 \ y \ x < y" by (rule power_less_imp_less_base) lemma power2_eq_imp_eq: "x\<^sup>2 = y\<^sup>2 \ 0 \ x \ 0 \ y \ x = y" unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp lemma power_Suc_le_self: "0 \ a \ a \ 1 \ a ^ Suc n \ a" using power_decreasing [of 1 "Suc n" a] by simp lemma power2_eq_iff_nonneg [simp]: assumes "0 \ x" "0 \ y" shows "(x ^ 2 = y ^ 2) \ x = y" using assms power2_eq_imp_eq by blast lemma of_nat_less_numeral_power_cancel_iff[simp]: "of_nat x < numeral i ^ n \ x < numeral i ^ n" using of_nat_less_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] . lemma of_nat_le_numeral_power_cancel_iff[simp]: "of_nat x \ numeral i ^ n \ x \ numeral i ^ n" using of_nat_le_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] . lemma numeral_power_less_of_nat_cancel_iff[simp]: "numeral i ^ n < of_nat x \ numeral i ^ n < x" using of_nat_less_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] . lemma numeral_power_le_of_nat_cancel_iff[simp]: "numeral i ^ n \ of_nat x \ numeral i ^ n \ x" using of_nat_le_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] . lemma of_nat_le_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w \ of_nat x \ b ^ w \ x" by (metis of_nat_le_iff of_nat_power) lemma of_nat_power_le_of_nat_cancel_iff[simp]: "of_nat x \ (of_nat b) ^ w \ x \ b ^ w" by (metis of_nat_le_iff of_nat_power) lemma of_nat_less_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w < of_nat x \ b ^ w < x" by (metis of_nat_less_iff of_nat_power) lemma of_nat_power_less_of_nat_cancel_iff[simp]: "of_nat x < (of_nat b) ^ w \ x < b ^ w" by (metis of_nat_less_iff of_nat_power) end text \Some @{typ nat}-specific lemmas:\ lemma mono_ge2_power_minus_self: assumes "k \ 2" shows "mono (\m. k ^ m - m)" unfolding mono_iff_le_Suc proof fix n have "k ^ n < k ^ Suc n" using power_strict_increasing_iff[of k "n" "Suc n"] assms by linarith thus "k ^ n - n \ k ^ Suc n - Suc n" by linarith qed lemma self_le_ge2_pow[simp]: assumes "k \ 2" shows "m \ k ^ m" proof (induction m) case 0 show ?case by simp next case (Suc m) hence "Suc m \ Suc (k ^ m)" by simp also have "... \ k^m + k^m" using one_le_power[of k m] assms by linarith also have "... \ k * k^m" by (metis mult_2 mult_le_mono1[OF assms]) finally show ?case by simp qed lemma diff_le_diff_pow[simp]: assumes "k \ 2" shows "m - n \ k ^ m - k ^ n" proof (cases "n \ m") case True thus ?thesis using monoD[OF mono_ge2_power_minus_self[OF assms] True] self_le_ge2_pow[OF assms, of m] by (simp add: le_diff_conv le_diff_conv2) qed auto context linordered_ring_strict begin lemma sum_squares_eq_zero_iff: "x * x + y * y = 0 \ x = 0 \ y = 0" by (simp add: add_nonneg_eq_0_iff) lemma sum_squares_le_zero_iff: "x * x + y * y \ 0 \ x = 0 \ y = 0" by (simp add: le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) lemma sum_squares_gt_zero_iff: "0 < x * x + y * y \ x \ 0 \ y \ 0" by (simp add: not_le [symmetric] sum_squares_le_zero_iff) end context linordered_idom begin lemma zero_le_power2 [simp]: "0 \ a\<^sup>2" by (simp add: power2_eq_square) lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \ a \ 0" by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff) lemma power2_less_0 [simp]: "\ a\<^sup>2 < 0" by (force simp add: power2_eq_square mult_less_0_iff) lemma power_abs: "\a ^ n\ = \a\ ^ n" \ \FIXME simp?\ by (induct n) (simp_all add: abs_mult) lemma power_sgn [simp]: "sgn (a ^ n) = sgn a ^ n" by (induct n) (simp_all add: sgn_mult) lemma abs_power_minus [simp]: "\(- a) ^ n\ = \a ^ n\" by (simp add: power_abs) lemma zero_less_power_abs_iff [simp]: "0 < \a\ ^ n \ a \ 0 \ n = 0" proof (induct n) case 0 show ?case by simp next case Suc then show ?case by (auto simp: zero_less_mult_iff) qed lemma zero_le_power_abs [simp]: "0 \ \a\ ^ n" by (rule zero_le_power [OF abs_ge_zero]) lemma power2_less_eq_zero_iff [simp]: "a\<^sup>2 \ 0 \ a = 0" by (simp add: le_less) lemma abs_power2 [simp]: "\a\<^sup>2\ = a\<^sup>2" by (simp add: power2_eq_square) lemma power2_abs [simp]: "\a\\<^sup>2 = a\<^sup>2" by (simp add: power2_eq_square) lemma odd_power_less_zero: "a < 0 \ a ^ Suc (2 * n) < 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" by (simp add: ac_simps power_add power2_eq_square) then show ?case by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg) qed lemma odd_0_le_power_imp_0_le: "0 \ a ^ Suc (2 * n) \ 0 \ a" using odd_power_less_zero [of a n] by (force simp add: linorder_not_less [symmetric]) lemma zero_le_even_power'[simp]: "0 \ a ^ (2 * n)" proof (induct n) case 0 show ?case by simp next case (Suc n) have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" by (simp add: ac_simps power_add power2_eq_square) then show ?case by (simp add: Suc zero_le_mult_iff) qed lemma sum_power2_ge_zero: "0 \ x\<^sup>2 + y\<^sup>2" by (intro add_nonneg_nonneg zero_le_power2) lemma not_sum_power2_lt_zero: "\ x\<^sup>2 + y\<^sup>2 < 0" unfolding not_less by (rule sum_power2_ge_zero) lemma sum_power2_eq_zero_iff: "x\<^sup>2 + y\<^sup>2 = 0 \ x = 0 \ y = 0" unfolding power2_eq_square by (simp add: add_nonneg_eq_0_iff) lemma sum_power2_le_zero_iff: "x\<^sup>2 + y\<^sup>2 \ 0 \ x = 0 \ y = 0" by (simp add: le_less sum_power2_eq_zero_iff not_sum_power2_lt_zero) lemma sum_power2_gt_zero_iff: "0 < x\<^sup>2 + y\<^sup>2 \ x \ 0 \ y \ 0" unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff) lemma abs_le_square_iff: "\x\ \ \y\ \ x\<^sup>2 \ y\<^sup>2" (is "?lhs \ ?rhs") proof assume ?lhs then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono) simp then show ?rhs by simp next assume ?rhs then show ?lhs by (auto intro!: power2_le_imp_le [OF _ abs_ge_zero]) qed lemma power2_le_iff_abs_le: "y \ 0 \ x\<^sup>2 \ y\<^sup>2 \ \x\ \ y" by (metis abs_le_square_iff abs_of_nonneg) lemma abs_square_le_1:"x\<^sup>2 \ 1 \ \x\ \ 1" using abs_le_square_iff [of x 1] by simp lemma abs_square_eq_1: "x\<^sup>2 = 1 \ \x\ = 1" by (auto simp add: abs_if power2_eq_1_iff) lemma abs_square_less_1: "x\<^sup>2 < 1 \ \x\ < 1" using abs_square_eq_1 [of x] abs_square_le_1 [of x] by (auto simp add: le_less) lemma square_le_1: assumes "- 1 \ x" "x \ 1" shows "x\<^sup>2 \ 1" using assms by (metis add.inverse_inverse linear mult_le_one neg_equal_0_iff_equal neg_le_iff_le power2_eq_square power_minus_Bit0) end subsection \Miscellaneous rules\ lemma (in linordered_semidom) self_le_power: "1 \ a \ 0 < n \ a \ a ^ n" using power_increasing [of 1 n a] power_one_right [of a] by auto lemma (in power) power_eq_if: "p ^ m = (if m=0 then 1 else p * (p ^ (m - 1)))" unfolding One_nat_def by (cases m) simp_all lemma (in comm_semiring_1) power2_sum: "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) context comm_ring_1 begin lemma power2_diff: "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) lemma power2_commute: "(x - y)\<^sup>2 = (y - x)\<^sup>2" by (simp add: algebra_simps power2_eq_square) lemma minus_power_mult_self: "(- a) ^ n * (- a) ^ n = a ^ (2 * n)" by (simp add: power_mult_distrib [symmetric]) (simp add: power2_eq_square [symmetric] power_mult [symmetric]) lemma minus_one_mult_self [simp]: "(- 1) ^ n * (- 1) ^ n = 1" using minus_power_mult_self [of 1 n] by simp lemma left_minus_one_mult_self [simp]: "(- 1) ^ n * ((- 1) ^ n * a) = a" by (simp add: mult.assoc [symmetric]) end text \Simprules for comparisons where common factors can be cancelled.\ lemmas zero_compare_simps = add_strict_increasing add_strict_increasing2 add_increasing zero_le_mult_iff zero_le_divide_iff zero_less_mult_iff zero_less_divide_iff mult_le_0_iff divide_le_0_iff mult_less_0_iff divide_less_0_iff zero_le_power2 power2_less_0 subsection \Exponentiation for the Natural Numbers\ lemma nat_one_le_power [simp]: "Suc 0 \ i \ Suc 0 \ i ^ n" by (rule one_le_power [of i n, unfolded One_nat_def]) lemma nat_zero_less_power_iff [simp]: "x ^ n > 0 \ x > 0 \ n = 0" for x :: nat by (induct n) auto lemma nat_power_eq_Suc_0_iff [simp]: "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" by (induct m) auto lemma power_Suc_0 [simp]: "Suc 0 ^ n = Suc 0" by simp text \ Valid for the naturals, but what if \0 < i < 1\? Premises cannot be weakened: consider the case where \i = 0\, \m = 1\ and \n = 0\. \ lemma nat_power_less_imp_less: fixes i :: nat assumes nonneg: "0 < i" assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") case True with less power_one [where 'a = nat] show ?thesis by simp next case False with nonneg have "1 < i" by auto from power_strict_increasing_iff [OF this] less show ?thesis .. qed lemma power_gt_expt: "n > Suc 0 \ n^k > k" by (induction k) (auto simp: less_trans_Suc n_less_m_mult_n) lemma less_exp [simp]: \n < 2 ^ n\ by (simp add: power_gt_expt) lemma power_dvd_imp_le: fixes i :: nat assumes "i ^ m dvd i ^ n" "1 < i" shows "m \ n" using assms by (auto intro: power_le_imp_le_exp [OF \1 < i\ dvd_imp_le]) lemma dvd_power_iff_le: fixes k::nat shows "2 \ k \ ((k ^ m) dvd (k ^ n) \ m \ n)" using le_imp_power_dvd power_dvd_imp_le by force lemma power2_nat_le_eq_le: "m\<^sup>2 \ n\<^sup>2 \ m \ n" for m n :: nat by (auto intro: power2_le_imp_le power_mono) lemma power2_nat_le_imp_le: fixes m n :: nat assumes "m\<^sup>2 \ n" shows "m \ n" proof (cases m) case 0 then show ?thesis by simp next case (Suc k) show ?thesis proof (rule ccontr) assume "\ ?thesis" then have "n < m" by simp with assms Suc show False by (simp add: power2_eq_square) qed qed lemma ex_power_ivl1: fixes b k :: nat assumes "b \ 2" shows "k \ 1 \ \n. b^n \ k \ k < b^(n+1)" (is "_ \ \n. ?P k n") proof(induction k) case 0 thus ?case by simp next case (Suc k) show ?case proof cases assume "k=0" hence "?P (Suc k) 0" using assms by simp thus ?case .. next assume "k\0" with Suc obtain n where IH: "?P k n" by auto show ?case proof (cases "k = b^(n+1) - 1") case True hence "?P (Suc k) (n+1)" using assms by (simp add: power_less_power_Suc) thus ?thesis .. next case False hence "?P (Suc k) n" using IH by auto thus ?thesis .. qed qed qed lemma ex_power_ivl2: fixes b k :: nat assumes "b \ 2" "k \ 2" shows "\n. b^n < k \ k \ b^(n+1)" proof - have "1 \ k - 1" using assms(2) by arith from ex_power_ivl1[OF assms(1) this] obtain n where "b ^ n \ k - 1 \ k - 1 < b ^ (n + 1)" .. hence "b^n < k \ k \ b^(n+1)" using assms by auto thus ?thesis .. qed subsubsection \Cardinality of the Powerset\ lemma card_UNIV_bool [simp]: "card (UNIV :: bool set) = 2" unfolding UNIV_bool by simp lemma card_Pow: "finite A \ card (Pow A) = 2 ^ card A" proof (induct rule: finite_induct) case empty show ?case by simp next case (insert x A) from \x \ A\ have disjoint: "Pow A \ insert x ` Pow A = {}" by blast from \x \ A\ have inj_on: "inj_on (insert x) (Pow A)" unfolding inj_on_def by auto have "card (Pow (insert x A)) = card (Pow A \ insert x ` Pow A)" by (simp only: Pow_insert) also have "\ = card (Pow A) + card (insert x ` Pow A)" by (rule card_Un_disjoint) (use \finite A\ disjoint in simp_all) also from inj_on have "card (insert x ` Pow A) = card (Pow A)" by (rule card_image) also have "\ + \ = 2 * \" by (simp add: mult_2) also from insert(3) have "\ = 2 ^ Suc (card A)" by simp also from insert(1,2) have "Suc (card A) = card (insert x A)" by (rule card_insert_disjoint [symmetric]) finally show ?case . qed subsection \Code generator tweak\ code_identifier code_module Power \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Presburger.thy b/src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy +++ b/src/HOL/Presburger.thy @@ -1,510 +1,566 @@ (* Title: HOL/Presburger.thy Author: Amine Chaieb, TU Muenchen *) section \Decision Procedure for Presburger Arithmetic\ theory Presburger imports Groebner_Basis Set_Interval keywords "try0" :: diag begin ML_file \Tools/Qelim/qelim.ML\ ML_file \Tools/Qelim/cooper_procedure.ML\ subsection\The \-\\ and \+\\ Properties\ lemma minf: "\\(z ::'a::linorder).\xz.\x \ \z.\x Q x) = (P' x \ Q' x)" "\\(z ::'a::linorder).\xz.\x \ \z.\x Q x) = (P' x \ Q' x)" "\(z ::'a::{linorder}).\x(z ::'a::{linorder}).\x t) = True" "\(z ::'a::{linorder}).\x(z ::'a::{linorder}).\x t) = True" "\(z ::'a::{linorder}).\x t) = False" "\(z ::'a::{linorder}).\x t) = False" "\z.\(x::'b::{linorder,plus,Rings.dvd})z.\(x::'b::{linorder,plus,Rings.dvd}) d dvd x + s) = (\ d dvd x + s)" "\z.\xxxx < min z1 z2. (P x \ Q x) = (P' x \ Q' x)" + by simp + then show "\z. \x Q x) = (P' x \ Q' x)" + by blast +next + fix z1 z2 + assume "\xxx < min z1 z2. (P x \ Q x) = (P' x \ Q' x)" + by simp + then show "\z. \x Q x) = (P' x \ Q' x)" + by blast +next + have "\x t" + by fastforce + then show "\z. \x t) = True" + by auto +next + have "\x t < x" + by fastforce + then show "\z. \xx t \ x" + by fastforce + then show "\z. \x x) = False" + by auto +qed auto lemma pinf: "\\(z ::'a::linorder).\x>z. P x = P' x; \z.\x>z. Q x = Q' x\ \ \z.\x>z. (P x \ Q x) = (P' x \ Q' x)" "\\(z ::'a::linorder).\x>z. P x = P' x; \z.\x>z. Q x = Q' x\ \ \z.\x>z. (P x \ Q x) = (P' x \ Q' x)" "\(z ::'a::{linorder}).\x>z.(x = t) = False" "\(z ::'a::{linorder}).\x>z.(x \ t) = True" "\(z ::'a::{linorder}).\x>z.(x < t) = False" "\(z ::'a::{linorder}).\x>z.(x \ t) = False" "\(z ::'a::{linorder}).\x>z.(x > t) = True" "\(z ::'a::{linorder}).\x>z.(x \ t) = True" "\z.\(x::'b::{linorder,plus,Rings.dvd})>z. (d dvd x + s) = (d dvd x + s)" "\z.\(x::'b::{linorder,plus,Rings.dvd})>z. (\ d dvd x + s) = (\ d dvd x + s)" "\z.\x>z. F = F" - by ((erule exE, erule exE,rule_tac x="max z za" in exI,simp)+,(rule_tac x="t" in exI,fastforce)+) simp_all +proof safe + fix z1 z2 + assume "\x>z1. P x = P' x" and "\x>z2. Q x = Q' x" + then have "\x > max z1 z2. (P x \ Q x) = (P' x \ Q' x)" + by simp + then show "\z. \x>z. (P x \ Q x) = (P' x \ Q' x)" + by blast +next + fix z1 z2 + assume "\x>z1. P x = P' x" and "\x>z2. Q x = Q' x" + then have "\x > max z1 z2. (P x \ Q x) = (P' x \ Q' x)" + by simp + then show "\z. \x>z. (P x \ Q x) = (P' x \ Q' x)" + by blast +next + have "\x>t. \ x < t" + by fastforce + then show "\z. \x>z. x < t = False" + by blast +next + have "\x>t. \ x \ t" + by fastforce + then show "\z. \x>z. x \ t = False" + by blast +next + have "\x>t. t \ x" + by fastforce + then show "\z. \x>z. t \ x = True" + by blast +qed auto lemma inf_period: "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" "(d::'a::{comm_ring,Rings.dvd}) dvd D \ \x k. (d dvd x + t) = (d dvd (x - k*D) + t)" "(d::'a::{comm_ring,Rings.dvd}) dvd D \ \x k. (\d dvd x + t) = (\d dvd (x - k*D) + t)" "\x k. F = F" apply (auto elim!: dvdE simp add: algebra_simps) unfolding mult.assoc [symmetric] distrib_right [symmetric] left_diff_distrib [symmetric] unfolding dvd_def mult.commute [of d] by auto subsection\The A and B sets\ lemma bset: "\\x.(\j \ {1 .. D}. \b\B. x \ b + j)\ P x \ P(x - D) ; \x.(\j\{1 .. D}. \b\B. x \ b + j)\ Q x \ Q(x - D)\ \ \x.(\j\{1 .. D}. \b\B. x \ b + j) \ (P x \ Q x) \ (P(x - D) \ Q (x - D))" "\\x.(\j\{1 .. D}. \b\B. x \ b + j)\ P x \ P(x - D) ; \x.(\j\{1 .. D}. \b\B. x \ b + j)\ Q x \ Q(x - D)\ \ \x.(\j\{1 .. D}. \b\B. x \ b + j)\ (P x \ Q x) \ (P(x - D) \ Q (x - D))" "\D>0; t - 1\ B\ \ (\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x = t) \ (x - D = t))" "\D>0 ; t \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" "D>0 \ (\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x < t) \ (x - D < t))" "D>0 \ (\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" "\D>0 ; t \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x > t) \ (x - D > t))" "\D>0 ; t - 1 \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" "d dvd D \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (d dvd x+t) \ (d dvd (x - D) + t))" "d dvd D \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (\d dvd x+t) \ (\ d dvd (x - D) + t))" "\x.(\j\{1 .. D}. \b\B. x \ b + j) \ F \ F" proof (blast, blast) assume dp: "D > 0" and tB: "t - 1\ B" show "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x = t) \ (x - D = t))" apply (rule allI, rule impI,erule ballE[where x="1"],erule ballE[where x="t - 1"]) apply algebra using dp tB by simp_all next assume dp: "D > 0" and tB: "t \ B" show "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" apply (rule allI, rule impI,erule ballE[where x="D"],erule ballE[where x="t"]) apply algebra using dp tB by simp_all next assume dp: "D > 0" thus "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x < t) \ (x - D < t))" by arith next assume dp: "D > 0" thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t)" by arith next assume dp: "D > 0" and tB:"t \ B" {fix x assume nob: "\j\{1 .. D}. \b\B. x \ b + j" and g: "x > t" and ng: "\ (x - D) > t" hence "x -t \ D" and "1 \ x - t" by simp+ hence "\j \ {1 .. D}. x - t = j" by auto hence "\j \ {1 .. D}. x = t + j" by (simp add: algebra_simps) with nob tB have "False" by simp} thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x > t) \ (x - D > t)" by blast next assume dp: "D > 0" and tB:"t - 1\ B" {fix x assume nob: "\j\{1 .. D}. \b\B. x \ b + j" and g: "x \ t" and ng: "\ (x - D) \ t" hence "x - (t - 1) \ D" and "1 \ x - (t - 1)" by simp+ hence "\j \ {1 .. D}. x - (t - 1) = j" by auto hence "\j \ {1 .. D}. x = (t - 1) + j" by (simp add: algebra_simps) with nob tB have "False" by simp} thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t)" by blast next assume d: "d dvd D" {fix x assume H: "d dvd x + t" with d have "d dvd (x - D) + t" by algebra} thus "\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (d dvd x+t) \ (d dvd (x - D) + t)" by simp next assume d: "d dvd D" {fix x assume H: "\(d dvd x + t)" with d have "\ d dvd (x - D) + t" by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: algebra_simps)} thus "\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (\d dvd x+t) \ (\d dvd (x - D) + t)" by auto qed blast lemma aset: "\\x.(\j\{1 .. D}. \b\A. x \ b - j)\ P x \ P(x + D) ; \x.(\j\{1 .. D}. \b\A. x \ b - j)\ Q x \ Q(x + D)\ \ \x.(\j\{1 .. D}. \b\A. x \ b - j) \ (P x \ Q x) \ (P(x + D) \ Q (x + D))" "\\x.(\j\{1 .. D}. \b\A. x \ b - j)\ P x \ P(x + D) ; \x.(\j\{1 .. D}. \b\A. x \ b - j)\ Q x \ Q(x + D)\ \ \x.(\j\{1 .. D}. \b\A. x \ b - j)\ (P x \ Q x) \ (P(x + D) \ Q (x + D))" "\D>0; t + 1\ A\ \ (\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x = t) \ (x + D = t))" "\D>0 ; t \ A\ \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" "\D>0; t\ A\ \(\(x::int). (\j\{1 .. D}. \b\A. x \ b - j)\ (x < t) \ (x + D < t))" "\D>0; t + 1 \ A\ \ (\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" "D>0 \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x > t) \ (x + D > t))" "D>0 \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" "d dvd D \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (d dvd x+t) \ (d dvd (x + D) + t))" "d dvd D \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (\d dvd x+t) \ (\ d dvd (x + D) + t))" "\x.(\j\{1 .. D}. \b\A. x \ b - j) \ F \ F" proof (blast, blast) assume dp: "D > 0" and tA: "t + 1 \ A" show "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x = t) \ (x + D = t))" apply (rule allI, rule impI,erule ballE[where x="1"],erule ballE[where x="t + 1"]) using dp tA by simp_all next assume dp: "D > 0" and tA: "t \ A" show "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" apply (rule allI, rule impI,erule ballE[where x="D"],erule ballE[where x="t"]) using dp tA by simp_all next assume dp: "D > 0" thus "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x > t) \ (x + D > t))" by arith next assume dp: "D > 0" thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t)" by arith next assume dp: "D > 0" and tA:"t \ A" {fix x assume nob: "\j\{1 .. D}. \b\A. x \ b - j" and g: "x < t" and ng: "\ (x + D) < t" hence "t - x \ D" and "1 \ t - x" by simp+ hence "\j \ {1 .. D}. t - x = j" by auto hence "\j \ {1 .. D}. x = t - j" by (auto simp add: algebra_simps) with nob tA have "False" by simp} thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x < t) \ (x + D < t)" by blast next assume dp: "D > 0" and tA:"t + 1\ A" {fix x assume nob: "\j\{1 .. D}. \b\A. x \ b - j" and g: "x \ t" and ng: "\ (x + D) \ t" hence "(t + 1) - x \ D" and "1 \ (t + 1) - x" by (simp_all add: algebra_simps) hence "\j \ {1 .. D}. (t + 1) - x = j" by auto hence "\j \ {1 .. D}. x = (t + 1) - j" by (auto simp add: algebra_simps) with nob tA have "False" by simp} thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t)" by blast next assume d: "d dvd D" - {fix x assume H: "d dvd x + t" with d have "d dvd (x + D) + t" - by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: algebra_simps)} + have "\x. d dvd x + t \ d dvd x + D + t" + proof - + fix x + assume H: "d dvd x + t" + then obtain ka where "x + t = d * ka" + unfolding dvd_def by blast + moreover from d obtain k where *:"D = d * k" + unfolding dvd_def by blast + ultimately have "x + d * k + t = d * (ka + k)" + by (simp add: algebra_simps) + then show "d dvd (x + D) + t" + using * unfolding dvd_def by blast + qed thus "\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (d dvd x+t) \ (d dvd (x + D) + t)" by simp next assume d: "d dvd D" {fix x assume H: "\(d dvd x + t)" with d have "\d dvd (x + D) + t" by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: algebra_simps)} thus "\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (\d dvd x+t) \ (\d dvd (x + D) + t)" by auto qed blast subsection\Cooper's Theorem \-\\ and \+\\ Version\ subsubsection\First some trivial facts about periodic sets or predicates\ lemma periodic_finite_ex: assumes dpos: "(0::int) < d" and modd: "\x k. P x = P(x - k*d)" shows "(\x. P x) = (\j \ {1..d}. P j)" (is "?LHS = ?RHS") proof assume ?LHS then obtain x where P: "P x" .. have "x mod d = x - (x div d)*d" by(simp add:mult_div_mod_eq [symmetric] ac_simps eq_diff_eq) hence Pmod: "P x = P(x mod d)" using modd by simp show ?RHS proof (cases) assume "x mod d = 0" hence "P 0" using P Pmod by simp moreover have "P 0 = P(0 - (-1)*d)" using modd by blast ultimately have "P d" by simp moreover have "d \ {1..d}" using dpos by simp ultimately show ?RHS .. next assume not0: "x mod d \ 0" have "P(x mod d)" using dpos P Pmod by simp moreover have "x mod d \ {1..d}" proof - from dpos have "0 \ x mod d" by(rule pos_mod_sign) moreover from dpos have "x mod d < d" by(rule pos_mod_bound) ultimately show ?thesis using not0 by simp qed ultimately show ?RHS .. qed qed auto subsubsection\The \-\\ Version\ lemma decr_lemma: "0 < (d::int) \ x - (\x - z\ + 1) * d < z" by (induct rule: int_gr_induct) (simp_all add: int_distrib) lemma incr_lemma: "0 < (d::int) \ z < x + (\x - z\ + 1) * d" by (induct rule: int_gr_induct) (simp_all add: int_distrib) lemma decr_mult_lemma: assumes dpos: "(0::int) < d" and minus: "\x. P x \ P(x - d)" and knneg: "0 <= k" shows "\x. P x \ P(x - k*d)" using knneg proof (induct rule:int_ge_induct) case base thus ?case by simp next case (step i) {fix x have "P x \ P (x - i * d)" using step.hyps by blast also have "\ \ P(x - (i + 1) * d)" using minus[THEN spec, of "x - i * d"] by (simp add: algebra_simps) ultimately have "P x \ P(x - (i + 1) * d)" by blast} thus ?case .. qed lemma minusinfinity: assumes dpos: "0 < d" and P1eqP1: "\x k. P1 x = P1(x - k*d)" and ePeqP1: "\z::int. \x. x < z \ (P x = P1 x)" shows "(\x. P1 x) \ (\x. P x)" proof assume eP1: "\x. P1 x" then obtain x where P1: "P1 x" .. from ePeqP1 obtain z where P1eqP: "\x. x < z \ (P x = P1 x)" .. let ?w = "x - (\x - z\ + 1) * d" from dpos have w: "?w < z" by(rule decr_lemma) have "P1 x = P1 ?w" using P1eqP1 by blast also have "\ = P(?w)" using w P1eqP by blast finally have "P ?w" using P1 by blast thus "\x. P x" .. qed lemma cpmi: assumes dp: "0 < D" and p1:"\z. \ x< z. P x = P' x" and nb:"\x.(\ j\ {1..D}. \(b::int) \ B. x \ b+j) \ P (x) \ P (x - D)" and pd: "\ x k. P' x = P' (x-k*D)" shows "(\x. P x) = ((\j \ {1..D} . P' j) \ (\j \ {1..D}. \ b \ B. P (b+j)))" (is "?L = (?R1 \ ?R2)") proof- {assume "?R2" hence "?L" by blast} moreover {assume H:"?R1" hence "?L" using minusinfinity[OF dp pd p1] periodic_finite_ex[OF dp pd] by simp} moreover { fix x assume P: "P x" and H: "\ ?R2" {fix y assume "\ (\j\{1..D}. \b\B. P (b + j))" and P: "P y" hence "\(\(j::int) \ {1..D}. \(b::int) \ B. y = b+j)" by auto with nb P have "P (y - D)" by auto } hence "\x. \(\(j::int) \ {1..D}. \(b::int) \ B. P(b+j)) \ P (x) \ P (x - D)" by blast with H P have th: " \x. P x \ P (x - D)" by auto from p1 obtain z where z: "\x. x < z \ (P x = P' x)" by blast let ?y = "x - (\x - z\ + 1)*D" have zp: "0 <= (\x - z\ + 1)" by arith from dp have yz: "?y < z" using decr_lemma[OF dp] by simp from z[rule_format, OF yz] decr_mult_lemma[OF dp th zp, rule_format, OF P] have th2: " P' ?y" by auto with periodic_finite_ex[OF dp pd] have "?R1" by blast} ultimately show ?thesis by blast qed subsubsection \The \+\\ Version\ lemma plusinfinity: assumes dpos: "(0::int) < d" and P1eqP1: "\x k. P' x = P'(x - k*d)" and ePeqP1: "\ z. \ x>z. P x = P' x" shows "(\ x. P' x) \ (\ x. P x)" proof assume eP1: "\x. P' x" then obtain x where P1: "P' x" .. from ePeqP1 obtain z where P1eqP: "\x>z. P x = P' x" .. let ?w' = "x + (\x - z\ + 1) * d" let ?w = "x - (- (\x - z\ + 1)) * d" have ww'[simp]: "?w = ?w'" by (simp add: algebra_simps) from dpos have w: "?w > z" by(simp only: ww' incr_lemma) hence "P' x = P' ?w" using P1eqP1 by blast also have "\ = P(?w)" using w P1eqP by blast finally have "P ?w" using P1 by blast thus "\x. P x" .. qed lemma incr_mult_lemma: assumes dpos: "(0::int) < d" and plus: "\x::int. P x \ P(x + d)" and knneg: "0 <= k" shows "\x. P x \ P(x + k*d)" using knneg proof (induct rule:int_ge_induct) case base thus ?case by simp next case (step i) {fix x have "P x \ P (x + i * d)" using step.hyps by blast also have "\ \ P(x + (i + 1) * d)" using plus[THEN spec, of "x + i * d"] by (simp add:int_distrib ac_simps) ultimately have "P x \ P(x + (i + 1) * d)" by blast} thus ?case .. qed lemma cppi: assumes dp: "0 < D" and p1:"\z. \ x> z. P x = P' x" and nb:"\x.(\ j\ {1..D}. \(b::int) \ A. x \ b - j) \ P (x) \ P (x + D)" and pd: "\ x k. P' x= P' (x-k*D)" shows "(\x. P x) = ((\j \ {1..D} . P' j) \ (\ j \ {1..D}. \ b\ A. P (b - j)))" (is "?L = (?R1 \ ?R2)") proof- {assume "?R2" hence "?L" by blast} moreover {assume H:"?R1" hence "?L" using plusinfinity[OF dp pd p1] periodic_finite_ex[OF dp pd] by simp} moreover { fix x assume P: "P x" and H: "\ ?R2" {fix y assume "\ (\j\{1..D}. \b\A. P (b - j))" and P: "P y" hence "\(\(j::int) \ {1..D}. \(b::int) \ A. y = b - j)" by auto with nb P have "P (y + D)" by auto } hence "\x. \(\(j::int) \ {1..D}. \(b::int) \ A. P(b-j)) \ P (x) \ P (x + D)" by blast with H P have th: " \x. P x \ P (x + D)" by auto from p1 obtain z where z: "\x. x > z \ (P x = P' x)" by blast let ?y = "x + (\x - z\ + 1)*D" have zp: "0 <= (\x - z\ + 1)" by arith from dp have yz: "?y > z" using incr_lemma[OF dp] by simp from z[rule_format, OF yz] incr_mult_lemma[OF dp th zp, rule_format, OF P] have th2: " P' ?y" by auto with periodic_finite_ex[OF dp pd] have "?R1" by blast} ultimately show ?thesis by blast qed lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})" apply(simp add:atLeastAtMost_def atLeast_def atMost_def) apply(fastforce) done theorem unity_coeff_ex: "(\(x::'a::{semiring_0,Rings.dvd}). P (l * x)) \ (\x. l dvd (x + 0) \ P x)" - apply (rule eq_reflection [symmetric]) - apply (rule iffI) - defer - apply (erule exE) - apply (rule_tac x = "l * x" in exI) - apply (simp add: dvd_def) - apply (rule_tac x = x in exI, simp) - apply (erule exE) - apply (erule conjE) - apply simp - apply (erule dvdE) - apply (rule_tac x = k in exI) - apply simp - done + unfolding dvd_def by (rule eq_reflection, rule iffI) auto lemma zdvd_mono: fixes k m t :: int assumes "k \ 0" shows "m dvd t \ k * m dvd k * t" using assms by simp lemma uminus_dvd_conv: fixes d t :: int shows "d dvd t \ - d dvd t" and "d dvd t \ d dvd - t" by simp_all text \\bigskip Theorems for transforming predicates on nat to predicates on \int\\ lemma zdiff_int_split: "P (int (x - y)) = ((y \ x \ P (int x - int y)) \ (x < y \ P 0))" by (cases "y \ x") (simp_all add: of_nat_diff) text \ \medskip Specific instances of congruence rules, to prevent simplifier from looping.\ theorem imp_le_cong: "\x = x'; 0 \ x' \ P = P'\ \ (0 \ (x::int) \ P) = (0 \ x' \ P')" by simp theorem conj_le_cong: "\x = x'; 0 \ x' \ P = P'\ \ (0 \ (x::int) \ P) = (0 \ x' \ P')" by (simp cong: conj_cong) ML_file \Tools/Qelim/cooper.ML\ method_setup presburger = \ let fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () fun simple_keyword k = Scan.lift (Args.$$$ k) >> K () val addN = "add" val delN = "del" val elimN = "elim" val any_keyword = keyword addN || keyword delN || simple_keyword elimN val thms = Scan.repeats (Scan.unless any_keyword Attrib.multi_thm) in Scan.optional (simple_keyword elimN >> K false) true -- Scan.optional (keyword addN |-- thms) [] -- Scan.optional (keyword delN |-- thms) [] >> (fn ((elim, add_ths), del_ths) => fn ctxt => SIMPLE_METHOD' (Cooper.tac elim add_ths del_ths ctxt)) end \ "Cooper's algorithm for Presburger arithmetic" declare mod_eq_0_iff_dvd [presburger] declare mod_by_Suc_0 [presburger] declare mod_0 [presburger] declare mod_by_1 [presburger] declare mod_self [presburger] declare div_by_0 [presburger] declare mod_by_0 [presburger] declare mod_div_trivial [presburger] declare mult_div_mod_eq [presburger] declare div_mult_mod_eq [presburger] declare mod_mult_self1 [presburger] declare mod_mult_self2 [presburger] declare mod2_Suc_Suc [presburger] declare not_mod_2_eq_0_eq_1 [presburger] declare nat_zero_less_power_iff [presburger] lemma [presburger, algebra]: "m mod 2 = (1::nat) \ \ 2 dvd m " by presburger lemma [presburger, algebra]: "m mod 2 = Suc 0 \ \ 2 dvd m " by presburger lemma [presburger, algebra]: "m mod (Suc (Suc 0)) = (1::nat) \ \ 2 dvd m " by presburger lemma [presburger, algebra]: "m mod (Suc (Suc 0)) = Suc 0 \ \ 2 dvd m " by presburger lemma [presburger, algebra]: "m mod 2 = (1::int) \ \ 2 dvd m " by presburger context semiring_parity begin declare even_mult_iff [presburger] declare even_power [presburger] lemma [presburger]: "even (a + b) \ even a \ even b \ odd a \ odd b" by auto end context ring_parity begin declare even_minus [presburger] end context linordered_idom begin declare zero_le_power_eq [presburger] declare zero_less_power_eq [presburger] declare power_less_zero_eq [presburger] declare power_le_zero_eq [presburger] end declare even_Suc [presburger] lemma [presburger]: "Suc n div Suc (Suc 0) = n div Suc (Suc 0) \ even n" by presburger declare even_diff_nat [presburger] lemma [presburger]: fixes k :: int shows "(k + 1) div 2 = k div 2 \ even k" by presburger lemma [presburger]: fixes k :: int shows "(k + 1) div 2 = k div 2 + 1 \ odd k" by presburger lemma [presburger]: "even n \ even (int n)" by simp subsection \Nice facts about division by \<^term>\4\\ lemma even_even_mod_4_iff: "even (n::nat) \ even (n mod 4)" by presburger lemma odd_mod_4_div_2: "n mod 4 = (3::nat) \ odd ((n - Suc 0) div 2)" by presburger lemma even_mod_4_div_2: "n mod 4 = Suc 0 \ even ((n - Suc 0) div 2)" by presburger subsection \Try0\ ML_file \Tools/try0.ML\ end diff --git a/src/HOL/Product_Type.thy b/src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy +++ b/src/HOL/Product_Type.thy @@ -1,1353 +1,1376 @@ (* Title: HOL/Product_Type.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section \Cartesian products\ theory Product_Type imports Typedef Inductive Fun keywords "inductive_set" "coinductive_set" :: thy_defn begin subsection \\<^typ>\bool\ is a datatype\ free_constructors (discs_sels) case_bool for True | False by auto text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype True False by (auto intro: bool_induct) setup \Sign.parent_path\ text \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "bool"\ lemmas induct = old.bool.induct lemmas inducts = old.bool.inducts lemmas rec = old.bool.rec lemmas simps = bool.distinct bool.case bool.rec setup \Sign.parent_path\ declare case_split [cases type: bool] \ \prefer plain propositional version\ lemma [code]: "HOL.equal False P \ \ P" and [code]: "HOL.equal True P \ P" and [code]: "HOL.equal P False \ \ P" and [code]: "HOL.equal P True \ P" and [code nbe]: "HOL.equal P P \ True" by (simp_all add: equal) lemma If_case_cert: assumes "CASE \ (\b. If b f g)" shows "(CASE True \ f) &&& (CASE False \ g)" using assms by simp_all setup \Code.declare_case_global @{thm If_case_cert}\ code_printing constant "HOL.equal :: bool \ bool \ bool" \ (Haskell) infix 4 "==" | class_instance "bool" :: "equal" \ (Haskell) - subsection \The \unit\ type\ typedef unit = "{True}" by auto definition Unity :: unit ("'(')") where "() = Abs_unit True" lemma unit_eq [no_atp]: "u = ()" by (induct u) (simp add: Unity_def) text \ Simplification procedure for @{thm [source] unit_eq}. Cannot use this rule directly --- it loops! \ simproc_setup unit_eq ("x::unit") = \ fn _ => fn _ => fn ct => if HOLogic.is_unit (Thm.term_of ct) then NONE else SOME (mk_meta_eq @{thm unit_eq}) \ free_constructors case_unit for "()" by auto text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "()" by simp setup \Sign.parent_path\ text \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "unit"\ lemmas induct = old.unit.induct lemmas inducts = old.unit.inducts lemmas rec = old.unit.rec lemmas simps = unit.case unit.rec setup \Sign.parent_path\ lemma unit_all_eq1: "(\x::unit. PROP P x) \ PROP P ()" by simp lemma unit_all_eq2: "(\x::unit. PROP P) \ PROP P" by (rule triv_forall_equality) text \ This rewrite counters the effect of simproc \unit_eq\ on @{term [source] "\u::unit. f u"}, replacing it by @{term [source] f} rather than by @{term [source] "\u. f ()"}. \ lemma unit_abs_eta_conv [simp]: "(\u::unit. f ()) = f" by (rule ext) simp lemma UNIV_unit: "UNIV = {()}" by auto instantiation unit :: default begin definition "default = ()" instance .. end instantiation unit :: "{complete_boolean_algebra,complete_linorder,wellorder}" begin definition less_eq_unit :: "unit \ unit \ bool" where "(_::unit) \ _ \ True" lemma less_eq_unit [iff]: "u \ v" for u v :: unit by (simp add: less_eq_unit_def) definition less_unit :: "unit \ unit \ bool" where "(_::unit) < _ \ False" lemma less_unit [iff]: "\ u < v" for u v :: unit by (simp_all add: less_eq_unit_def less_unit_def) definition bot_unit :: unit where [code_unfold]: "\ = ()" definition top_unit :: unit where [code_unfold]: "\ = ()" definition inf_unit :: "unit \ unit \ unit" where [simp]: "_ \ _ = ()" definition sup_unit :: "unit \ unit \ unit" where [simp]: "_ \ _ = ()" definition Inf_unit :: "unit set \ unit" where [simp]: "\_ = ()" definition Sup_unit :: "unit set \ unit" where [simp]: "\_ = ()" definition uminus_unit :: "unit \ unit" where [simp]: "- _ = ()" declare less_eq_unit_def [abs_def, code_unfold] less_unit_def [abs_def, code_unfold] inf_unit_def [abs_def, code_unfold] sup_unit_def [abs_def, code_unfold] Inf_unit_def [abs_def, code_unfold] Sup_unit_def [abs_def, code_unfold] uminus_unit_def [abs_def, code_unfold] instance by intro_classes auto end lemma [code]: "HOL.equal u v \ True" for u v :: unit - unfolding equal unit_eq [of u] unit_eq [of v] by rule+ + unfolding equal unit_eq [of u] unit_eq [of v] by (rule iffI TrueI refl)+ code_printing type_constructor unit \ (SML) "unit" and (OCaml) "unit" and (Haskell) "()" and (Scala) "Unit" | constant Unity \ (SML) "()" and (OCaml) "()" and (Haskell) "()" and (Scala) "()" | class_instance unit :: equal \ (Haskell) - | constant "HOL.equal :: unit \ unit \ bool" \ (Haskell) infix 4 "==" code_reserved SML unit code_reserved OCaml unit code_reserved Scala Unit subsection \The product type\ subsubsection \Type definition\ definition Pair_Rep :: "'a \ 'b \ 'a \ 'b \ bool" where "Pair_Rep a b = (\x y. x = a \ y = b)" definition "prod = {f. \a b. f = Pair_Rep (a::'a) (b::'b)}" typedef ('a, 'b) prod ("(_ \/ _)" [21, 20] 20) = "prod :: ('a \ 'b \ bool) set" unfolding prod_def by auto type_notation (ASCII) prod (infixr "*" 20) definition Pair :: "'a \ 'b \ 'a \ 'b" where "Pair a b = Abs_prod (Pair_Rep a b)" lemma prod_cases: "(\a b. P (Pair a b)) \ P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) free_constructors case_prod for Pair fst snd proof - fix P :: bool and p :: "'a \ 'b" show "(\x1 x2. p = Pair x1 x2 \ P) \ P" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) next fix a c :: 'a and b d :: 'b have "Pair_Rep a b = Pair_Rep c d \ a = c \ b = d" by (auto simp add: Pair_Rep_def fun_eq_iff) moreover have "Pair_Rep a b \ prod" and "Pair_Rep c d \ prod" by (auto simp add: prod_def) ultimately show "Pair a b = Pair c d \ a = c \ b = d" by (simp add: Pair_def Abs_prod_inject) qed text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype Pair by (erule prod_cases) (rule prod.inject) setup \Sign.parent_path\ text \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "prod"\ declare old.prod.inject [iff del] lemmas induct = old.prod.induct lemmas inducts = old.prod.inducts lemmas rec = old.prod.rec lemmas simps = prod.inject prod.case prod.rec setup \Sign.parent_path\ declare prod.case [nitpick_simp del] declare old.prod.case_cong_weak [cong del] declare prod.case_eq_if [mono] declare prod.split [no_atp] declare prod.split_asm [no_atp] text \ @{thm [source] prod.split} could be declared as \[split]\ done after the Splitter has been speeded up significantly; precompute the constants involved and don't do anything unless the current goal contains one of those constants. \ subsubsection \Tuple syntax\ text \ Patterns -- extends pre-defined type \<^typ>\pttrn\ used in abstractions. \ nonterminal tuple_args and patterns syntax "_tuple" :: "'a \ tuple_args \ 'a \ 'b" ("(1'(_,/ _'))") "_tuple_arg" :: "'a \ tuple_args" ("_") "_tuple_args" :: "'a \ tuple_args \ tuple_args" ("_,/ _") "_pattern" :: "pttrn \ patterns \ pttrn" ("'(_,/ _')") "" :: "pttrn \ patterns" ("_") "_patterns" :: "pttrn \ patterns \ patterns" ("_,/ _") "_unit" :: pttrn ("'(')") translations "(x, y)" \ "CONST Pair x y" "_pattern x y" \ "CONST Pair x y" "_patterns x y" \ "CONST Pair x y" "_tuple x (_tuple_args y z)" \ "_tuple x (_tuple_arg (_tuple y z))" "\(x, y, zs). b" \ "CONST case_prod (\x (y, zs). b)" "\(x, y). b" \ "CONST case_prod (\x y. b)" "_abs (CONST Pair x y) t" \ "\(x, y). t" \ \This rule accommodates tuples in \case C \ (x, y) \ \ \\: The \(x, y)\ is parsed as \Pair x y\ because it is \logic\, not \pttrn\.\ "\(). b" \ "CONST case_unit b" "_abs (CONST Unity) t" \ "\(). t" text \print \<^term>\case_prod f\ as \<^term>\\(x, y). f x y\ and \<^term>\case_prod (\x. f x)\ as \<^term>\\(x, y). f x y\\ typed_print_translation \ let fun case_prod_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match | case_prod_guess_names_tr' T [Abs (x, xT, t)] = (case (head_of t) of Const (\<^const_syntax>\case_prod\, _) => raise Match | _ => let val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0); val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ y) $ t'' end) | case_prod_guess_names_tr' T [t] = (case head_of t of Const (\<^const_syntax>\case_prod\, _) => raise Match | _ => let val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0); val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ y) $ t'' end) | case_prod_guess_names_tr' _ _ = raise Match; in [(\<^const_syntax>\case_prod\, K case_prod_guess_names_tr')] end \ text \Reconstruct pattern from (nested) \<^const>\case_prod\s, avoiding eta-contraction of body; required for enclosing "let", if "let" does not avoid eta-contraction, which has been observed to occur.\ print_translation \ let fun case_prod_tr' [Abs (x, T, t as (Abs abs))] = (* case_prod (\x y. t) \ \(x, y) t *) let val (y, t') = Syntax_Trans.atomic_abs_tr' abs; val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ y) $ t'' end | case_prod_tr' [Abs (x, T, (s as Const (\<^const_syntax>\case_prod\, _) $ t))] = (* case_prod (\x. (case_prod (\y z. t))) \ \(x, y, z). t *) let val Const (\<^syntax_const>\_abs\, _) $ (Const (\<^syntax_const>\_pattern\, _) $ y $ z) $ t' = case_prod_tr' [t]; val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ (Syntax.const \<^syntax_const>\_patterns\ $ y $ z)) $ t'' end | case_prod_tr' [Const (\<^const_syntax>\case_prod\, _) $ t] = (* case_prod (case_prod (\x y z. t)) \ \((x, y), z). t *) case_prod_tr' [(case_prod_tr' [t])] (* inner case_prod_tr' creates next pattern *) | case_prod_tr' [Const (\<^syntax_const>\_abs\, _) $ x_y $ Abs abs] = (* case_prod (\pttrn z. t) \ \(pttrn, z). t *) let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x_y $ z) $ t end | case_prod_tr' _ = raise Match; in [(\<^const_syntax>\case_prod\, K case_prod_tr')] end \ subsubsection \Code generator setup\ code_printing type_constructor prod \ (SML) infix 2 "*" and (OCaml) infix 2 "*" and (Haskell) "!((_),/ (_))" and (Scala) "((_),/ (_))" | constant Pair \ (SML) "!((_),/ (_))" and (OCaml) "!((_),/ (_))" and (Haskell) "!((_),/ (_))" and (Scala) "!((_),/ (_))" | class_instance prod :: equal \ (Haskell) - | constant "HOL.equal :: 'a \ 'b \ 'a \ 'b \ bool" \ (Haskell) infix 4 "==" | constant fst \ (Haskell) "fst" | constant snd \ (Haskell) "snd" subsubsection \Fundamental operations and properties\ lemma Pair_inject: "(a, b) = (a', b') \ (a = a' \ b = b' \ R) \ R" by simp lemma surj_pair [simp]: "\x y. p = (x, y)" by (cases p) simp lemma fst_eqD: "fst (x, y) = a \ x = a" by simp lemma snd_eqD: "snd (x, y) = a \ y = a" by simp lemma case_prod_unfold [nitpick_unfold]: "case_prod = (\c p. c (fst p) (snd p))" by (simp add: fun_eq_iff split: prod.split) lemma case_prod_conv [simp, code]: "(case (a, b) of (c, d) \ f c d) = f a b" by (fact prod.case) lemmas surjective_pairing = prod.collapse [symmetric] lemma prod_eq_iff: "s = t \ fst s = fst t \ snd s = snd t" by (cases s, cases t) simp lemma prod_eqI [intro?]: "fst p = fst q \ snd p = snd q \ p = q" by (simp add: prod_eq_iff) lemma case_prodI: "f a b \ case (a, b) of (c, d) \ f c d" by (rule prod.case [THEN iffD2]) lemma case_prodD: "(case (a, b) of (c, d) \ f c d) \ f a b" by (rule prod.case [THEN iffD1]) lemma case_prod_Pair [simp]: "case_prod Pair = id" by (simp add: fun_eq_iff split: prod.split) lemma case_prod_eta: "(\(x, y). f (x, y)) = f" \ \Subsumes the old \split_Pair\ when \<^term>\f\ is the identity function.\ by (simp add: fun_eq_iff split: prod.split) (* This looks like a sensible simp-rule but appears to do more harm than good: lemma case_prod_const [simp]: "(\(_,_). c) = (\_. c)" by(rule case_prod_eta) *) lemma case_prod_comp: "(case x of (a, b) \ (f \ g) a b) = f (g (fst x)) (snd x)" by (cases x) simp lemma The_case_prod: "The (case_prod P) = (THE xy. P (fst xy) (snd xy))" by (simp add: case_prod_unfold) lemma cond_case_prod_eta: "(\x y. f x y = g (x, y)) \ (\(x, y). f x y) = g" by (simp add: case_prod_eta) lemma split_paired_all [no_atp]: "(\x. PROP P x) \ (\a b. PROP P (a, b))" proof fix a b assume "\x. PROP P x" then show "PROP P (a, b)" . next fix x assume "\a b. PROP P (a, b)" from \PROP P (fst x, snd x)\ show "PROP P x" by simp qed text \ The rule @{thm [source] split_paired_all} does not work with the Simplifier because it also affects premises in congrence rules, where this can lead to premises of the form \\a b. \ = ?P(a, b)\ which cannot be solved by reflexivity. \ lemmas split_tupled_all = split_paired_all unit_all_eq2 ML \ (* replace parameters of product type by individual component parameters *) local (* filtering with exists_paired_all is an essential optimization *) fun exists_paired_all (Const (\<^const_name>\Pure.all\, _) $ Abs (_, T, t)) = can HOLogic.dest_prodT T orelse exists_paired_all t | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u | exists_paired_all (Abs (_, _, t)) = exists_paired_all t | exists_paired_all _ = false; val ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}] addsimprocs [\<^simproc>\unit_eq\]); in fun split_all_tac ctxt = SUBGOAL (fn (t, i) => if exists_paired_all t then safe_full_simp_tac (put_simpset ss ctxt) i else no_tac); fun unsafe_split_all_tac ctxt = SUBGOAL (fn (t, i) => if exists_paired_all t then full_simp_tac (put_simpset ss ctxt) i else no_tac); fun split_all ctxt th = if exists_paired_all (Thm.prop_of th) then full_simplify (put_simpset ss ctxt) th else th; end; \ setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\ lemma split_paired_All [simp, no_atp]: "(\x. P x) \ (\a b. P (a, b))" \ \\[iff]\ is not a good idea because it makes \blast\ loop\ by fast lemma split_paired_Ex [simp, no_atp]: "(\x. P x) \ (\a b. P (a, b))" by fast lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))" \ \Can't be added to simpset: loops!\ by (simp add: case_prod_eta) text \ Simplification procedure for @{thm [source] cond_case_prod_eta}. Using @{thm [source] case_prod_eta} as a rewrite rule is not general enough, and using @{thm [source] cond_case_prod_eta} directly would render some existing proofs very inefficient; similarly for \prod.case_eq_if\. \ ML \ local val cond_case_prod_eta_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms cond_case_prod_eta}); fun Pair_pat k 0 (Bound m) = (m = k) | Pair_pat k i (Const (\<^const_name>\Pair\, _) $ Bound m $ t) = i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t | Pair_pat _ _ _ = false; fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t | no_args k i (t $ u) = no_args k i t andalso no_args k i u | no_args k i (Bound m) = m < k orelse m > k + i | no_args _ _ _ = true; fun split_pat tp i (Abs (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE | split_pat tp i (Const (\<^const_name>\case_prod\, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t | split_pat tp i _ = NONE; fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))) (K (simp_tac (put_simpset cond_case_prod_eta_ss ctxt) 1))); fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t | beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u) | beta_term_pat k i t = no_args k i t; fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg | eta_term_pat _ _ _ = false; fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) | subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg else (subst arg k i t $ subst arg k i u) | subst arg k i t = t; in fun beta_proc ctxt (s as Const (\<^const_name>\case_prod\, _) $ Abs (_, _, t) $ arg) = (case split_pat beta_term_pat 1 t of SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f)) | NONE => NONE) | beta_proc _ _ = NONE; fun eta_proc ctxt (s as Const (\<^const_name>\case_prod\, _) $ Abs (_, _, t)) = (case split_pat eta_term_pat 1 t of SOME (_, ft) => SOME (metaeq ctxt s (let val f $ _ = ft in f end)) | NONE => NONE) | eta_proc _ _ = NONE; end; \ simproc_setup case_prod_beta ("case_prod f z") = \fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct)\ simproc_setup case_prod_eta ("case_prod f") = \fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct)\ lemma case_prod_beta': "(\(x,y). f x y) = (\x. f (fst x) (snd x))" by (auto simp: fun_eq_iff) text \ \<^medskip> \<^const>\case_prod\ used as a logical connective or set former. \<^medskip> These rules are for use with \blast\; could instead call \simp\ using @{thm [source] prod.split} as rewrite.\ lemma case_prodI2: "\p. (\a b. p = (a, b) \ c a b) \ case p of (a, b) \ c a b" by (simp add: split_tupled_all) lemma case_prodI2': "\p. (\a b. (a, b) = p \ c a b x) \ (case p of (a, b) \ c a b) x" by (simp add: split_tupled_all) lemma case_prodE [elim!]: "(case p of (a, b) \ c a b) \ (\x y. p = (x, y) \ c x y \ Q) \ Q" by (induct p) simp lemma case_prodE' [elim!]: "(case p of (a, b) \ c a b) z \ (\x y. p = (x, y) \ c x y z \ Q) \ Q" by (induct p) simp lemma case_prodE2: assumes q: "Q (case z of (a, b) \ P a b)" and r: "\x y. z = (x, y) \ Q (P x y) \ R" shows R proof (rule r) show "z = (fst z, snd z)" by simp then show "Q (P (fst z) (snd z))" using q by (simp add: case_prod_unfold) qed lemma case_prodD': "(case (a, b) of (c, d) \ R c d) c \ R a b c" by simp lemma mem_case_prodI: "z \ c a b \ z \ (case (a, b) of (d, e) \ c d e)" by simp lemma mem_case_prodI2 [intro!]: "\p. (\a b. p = (a, b) \ z \ c a b) \ z \ (case p of (a, b) \ c a b)" by (simp only: split_tupled_all) simp declare mem_case_prodI [intro!] \ \postponed to maintain traditional declaration order!\ declare case_prodI2' [intro!] \ \postponed to maintain traditional declaration order!\ declare case_prodI2 [intro!] \ \postponed to maintain traditional declaration order!\ declare case_prodI [intro!] \ \postponed to maintain traditional declaration order!\ lemma mem_case_prodE [elim!]: assumes "z \ case_prod c p" obtains x y where "p = (x, y)" and "z \ c x y" using assms by (rule case_prodE2) ML \ local (* filtering with exists_p_split is an essential optimization *) fun exists_p_split (Const (\<^const_name>\case_prod\,_) $ _ $ (Const (\<^const_name>\Pair\,_)$_$_)) = true | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u | exists_p_split (Abs (_, _, t)) = exists_p_split t | exists_p_split _ = false; in fun split_conv_tac ctxt = SUBGOAL (fn (t, i) => if exists_p_split t then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms case_prod_conv}) i else no_tac); end; \ (* This prevents applications of splitE for already splitted arguments leading to quite time-consuming computations (in particular for nested tuples) *) setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac))\ lemma split_eta_SetCompr [simp, no_atp]: "(\u. \x y. u = (x, y) \ P (x, y)) = P" by (rule ext) fast lemma split_eta_SetCompr2 [simp, no_atp]: "(\u. \x y. u = (x, y) \ P x y) = case_prod P" by (rule ext) fast lemma split_part [simp]: "(\(a,b). P \ Q a b) = (\ab. P \ case_prod Q ab)" \ \Allows simplifications of nested splits in case of independent predicates.\ by (rule ext) blast (* Do NOT make this a simp rule as it a) only helps in special situations b) can lead to nontermination in the presence of split_def *) lemma split_comp_eq: fixes f :: "'a \ 'b \ 'c" and g :: "'d \ 'a" shows "(\u. f (g (fst u)) (snd u)) = case_prod (\x. f (g x))" by (rule ext) auto lemma pair_imageI [intro]: "(a, b) \ A \ f a b \ (\(a, b). f a b) ` A" by (rule image_eqI [where x = "(a, b)"]) auto lemma Collect_const_case_prod[simp]: "{(a,b). P} = (if P then UNIV else {})" by auto lemma The_split_eq [simp]: "(THE (x',y'). x = x' \ y = y') = (x, y)" by blast (* the following would be slightly more general, but cannot be used as rewrite rule: ### Cannot add premise as rewrite rule because it contains (type) unknowns: ### ?y = .x Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" by (rtac some_equality 1) by ( Simp_tac 1) by (split_all_tac 1) by (Asm_full_simp_tac 1) qed "The_split_eq"; *) lemma case_prod_beta: "case_prod f p = f (fst p) (snd p)" by (fact prod.case_eq_if) lemma prod_cases3 [cases type]: obtains (fields) a b c where "y = (a, b, c)" - by (cases y, case_tac b) blast +proof (cases y) + case (Pair a b) + with that show ?thesis + by (cases b) blast +qed + lemma prod_induct3 [case_names fields, induct type]: "(\a b c. P (a, b, c)) \ P x" by (cases x) blast lemma prod_cases4 [cases type]: obtains (fields) a b c d where "y = (a, b, c, d)" - by (cases y, case_tac c) blast +proof (cases y) + case (fields a b c) + with that show ?thesis + by (cases c) blast +qed lemma prod_induct4 [case_names fields, induct type]: "(\a b c d. P (a, b, c, d)) \ P x" by (cases x) blast lemma prod_cases5 [cases type]: obtains (fields) a b c d e where "y = (a, b, c, d, e)" - by (cases y, case_tac d) blast +proof (cases y) + case (fields a b c d) + with that show ?thesis + by (cases d) blast +qed lemma prod_induct5 [case_names fields, induct type]: "(\a b c d e. P (a, b, c, d, e)) \ P x" by (cases x) blast lemma prod_cases6 [cases type]: obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)" - by (cases y, case_tac e) blast +proof (cases y) + case (fields a b c d e) + with that show ?thesis + by (cases e) blast +qed lemma prod_induct6 [case_names fields, induct type]: "(\a b c d e f. P (a, b, c, d, e, f)) \ P x" by (cases x) blast lemma prod_cases7 [cases type]: obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)" - by (cases y, case_tac f) blast +proof (cases y) + case (fields a b c d e f) + with that show ?thesis + by (cases f) blast +qed + lemma prod_induct7 [case_names fields, induct type]: "(\a b c d e f g. P (a, b, c, d, e, f, g)) \ P x" by (cases x) blast definition internal_case_prod :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "internal_case_prod \ case_prod" lemma internal_case_prod_conv: "internal_case_prod c (a, b) = c a b" by (simp only: internal_case_prod_def case_prod_conv) ML_file \Tools/split_rule.ML\ hide_const internal_case_prod subsubsection \Derived operations\ definition curry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "curry = (\c x y. c (x, y))" lemma curry_conv [simp, code]: "curry f a b = f (a, b)" by (simp add: curry_def) lemma curryI [intro!]: "f (a, b) \ curry f a b" by (simp add: curry_def) lemma curryD [dest!]: "curry f a b \ f (a, b)" by (simp add: curry_def) lemma curryE: "curry f a b \ (f (a, b) \ Q) \ Q" by (simp add: curry_def) lemma curry_case_prod [simp]: "curry (case_prod f) = f" by (simp add: curry_def case_prod_unfold) lemma case_prod_curry [simp]: "case_prod (curry f) = f" by (simp add: curry_def case_prod_unfold) lemma curry_K: "curry (\x. c) = (\x y. c)" by (simp add: fun_eq_iff) text \The composition-uncurry combinator.\ definition scomp :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "\\" 60) where "f \\ g = (\x. case_prod g (f x))" no_notation scomp (infixl "\\" 60) bundle state_combinator_syntax begin notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) end context includes state_combinator_syntax begin lemma scomp_unfold: "(\\) = (\f g x. g (fst (f x)) (snd (f x)))" by (simp add: fun_eq_iff scomp_def case_prod_unfold) lemma scomp_apply [simp]: "(f \\ g) x = case_prod g (f x)" by (simp add: scomp_unfold case_prod_unfold) lemma Pair_scomp: "Pair x \\ f = f x" by (simp add: fun_eq_iff) lemma scomp_Pair: "x \\ Pair = x" by (simp add: fun_eq_iff) lemma scomp_scomp: "(f \\ g) \\ h = f \\ (\x. g x \\ h)" by (simp add: fun_eq_iff scomp_unfold) lemma scomp_fcomp: "(f \\ g) \> h = f \\ (\x. g x \> h)" by (simp add: fun_eq_iff scomp_unfold fcomp_def) lemma fcomp_scomp: "(f \> g) \\ h = f \> (g \\ h)" by (simp add: fun_eq_iff scomp_unfold) end code_printing constant scomp \ (Eval) infixl 3 "#->" text \ \<^term>\map_prod\ --- action of the product functor upon functions. \ definition map_prod :: "('a \ 'c) \ ('b \ 'd) \ 'a \ 'b \ 'c \ 'd" where "map_prod f g = (\(x, y). (f x, g y))" lemma map_prod_simp [simp, code]: "map_prod f g (a, b) = (f a, g b)" by (simp add: map_prod_def) functor map_prod: map_prod by (auto simp add: split_paired_all) lemma fst_map_prod [simp]: "fst (map_prod f g x) = f (fst x)" by (cases x) simp_all lemma snd_map_prod [simp]: "snd (map_prod f g x) = g (snd x)" by (cases x) simp_all lemma fst_comp_map_prod [simp]: "fst \ map_prod f g = f \ fst" by (rule ext) simp_all lemma snd_comp_map_prod [simp]: "snd \ map_prod f g = g \ snd" by (rule ext) simp_all lemma map_prod_compose: "map_prod (f1 \ f2) (g1 \ g2) = (map_prod f1 g1 \ map_prod f2 g2)" by (rule ext) (simp add: map_prod.compositionality comp_def) lemma map_prod_ident [simp]: "map_prod (\x. x) (\y. y) = (\z. z)" by (rule ext) (simp add: map_prod.identity) lemma map_prod_imageI [intro]: "(a, b) \ R \ (f a, g b) \ map_prod f g ` R" by (rule image_eqI) simp_all lemma prod_fun_imageE [elim!]: assumes major: "c \ map_prod f g ` R" and cases: "\x y. c = (f x, g y) \ (x, y) \ R \ P" shows P - apply (rule major [THEN imageE]) - apply (case_tac x) - apply (rule cases) - apply simp_all - done +proof (rule major [THEN imageE]) + fix x + assume "c = map_prod f g x" "x \ R" + then show P + using cases by (cases x) simp +qed definition apfst :: "('a \ 'c) \ 'a \ 'b \ 'c \ 'b" where "apfst f = map_prod f id" definition apsnd :: "('b \ 'c) \ 'a \ 'b \ 'a \ 'c" where "apsnd f = map_prod id f" lemma apfst_conv [simp, code]: "apfst f (x, y) = (f x, y)" by (simp add: apfst_def) lemma apsnd_conv [simp, code]: "apsnd f (x, y) = (x, f y)" by (simp add: apsnd_def) lemma fst_apfst [simp]: "fst (apfst f x) = f (fst x)" by (cases x) simp lemma fst_comp_apfst [simp]: "fst \ apfst f = f \ fst" by (simp add: fun_eq_iff) lemma fst_apsnd [simp]: "fst (apsnd f x) = fst x" by (cases x) simp lemma fst_comp_apsnd [simp]: "fst \ apsnd f = fst" by (simp add: fun_eq_iff) lemma snd_apfst [simp]: "snd (apfst f x) = snd x" by (cases x) simp lemma snd_comp_apfst [simp]: "snd \ apfst f = snd" by (simp add: fun_eq_iff) lemma snd_apsnd [simp]: "snd (apsnd f x) = f (snd x)" by (cases x) simp lemma snd_comp_apsnd [simp]: "snd \ apsnd f = f \ snd" by (simp add: fun_eq_iff) lemma apfst_compose: "apfst f (apfst g x) = apfst (f \ g) x" by (cases x) simp lemma apsnd_compose: "apsnd f (apsnd g x) = apsnd (f \ g) x" by (cases x) simp lemma apfst_apsnd [simp]: "apfst f (apsnd g x) = (f (fst x), g (snd x))" by (cases x) simp lemma apsnd_apfst [simp]: "apsnd f (apfst g x) = (g (fst x), f (snd x))" by (cases x) simp lemma apfst_id [simp]: "apfst id = id" by (simp add: fun_eq_iff) lemma apsnd_id [simp]: "apsnd id = id" by (simp add: fun_eq_iff) lemma apfst_eq_conv [simp]: "apfst f x = apfst g x \ f (fst x) = g (fst x)" by (cases x) simp lemma apsnd_eq_conv [simp]: "apsnd f x = apsnd g x \ f (snd x) = g (snd x)" by (cases x) simp lemma apsnd_apfst_commute: "apsnd f (apfst g p) = apfst g (apsnd f p)" by simp context begin local_setup \Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")\ definition swap :: "'a \ 'b \ 'b \ 'a" where "swap p = (snd p, fst p)" end lemma swap_simp [simp]: "prod.swap (x, y) = (y, x)" by (simp add: prod.swap_def) lemma swap_swap [simp]: "prod.swap (prod.swap p) = p" by (cases p) simp lemma swap_comp_swap [simp]: "prod.swap \ prod.swap = id" by (simp add: fun_eq_iff) lemma pair_in_swap_image [simp]: "(y, x) \ prod.swap ` A \ (x, y) \ A" by (auto intro!: image_eqI) lemma inj_swap [simp]: "inj_on prod.swap A" by (rule inj_onI) auto lemma swap_inj_on: "inj_on (\(i, j). (j, i)) A" by (rule inj_onI) auto lemma surj_swap [simp]: "surj prod.swap" by (rule surjI [of _ prod.swap]) simp lemma bij_swap [simp]: "bij prod.swap" by (simp add: bij_def) lemma case_swap [simp]: "(case prod.swap p of (y, x) \ f x y) = (case p of (x, y) \ f x y)" by (cases p) simp lemma fst_swap [simp]: "fst (prod.swap x) = snd x" by (cases x) simp lemma snd_swap [simp]: "snd (prod.swap x) = fst x" by (cases x) simp text \Disjoint union of a family of sets -- Sigma.\ definition Sigma :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" where "Sigma A B \ \x\A. \y\B x. {Pair x y}" abbreviation Times :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 80) where "A \ B \ Sigma A (\_. B)" hide_const (open) Times bundle no_Set_Product_syntax begin no_notation Product_Type.Times (infixr "\" 80) end bundle Set_Product_syntax begin notation Product_Type.Times (infixr "\" 80) end syntax "_Sigma" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) translations "SIGMA x:A. B" \ "CONST Sigma A (\x. B)" lemma SigmaI [intro!]: "a \ A \ b \ B a \ (a, b) \ Sigma A B" unfolding Sigma_def by blast lemma SigmaE [elim!]: "c \ Sigma A B \ (\x y. x \ A \ y \ B x \ c = (x, y) \ P) \ P" \ \The general elimination rule.\ unfolding Sigma_def by blast text \ Elimination of \<^term>\(a, b) \ A \ B\ -- introduces no eigenvariables. \ lemma SigmaD1: "(a, b) \ Sigma A B \ a \ A" by blast lemma SigmaD2: "(a, b) \ Sigma A B \ b \ B a" by blast lemma SigmaE2: "(a, b) \ Sigma A B \ (a \ A \ b \ B a \ P) \ P" by blast lemma Sigma_cong: "A = B \ (\x. x \ B \ C x = D x) \ (SIGMA x:A. C x) = (SIGMA x:B. D x)" by auto lemma Sigma_mono: "A \ C \ (\x. x \ A \ B x \ D x) \ Sigma A B \ Sigma C D" by blast lemma Sigma_empty1 [simp]: "Sigma {} B = {}" by blast lemma Sigma_empty2 [simp]: "A \ {} = {}" by blast lemma UNIV_Times_UNIV [simp]: "UNIV \ UNIV = UNIV" by auto lemma Compl_Times_UNIV1 [simp]: "- (UNIV \ A) = UNIV \ (-A)" by auto lemma Compl_Times_UNIV2 [simp]: "- (A \ UNIV) = (-A) \ UNIV" by auto lemma mem_Sigma_iff [iff]: "(a, b) \ Sigma A B \ a \ A \ b \ B a" by blast lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" by (induct x) simp lemma Sigma_empty_iff: "(SIGMA i:I. X i) = {} \ (\i\I. X i = {})" by auto lemma Times_subset_cancel2: "x \ C \ A \ C \ B \ C \ A \ B" by blast lemma Times_eq_cancel2: "x \ C \ A \ C = B \ C \ A = B" by (blast elim: equalityE) lemma Collect_case_prod_Sigma: "{(x, y). P x \ Q x y} = (SIGMA x:Collect P. Collect (Q x))" by blast lemma Collect_case_prod [simp]: "{(a, b). P a \ Q b} = Collect P \ Collect Q " by (fact Collect_case_prod_Sigma) lemma Collect_case_prodD: "x \ Collect (case_prod A) \ A (fst x) (snd x)" by auto lemma Collect_case_prod_mono: "A \ B \ Collect (case_prod A) \ Collect (case_prod B)" by auto (auto elim!: le_funE) lemma Collect_split_mono_strong: "X = fst ` A \ Y = snd ` A \ \a\X. \b \ Y. P a b \ Q a b \ A \ Collect (case_prod P) \ A \ Collect (case_prod Q)" by fastforce lemma UN_Times_distrib: "(\(a, b)\A \ B. E a \ F b) = \(E ` A) \ \(F ` B)" \ \Suggested by Pierre Chartier\ by blast lemma split_paired_Ball_Sigma [simp, no_atp]: "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" by blast lemma split_paired_Bex_Sigma [simp, no_atp]: "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" by blast lemma Sigma_Un_distrib1: "Sigma (I \ J) C = Sigma I C \ Sigma J C" by blast lemma Sigma_Un_distrib2: "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" by blast lemma Sigma_Int_distrib1: "Sigma (I \ J) C = Sigma I C \ Sigma J C" by blast lemma Sigma_Int_distrib2: "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" by blast lemma Sigma_Diff_distrib1: "Sigma (I - J) C = Sigma I C - Sigma J C" by blast lemma Sigma_Diff_distrib2: "(SIGMA i:I. A i - B i) = Sigma I A - Sigma I B" by blast lemma Sigma_Union: "Sigma (\X) B = (\A\X. Sigma A B)" by blast lemma Pair_vimage_Sigma: "Pair x -` Sigma A f = (if x \ A then f x else {})" by auto text \ Non-dependent versions are needed to avoid the need for higher-order matching, especially when the rules are re-oriented. \ lemma Times_Un_distrib1: "(A \ B) \ C = A \ C \ B \ C " by (fact Sigma_Un_distrib1) lemma Times_Int_distrib1: "(A \ B) \ C = A \ C \ B \ C " by (fact Sigma_Int_distrib1) lemma Times_Diff_distrib1: "(A - B) \ C = A \ C - B \ C " by (fact Sigma_Diff_distrib1) lemma Times_empty [simp]: "A \ B = {} \ A = {} \ B = {}" by auto lemma times_subset_iff: "A \ C \ B \ D \ A={} \ C={} \ A \ B \ C \ D" by blast lemma times_eq_iff: "A \ B = C \ D \ A = C \ B = D \ (A = {} \ B = {}) \ (C = {} \ D = {})" by auto lemma fst_image_times [simp]: "fst ` (A \ B) = (if B = {} then {} else A)" by force lemma snd_image_times [simp]: "snd ` (A \ B) = (if A = {} then {} else B)" by force lemma fst_image_Sigma: "fst ` (Sigma A B) = {x \ A. B(x) \ {}}" by force lemma snd_image_Sigma: "snd ` (Sigma A B) = (\ x \ A. B x)" by force lemma vimage_fst: "fst -` A = A \ UNIV" by auto lemma vimage_snd: "snd -` A = UNIV \ A" by auto lemma insert_Times_insert [simp]: "insert a A \ insert b B = insert (a,b) (A \ insert b B \ insert a A \ B)" by blast lemma vimage_Times: "f -` (A \ B) = (fst \ f) -` A \ (snd \ f) -` B" proof (rule set_eqI) show "x \ f -` (A \ B) \ x \ (fst \ f) -` A \ (snd \ f) -` B" for x by (cases "f x") (auto split: prod.split) qed lemma Times_Int_Times: "A \ B \ C \ D = (A \ C) \ (B \ D)" by auto lemma image_paired_Times: "(\(x,y). (f x, g y)) ` (A \ B) = (f ` A) \ (g ` B)" by auto lemma product_swap: "prod.swap ` (A \ B) = B \ A" by (auto simp add: set_eq_iff) lemma swap_product: "(\(i, j). (j, i)) ` (A \ B) = B \ A" by (auto simp add: set_eq_iff) lemma image_split_eq_Sigma: "(\x. (f x, g x)) ` A = Sigma (f ` A) (\x. g ` (f -` {x} \ A))" proof (safe intro!: imageI) fix a b assume *: "a \ A" "b \ A" and eq: "f a = f b" show "(f b, g a) \ (\x. (f x, g x)) ` A" using * eq[symmetric] by auto qed simp_all lemma subset_fst_snd: "A \ (fst ` A \ snd ` A)" by force lemma inj_on_apfst [simp]: "inj_on (apfst f) (A \ UNIV) \ inj_on f A" by (auto simp add: inj_on_def) lemma inj_apfst [simp]: "inj (apfst f) \ inj f" using inj_on_apfst[of f UNIV] by simp lemma inj_on_apsnd [simp]: "inj_on (apsnd f) (UNIV \ A) \ inj_on f A" by (auto simp add: inj_on_def) lemma inj_apsnd [simp]: "inj (apsnd f) \ inj f" using inj_on_apsnd[of f UNIV] by simp context begin qualified definition product :: "'a set \ 'b set \ ('a \ 'b) set" where [code_abbrev]: "product A B = A \ B" lemma member_product: "x \ Product_Type.product A B \ x \ A \ B" by (simp add: product_def) end text \The following \<^const>\map_prod\ lemmas are due to Joachim Breitner:\ lemma map_prod_inj_on: assumes "inj_on f A" and "inj_on g B" shows "inj_on (map_prod f g) (A \ B)" proof (rule inj_onI) fix x :: "'a \ 'c" fix y :: "'a \ 'c" assume "x \ A \ B" then have "fst x \ A" and "snd x \ B" by auto assume "y \ A \ B" then have "fst y \ A" and "snd y \ B" by auto assume "map_prod f g x = map_prod f g y" then have "fst (map_prod f g x) = fst (map_prod f g y)" by auto then have "f (fst x) = f (fst y)" by (cases x, cases y) auto with \inj_on f A\ and \fst x \ A\ and \fst y \ A\ have "fst x = fst y" by (auto dest: inj_onD) moreover from \map_prod f g x = map_prod f g y\ have "snd (map_prod f g x) = snd (map_prod f g y)" by auto then have "g (snd x) = g (snd y)" by (cases x, cases y) auto with \inj_on g B\ and \snd x \ B\ and \snd y \ B\ have "snd x = snd y" by (auto dest: inj_onD) ultimately show "x = y" by (rule prod_eqI) qed lemma map_prod_surj: fixes f :: "'a \ 'b" and g :: "'c \ 'd" assumes "surj f" and "surj g" shows "surj (map_prod f g)" unfolding surj_def proof fix y :: "'b \ 'd" from \surj f\ obtain a where "fst y = f a" by (auto elim: surjE) moreover from \surj g\ obtain b where "snd y = g b" by (auto elim: surjE) ultimately have "(fst y, snd y) = map_prod f g (a,b)" by auto then show "\x. y = map_prod f g x" by auto qed lemma map_prod_surj_on: assumes "f ` A = A'" and "g ` B = B'" shows "map_prod f g ` (A \ B) = A' \ B'" unfolding image_def proof (rule set_eqI, rule iffI) fix x :: "'a \ 'c" assume "x \ {y::'a \ 'c. \x::'b \ 'd\A \ B. y = map_prod f g x}" then obtain y where "y \ A \ B" and "x = map_prod f g y" by blast from \image f A = A'\ and \y \ A \ B\ have "f (fst y) \ A'" by auto moreover from \image g B = B'\ and \y \ A \ B\ have "g (snd y) \ B'" by auto ultimately have "(f (fst y), g (snd y)) \ (A' \ B')" by auto with \x = map_prod f g y\ show "x \ A' \ B'" by (cases y) auto next fix x :: "'a \ 'c" assume "x \ A' \ B'" then have "fst x \ A'" and "snd x \ B'" by auto from \image f A = A'\ and \fst x \ A'\ have "fst x \ image f A" by auto then obtain a where "a \ A" and "fst x = f a" by (rule imageE) moreover from \image g B = B'\ and \snd x \ B'\ obtain b where "b \ B" and "snd x = g b" by auto ultimately have "(fst x, snd x) = map_prod f g (a, b)" by auto moreover from \a \ A\ and \b \ B\ have "(a , b) \ A \ B" by auto ultimately have "\y \ A \ B. x = map_prod f g y" by auto then show "x \ {x. \y \ A \ B. x = map_prod f g y}" by auto qed subsection \Simproc for rewriting a set comprehension into a pointfree expression\ ML_file \Tools/set_comprehension_pointfree.ML\ setup \ Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs [Simplifier.make_simproc \<^context> "set comprehension" {lhss = [\<^term>\Collect P\], proc = K Set_Comprehension_Pointfree.code_simproc}]) \ subsection \Lemmas about disjointness\ lemma disjnt_Times1_iff [simp]: "disjnt (C \ A) (C \ B) \ C = {} \ disjnt A B" by (auto simp: disjnt_def) lemma disjnt_Times2_iff [simp]: "disjnt (A \ C) (B \ C) \ C = {} \ disjnt A B" by (auto simp: disjnt_def) lemma disjnt_Sigma_iff: "disjnt (Sigma A C) (Sigma B C) \ (\i \ A\B. C i = {}) \ disjnt A B" by (auto simp: disjnt_def) subsection \Inductively defined sets\ (* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *) simproc_setup Collect_mem ("Collect t") = \ fn _ => fn ctxt => fn ct => (case Thm.term_of ct of S as Const (\<^const_name>\Collect\, Type (\<^type_name>\fun\, [_, T])) $ t => let val (u, _, ps) = HOLogic.strip_ptupleabs t in (case u of (c as Const (\<^const_name>\Set.member\, _)) $ q $ S' => (case try (HOLogic.strip_ptuple ps) q of NONE => NONE | SOME ts => if not (Term.is_open S') andalso ts = map Bound (length ps downto 0) then let val simp = full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_paired_all}, @{thm case_prod_conv}]) 1 in SOME (Goal.prove ctxt [] [] (Const (\<^const_name>\Pure.eq\, T --> T --> propT) $ S $ S') (K (EVERY [resolve_tac ctxt [eq_reflection] 1, resolve_tac ctxt @{thms subset_antisym} 1, resolve_tac ctxt @{thms subsetI} 1, dresolve_tac ctxt @{thms CollectD} 1, simp, resolve_tac ctxt @{thms subsetI} 1, resolve_tac ctxt @{thms CollectI} 1, simp]))) end else NONE) | _ => NONE) end | _ => NONE) \ ML_file \Tools/inductive_set.ML\ subsection \Legacy theorem bindings and duplicates\ lemmas fst_conv = prod.sel(1) lemmas snd_conv = prod.sel(2) lemmas split_def = case_prod_unfold lemmas split_beta' = case_prod_beta' lemmas split_beta = prod.case_eq_if lemmas split_conv = case_prod_conv lemmas split = case_prod_conv hide_const (open) prod end diff --git a/src/HOL/Quotient.thy b/src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy +++ b/src/HOL/Quotient.thy @@ -1,883 +1,883 @@ (* Title: HOL/Quotient.thy Author: Cezary Kaliszyk and Christian Urban *) section \Definition of Quotient Types\ theory Quotient imports Lifting keywords "print_quotmapsQ3" "print_quotientsQ3" "print_quotconsts" :: diag and "quotient_type" :: thy_goal_defn and "/" and "quotient_definition" :: thy_goal_defn and "copy_bnf" :: thy_defn and "lift_bnf" :: thy_goal_defn begin text \ Basic definition for equivalence relations that are represented by predicates. \ text \Composition of Relations\ abbreviation rel_conj :: "('a \ 'b \ bool) \ ('b \ 'a \ bool) \ 'a \ 'b \ bool" (infixr "OOO" 75) where "r1 OOO r2 \ r1 OO r2 OO r1" lemma eq_comp_r: shows "((=) OOO R) = R" by (auto simp add: fun_eq_iff) context includes lifting_syntax begin subsection \Quotient Predicate\ definition "Quotient3 R Abs Rep \ (\a. Abs (Rep a) = a) \ (\a. R (Rep a) (Rep a)) \ (\r s. R r s \ R r r \ R s s \ Abs r = Abs s)" lemma Quotient3I: assumes "\a. Abs (Rep a) = a" and "\a. R (Rep a) (Rep a)" and "\r s. R r s \ R r r \ R s s \ Abs r = Abs s" shows "Quotient3 R Abs Rep" using assms unfolding Quotient3_def by blast context fixes R Abs Rep assumes a: "Quotient3 R Abs Rep" begin lemma Quotient3_abs_rep: "Abs (Rep a) = a" using a unfolding Quotient3_def by simp lemma Quotient3_rep_reflp: "R (Rep a) (Rep a)" using a unfolding Quotient3_def by blast lemma Quotient3_rel: "R r r \ R s s \ Abs r = Abs s \ R r s" \ \orientation does not loop on rewriting\ using a unfolding Quotient3_def by blast lemma Quotient3_refl1: "R r s \ R r r" using a unfolding Quotient3_def by fast lemma Quotient3_refl2: "R r s \ R s s" using a unfolding Quotient3_def by fast lemma Quotient3_rel_rep: "R (Rep a) (Rep b) \ a = b" using a unfolding Quotient3_def by metis lemma Quotient3_rep_abs: "R r r \ R (Rep (Abs r)) r" using a unfolding Quotient3_def by blast lemma Quotient3_rel_abs: "R r s \ Abs r = Abs s" using a unfolding Quotient3_def by blast lemma Quotient3_symp: "symp R" using a unfolding Quotient3_def using sympI by metis lemma Quotient3_transp: "transp R" using a unfolding Quotient3_def using transpI by (metis (full_types)) lemma Quotient3_part_equivp: "part_equivp R" by (metis Quotient3_rep_reflp Quotient3_symp Quotient3_transp part_equivpI) lemma abs_o_rep: "Abs \ Rep = id" unfolding fun_eq_iff by (simp add: Quotient3_abs_rep) lemma equals_rsp: assumes b: "R xa xb" "R ya yb" shows "R xa ya = R xb yb" using b Quotient3_symp Quotient3_transp by (blast elim: sympE transpE) lemma rep_abs_rsp: assumes b: "R x1 x2" shows "R x1 (Rep (Abs x2))" using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp by metis lemma rep_abs_rsp_left: assumes b: "R x1 x2" shows "R (Rep (Abs x1)) x2" using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp by metis end lemma identity_quotient3: "Quotient3 (=) id id" unfolding Quotient3_def id_def by blast lemma fun_quotient3: assumes q1: "Quotient3 R1 abs1 rep1" and q2: "Quotient3 R2 abs2 rep2" shows "Quotient3 (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)" proof - have "(rep1 ---> abs2) ((abs1 ---> rep2) a) = a" for a using q1 q2 by (simp add: Quotient3_def fun_eq_iff) moreover have "(R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" for a by (rule rel_funI) - (insert q1 q2 Quotient3_rel_abs [of R1 abs1 rep1] Quotient3_rel_rep [of R2 abs2 rep2], - simp (no_asm) add: Quotient3_def, simp) + (use q1 q2 Quotient3_rel_abs [of R1 abs1 rep1] Quotient3_rel_rep [of R2 abs2 rep2] + in \simp (no_asm) add: Quotient3_def, simp\) moreover have "(R1 ===> R2) r s = ((R1 ===> R2) r r \ (R1 ===> R2) s s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s)" for r s proof - have "(R1 ===> R2) r s \ (R1 ===> R2) r r" unfolding rel_fun_def using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] by (metis (full_types) part_equivp_def) moreover have "(R1 ===> R2) r s \ (R1 ===> R2) s s" unfolding rel_fun_def using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] by (metis (full_types) part_equivp_def) moreover have "(R1 ===> R2) r s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s" by (auto simp add: rel_fun_def fun_eq_iff) (use q1 q2 in \unfold Quotient3_def, metis\) moreover have "((R1 ===> R2) r r \ (R1 ===> R2) s s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s) \ (R1 ===> R2) r s" by (auto simp add: rel_fun_def fun_eq_iff) (use q1 q2 in \unfold Quotient3_def, metis map_fun_apply\) ultimately show ?thesis by blast qed ultimately show ?thesis by (intro Quotient3I) (assumption+) qed lemma lambda_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) = (\x. f x)" unfolding fun_eq_iff using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by simp lemma lambda_prs1: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep1 ---> Abs2) (\x. (Abs1 ---> Rep2) f x) = (\x. f x)" unfolding fun_eq_iff using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by simp text\ In the following theorem R1 can be instantiated with anything, but we know some of the types of the Rep and Abs functions; so by solving Quotient assumptions we can get a unique R1 that will be provable; which is why we need to use \apply_rsp\ and not the primed version\ lemma apply_rspQ3: fixes f g::"'a \ 'c" assumes q: "Quotient3 R1 Abs1 Rep1" and a: "(R1 ===> R2) f g" "R1 x y" shows "R2 (f x) (g y)" using a by (auto elim: rel_funE) lemma apply_rspQ3'': assumes "Quotient3 R Abs Rep" and "(R ===> S) f f" shows "S (f (Rep x)) (f (Rep x))" proof - from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient3_rep_reflp) then show ?thesis using assms(2) by (auto intro: apply_rsp') qed subsection \lemmas for regularisation of ball and bex\ lemma ball_reg_eqv: fixes P :: "'a \ bool" assumes a: "equivp R" shows "Ball (Respects R) P = (All P)" using a unfolding equivp_def by (auto simp add: in_respects) lemma bex_reg_eqv: fixes P :: "'a \ bool" assumes a: "equivp R" shows "Bex (Respects R) P = (Ex P)" using a unfolding equivp_def by (auto simp add: in_respects) lemma ball_reg_right: assumes a: "\x. x \ R \ P x \ Q x" shows "All P \ Ball R Q" using a by fast lemma bex_reg_left: assumes a: "\x. x \ R \ Q x \ P x" shows "Bex R Q \ Ex P" using a by fast lemma ball_reg_left: assumes a: "equivp R" shows "(\x. (Q x \ P x)) \ Ball (Respects R) Q \ All P" using a by (metis equivp_reflp in_respects) lemma bex_reg_right: assumes a: "equivp R" shows "(\x. (Q x \ P x)) \ Ex Q \ Bex (Respects R) P" using a by (metis equivp_reflp in_respects) lemma ball_reg_eqv_range: fixes P::"'a \ bool" and x::"'a" assumes a: "equivp R2" shows "(Ball (Respects (R1 ===> R2)) (\f. P (f x)) = All (\f. P (f x)))" proof (intro allI iffI) fix f assume "\f \ Respects (R1 ===> R2). P (f x)" moreover have "(\y. f x) \ Respects (R1 ===> R2)" using a equivp_reflp_symp_transp[of "R2"] by(auto simp add: in_respects rel_fun_def elim: equivpE reflpE) ultimately show "P (f x)" by auto qed auto lemma bex_reg_eqv_range: assumes a: "equivp R2" shows "(Bex (Respects (R1 ===> R2)) (\f. P (f x)) = Ex (\f. P (f x)))" proof - { fix f assume "P (f x)" have "(\y. f x) \ Respects (R1 ===> R2)" using a equivp_reflp_symp_transp[of "R2"] by (auto simp add: Respects_def in_respects rel_fun_def elim: equivpE reflpE) } then show ?thesis by auto qed (* Next four lemmas are unused *) lemma all_reg: assumes a: "\x :: 'a. (P x \ Q x)" and b: "All P" shows "All Q" using a b by fast lemma ex_reg: assumes a: "\x :: 'a. (P x \ Q x)" and b: "Ex P" shows "Ex Q" using a b by fast lemma ball_reg: assumes a: "\x :: 'a. (x \ R \ P x \ Q x)" and b: "Ball R P" shows "Ball R Q" using a b by fast lemma bex_reg: assumes a: "\x :: 'a. (x \ R \ P x \ Q x)" and b: "Bex R P" shows "Bex R Q" using a b by fast lemma ball_all_comm: assumes "\y. (\x\P. A x y) \ (\x. B x y)" shows "(\x\P. \y. A x y) \ (\x. \y. B x y)" using assms by auto lemma bex_ex_comm: assumes "(\y. \x. A x y) \ (\y. \x\P. B x y)" shows "(\x. \y. A x y) \ (\x\P. \y. B x y)" using assms by auto subsection \Bounded abstraction\ definition Babs :: "'a set \ ('a \ 'b) \ 'a \ 'b" where "x \ p \ Babs p m x = m x" lemma babs_rsp: assumes q: "Quotient3 R1 Abs1 Rep1" - and a: "(R1 ===> R2) f g" - shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" -proof (clarsimp simp add: Babs_def in_respects rel_fun_def) + and a: "(R1 ===> R2) f g" + shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" +proof fix x y assume "R1 x y" then have "x \ Respects R1 \ y \ Respects R1" unfolding in_respects rel_fun_def using Quotient3_rel[OF q]by metis then show "R2 (Babs (Respects R1) f x) (Babs (Respects R1) g y)" using \R1 x y\ a by (simp add: Babs_def rel_fun_def) qed lemma babs_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f" proof - { fix x have "Rep1 x \ Respects R1" by (simp add: in_respects Quotient3_rel_rep[OF q1]) then have "Abs2 (Babs (Respects R1) ((Abs1 ---> Rep2) f) (Rep1 x)) = f x" by (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2]) } then show ?thesis by force qed lemma babs_simp: assumes q: "Quotient3 R1 Abs Rep" shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding rel_fun_def by (metis Babs_def in_respects Quotient3_rel[OF q]) qed (simp add: babs_rsp[OF q]) text \If a user proves that a particular functional relation is an equivalence, this may be useful in regularising\ lemma babs_reg_eqv: shows "equivp R \ Babs (Respects R) P = P" by (simp add: fun_eq_iff Babs_def in_respects equivp_reflp) (* 3 lemmas needed for proving repabs_inj *) lemma ball_rsp: assumes a: "(R ===> (=)) f g" shows "Ball (Respects R) f = Ball (Respects R) g" using a by (auto simp add: Ball_def in_respects elim: rel_funE) lemma bex_rsp: assumes a: "(R ===> (=)) f g" shows "(Bex (Respects R) f = Bex (Respects R) g)" using a by (auto simp add: Bex_def in_respects elim: rel_funE) lemma bex1_rsp: assumes a: "(R ===> (=)) f g" shows "Ex1 (\x. x \ Respects R \ f x) = Ex1 (\x. x \ Respects R \ g x)" using a by (auto elim: rel_funE simp add: Ex1_def in_respects) text \Two lemmas needed for cleaning of quantifiers\ lemma all_prs: assumes a: "Quotient3 R absf repf" shows "Ball (Respects R) ((absf ---> id) f) = All f" using a unfolding Quotient3_def Ball_def in_respects id_apply comp_def map_fun_def by metis lemma ex_prs: assumes a: "Quotient3 R absf repf" shows "Bex (Respects R) ((absf ---> id) f) = Ex f" using a unfolding Quotient3_def Bex_def in_respects id_apply comp_def map_fun_def by metis subsection \\Bex1_rel\ quantifier\ definition Bex1_rel :: "('a \ 'a \ bool) \ ('a \ bool) \ bool" where "Bex1_rel R P \ (\x \ Respects R. P x) \ (\x \ Respects R. \y \ Respects R. ((P x \ P y) \ (R x y)))" lemma bex1_rel_aux: "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R x\ \ Bex1_rel R y" unfolding Bex1_rel_def by (metis in_respects) lemma bex1_rel_aux2: "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R y\ \ Bex1_rel R x" unfolding Bex1_rel_def by (metis in_respects) lemma bex1_rel_rsp: assumes a: "Quotient3 R absf repf" shows "((R ===> (=)) ===> (=)) (Bex1_rel R) (Bex1_rel R)" unfolding rel_fun_def by (metis bex1_rel_aux bex1_rel_aux2) lemma ex1_prs: assumes "Quotient3 R absf repf" shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f" (is "?lhs = ?rhs") using assms apply (auto simp add: Bex1_rel_def Respects_def) by (metis (full_types) Quotient3_def)+ lemma bex1_bexeq_reg: shows "(\!x\Respects R. P x) \ (Bex1_rel R (\x. P x))" by (auto simp add: Ex1_def Bex1_rel_def Bex_def Ball_def in_respects) lemma bex1_bexeq_reg_eqv: assumes a: "equivp R" shows "(\!x. P x) \ Bex1_rel R P" using equivp_reflp[OF a] by (metis (full_types) Bex1_rel_def in_respects) subsection \Various respects and preserve lemmas\ lemma quot_rel_rsp: assumes a: "Quotient3 R Abs Rep" shows "(R ===> R ===> (=)) R R" apply(rule rel_funI)+ by (meson assms equals_rsp) lemma o_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" and q3: "Quotient3 R3 Abs3 Rep3" shows "((Abs2 ---> Rep3) ---> (Abs1 ---> Rep2) ---> (Rep1 ---> Abs3)) (\) = (\)" and "(id ---> (Abs1 ---> id) ---> Rep1 ---> id) (\) = (\)" using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] Quotient3_abs_rep[OF q3] by (simp_all add: fun_eq_iff) lemma o_rsp: "((R2 ===> R3) ===> (R1 ===> R2) ===> (R1 ===> R3)) (\) (\)" "((=) ===> (R1 ===> (=)) ===> R1 ===> (=)) (\) (\)" by (force elim: rel_funE)+ lemma cond_prs: assumes a: "Quotient3 R absf repf" shows "absf (if a then repf b else repf c) = (if a then b else c)" using a unfolding Quotient3_def by auto lemma if_prs: assumes q: "Quotient3 R Abs Rep" shows "(id ---> Rep ---> Rep ---> Abs) If = If" using Quotient3_abs_rep[OF q] by (auto simp add: fun_eq_iff) lemma if_rsp: assumes q: "Quotient3 R Abs Rep" shows "((=) ===> R ===> R ===> R) If If" by force lemma let_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep2 ---> (Abs2 ---> Rep1) ---> Abs1) Let = Let" using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by (auto simp add: fun_eq_iff) lemma let_rsp: shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let" by (force elim: rel_funE) lemma id_rsp: shows "(R ===> R) id id" by auto lemma id_prs: assumes a: "Quotient3 R Abs Rep" shows "(Rep ---> Abs) id = id" by (simp add: fun_eq_iff Quotient3_abs_rep [OF a]) end locale quot_type = fixes R :: "'a \ 'a \ bool" and Abs :: "'a set \ 'b" and Rep :: "'b \ 'a set" assumes equivp: "part_equivp R" and rep_prop: "\y. \x. R x x \ Rep y = Collect (R x)" and rep_inverse: "\x. Abs (Rep x) = x" and abs_inverse: "\c. (\x. ((R x x) \ (c = Collect (R x)))) \ (Rep (Abs c)) = c" and rep_inject: "\x y. (Rep x = Rep y) = (x = y)" begin definition abs :: "'a \ 'b" where "abs x = Abs (Collect (R x))" definition rep :: "'b \ 'a" where "rep a = (SOME x. x \ Rep a)" lemma some_collect: assumes "R r r" shows "R (SOME x. x \ Collect (R r)) = R r" apply simp by (metis assms exE_some equivp[simplified part_equivp_def]) lemma Quotient: shows "Quotient3 R abs rep" unfolding Quotient3_def abs_def rep_def proof (intro conjI allI) fix a r s show x: "R (SOME x. x \ Rep a) (SOME x. x \ Rep a)" proof - obtain x where r: "R x x" and rep: "Rep a = Collect (R x)" using rep_prop[of a] by auto have "R (SOME x. x \ Rep a) x" using r rep some_collect by metis then have "R x (SOME x. x \ Rep a)" using part_equivp_symp[OF equivp] by fast then show "R (SOME x. x \ Rep a) (SOME x. x \ Rep a)" using part_equivp_transp[OF equivp] by (metis \R (SOME x. x \ Rep a) x\) qed have "Collect (R (SOME x. x \ Rep a)) = (Rep a)" by (metis some_collect rep_prop) then show "Abs (Collect (R (SOME x. x \ Rep a))) = a" using rep_inverse by auto have "R r r \ R s s \ Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" proof - assume "R r r" and "R s s" then have "Abs (Collect (R r)) = Abs (Collect (R s)) \ Collect (R r) = Collect (R s)" by (metis abs_inverse) also have "Collect (R r) = Collect (R s) \ (\A x. x \ A) (Collect (R r)) = (\A x. x \ A) (Collect (R s))" - by rule simp_all + by (rule iffI) simp_all finally show "Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" by simp qed then show "R r s \ R r r \ R s s \ (Abs (Collect (R r)) = Abs (Collect (R s)))" using equivp[simplified part_equivp_def] by metis qed end subsection \Quotient composition\ lemma OOO_quotient3: fixes R1 :: "'a \ 'a \ bool" fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" fixes R2' :: "'a \ 'a \ bool" fixes R2 :: "'b \ 'b \ bool" assumes R1: "Quotient3 R1 Abs1 Rep1" assumes R2: "Quotient3 R2 Abs2 Rep2" assumes Abs1: "\x y. R2' x y \ R1 x x \ R1 y y \ R2 (Abs1 x) (Abs1 y)" assumes Rep1: "\x y. R2 x y \ R2' (Rep1 x) (Rep1 y)" shows "Quotient3 (R1 OO R2' OO R1) (Abs2 \ Abs1) (Rep1 \ Rep2)" proof - have *: "(R1 OOO R2') r r \ (R1 OOO R2') s s \ (Abs2 \ Abs1) r = (Abs2 \ Abs1) s \ (R1 OOO R2') r s" for r s proof (intro iffI conjI; clarify) show "(R1 OOO R2') r s" if r: "R1 r a" "R2' a b" "R1 b r" and eq: "(Abs2 \ Abs1) r = (Abs2 \ Abs1) s" and s: "R1 s c" "R2' c d" "R1 d s" for a b c d proof - have "R1 r (Rep1 (Abs1 r))" using r Quotient3_refl1 R1 rep_abs_rsp by fastforce moreover have "R2' (Rep1 (Abs1 r)) (Rep1 (Abs1 s))" using that apply simp apply (metis (full_types) Rep1 Abs1 Quotient3_rel R2 Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1] Quotient3_rel_abs [OF R1]) done moreover have "R1 (Rep1 (Abs1 s)) s" by (metis s Quotient3_rel R1 rep_abs_rsp_left) ultimately show ?thesis by (metis relcomppI) qed next fix x y assume xy: "R1 r x" "R2' x y" "R1 y s" then have "R2 (Abs1 x) (Abs1 y)" by (iprover dest: Abs1 elim: Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1]) then have "R2' (Rep1 (Abs1 x)) (Rep1 (Abs1 x))" "R2' (Rep1 (Abs1 y)) (Rep1 (Abs1 y))" by (simp_all add: Quotient3_refl1 [OF R2] Quotient3_refl2 [OF R2] Rep1) with \R1 r x\ \R1 y s\ show "(R1 OOO R2') r r" "(R1 OOO R2') s s" by (metis (full_types) Quotient3_def R1 relcompp.relcompI)+ show "(Abs2 \ Abs1) r = (Abs2 \ Abs1) s" using xy by simp (metis (full_types) Abs1 Quotient3_rel R1 R2) qed show ?thesis apply (rule Quotient3I) using * apply (simp_all add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1]) apply (metis Quotient3_rep_reflp R1 R2 Rep1 relcompp.relcompI) done qed lemma OOO_eq_quotient3: fixes R1 :: "'a \ 'a \ bool" fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" assumes R1: "Quotient3 R1 Abs1 Rep1" assumes R2: "Quotient3 (=) Abs2 Rep2" shows "Quotient3 (R1 OOO (=)) (Abs2 \ Abs1) (Rep1 \ Rep2)" using assms by (rule OOO_quotient3) auto subsection \Quotient3 to Quotient\ lemma Quotient3_to_Quotient: assumes "Quotient3 R Abs Rep" and "T \ \x y. R x x \ Abs x = y" shows "Quotient R Abs Rep T" using assms unfolding Quotient3_def by (intro QuotientI) blast+ lemma Quotient3_to_Quotient_equivp: assumes q: "Quotient3 R Abs Rep" and T_def: "T \ \x y. Abs x = y" and eR: "equivp R" shows "Quotient R Abs Rep T" proof (intro QuotientI) fix a show "Abs (Rep a) = a" using q by(rule Quotient3_abs_rep) next fix a show "R (Rep a) (Rep a)" using q by(rule Quotient3_rep_reflp) next fix r s show "R r s = (R r r \ R s s \ Abs r = Abs s)" using q by(rule Quotient3_rel[symmetric]) next show "T = (\x y. R x x \ Abs x = y)" using T_def equivp_reflp[OF eR] by simp qed subsection \ML setup\ text \Auxiliary data for the quotient package\ named_theorems quot_equiv "equivalence relation theorems" and quot_respect "respectfulness theorems" and quot_preserve "preservation theorems" and id_simps "identity simp rules for maps" and quot_thm "quotient theorems" ML_file \Tools/Quotient/quotient_info.ML\ declare [[mapQ3 "fun" = (rel_fun, fun_quotient3)]] lemmas [quot_thm] = fun_quotient3 lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp id_rsp lemmas [quot_preserve] = if_prs o_prs let_prs id_prs lemmas [quot_equiv] = identity_equivp text \Lemmas about simplifying id's.\ lemmas [id_simps] = id_def[symmetric] map_fun_id id_apply id_o o_id eq_comp_r vimage_id text \Translation functions for the lifting process.\ ML_file \Tools/Quotient/quotient_term.ML\ text \Definitions of the quotient types.\ ML_file \Tools/Quotient/quotient_type.ML\ text \Definitions for quotient constants.\ ML_file \Tools/Quotient/quotient_def.ML\ text \ An auxiliary constant for recording some information about the lifted theorem in a tactic. \ definition Quot_True :: "'a \ bool" where "Quot_True x \ True" lemma shows QT_all: "Quot_True (All P) \ Quot_True P" and QT_ex: "Quot_True (Ex P) \ Quot_True P" and QT_ex1: "Quot_True (Ex1 P) \ Quot_True P" and QT_lam: "Quot_True (\x. P x) \ (\x. Quot_True (P x))" and QT_ext: "(\x. Quot_True (a x) \ f x = g x) \ (Quot_True a \ f = g)" by (simp_all add: Quot_True_def ext) lemma QT_imp: "Quot_True a \ Quot_True b" by (simp add: Quot_True_def) context includes lifting_syntax begin text \Tactics for proving the lifted theorems\ ML_file \Tools/Quotient/quotient_tacs.ML\ end subsection \Methods / Interface\ method_setup lifting = \Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms))\ \lift theorems to quotient types\ method_setup lifting_setup = \Attrib.thm >> (fn thm => fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm))\ \set up the three goals for the quotient lifting procedure\ method_setup descending = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt []))\ \decend theorems to the raw level\ method_setup descending_setup = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt []))\ \set up the three goals for the decending theorems\ method_setup partiality_descending = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt []))\ \decend theorems to the raw level\ method_setup partiality_descending_setup = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt []))\ \set up the three goals for the decending theorems\ method_setup regularize = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt))\ \prove the regularization goals from the quotient lifting procedure\ method_setup injection = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt))\ \prove the rep/abs injection goals from the quotient lifting procedure\ method_setup cleaning = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt))\ \prove the cleaning goals from the quotient lifting procedure\ attribute_setup quot_lifted = \Scan.succeed Quotient_Tacs.lifted_attrib\ \lift theorems to quotient types\ no_notation rel_conj (infixr "OOO" 75) section \Lifting of BNFs\ lemma sum_insert_Inl_unit: "x \ A \ (\y. x = Inr y \ Inr y \ B) \ x \ insert (Inl ()) B" by (cases x) (simp_all) lemma lift_sum_unit_vimage_commute: "insert (Inl ()) (Inr ` f -` A) = map_sum id f -` insert (Inl ()) (Inr ` A)" by (auto simp: map_sum_def split: sum.splits) lemma insert_Inl_int_map_sum_unit: "insert (Inl ()) A \ range (map_sum id f) \ {}" by (auto simp: map_sum_def split: sum.splits) lemma image_map_sum_unit_subset: "A \ insert (Inl ()) (Inr ` B) \ map_sum id f ` A \ insert (Inl ()) (Inr ` f ` B)" by auto lemma subset_lift_sum_unitD: "A \ insert (Inl ()) (Inr ` B) \ Inr x \ A \ x \ B" unfolding insert_def by auto lemma UNIV_sum_unit_conv: "insert (Inl ()) (range Inr) = UNIV" unfolding UNIV_sum UNIV_unit image_insert image_empty Un_insert_left sup_bot.left_neutral.. lemma subset_vimage_image_subset: "A \ f -` B \ f ` A \ B" by auto lemma relcompp_mem_Grp_neq_bot: "A \ range f \ {} \ (\x y. x \ A \ y \ A) OO (Grp UNIV f)\\ \ bot" unfolding Grp_def relcompp_apply fun_eq_iff by blast lemma comp_projr_Inr: "projr \ Inr = id" by auto lemma in_rel_sum_in_image_projr: "B \ {(x,y). rel_sum ((=) :: unit \ unit \ bool) A x y} \ Inr ` C = fst ` B \ snd ` B = Inr ` D \ map_prod projr projr ` B \ {(x,y). A x y}" by (force simp: projr_def image_iff dest!: spec[of _ "Inl ()"] split: sum.splits) lemma subset_rel_sumI: "B \ {(x,y). A x y} \ rel_sum ((=) :: unit => unit => bool) A (if x \ B then Inr (fst x) else Inl ()) (if x \ B then Inr (snd x) else Inl ())" by auto lemma relcompp_eq_Grp_neq_bot: "(=) OO (Grp UNIV f)\\ \ bot" unfolding Grp_def relcompp_apply fun_eq_iff by blast lemma rel_fun_rel_OO1: "(rel_fun Q (rel_fun R (=))) A B \ conversep Q OO A OO R \ B" by (auto simp: rel_fun_def) lemma rel_fun_rel_OO2: "(rel_fun Q (rel_fun R (=))) A B \ Q OO B OO conversep R \ A" by (auto simp: rel_fun_def) lemma rel_sum_eq2_nonempty: "rel_sum (=) A OO rel_sum (=) B \ bot" by (auto simp: fun_eq_iff relcompp_apply intro!: exI[of _ "Inl _"]) lemma rel_sum_eq3_nonempty: "rel_sum (=) A OO (rel_sum (=) B OO rel_sum (=) C) \ bot" by (auto simp: fun_eq_iff relcompp_apply intro!: exI[of _ "Inl _"]) lemma hypsubst: "A = B \ x \ B \ (x \ A \ P) \ P" by simp lemma Quotient_crel_quotient: "Quotient R Abs Rep T \ equivp R \ T \ (\x y. Abs x = y)" by (drule Quotient_cr_rel) (auto simp: fun_eq_iff equivp_reflp intro!: eq_reflection) lemma Quotient_crel_typedef: "Quotient (eq_onp P) Abs Rep T \ T \ (\x y. x = Rep y)" unfolding Quotient_def by (auto 0 4 simp: fun_eq_iff eq_onp_def intro: sym intro!: eq_reflection) lemma Quotient_crel_typecopy: "Quotient (=) Abs Rep T \ T \ (\x y. x = Rep y)" by (subst (asm) eq_onp_True[symmetric]) (rule Quotient_crel_typedef) lemma equivp_add_relconj: assumes equiv: "equivp R" "equivp R'" and le: "S OO T OO U \ R OO STU OO R'" shows "R OO S OO T OO U OO R' \ R OO STU OO R'" proof - have trans: "R OO R \ R" "R' OO R' \ R'" using equiv unfolding equivp_reflp_symp_transp transp_relcompp by blast+ have "R OO S OO T OO U OO R' = R OO (S OO T OO U) OO R'" unfolding relcompp_assoc .. also have "\ \ R OO (R OO STU OO R') OO R'" by (intro le relcompp_mono order_refl) also have "\ \ (R OO R) OO STU OO (R' OO R')" unfolding relcompp_assoc .. also have "\ \ R OO STU OO R'" by (intro trans relcompp_mono order_refl) finally show ?thesis . qed lemma Grp_conversep_eq_onp: "((BNF_Def.Grp UNIV f)\\ OO BNF_Def.Grp UNIV f) = eq_onp (\x. x \ range f)" by (auto simp: fun_eq_iff Grp_def eq_onp_def image_iff) lemma Grp_conversep_nonempty: "(BNF_Def.Grp UNIV f)\\ OO BNF_Def.Grp UNIV f \ bot" by (auto simp: fun_eq_iff Grp_def) lemma relcomppI2: "r a b \ s b c \ t c d \ (r OO s OO t) a d" by (auto) lemma rel_conj_eq_onp: "equivp R \ rel_conj R (eq_onp P) \ R" by (auto simp: eq_onp_def transp_def equivp_def) lemma Quotient_Quotient3: "Quotient R Abs Rep T \ Quotient3 R Abs Rep" unfolding Quotient_def Quotient3_def by blast lemma Quotient_reflp_imp_equivp: "Quotient R Abs Rep T \ reflp R \ equivp R" using Quotient_symp Quotient_transp equivpI by blast lemma Quotient_eq_onp_typedef: "Quotient (eq_onp P) Abs Rep cr \ type_definition Rep Abs {x. P x}" unfolding Quotient_def eq_onp_def by unfold_locales auto lemma Quotient_eq_onp_type_copy: "Quotient (=) Abs Rep cr \ type_definition Rep Abs UNIV" unfolding Quotient_def eq_onp_def by unfold_locales auto ML_file "Tools/BNF/bnf_lift.ML" hide_fact sum_insert_Inl_unit lift_sum_unit_vimage_commute insert_Inl_int_map_sum_unit image_map_sum_unit_subset subset_lift_sum_unitD UNIV_sum_unit_conv subset_vimage_image_subset relcompp_mem_Grp_neq_bot comp_projr_Inr in_rel_sum_in_image_projr subset_rel_sumI relcompp_eq_Grp_neq_bot rel_fun_rel_OO1 rel_fun_rel_OO2 rel_sum_eq2_nonempty rel_sum_eq3_nonempty hypsubst equivp_add_relconj Grp_conversep_eq_onp Grp_conversep_nonempty relcomppI2 rel_conj_eq_onp Quotient_reflp_imp_equivp Quotient_Quotient3 end diff --git a/src/HOL/Relation.thy b/src/HOL/Relation.thy --- a/src/HOL/Relation.thy +++ b/src/HOL/Relation.thy @@ -1,1336 +1,1343 @@ (* Title: HOL/Relation.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer, TU Muenchen *) section \Relations -- as sets of pairs, and binary predicates\ theory Relation imports Finite_Set begin text \A preliminary: classical rules for reasoning on predicates\ declare predicate1I [Pure.intro!, intro!] declare predicate1D [Pure.dest, dest] declare predicate2I [Pure.intro!, intro!] declare predicate2D [Pure.dest, dest] declare bot1E [elim!] declare bot2E [elim!] declare top1I [intro!] declare top2I [intro!] declare inf1I [intro!] declare inf2I [intro!] declare inf1E [elim!] declare inf2E [elim!] declare sup1I1 [intro?] declare sup2I1 [intro?] declare sup1I2 [intro?] declare sup2I2 [intro?] declare sup1E [elim!] declare sup2E [elim!] declare sup1CI [intro!] declare sup2CI [intro!] declare Inf1_I [intro!] declare INF1_I [intro!] declare Inf2_I [intro!] declare INF2_I [intro!] declare Inf1_D [elim] declare INF1_D [elim] declare Inf2_D [elim] declare INF2_D [elim] declare Inf1_E [elim] declare INF1_E [elim] declare Inf2_E [elim] declare INF2_E [elim] declare Sup1_I [intro] declare SUP1_I [intro] declare Sup2_I [intro] declare SUP2_I [intro] declare Sup1_E [elim!] declare SUP1_E [elim!] declare Sup2_E [elim!] declare SUP2_E [elim!] subsection \Fundamental\ subsubsection \Relations as sets of pairs\ type_synonym 'a rel = "('a \ 'a) set" lemma subrelI: "(\x y. (x, y) \ r \ (x, y) \ s) \ r \ s" \ \Version of @{thm [source] subsetI} for binary relations\ by auto lemma lfp_induct2: "(a, b) \ lfp f \ mono f \ (\a b. (a, b) \ f (lfp f \ {(x, y). P x y}) \ P a b) \ P a b" \ \Version of @{thm [source] lfp_induct} for binary relations\ using lfp_induct_set [of "(a, b)" f "case_prod P"] by auto subsubsection \Conversions between set and predicate relations\ lemma pred_equals_eq [pred_set_conv]: "(\x. x \ R) = (\x. x \ S) \ R = S" by (simp add: set_eq_iff fun_eq_iff) lemma pred_equals_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) = (\x y. (x, y) \ S) \ R = S" by (simp add: set_eq_iff fun_eq_iff) lemma pred_subset_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) \ R \ S" by (simp add: subset_iff le_fun_def) lemma pred_subset_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) \ R \ S" by (simp add: subset_iff le_fun_def) lemma bot_empty_eq [pred_set_conv]: "\ = (\x. x \ {})" by (auto simp add: fun_eq_iff) lemma bot_empty_eq2 [pred_set_conv]: "\ = (\x y. (x, y) \ {})" by (auto simp add: fun_eq_iff) lemma top_empty_eq [pred_set_conv]: "\ = (\x. x \ UNIV)" by (auto simp add: fun_eq_iff) lemma top_empty_eq2 [pred_set_conv]: "\ = (\x y. (x, y) \ UNIV)" by (auto simp add: fun_eq_iff) lemma inf_Int_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: inf_fun_def) lemma inf_Int_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" by (simp add: inf_fun_def) lemma sup_Un_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: sup_fun_def) lemma sup_Un_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" by (simp add: sup_fun_def) lemma INF_INT_eq [pred_set_conv]: "(\i\S. (\x. x \ r i)) = (\x. x \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma INF_INT_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma SUP_UN_eq [pred_set_conv]: "(\i\S. (\x. x \ r i)) = (\x. x \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma SUP_UN_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" by (simp add: fun_eq_iff) lemma Inf_INT_eq [pred_set_conv]: "\S = (\x. x \ (\(Collect ` S)))" by (simp add: fun_eq_iff) lemma INF_Int_eq [pred_set_conv]: "(\i\S. (\x. x \ i)) = (\x. x \ \S)" by (simp add: fun_eq_iff) lemma Inf_INT_eq2 [pred_set_conv]: "\S = (\x y. (x, y) \ (\(Collect ` case_prod ` S)))" by (simp add: fun_eq_iff) lemma INF_Int_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ i)) = (\x y. (x, y) \ \S)" by (simp add: fun_eq_iff) lemma Sup_SUP_eq [pred_set_conv]: "\S = (\x. x \ \(Collect ` S))" by (simp add: fun_eq_iff) lemma SUP_Sup_eq [pred_set_conv]: "(\i\S. (\x. x \ i)) = (\x. x \ \S)" by (simp add: fun_eq_iff) lemma Sup_SUP_eq2 [pred_set_conv]: "\S = (\x y. (x, y) \ (\(Collect ` case_prod ` S)))" by (simp add: fun_eq_iff) lemma SUP_Sup_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ i)) = (\x y. (x, y) \ \S)" by (simp add: fun_eq_iff) subsection \Properties of relations\ subsubsection \Reflexivity\ definition refl_on :: "'a set \ 'a rel \ bool" where "refl_on A r \ r \ A \ A \ (\x\A. (x, x) \ r)" abbreviation refl :: "'a rel \ bool" \ \reflexivity over a type\ where "refl \ refl_on UNIV" definition reflp_on :: "'a set \ ('a \ 'a \ bool) \ bool" where "reflp_on A R \ (\x\A. R x x)" abbreviation reflp :: "('a \ 'a \ bool) \ bool" where "reflp \ reflp_on UNIV" lemma reflp_def[no_atp]: "reflp R \ (\x. R x x)" by (simp add: reflp_on_def) text \@{thm [source] reflp_def} is for backward compatibility.\ lemma reflp_refl_eq [pred_set_conv]: "reflp (\x y. (x, y) \ r) \ refl r" by (simp add: refl_on_def reflp_def) lemma refl_onI [intro?]: "r \ A \ A \ (\x. x \ A \ (x, x) \ r) \ refl_on A r" unfolding refl_on_def by (iprover intro!: ballI) lemma reflp_onI: "(\x y. x \ A \ R x x) \ reflp_on A R" by (simp add: reflp_on_def) lemma reflpI[intro?]: "(\x. R x x) \ reflp R" by (rule reflp_onI) lemma refl_onD: "refl_on A r \ a \ A \ (a, a) \ r" unfolding refl_on_def by blast lemma refl_onD1: "refl_on A r \ (x, y) \ r \ x \ A" unfolding refl_on_def by blast lemma refl_onD2: "refl_on A r \ (x, y) \ r \ y \ A" unfolding refl_on_def by blast lemma reflp_onD: "reflp_on A R \ x \ A \ R x x" by (simp add: reflp_on_def) lemma reflpD[dest?]: "reflp R \ R x x" by (simp add: reflp_onD) lemma reflpE: assumes "reflp r" obtains "r x x" using assms by (auto dest: refl_onD simp add: reflp_def) lemma reflp_on_subset: "reflp_on A R \ B \ A \ reflp_on B R" by (auto intro: reflp_onI dest: reflp_onD) lemma refl_on_Int: "refl_on A r \ refl_on B s \ refl_on (A \ B) (r \ s)" unfolding refl_on_def by blast lemma reflp_on_inf: "reflp_on A R \ reflp_on B S \ reflp_on (A \ B) (R \ S)" by (auto intro: reflp_onI dest: reflp_onD) lemma reflp_inf: "reflp r \ reflp s \ reflp (r \ s)" by (rule reflp_on_inf[of UNIV _ UNIV, unfolded Int_absorb]) lemma refl_on_Un: "refl_on A r \ refl_on B s \ refl_on (A \ B) (r \ s)" unfolding refl_on_def by blast lemma reflp_on_sup: "reflp_on A R \ reflp_on B S \ reflp_on (A \ B) (R \ S)" by (auto intro: reflp_onI dest: reflp_onD) lemma reflp_sup: "reflp r \ reflp s \ reflp (r \ s)" by (rule reflp_on_sup[of UNIV _ UNIV, unfolded Un_absorb]) lemma refl_on_INTER: "\x\S. refl_on (A x) (r x) \ refl_on (\(A ` S)) (\(r ` S))" unfolding refl_on_def by fast lemma reflp_on_Inf: "\x\S. reflp_on (A x) (R x) \ reflp_on (\(A ` S)) (\(R ` S))" by (auto intro: reflp_onI dest: reflp_onD) lemma refl_on_UNION: "\x\S. refl_on (A x) (r x) \ refl_on (\(A ` S)) (\(r ` S))" unfolding refl_on_def by blast lemma reflp_on_Sup: "\x\S. reflp_on (A x) (R x) \ reflp_on (\(A ` S)) (\(R ` S))" by (auto intro: reflp_onI dest: reflp_onD) lemma refl_on_empty [simp]: "refl_on {} {}" by (simp add: refl_on_def) lemma reflp_on_empty [simp]: "reflp_on {} R" by (auto intro: reflp_onI) lemma refl_on_singleton [simp]: "refl_on {x} {(x, x)}" by (blast intro: refl_onI) lemma refl_on_def' [nitpick_unfold, code]: "refl_on A r \ (\(x, y) \ r. x \ A \ y \ A) \ (\x \ A. (x, x) \ r)" by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2) lemma reflp_equality [simp]: "reflp (=)" by (simp add: reflp_def) lemma reflp_on_mono: "reflp_on A R \ (\x y. x \ A \ y \ A \ R x y \ Q x y) \ reflp_on A Q" by (auto intro: reflp_onI dest: reflp_onD) lemma reflp_mono: "reflp R \ (\x y. R x y \ Q x y) \ reflp Q" by (rule reflp_on_mono[of UNIV R Q]) simp_all subsubsection \Irreflexivity\ definition irrefl :: "'a rel \ bool" where "irrefl r \ (\a. (a, a) \ r)" definition irreflp :: "('a \ 'a \ bool) \ bool" where "irreflp R \ (\a. \ R a a)" lemma irreflp_irrefl_eq [pred_set_conv]: "irreflp (\a b. (a, b) \ R) \ irrefl R" by (simp add: irrefl_def irreflp_def) lemma irreflI [intro?]: "(\a. (a, a) \ R) \ irrefl R" by (simp add: irrefl_def) lemma irreflpI [intro?]: "(\a. \ R a a) \ irreflp R" by (fact irreflI [to_pred]) lemma irrefl_distinct [code]: "irrefl r \ (\(a, b) \ r. a \ b)" by (auto simp add: irrefl_def) lemma (in preorder) irreflp_less[simp]: "irreflp (<)" by (simp add: irreflpI) lemma (in preorder) irreflp_greater[simp]: "irreflp (>)" by (simp add: irreflpI) subsubsection \Asymmetry\ inductive asym :: "'a rel \ bool" where asymI: "(\a b. (a, b) \ R \ (b, a) \ R) \ asym R" inductive asymp :: "('a \ 'a \ bool) \ bool" where asympI: "(\a b. R a b \ \ R b a) \ asymp R" lemma asymp_asym_eq [pred_set_conv]: "asymp (\a b. (a, b) \ R) \ asym R" by (auto intro!: asymI asympI elim: asym.cases asymp.cases simp add: irreflp_irrefl_eq) lemma asymD: "\asym R; (x,y) \ R\ \ (y,x) \ R" by (simp add: asym.simps) lemma asympD: "asymp R \ R x y \ \ R y x" by (rule asymD[to_pred]) lemma asym_iff: "asym R \ (\x y. (x,y) \ R \ (y,x) \ R)" by (blast intro: asymI dest: asymD) lemma (in preorder) asymp_less[simp]: "asymp (<)" by (auto intro: asympI dual_order.asym) lemma (in preorder) asymp_greater[simp]: "asymp (>)" by (auto intro: asympI dual_order.asym) subsubsection \Symmetry\ definition sym :: "'a rel \ bool" where "sym r \ (\x y. (x, y) \ r \ (y, x) \ r)" definition symp :: "('a \ 'a \ bool) \ bool" where "symp r \ (\x y. r x y \ r y x)" lemma symp_sym_eq [pred_set_conv]: "symp (\x y. (x, y) \ r) \ sym r" by (simp add: sym_def symp_def) lemma symI [intro?]: "(\a b. (a, b) \ r \ (b, a) \ r) \ sym r" by (unfold sym_def) iprover lemma sympI [intro?]: "(\a b. r a b \ r b a) \ symp r" by (fact symI [to_pred]) lemma symE: assumes "sym r" and "(b, a) \ r" obtains "(a, b) \ r" using assms by (simp add: sym_def) lemma sympE: assumes "symp r" and "r b a" obtains "r a b" using assms by (rule symE [to_pred]) lemma symD [dest?]: assumes "sym r" and "(b, a) \ r" shows "(a, b) \ r" using assms by (rule symE) lemma sympD [dest?]: assumes "symp r" and "r b a" shows "r a b" using assms by (rule symD [to_pred]) lemma sym_Int: "sym r \ sym s \ sym (r \ s)" by (fast intro: symI elim: symE) lemma symp_inf: "symp r \ symp s \ symp (r \ s)" by (fact sym_Int [to_pred]) lemma sym_Un: "sym r \ sym s \ sym (r \ s)" by (fast intro: symI elim: symE) lemma symp_sup: "symp r \ symp s \ symp (r \ s)" by (fact sym_Un [to_pred]) lemma sym_INTER: "\x\S. sym (r x) \ sym (\(r ` S))" by (fast intro: symI elim: symE) lemma symp_INF: "\x\S. symp (r x) \ symp (\(r ` S))" by (fact sym_INTER [to_pred]) lemma sym_UNION: "\x\S. sym (r x) \ sym (\(r ` S))" by (fast intro: symI elim: symE) lemma symp_SUP: "\x\S. symp (r x) \ symp (\(r ` S))" by (fact sym_UNION [to_pred]) subsubsection \Antisymmetry\ definition antisym :: "'a rel \ bool" where "antisym r \ (\x y. (x, y) \ r \ (y, x) \ r \ x = y)" definition antisymp :: "('a \ 'a \ bool) \ bool" where "antisymp r \ (\x y. r x y \ r y x \ x = y)" lemma antisymp_antisym_eq [pred_set_conv]: "antisymp (\x y. (x, y) \ r) \ antisym r" by (simp add: antisym_def antisymp_def) lemma antisymI [intro?]: "(\x y. (x, y) \ r \ (y, x) \ r \ x = y) \ antisym r" unfolding antisym_def by iprover lemma antisympI [intro?]: "(\x y. r x y \ r y x \ x = y) \ antisymp r" by (fact antisymI [to_pred]) lemma antisymD [dest?]: "antisym r \ (a, b) \ r \ (b, a) \ r \ a = b" unfolding antisym_def by iprover lemma antisympD [dest?]: "antisymp r \ r a b \ r b a \ a = b" by (fact antisymD [to_pred]) lemma antisym_subset: "r \ s \ antisym s \ antisym r" unfolding antisym_def by blast lemma antisymp_less_eq: "r \ s \ antisymp s \ antisymp r" by (fact antisym_subset [to_pred]) lemma antisym_empty [simp]: "antisym {}" unfolding antisym_def by blast lemma antisym_bot [simp]: "antisymp \" by (fact antisym_empty [to_pred]) lemma antisymp_equality [simp]: "antisymp HOL.eq" by (auto intro: antisympI) lemma antisym_singleton [simp]: "antisym {x}" by (blast intro: antisymI) subsubsection \Transitivity\ definition trans :: "'a rel \ bool" where "trans r \ (\x y z. (x, y) \ r \ (y, z) \ r \ (x, z) \ r)" definition transp :: "('a \ 'a \ bool) \ bool" where "transp r \ (\x y z. r x y \ r y z \ r x z)" lemma transp_trans_eq [pred_set_conv]: "transp (\x y. (x, y) \ r) \ trans r" by (simp add: trans_def transp_def) lemma transI [intro?]: "(\x y z. (x, y) \ r \ (y, z) \ r \ (x, z) \ r) \ trans r" by (unfold trans_def) iprover lemma transpI [intro?]: "(\x y z. r x y \ r y z \ r x z) \ transp r" by (fact transI [to_pred]) lemma transE: assumes "trans r" and "(x, y) \ r" and "(y, z) \ r" obtains "(x, z) \ r" using assms by (unfold trans_def) iprover lemma transpE: assumes "transp r" and "r x y" and "r y z" obtains "r x z" using assms by (rule transE [to_pred]) lemma transD [dest?]: assumes "trans r" and "(x, y) \ r" and "(y, z) \ r" shows "(x, z) \ r" using assms by (rule transE) lemma transpD [dest?]: assumes "transp r" and "r x y" and "r y z" shows "r x z" using assms by (rule transD [to_pred]) lemma trans_Int: "trans r \ trans s \ trans (r \ s)" by (fast intro: transI elim: transE) lemma transp_inf: "transp r \ transp s \ transp (r \ s)" by (fact trans_Int [to_pred]) lemma trans_INTER: "\x\S. trans (r x) \ trans (\(r ` S))" by (fast intro: transI elim: transD) lemma transp_INF: "\x\S. transp (r x) \ transp (\(r ` S))" by (fact trans_INTER [to_pred]) lemma trans_join [code]: "trans r \ (\(x, y1) \ r. \(y2, z) \ r. y1 = y2 \ (x, z) \ r)" by (auto simp add: trans_def) lemma transp_trans: "transp r \ trans {(x, y). r x y}" by (simp add: trans_def transp_def) lemma transp_equality [simp]: "transp (=)" by (auto intro: transpI) lemma trans_empty [simp]: "trans {}" by (blast intro: transI) lemma transp_empty [simp]: "transp (\x y. False)" using trans_empty[to_pred] by (simp add: bot_fun_def) lemma trans_singleton [simp]: "trans {(a, a)}" by (blast intro: transI) lemma transp_singleton [simp]: "transp (\x y. x = a \ y = a)" by (simp add: transp_def) context preorder begin lemma transp_le[simp]: "transp (\)" by(auto simp add: transp_def intro: order_trans) lemma transp_less[simp]: "transp (<)" by(auto simp add: transp_def intro: less_trans) lemma transp_ge[simp]: "transp (\)" by(auto simp add: transp_def intro: order_trans) lemma transp_gr[simp]: "transp (>)" by(auto simp add: transp_def intro: less_trans) end subsubsection \Totality\ definition total_on :: "'a set \ 'a rel \ bool" where "total_on A r \ (\x\A. \y\A. x \ y \ (x, y) \ r \ (y, x) \ r)" lemma total_onI [intro?]: "(\x y. \x \ A; y \ A; x \ y\ \ (x, y) \ r \ (y, x) \ r) \ total_on A r" unfolding total_on_def by blast abbreviation "total \ total_on UNIV" definition totalp_on where "totalp_on A R \ (\x \ A. \y \ A. x \ y \ R x y \ R y x)" abbreviation totalp where "totalp \ totalp_on UNIV" lemma totalp_on_refl_on_eq[pred_set_conv]: "totalp_on A (\x y. (x, y) \ r) \ total_on A r" by (simp add: totalp_on_def total_on_def) lemma totalp_onI: "(\x y. x \ A \ y \ A \ x \ y \ R x y \ R y x) \ totalp_on A R" by (simp add: totalp_on_def) lemma totalpI: "(\x y. x \ y \ R x y \ R y x) \ totalp R" by (rule totalp_onI) lemma totalp_onD: "totalp_on A R \ x \ A \ y \ A \ x \ y \ R x y \ R y x" by (simp add: totalp_on_def) lemma totalpD: "totalp R \ x \ y \ R x y \ R y x" by (simp add: totalp_onD) lemma total_on_subset: "total_on A r \ B \ A \ total_on B r" by (auto simp: total_on_def) lemma totalp_on_subset: "totalp_on A R \ B \ A \ totalp_on B R" by (auto intro: totalp_onI dest: totalp_onD) lemma total_on_empty [simp]: "total_on {} r" by (simp add: total_on_def) lemma totalp_on_empty [simp]: "totalp_on {} R" by (auto intro: totalp_onI) lemma total_on_singleton [simp]: "total_on {x} {(x, x)}" unfolding total_on_def by blast subsubsection \Single valued relations\ definition single_valued :: "('a \ 'b) set \ bool" where "single_valued r \ (\x y. (x, y) \ r \ (\z. (x, z) \ r \ y = z))" definition single_valuedp :: "('a \ 'b \ bool) \ bool" where "single_valuedp r \ (\x y. r x y \ (\z. r x z \ y = z))" lemma single_valuedp_single_valued_eq [pred_set_conv]: "single_valuedp (\x y. (x, y) \ r) \ single_valued r" by (simp add: single_valued_def single_valuedp_def) lemma single_valuedp_iff_Uniq: "single_valuedp r \ (\x. \\<^sub>\\<^sub>1y. r x y)" unfolding Uniq_def single_valuedp_def by auto lemma single_valuedI: "(\x y. (x, y) \ r \ (\z. (x, z) \ r \ y = z)) \ single_valued r" unfolding single_valued_def by blast lemma single_valuedpI: "(\x y. r x y \ (\z. r x z \ y = z)) \ single_valuedp r" by (fact single_valuedI [to_pred]) lemma single_valuedD: "single_valued r \ (x, y) \ r \ (x, z) \ r \ y = z" by (simp add: single_valued_def) lemma single_valuedpD: "single_valuedp r \ r x y \ r x z \ y = z" by (fact single_valuedD [to_pred]) lemma single_valued_empty [simp]: "single_valued {}" by (simp add: single_valued_def) lemma single_valuedp_bot [simp]: "single_valuedp \" by (fact single_valued_empty [to_pred]) lemma single_valued_subset: "r \ s \ single_valued s \ single_valued r" unfolding single_valued_def by blast lemma single_valuedp_less_eq: "r \ s \ single_valuedp s \ single_valuedp r" by (fact single_valued_subset [to_pred]) subsection \Relation operations\ subsubsection \The identity relation\ definition Id :: "'a rel" where "Id = {p. \x. p = (x, x)}" lemma IdI [intro]: "(a, a) \ Id" by (simp add: Id_def) lemma IdE [elim!]: "p \ Id \ (\x. p = (x, x) \ P) \ P" unfolding Id_def by (iprover elim: CollectE) lemma pair_in_Id_conv [iff]: "(a, b) \ Id \ a = b" unfolding Id_def by blast lemma refl_Id: "refl Id" by (simp add: refl_on_def) lemma antisym_Id: "antisym Id" \ \A strange result, since \Id\ is also symmetric.\ by (simp add: antisym_def) lemma sym_Id: "sym Id" by (simp add: sym_def) lemma trans_Id: "trans Id" by (simp add: trans_def) lemma single_valued_Id [simp]: "single_valued Id" by (unfold single_valued_def) blast lemma irrefl_diff_Id [simp]: "irrefl (r - Id)" by (simp add: irrefl_def) lemma trans_diff_Id: "trans r \ antisym r \ trans (r - Id)" unfolding antisym_def trans_def by blast lemma total_on_diff_Id [simp]: "total_on A (r - Id) = total_on A r" by (simp add: total_on_def) lemma Id_fstsnd_eq: "Id = {x. fst x = snd x}" by force subsubsection \Diagonal: identity over a set\ definition Id_on :: "'a set \ 'a rel" where "Id_on A = (\x\A. {(x, x)})" lemma Id_on_empty [simp]: "Id_on {} = {}" by (simp add: Id_on_def) lemma Id_on_eqI: "a = b \ a \ A \ (a, b) \ Id_on A" by (simp add: Id_on_def) lemma Id_onI [intro!]: "a \ A \ (a, a) \ Id_on A" by (rule Id_on_eqI) (rule refl) lemma Id_onE [elim!]: "c \ Id_on A \ (\x. x \ A \ c = (x, x) \ P) \ P" \ \The general elimination rule.\ unfolding Id_on_def by (iprover elim!: UN_E singletonE) lemma Id_on_iff: "(x, y) \ Id_on A \ x = y \ x \ A" by blast lemma Id_on_def' [nitpick_unfold]: "Id_on {x. A x} = Collect (\(x, y). x = y \ A x)" by auto lemma Id_on_subset_Times: "Id_on A \ A \ A" by blast lemma refl_on_Id_on: "refl_on A (Id_on A)" by (rule refl_onI [OF Id_on_subset_Times Id_onI]) lemma antisym_Id_on [simp]: "antisym (Id_on A)" unfolding antisym_def by blast lemma sym_Id_on [simp]: "sym (Id_on A)" by (rule symI) clarify lemma trans_Id_on [simp]: "trans (Id_on A)" by (fast intro: transI elim: transD) lemma single_valued_Id_on [simp]: "single_valued (Id_on A)" unfolding single_valued_def by blast subsubsection \Composition\ inductive_set relcomp :: "('a \ 'b) set \ ('b \ 'c) set \ ('a \ 'c) set" (infixr "O" 75) for r :: "('a \ 'b) set" and s :: "('b \ 'c) set" where relcompI [intro]: "(a, b) \ r \ (b, c) \ s \ (a, c) \ r O s" notation relcompp (infixr "OO" 75) lemmas relcomppI = relcompp.intros text \ For historic reasons, the elimination rules are not wholly corresponding. Feel free to consolidate this. \ inductive_cases relcompEpair: "(a, c) \ r O s" inductive_cases relcomppE [elim!]: "(r OO s) a c" lemma relcompE [elim!]: "xz \ r O s \ (\x y z. xz = (x, z) \ (x, y) \ r \ (y, z) \ s \ P) \ P" apply (cases xz) apply simp apply (erule relcompEpair) apply iprover done lemma R_O_Id [simp]: "R O Id = R" by fast lemma Id_O_R [simp]: "Id O R = R" by fast lemma relcomp_empty1 [simp]: "{} O R = {}" by blast lemma relcompp_bot1 [simp]: "\ OO R = \" by (fact relcomp_empty1 [to_pred]) lemma relcomp_empty2 [simp]: "R O {} = {}" by blast lemma relcompp_bot2 [simp]: "R OO \ = \" by (fact relcomp_empty2 [to_pred]) lemma O_assoc: "(R O S) O T = R O (S O T)" by blast lemma relcompp_assoc: "(r OO s) OO t = r OO (s OO t)" by (fact O_assoc [to_pred]) lemma trans_O_subset: "trans r \ r O r \ r" by (unfold trans_def) blast lemma transp_relcompp_less_eq: "transp r \ r OO r \ r " by (fact trans_O_subset [to_pred]) lemma relcomp_mono: "r' \ r \ s' \ s \ r' O s' \ r O s" by blast lemma relcompp_mono: "r' \ r \ s' \ s \ r' OO s' \ r OO s " by (fact relcomp_mono [to_pred]) lemma relcomp_subset_Sigma: "r \ A \ B \ s \ B \ C \ r O s \ A \ C" by blast lemma relcomp_distrib [simp]: "R O (S \ T) = (R O S) \ (R O T)" by auto lemma relcompp_distrib [simp]: "R OO (S \ T) = R OO S \ R OO T" by (fact relcomp_distrib [to_pred]) lemma relcomp_distrib2 [simp]: "(S \ T) O R = (S O R) \ (T O R)" by auto lemma relcompp_distrib2 [simp]: "(S \ T) OO R = S OO R \ T OO R" by (fact relcomp_distrib2 [to_pred]) lemma relcomp_UNION_distrib: "s O \(r ` I) = (\i\I. s O r i) " by auto lemma relcompp_SUP_distrib: "s OO \(r ` I) = (\i\I. s OO r i)" by (fact relcomp_UNION_distrib [to_pred]) lemma relcomp_UNION_distrib2: "\(r ` I) O s = (\i\I. r i O s) " by auto lemma relcompp_SUP_distrib2: "\(r ` I) OO s = (\i\I. r i OO s)" by (fact relcomp_UNION_distrib2 [to_pred]) lemma single_valued_relcomp: "single_valued r \ single_valued s \ single_valued (r O s)" unfolding single_valued_def by blast lemma relcomp_unfold: "r O s = {(x, z). \y. (x, y) \ r \ (y, z) \ s}" by (auto simp add: set_eq_iff) lemma relcompp_apply: "(R OO S) a c \ (\b. R a b \ S b c)" unfolding relcomp_unfold [to_pred] .. lemma eq_OO: "(=) OO R = R" by blast lemma OO_eq: "R OO (=) = R" by blast subsubsection \Converse\ inductive_set converse :: "('a \ 'b) set \ ('b \ 'a) set" ("(_\)" [1000] 999) for r :: "('a \ 'b) set" where "(a, b) \ r \ (b, a) \ r\" notation conversep ("(_\\)" [1000] 1000) notation (ASCII) converse ("(_^-1)" [1000] 999) and conversep ("(_^--1)" [1000] 1000) lemma converseI [sym]: "(a, b) \ r \ (b, a) \ r\" by (fact converse.intros) lemma conversepI (* CANDIDATE [sym] *): "r a b \ r\\ b a" by (fact conversep.intros) lemma converseD [sym]: "(a, b) \ r\ \ (b, a) \ r" by (erule converse.cases) iprover lemma conversepD (* CANDIDATE [sym] *): "r\\ b a \ r a b" by (fact converseD [to_pred]) lemma converseE [elim!]: "yx \ r\ \ (\x y. yx = (y, x) \ (x, y) \ r \ P) \ P" \ \More general than \converseD\, as it ``splits'' the member of the relation.\ apply (cases yx) apply simp apply (erule converse.cases) apply iprover done lemmas conversepE [elim!] = conversep.cases lemma converse_iff [iff]: "(a, b) \ r\ \ (b, a) \ r" by (auto intro: converseI) lemma conversep_iff [iff]: "r\\ a b = r b a" by (fact converse_iff [to_pred]) lemma converse_converse [simp]: "(r\)\ = r" by (simp add: set_eq_iff) lemma conversep_conversep [simp]: "(r\\)\\ = r" by (fact converse_converse [to_pred]) lemma converse_empty[simp]: "{}\ = {}" by auto lemma converse_UNIV[simp]: "UNIV\ = UNIV" by auto lemma converse_relcomp: "(r O s)\ = s\ O r\" by blast lemma converse_relcompp: "(r OO s)\\ = s\\ OO r\\" by (iprover intro: order_antisym conversepI relcomppI elim: relcomppE dest: conversepD) lemma converse_Int: "(r \ s)\ = r\ \ s\" by blast lemma converse_meet: "(r \ s)\\ = r\\ \ s\\" by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD) lemma converse_Un: "(r \ s)\ = r\ \ s\" by blast lemma converse_join: "(r \ s)\\ = r\\ \ s\\" by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD) lemma converse_INTER: "(\(r ` S))\ = (\x\S. (r x)\)" by fast lemma converse_UNION: "(\(r ` S))\ = (\x\S. (r x)\)" by blast lemma converse_mono[simp]: "r\ \ s \ \ r \ s" by auto lemma conversep_mono[simp]: "r\\ \ s \\ \ r \ s" by (fact converse_mono[to_pred]) lemma converse_inject[simp]: "r\ = s \ \ r = s" by auto lemma conversep_inject[simp]: "r\\ = s \\ \ r = s" by (fact converse_inject[to_pred]) lemma converse_subset_swap: "r \ s \ \ r \ \ s" by auto lemma conversep_le_swap: "r \ s \\ \ r \\ \ s" by (fact converse_subset_swap[to_pred]) lemma converse_Id [simp]: "Id\ = Id" by blast lemma converse_Id_on [simp]: "(Id_on A)\ = Id_on A" by blast lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r" by (auto simp: refl_on_def) lemma sym_converse [simp]: "sym (converse r) = sym r" unfolding sym_def by blast lemma antisym_converse [simp]: "antisym (converse r) = antisym r" unfolding antisym_def by blast lemma trans_converse [simp]: "trans (converse r) = trans r" unfolding trans_def by blast lemma sym_conv_converse_eq: "sym r \ r\ = r" unfolding sym_def by fast lemma sym_Un_converse: "sym (r \ r\)" unfolding sym_def by blast lemma sym_Int_converse: "sym (r \ r\)" unfolding sym_def by blast lemma total_on_converse [simp]: "total_on A (r\) = total_on A r" by (auto simp: total_on_def) lemma finite_converse [iff]: "finite (r\) = finite r" unfolding converse_def conversep_iff using [[simproc add: finite_Collect]] by (auto elim: finite_imageD simp: inj_on_def) lemma card_inverse[simp]: "card (R\) = card R" proof - have *: "\R. prod.swap ` R = R\" by auto { assume "\finite R" hence ?thesis by auto } moreover { assume "finite R" with card_image_le[of R prod.swap] card_image_le[of "R\" prod.swap] have ?thesis by (auto simp: *) } ultimately show ?thesis by blast qed lemma conversep_noteq [simp]: "(\)\\ = (\)" by (auto simp add: fun_eq_iff) lemma conversep_eq [simp]: "(=)\\ = (=)" by (auto simp add: fun_eq_iff) lemma converse_unfold [code]: "r\ = {(y, x). (x, y) \ r}" by (simp add: set_eq_iff) subsubsection \Domain, range and field\ inductive_set Domain :: "('a \ 'b) set \ 'a set" for r :: "('a \ 'b) set" where DomainI [intro]: "(a, b) \ r \ a \ Domain r" lemmas DomainPI = Domainp.DomainI inductive_cases DomainE [elim!]: "a \ Domain r" inductive_cases DomainpE [elim!]: "Domainp r a" inductive_set Range :: "('a \ 'b) set \ 'b set" for r :: "('a \ 'b) set" where RangeI [intro]: "(a, b) \ r \ b \ Range r" lemmas RangePI = Rangep.RangeI inductive_cases RangeE [elim!]: "b \ Range r" inductive_cases RangepE [elim!]: "Rangep r b" definition Field :: "'a rel \ 'a set" where "Field r = Domain r \ Range r" lemma FieldI1: "(i, j) \ R \ i \ Field R" unfolding Field_def by blast lemma FieldI2: "(i, j) \ R \ j \ Field R" unfolding Field_def by auto lemma Domain_fst [code]: "Domain r = fst ` r" by force lemma Range_snd [code]: "Range r = snd ` r" by force lemma fst_eq_Domain: "fst ` R = Domain R" by force lemma snd_eq_Range: "snd ` R = Range R" by force lemma range_fst [simp]: "range fst = UNIV" by (auto simp: fst_eq_Domain) lemma range_snd [simp]: "range snd = UNIV" by (auto simp: snd_eq_Range) lemma Domain_empty [simp]: "Domain {} = {}" by auto lemma Range_empty [simp]: "Range {} = {}" by auto lemma Field_empty [simp]: "Field {} = {}" by (simp add: Field_def) lemma Domain_empty_iff: "Domain r = {} \ r = {}" by auto lemma Range_empty_iff: "Range r = {} \ r = {}" by auto lemma Domain_insert [simp]: "Domain (insert (a, b) r) = insert a (Domain r)" by blast lemma Range_insert [simp]: "Range (insert (a, b) r) = insert b (Range r)" by blast lemma Field_insert [simp]: "Field (insert (a, b) r) = {a, b} \ Field r" by (auto simp add: Field_def) lemma Domain_iff: "a \ Domain r \ (\y. (a, y) \ r)" by blast lemma Range_iff: "a \ Range r \ (\y. (y, a) \ r)" by blast lemma Domain_Id [simp]: "Domain Id = UNIV" by blast lemma Range_Id [simp]: "Range Id = UNIV" by blast lemma Domain_Id_on [simp]: "Domain (Id_on A) = A" by blast lemma Range_Id_on [simp]: "Range (Id_on A) = A" by blast lemma Domain_Un_eq: "Domain (A \ B) = Domain A \ Domain B" by blast lemma Range_Un_eq: "Range (A \ B) = Range A \ Range B" by blast lemma Field_Un [simp]: "Field (r \ s) = Field r \ Field s" by (auto simp: Field_def) lemma Domain_Int_subset: "Domain (A \ B) \ Domain A \ Domain B" by blast lemma Range_Int_subset: "Range (A \ B) \ Range A \ Range B" by blast lemma Domain_Diff_subset: "Domain A - Domain B \ Domain (A - B)" by blast lemma Range_Diff_subset: "Range A - Range B \ Range (A - B)" by blast lemma Domain_Union: "Domain (\S) = (\A\S. Domain A)" by blast lemma Range_Union: "Range (\S) = (\A\S. Range A)" by blast lemma Field_Union [simp]: "Field (\R) = \(Field ` R)" by (auto simp: Field_def) lemma Domain_converse [simp]: "Domain (r\) = Range r" by auto lemma Range_converse [simp]: "Range (r\) = Domain r" by blast lemma Field_converse [simp]: "Field (r\) = Field r" by (auto simp: Field_def) lemma Domain_Collect_case_prod [simp]: "Domain {(x, y). P x y} = {x. \y. P x y}" by auto lemma Range_Collect_case_prod [simp]: "Range {(x, y). P x y} = {y. \x. P x y}" by auto lemma finite_Domain: "finite r \ finite (Domain r)" by (induct set: finite) auto lemma finite_Range: "finite r \ finite (Range r)" by (induct set: finite) auto lemma finite_Field: "finite r \ finite (Field r)" by (simp add: Field_def finite_Domain finite_Range) lemma Domain_mono: "r \ s \ Domain r \ Domain s" by blast lemma Range_mono: "r \ s \ Range r \ Range s" by blast lemma mono_Field: "r \ s \ Field r \ Field s" by (auto simp: Field_def Domain_def Range_def) lemma Domain_unfold: "Domain r = {x. \y. (x, y) \ r}" by blast lemma Field_square [simp]: "Field (x \ x) = x" unfolding Field_def by blast subsubsection \Image of a set under a relation\ definition Image :: "('a \ 'b) set \ 'a set \ 'b set" (infixr "``" 90) where "r `` s = {y. \x\s. (x, y) \ r}" lemma Image_iff: "b \ r``A \ (\x\A. (x, b) \ r)" by (simp add: Image_def) lemma Image_singleton: "r``{a} = {b. (a, b) \ r}" by (simp add: Image_def) lemma Image_singleton_iff [iff]: "b \ r``{a} \ (a, b) \ r" by (rule Image_iff [THEN trans]) simp lemma ImageI [intro]: "(a, b) \ r \ a \ A \ b \ r``A" unfolding Image_def by blast lemma ImageE [elim!]: "b \ r `` A \ (\x. (x, b) \ r \ x \ A \ P) \ P" unfolding Image_def by (iprover elim!: CollectE bexE) lemma rev_ImageI: "a \ A \ (a, b) \ r \ b \ r `` A" \ \This version's more effective when we already have the required \a\\ by blast lemma Image_empty1 [simp]: "{} `` X = {}" by auto lemma Image_empty2 [simp]: "R``{} = {}" by blast lemma Image_Id [simp]: "Id `` A = A" by blast lemma Image_Id_on [simp]: "Id_on A `` B = A \ B" by blast lemma Image_Int_subset: "R `` (A \ B) \ R `` A \ R `` B" by blast lemma Image_Int_eq: "single_valued (converse R) \ R `` (A \ B) = R `` A \ R `` B" by (auto simp: single_valued_def) lemma Image_Un: "R `` (A \ B) = R `` A \ R `` B" by blast lemma Un_Image: "(R \ S) `` A = R `` A \ S `` A" by blast lemma Image_subset: "r \ A \ B \ r``C \ B" by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2) lemma Image_eq_UN: "r``B = (\y\ B. r``{y})" \ \NOT suitable for rewriting\ by blast lemma Image_mono: "r' \ r \ A' \ A \ (r' `` A') \ (r `` A)" by blast lemma Image_UN: "r `` (\(B ` A)) = (\x\A. r `` (B x))" by blast lemma UN_Image: "(\i\I. X i) `` S = (\i\I. X i `` S)" by auto lemma Image_INT_subset: "(r `` (\(B ` A))) \ (\x\A. r `` (B x))" by blast text \Converse inclusion requires some assumptions\ -lemma Image_INT_eq: "single_valued (r\) \ A \ {} \ r `` (\(B ` A)) = (\x\A. r `` B x)" - apply (rule equalityI) - apply (rule Image_INT_subset) - apply (auto simp add: single_valued_def) - apply blast - done +lemma Image_INT_eq: + assumes "single_valued (r\)" + and "A \ {}" + shows "r `` (\(B ` A)) = (\x\A. r `` B x)" +proof(rule equalityI, rule Image_INT_subset) + show "(\x\A. r `` B x) \ r `` \ (B ` A)" + proof + fix x + assume "x \ (\x\A. r `` B x)" + then show "x \ r `` \ (B ` A)" + using assms unfolding single_valued_def by simp blast + qed +qed lemma Image_subset_eq: "r``A \ B \ A \ - ((r\) `` (- B))" by blast lemma Image_Collect_case_prod [simp]: "{(x, y). P x y} `` A = {y. \x\A. P x y}" by auto lemma Sigma_Image: "(SIGMA x:A. B x) `` X = (\x\X \ A. B x)" by auto lemma relcomp_Image: "(X O Y) `` Z = Y `` (X `` Z)" by auto lemma finite_Image[simp]: assumes "finite R" shows "finite (R `` A)" by(rule finite_subset[OF _ finite_Range[OF assms]]) auto subsubsection \Inverse image\ definition inv_image :: "'b rel \ ('a \ 'b) \ 'a rel" where "inv_image r f = {(x, y). (f x, f y) \ r}" definition inv_imagep :: "('b \ 'b \ bool) \ ('a \ 'b) \ 'a \ 'a \ bool" where "inv_imagep r f = (\x y. r (f x) (f y))" lemma [pred_set_conv]: "inv_imagep (\x y. (x, y) \ r) f = (\x y. (x, y) \ inv_image r f)" by (simp add: inv_image_def inv_imagep_def) lemma sym_inv_image: "sym r \ sym (inv_image r f)" unfolding sym_def inv_image_def by blast lemma trans_inv_image: "trans r \ trans (inv_image r f)" unfolding trans_def inv_image_def by (simp (no_asm)) blast lemma total_inv_image: "\inj f; total r\ \ total (inv_image r f)" unfolding inv_image_def total_on_def by (auto simp: inj_eq) lemma asym_inv_image: "asym R \ asym (inv_image R f)" by (simp add: inv_image_def asym_iff) lemma in_inv_image[simp]: "(x, y) \ inv_image r f \ (f x, f y) \ r" by (auto simp: inv_image_def) lemma converse_inv_image[simp]: "(inv_image R f)\ = inv_image (R\) f" unfolding inv_image_def converse_unfold by auto lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" by (simp add: inv_imagep_def) subsubsection \Powerset\ definition Powp :: "('a \ bool) \ 'a set \ bool" where "Powp A = (\B. \x \ B. A x)" lemma Powp_Pow_eq [pred_set_conv]: "Powp (\x. x \ A) = (\x. x \ Pow A)" by (auto simp add: Powp_def fun_eq_iff) lemmas Powp_mono [mono] = Pow_mono [to_pred] subsubsection \Expressing relation operations via \<^const>\Finite_Set.fold\\ lemma Id_on_fold: assumes "finite A" shows "Id_on A = Finite_Set.fold (\x. Set.insert (Pair x x)) {} A" proof - interpret comp_fun_commute "\x. Set.insert (Pair x x)" by standard auto from assms show ?thesis unfolding Id_on_def by (induct A) simp_all qed lemma comp_fun_commute_Image_fold: "comp_fun_commute (\(x,y) A. if x \ S then Set.insert y A else A)" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff comp_fun_commute split: prod.split) qed lemma Image_fold: assumes "finite R" shows "R `` S = Finite_Set.fold (\(x,y) A. if x \ S then Set.insert y A else A) {} R" proof - interpret comp_fun_commute "(\(x,y) A. if x \ S then Set.insert y A else A)" by (rule comp_fun_commute_Image_fold) have *: "\x F. Set.insert x F `` S = (if fst x \ S then Set.insert (snd x) (F `` S) else (F `` S))" by (force intro: rev_ImageI) show ?thesis using assms by (induct R) (auto simp: *) qed lemma insert_relcomp_union_fold: assumes "finite S" shows "{x} O S \ X = Finite_Set.fold (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') X S" proof - interpret comp_fun_commute "\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A'" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show "comp_fun_commute (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A')" by standard (auto simp add: fun_eq_iff split: prod.split) qed have *: "{x} O S = {(x', z). x' = fst x \ (snd x, z) \ S}" by (auto simp: relcomp_unfold intro!: exI) show ?thesis unfolding * using \finite S\ by (induct S) (auto split: prod.split) qed lemma insert_relcomp_fold: assumes "finite S" shows "Set.insert x R O S = Finite_Set.fold (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') (R O S) S" proof - have "Set.insert x R O S = ({x} O S) \ (R O S)" by auto then show ?thesis by (auto simp: insert_relcomp_union_fold [OF assms]) qed lemma comp_fun_commute_relcomp_fold: assumes "finite S" shows "comp_fun_commute (\(x,y) A. Finite_Set.fold (\(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S)" proof - have *: "\a b A. Finite_Set.fold (\(w, z) A'. if b = w then Set.insert (a, z) A' else A') A S = {(a,b)} O S \ A" by (auto simp: insert_relcomp_union_fold[OF assms] cong: if_cong) show ?thesis by standard (auto simp: *) qed lemma relcomp_fold: assumes "finite R" "finite S" shows "R O S = Finite_Set.fold (\(x,y) A. Finite_Set.fold (\(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S) {} R" proof - interpret commute_relcomp_fold: comp_fun_commute "(\(x, y) A. Finite_Set.fold (\(w, z) A'. if y = w then insert (x, z) A' else A') A S)" by (fact comp_fun_commute_relcomp_fold[OF \finite S\]) from assms show ?thesis by (induct R) (auto simp: comp_fun_commute_relcomp_fold insert_relcomp_fold cong: if_cong) qed end diff --git a/src/HOL/Rings.thy b/src/HOL/Rings.thy --- a/src/HOL/Rings.thy +++ b/src/HOL/Rings.thy @@ -1,2753 +1,2755 @@ (* Title: HOL/Rings.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Rings\ theory Rings imports Groups Set Fun begin subsection \Semirings and rings\ class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c" begin text \For the \combine_numerals\ simproc\ lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c" by (simp add: distrib_right ac_simps) end class mult_zero = times + zero + assumes mult_zero_left [simp]: "0 * a = 0" assumes mult_zero_right [simp]: "a * 0 = 0" begin lemma mult_not_zero: "a * b \ 0 \ a \ 0 \ b \ 0" by auto end class semiring_0 = semiring + comm_monoid_add + mult_zero class semiring_0_cancel = semiring + cancel_comm_monoid_add begin subclass semiring_0 proof fix a :: 'a have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric]) then show "0 * a = 0" by (simp only: add_left_cancel) have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric]) then show "a * 0 = 0" by (simp only: add_left_cancel) qed end class comm_semiring = ab_semigroup_add + ab_semigroup_mult + assumes distrib: "(a + b) * c = a * c + b * c" begin subclass semiring proof fix a b c :: 'a show "(a + b) * c = a * c + b * c" by (simp add: distrib) have "a * (b + c) = (b + c) * a" by (simp add: ac_simps) also have "\ = b * a + c * a" by (simp only: distrib) also have "\ = a * b + a * c" by (simp add: ac_simps) finally show "a * (b + c) = a * b + a * c" by blast qed end class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero begin subclass semiring_0 .. end class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass comm_semiring_0 .. end class zero_neq_one = zero + one + assumes zero_neq_one [simp]: "0 \ 1" begin lemma one_neq_zero [simp]: "1 \ 0" by (rule not_sym) (rule zero_neq_one) definition of_bool :: "bool \ 'a" where "of_bool p = (if p then 1 else 0)" lemma of_bool_eq [simp, code]: "of_bool False = 0" "of_bool True = 1" by (simp_all add: of_bool_def) lemma of_bool_eq_iff: "of_bool p = of_bool q \ p = q" by (simp add: of_bool_def) lemma split_of_bool [split]: "P (of_bool p) \ (p \ P 1) \ (\ p \ P 0)" by (cases p) simp_all lemma split_of_bool_asm: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" by (cases p) simp_all lemma of_bool_eq_0_iff [simp]: \of_bool P = 0 \ \ P\ by simp lemma of_bool_eq_1_iff [simp]: \of_bool P = 1 \ P\ by simp end class semiring_1 = zero_neq_one + semiring_0 + monoid_mult begin lemma of_bool_conj: "of_bool (P \ Q) = of_bool P * of_bool Q" by auto end lemma lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto subsection \Abstract divisibility\ class dvd = times begin definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" unfolding dvd_def .. lemma dvdE [elim]: "b dvd a \ (\k. a = b * k \ P) \ P" unfolding dvd_def by blast end context comm_monoid_mult begin subclass dvd . lemma dvd_refl [simp]: "a dvd a" proof show "a = a * 1" by simp qed lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - from assms obtain v where "b = a * v" by auto moreover from assms obtain w where "c = b * w" by auto ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed lemma subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b" by (auto simp add: subset_iff intro: dvd_trans) lemma strict_subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b \ \ b dvd a" by (auto simp add: subset_iff intro: dvd_trans) lemma one_dvd [simp]: "1 dvd a" by (auto intro: dvdI) lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c" - using that by rule (auto intro: mult.left_commute dvdI) + using that by (auto intro: mult.left_commute dvdI) lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b" using that dvd_mult [of a b c] by (simp add: ac_simps) lemma dvd_triv_right [simp]: "a dvd b * a" by (rule dvd_mult) (rule dvd_refl) lemma dvd_triv_left [simp]: "a dvd a * b" by (rule dvd_mult2) (rule dvd_refl) lemma mult_dvd_mono: assumes "a dvd b" and "c dvd d" shows "a * c dvd b * d" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \c dvd d\ obtain d' where "d = c * d'" .. ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps) then show ?thesis .. qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" by (simp add: dvd_def mult.assoc) blast lemma dvd_mult_right: "a * b dvd c \ b dvd c" using dvd_mult_left [of b a c] by (simp add: ac_simps) end class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult begin subclass semiring_1 .. lemma dvd_0_left_iff [simp]: "0 dvd a \ a = 0" by auto lemma dvd_0_right [iff]: "a dvd 0" proof show "0 = a * 0" by simp qed lemma dvd_0_left: "0 dvd a \ a = 0" by simp lemma dvd_add [simp]: assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \a dvd c\ obtain c' where "c = a * c'" .. ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left) then show ?thesis .. qed end class semiring_1_cancel = semiring + cancel_comm_monoid_add + zero_neq_one + monoid_mult begin subclass semiring_0_cancel .. subclass semiring_1 .. end class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult + assumes right_diff_distrib' [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" begin subclass semiring_1_cancel .. subclass comm_semiring_0_cancel .. subclass comm_semiring_1 .. lemma left_diff_distrib' [algebra_simps, algebra_split_simps]: "(b - c) * a = b * a - c * a" by (simp add: algebra_simps) lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \ a dvd b" proof - have "a dvd a * c + b \ a dvd b" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then obtain d where "a * c + b = a * d" .. then have "a * c + b - a * c = a * d - a * c" by simp then have "b = a * d - a * c" by simp then have "b = a * (d - c)" by (simp add: algebra_simps) then show ?Q .. qed then show "a dvd c * a + b \ a dvd b" by (simp add: ac_simps) qed lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \ a dvd b" using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps) lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \ a dvd b" using dvd_add_times_triv_left_iff [of a 1 b] by simp lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \ a dvd b" using dvd_add_times_triv_right_iff [of a b 1] by simp lemma dvd_add_right_iff: assumes "a dvd b" shows "a dvd b + c \ a dvd c" (is "?P \ ?Q") proof assume ?P then obtain d where "b + c = a * d" .. moreover from \a dvd b\ obtain e where "b = a * e" .. ultimately have "a * e + c = a * d" by simp then have "a * e + c - a * e = a * d - a * e" by simp then have "c = a * d - a * e" by simp then have "c = a * (d - e)" by (simp add: algebra_simps) then show ?Q .. next assume ?Q with assms show ?P by simp qed lemma dvd_add_left_iff: "a dvd c \ a dvd b + c \ a dvd b" using dvd_add_right_iff [of a c b] by (simp add: ac_simps) end class ring = semiring + ab_group_add begin subclass semiring_0_cancel .. text \Distribution rules\ lemma minus_mult_left: "- (a * b) = - a * b" by (rule minus_unique) (simp add: distrib_right [symmetric]) lemma minus_mult_right: "- (a * b) = a * - b" by (rule minus_unique) (simp add: distrib_left [symmetric]) text \Extract signs from products\ lemmas mult_minus_left [simp] = minus_mult_left [symmetric] lemmas mult_minus_right [simp] = minus_mult_right [symmetric] lemma minus_mult_minus [simp]: "- a * - b = a * b" by simp lemma minus_mult_commute: "- a * b = a * - b" by simp lemma right_diff_distrib [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" using distrib_left [of a b "-c "] by simp lemma left_diff_distrib [algebra_simps, algebra_split_simps]: "(a - b) * c = a * c - b * c" using distrib_right [of a "- b" c] by simp lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib lemma eq_add_iff1: "a * e + c = b * e + d \ (a - b) * e + c = d" by (simp add: algebra_simps) lemma eq_add_iff2: "a * e + c = b * e + d \ c = (b - a) * e + d" by (simp add: algebra_simps) end lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib class comm_ring = comm_semiring + ab_group_add begin subclass ring .. subclass comm_semiring_0_cancel .. lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)" by (simp add: algebra_simps) end class ring_1 = ring + zero_neq_one + monoid_mult begin subclass semiring_1_cancel .. lemma of_bool_not_iff [simp]: \of_bool (\ P) = 1 - of_bool P\ by simp lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)" by (simp add: algebra_simps) end class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult begin subclass ring_1 .. subclass comm_semiring_1_cancel by standard (simp add: algebra_simps) lemma dvd_minus_iff [simp]: "x dvd - y \ x dvd y" proof assume "x dvd - y" then have "x dvd - 1 * - y" by (rule dvd_mult) then show "x dvd y" by simp next assume "x dvd y" then have "x dvd - 1 * y" by (rule dvd_mult) then show "x dvd - y" by simp qed lemma minus_dvd_iff [simp]: "- x dvd y \ x dvd y" proof assume "- x dvd y" then obtain k where "y = - x * k" .. then have "y = x * - k" by simp then show "x dvd y" .. next assume "x dvd y" then obtain k where "y = x * k" .. then have "y = - x * - k" by simp then show "- x dvd y" .. qed lemma dvd_diff [simp]: "x dvd y \ x dvd z \ x dvd (y - z)" using dvd_add [of x y "- z"] by simp end subsection \Towards integral domains\ class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin lemma divisors_zero: assumes "a * b = 0" shows "a = 0 \ b = 0" proof (rule classical) assume "\ ?thesis" then have "a \ 0" and "b \ 0" by auto with no_zero_divisors have "a * b \ 0" by blast with assms show ?thesis by simp qed lemma mult_eq_0_iff [simp]: "a * b = 0 \ a = 0 \ b = 0" proof (cases "a = 0 \ b = 0") case False then have "a \ 0" and "b \ 0" by auto then show ?thesis using no_zero_divisors by simp next case True then show ?thesis by auto qed end class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors + assumes mult_cancel_right [simp]: "a * c = b * c \ c = 0 \ a = b" and mult_cancel_left [simp]: "c * a = c * b \ c = 0 \ a = b" begin lemma mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" by simp lemma mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" by simp end class ring_no_zero_divisors = ring + semiring_no_zero_divisors begin subclass semiring_no_zero_divisors_cancel proof fix a b c have "a * c = b * c \ (a - b) * c = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "a * c = b * c \ c = 0 \ a = b" . have "c * a = c * b \ c * (a - b) = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "c * a = c * b \ c = 0 \ a = b" . qed end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. lemma square_eq_1_iff: "x * x = 1 \ x = 1 \ x = - 1" proof - have "(x - 1) * (x + 1) = x * x - 1" by (simp add: algebra_simps) then have "x * x = 1 \ (x - 1) * (x + 1) = 0" by simp then show ?thesis by (simp add: eq_neg_iff_add_eq_0) qed lemma mult_cancel_right1 [simp]: "c = b * c \ c = 0 \ b = 1" using mult_cancel_right [of 1 c b] by auto lemma mult_cancel_right2 [simp]: "a * c = c \ c = 0 \ a = 1" using mult_cancel_right [of a c 1] by simp lemma mult_cancel_left1 [simp]: "c = c * b \ c = 0 \ b = 1" using mult_cancel_left [of c 1 b] by force lemma mult_cancel_left2 [simp]: "c * a = c \ c = 0 \ a = 1" using mult_cancel_left [of c a 1] by simp end class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. end class idom = comm_ring_1 + semiring_no_zero_divisors begin subclass semidom .. subclass ring_1_no_zero_divisors .. lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \ c = 0 \ a dvd b" proof - have "a * c dvd b * c \ (\k. b * c = (a * k) * c)" by (auto simp add: ac_simps) also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" by auto finally show ?thesis . qed lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \ c = 0 \ a dvd b" using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps) lemma square_eq_iff: "a * a = b * b \ a = b \ a = - b" proof assume "a * a = b * b" then have "(a - b) * (a + b) = 0" by (simp add: algebra_simps) then show "a = b \ a = - b" by (simp add: eq_neg_iff_add_eq_0) next assume "a = b \ a = - b" then show "a * a = b * b" by auto qed end class idom_abs_sgn = idom + abs + sgn + assumes sgn_mult_abs: "sgn a * \a\ = a" and sgn_sgn [simp]: "sgn (sgn a) = sgn a" and abs_abs [simp]: "\\a\\ = \a\" and abs_0 [simp]: "\0\ = 0" and sgn_0 [simp]: "sgn 0 = 0" and sgn_1 [simp]: "sgn 1 = 1" and sgn_minus_1: "sgn (- 1) = - 1" and sgn_mult: "sgn (a * b) = sgn a * sgn b" begin lemma sgn_eq_0_iff: "sgn a = 0 \ a = 0" proof - { assume "sgn a = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_eq_0_iff: "\a\ = 0 \ a = 0" proof - { assume "\a\ = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_mult_sgn: "\a\ * sgn a = a" using sgn_mult_abs [of a] by (simp add: ac_simps) lemma abs_1 [simp]: "\1\ = 1" using sgn_mult_abs [of 1] by simp lemma sgn_abs [simp]: "\sgn a\ = of_bool (a \ 0)" using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\sgn a\" 1] by (auto simp add: sgn_eq_0_iff) lemma abs_sgn [simp]: "sgn \a\ = of_bool (a \ 0)" using sgn_mult_abs [of "\a\"] mult_cancel_right [of "sgn \a\" "\a\" 1] by (auto simp add: abs_eq_0_iff) lemma abs_mult: "\a * b\ = \a\ * \b\" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False then have *: "sgn (a * b) \ 0" by (simp add: sgn_eq_0_iff) from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b] have "\a * b\ * sgn (a * b) = \a\ * sgn a * \b\ * sgn b" by (simp add: ac_simps) then have "\a * b\ * sgn (a * b) = \a\ * \b\ * sgn (a * b)" by (simp add: sgn_mult ac_simps) with * show ?thesis by simp qed lemma sgn_minus [simp]: "sgn (- a) = - sgn a" proof - from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a" by (simp only: sgn_mult) then show ?thesis by simp qed lemma abs_minus [simp]: "\- a\ = \a\" proof - have [simp]: "\- 1\ = 1" using sgn_mult_abs [of "- 1"] by simp then have "\- 1 * a\ = 1 * \a\" by (simp only: abs_mult) then show ?thesis by simp qed end subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a \ 'a \ 'a\)\ context semiring begin lemma [field_simps, field_split_simps]: shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \ a * (b + c) = a * b + a * c" and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \ (a + b) * c = a * c + b * c" by (rule distrib_left distrib_right)+ end context ring begin lemma [field_simps, field_split_simps]: shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \ (a - b) * c = a * c - b * c" and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \ a * (b - c) = a * b - a * c" by (rule left_diff_distrib right_diff_distrib)+ end setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a::divide \ 'a \ 'a\)\ text \Algebraic classes with division\ class semidom_divide = semidom + divide + assumes nonzero_mult_div_cancel_right [simp]: "b \ 0 \ (a * b) div b = a" assumes div_by_0 [simp]: "a div 0 = 0" begin lemma nonzero_mult_div_cancel_left [simp]: "a \ 0 \ (a * b) div a = b" using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps) subclass semiring_no_zero_divisors_cancel proof show *: "a * c = b * c \ c = 0 \ a = b" for a b c proof (cases "c = 0") case True then show ?thesis by simp next case False have "a = b" if "a * c = b * c" proof - from that have "a * c div c = b * c div c" by simp with False show ?thesis by simp qed then show ?thesis by auto qed show "c * a = c * b \ c = 0 \ a = b" for a b c using * [of a c b] by (simp add: ac_simps) qed lemma div_self [simp]: "a \ 0 \ a div a = 1" using nonzero_mult_div_cancel_left [of a 1] by simp lemma div_0 [simp]: "0 div a = 0" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * 0 div a = 0" by (rule nonzero_mult_div_cancel_left) then show ?thesis by simp qed lemma div_by_1 [simp]: "a div 1 = a" using nonzero_mult_div_cancel_left [of 1 a] by simp lemma dvd_div_eq_0_iff: assumes "b dvd a" shows "a div b = 0 \ a = 0" using assms by (elim dvdE, cases "b = 0") simp_all lemma dvd_div_eq_cancel: "a div c = b div c \ c dvd a \ c dvd b \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma dvd_div_eq_iff: "c dvd a \ c dvd b \ a div c = b div c \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma inj_on_mult: "inj_on ((*) a) A" if "a \ 0" proof (rule inj_onI) fix b c assume "a * b = a * c" then have "a * b div a = a * c div a" by (simp only:) with that show "b = c" by simp qed end class idom_divide = idom + semidom_divide begin lemma dvd_neg_div: assumes "b dvd a" shows "- a div b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False from assms obtain c where "a = b * c" .. then have "- a div b = (b * - c) div b" by simp from False also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed lemma dvd_div_neg: assumes "b dvd a" shows "a div - b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "- b \ 0" by simp from assms obtain c where "a = b * c" .. then have "a div - b = (- b * - c) div - b" by simp from \- b \ 0\ also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed end class algebraic_semidom = semidom_divide begin text \ Class \<^class>\algebraic_semidom\ enriches a integral domain by notions from algebra, like units in a ring. It is a separate class to avoid spoiling fields with notions which are degenerated there. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma div_dvd_iff_mult: assumes "b \ 0" and "b dvd a" shows "a div b dvd c \ a dvd c * b" proof - from \b dvd a\ obtain d where "a = b * d" .. with \b \ 0\ show ?thesis by (simp add: ac_simps) qed lemma dvd_div_iff_mult: assumes "c \ 0" and "c dvd b" shows "a dvd b div c \ a * c dvd b" proof - from \c dvd b\ obtain d where "b = c * d" .. with \c \ 0\ show ?thesis by (simp add: mult.commute [of a]) qed lemma div_dvd_div [simp]: assumes "a dvd b" and "a dvd c" shows "b div a dvd c div a \ b dvd c" proof (cases "a = 0") case True with assms show ?thesis by simp next case False moreover from assms obtain k l where "b = a * k" and "c = a * l" by blast ultimately show ?thesis by simp qed lemma div_add [simp]: assumes "c dvd a" and "c dvd b" shows "(a + b) div c = a div c + b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False moreover from assms obtain k l where "a = c * k" and "b = c * l" by blast moreover have "c * k + c * l = c * (k + l)" by (simp add: algebra_simps) ultimately show ?thesis by simp qed lemma div_mult_div_if_dvd: assumes "b dvd a" and "d dvd c" shows "(a div b) * (c div d) = (a * c) div (b * d)" proof (cases "b = 0 \ c = 0") case True with assms show ?thesis by auto next case False moreover from assms obtain k l where "a = b * k" and "c = d * l" by blast moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)" by (simp add: ac_simps) ultimately show ?thesis by simp qed lemma dvd_div_eq_mult: assumes "a \ 0" and "a dvd b" shows "b div a = c \ b = c * a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: assms) next assume ?lhs then have "b div a * a = c * a" by simp moreover from assms have "b div a * a = b" by (auto simp add: ac_simps) ultimately show ?rhs by simp qed lemma dvd_div_mult_self [simp]: "a dvd b \ b div a * a = b" by (cases "a = 0") (auto simp add: ac_simps) lemma dvd_mult_div_cancel [simp]: "a dvd b \ a * (b div a) = b" using dvd_div_mult_self [of a b] by (simp add: ac_simps) lemma div_mult_swap: assumes "c dvd b" shows "a * (b div c) = (a * b) div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from assms obtain d where "b = c * d" .. moreover from False have "a * divide (d * c) c = ((a * d) * c) div c" by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma dvd_div_mult: "c dvd b \ b div c * a = (b * a) div c" using div_mult_swap [of c b a] by (simp add: ac_simps) lemma dvd_div_mult2_eq: assumes "b * c dvd a" shows "a div (b * c) = a div b div c" proof - from assms obtain k where "a = b * c * k" .. then show ?thesis by (cases "b = 0 \ c = 0") (auto, simp add: ac_simps) qed lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" (is "?lhs \ ?rhs") proof - from assms have "a * c \ 0" by simp then have "?lhs \ b div a * (a * c) = d div c * (a * c)" by simp also have "\ \ (a * (b div a)) * c = (c * (d div c)) * a" by (simp add: ac_simps) also have "\ \ (a * b div a) * c = (c * d div c) * a" using assms by (simp add: div_mult_swap) also have "\ \ ?rhs" using assms by (simp add: ac_simps) finally show ?thesis . qed lemma dvd_mult_imp_div: assumes "a * c dvd b" shows "a dvd b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from \a * c dvd b\ obtain d where "b = a * c * d" .. with False show ?thesis by (simp add: mult.commute [of a] mult.assoc) qed lemma div_div_eq_right: assumes "c dvd b" "b dvd a" shows "a div (b div c) = a div b * c" proof (cases "c = 0 \ b = 0") case True then show ?thesis by auto next case False from assms obtain r s where "b = c * r" and "a = c * r * s" by blast moreover with False have "r \ 0" by auto ultimately show ?thesis using False by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c]) qed lemma div_div_div_same: assumes "d dvd b" "b dvd a" shows "(a div d) div (b div d) = a div b" proof (cases "b = 0 \ d = 0") case True with assms show ?thesis by auto next case False from assms obtain r s where "a = d * r * s" and "b = d * r" by blast with False show ?thesis by simp (simp add: ac_simps) qed text \Units: invertible elements in a ring\ abbreviation is_unit :: "'a \ bool" where "is_unit a \ a dvd 1" lemma not_is_unit_0 [simp]: "\ is_unit 0" by simp lemma unit_imp_dvd [dest]: "is_unit b \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_dvdE: assumes "is_unit a" obtains c where "a \ 0" and "b = a * c" proof - from assms have "a dvd b" by auto then obtain c where "b = a * c" .. moreover from assms have "a \ 0" by auto ultimately show thesis using that by blast qed lemma dvd_unit_imp_unit: "a dvd b \ is_unit b \ is_unit a" by (rule dvd_trans) lemma unit_div_1_unit [simp, intro]: assumes "is_unit a" shows "is_unit (1 div a)" proof - from assms have "1 = 1 div a * a" by simp then show "is_unit (1 div a)" by (rule dvdI) qed lemma is_unitE [elim?]: assumes "is_unit a" obtains b where "a \ 0" and "b \ 0" and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) define b where "b = 1 div a" then show "1 div a = b" by simp from assms b_def show "is_unit b" by simp with assms show "a \ 0" and "b \ 0" by auto from assms b_def show "a * b = 1" by simp then have "1 = a * b" .. with b_def \b \ 0\ show "1 div b = a" by simp from assms have "a dvd c" .. then obtain d where "c = a * d" .. with \a \ 0\ \a * b = 1\ show "c div a = c * b" by (simp add: mult.assoc mult.left_commute [of a]) qed lemma unit_prod [intro]: "is_unit a \ is_unit b \ is_unit (a * b)" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff: "is_unit (a * b) \ is_unit a \ is_unit b" by (auto dest: dvd_mult_left dvd_mult_right) lemma unit_div [intro]: "is_unit a \ is_unit b \ is_unit (a div b)" by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod) lemma mult_unit_dvd_iff: assumes "is_unit b" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" by (simp add: dvd_mult_left) next assume "a dvd c" then obtain k where "c = a * k" .. with assms have "c = (a * b) * (1 div b * k)" by (simp add: mult_ac) then show "a * b dvd c" by (rule dvdI) qed lemma mult_unit_dvd_iff': "is_unit a \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma dvd_mult_unit_iff: assumes "is_unit b" shows "a dvd c * b \ a dvd c" proof assume "a dvd c * b" with assms have "c * b dvd c * (b * (1 div b))" by (subst mult_assoc [symmetric]) simp also from assms have "b * (1 div b) = 1" by (rule is_unitE) simp finally have "c * b dvd c" by simp with \a dvd c * b\ show "a dvd c" by (rule dvd_trans) next assume "a dvd c" then show "a dvd c * b" by simp qed lemma dvd_mult_unit_iff': "is_unit b \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma div_unit_dvd_iff: "is_unit b \ a div b dvd c \ a dvd c" by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff) lemma dvd_div_unit_iff: "is_unit b \ a dvd c div b \ a dvd c" by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff) lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff' dvd_mult_unit_iff dvd_mult_unit_iff' div_unit_dvd_iff dvd_div_unit_iff (* FIXME consider named_theorems *) lemma unit_mult_div_div [simp]: "is_unit a \ b * (1 div a) = b div a" by (erule is_unitE [of _ b]) simp lemma unit_div_mult_self [simp]: "is_unit a \ b div a * a = b" by (rule dvd_div_mult_self) auto lemma unit_div_1_div_1 [simp]: "is_unit a \ 1 div (1 div a) = a" by (erule is_unitE) simp lemma unit_div_mult_swap: "is_unit c \ a * (b div c) = (a * b) div c" by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c]) lemma unit_div_commute: "is_unit b \ (a div b) * c = (a * c) div b" using unit_div_mult_swap [of b c a] by (simp add: ac_simps) lemma unit_eq_div1: "is_unit b \ a div b = c \ a = c * b" by (auto elim: is_unitE) lemma unit_eq_div2: "is_unit b \ a = c div b \ a * b = c" using unit_eq_div1 [of b c a] by auto lemma unit_mult_left_cancel: "is_unit a \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "is_unit a \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) lemma unit_div_cancel: assumes "is_unit a" shows "b div a = c div a \ b = c" proof - from assms have "is_unit (1 div a)" by simp then have "b * (1 div a) = c * (1 div a) \ b = c" by (rule unit_mult_right_cancel) with assms show ?thesis by simp qed lemma is_unit_div_mult2_eq: assumes "is_unit b" and "is_unit c" shows "a div (b * c) = a div b div c" proof - from assms have "is_unit (b * c)" by (simp add: unit_prod) then have "b * c dvd a" by (rule unit_imp_dvd) then show ?thesis by (rule dvd_div_mult2_eq) qed lemma is_unit_div_mult_cancel_left: assumes "a \ 0" and "is_unit b" shows "a div (a * b) = 1 div b" proof - from assms have "a div (a * b) = a div a div b" by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq) with assms show ?thesis by simp qed lemma is_unit_div_mult_cancel_right: assumes "a \ 0" and "is_unit b" shows "a div (b * a) = 1 div b" using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps) lemma unit_div_eq_0_iff: assumes "is_unit b" shows "a div b = 0 \ a = 0" - by (rule dvd_div_eq_0_iff) (insert assms, auto) + using assms by (simp add: dvd_div_eq_0_iff unit_imp_dvd) lemma div_mult_unit2: "is_unit c \ b dvd a \ a div (b * c) = a div b div c" by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff) text \Coprimality\ definition coprime :: "'a \ 'a \ bool" where "coprime a b \ (\c. c dvd a \ c dvd b \ is_unit c)" lemma coprimeI: assumes "\c. c dvd a \ c dvd b \ is_unit c" shows "coprime a b" using assms by (auto simp: coprime_def) lemma not_coprimeI: assumes "c dvd a" and "c dvd b" and "\ is_unit c" shows "\ coprime a b" using assms by (auto simp: coprime_def) lemma coprime_common_divisor: "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b" using that by (auto simp: coprime_def) lemma not_coprimeE: assumes "\ coprime a b" obtains c where "c dvd a" and "c dvd b" and "\ is_unit c" using assms by (auto simp: coprime_def) lemma coprime_imp_coprime: "coprime a b" if "coprime c d" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd c" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd d" proof (rule coprimeI) fix e assume "e dvd a" and "e dvd b" with that have "e dvd c" and "e dvd d" by (auto intro: dvd_trans) with \coprime c d\ show "is_unit e" by (rule coprime_common_divisor) qed lemma coprime_divisors: "coprime a b" if "a dvd c" "b dvd d" and "coprime c d" using \coprime c d\ proof (rule coprime_imp_coprime) fix e assume "e dvd a" then show "e dvd c" using \a dvd c\ by (rule dvd_trans) assume "e dvd b" then show "e dvd d" using \b dvd d\ by (rule dvd_trans) qed lemma coprime_self [simp]: "coprime a a \ is_unit a" (is "?P \ ?Q") proof assume ?P then show ?Q by (rule coprime_common_divisor) simp_all next assume ?Q show ?P by (rule coprimeI) (erule dvd_unit_imp_unit, rule \?Q\) qed lemma coprime_commute [ac_simps]: "coprime b a \ coprime a b" unfolding coprime_def by auto lemma is_unit_left_imp_coprime: "coprime a b" if "is_unit a" proof (rule coprimeI) fix c assume "c dvd a" with that show "is_unit c" by (auto intro: dvd_unit_imp_unit) qed lemma is_unit_right_imp_coprime: "coprime a b" if "is_unit b" using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps) lemma coprime_1_left [simp]: "coprime 1 a" by (rule coprimeI) lemma coprime_1_right [simp]: "coprime a 1" by (rule coprimeI) lemma coprime_0_left_iff [simp]: "coprime 0 a \ is_unit a" by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a]) lemma coprime_0_right_iff [simp]: "coprime a 0 \ is_unit a" using coprime_0_left_iff [of a] by (simp add: ac_simps) lemma coprime_mult_self_left_iff [simp]: "coprime (c * a) (c * b) \ is_unit c \ coprime a b" by (auto intro: coprime_common_divisor) (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+ lemma coprime_mult_self_right_iff [simp]: "coprime (a * c) (b * c) \ is_unit c \ coprime a b" using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps) lemma coprime_absorb_left: assumes "x dvd y" shows "coprime x y \ is_unit x" using assms coprime_common_divisor is_unit_left_imp_coprime by auto lemma coprime_absorb_right: assumes "y dvd x" shows "coprime x y \ is_unit y" using assms coprime_common_divisor is_unit_right_imp_coprime by auto end class unit_factor = fixes unit_factor :: "'a \ 'a" class semidom_divide_unit_factor = semidom_divide + unit_factor + assumes unit_factor_0 [simp]: "unit_factor 0 = 0" and is_unit_unit_factor: "a dvd 1 \ unit_factor a = a" and unit_factor_is_unit: "a \ 0 \ unit_factor a dvd 1" and unit_factor_mult_unit_left: "a dvd 1 \ unit_factor (a * b) = a * unit_factor b" \ \This fine-grained hierarchy will later on allow lean normalization of polynomials\ begin lemma unit_factor_mult_unit_right: "a dvd 1 \ unit_factor (b * a) = unit_factor b * a" using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac) lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right end class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor + fixes normalize :: "'a \ 'a" assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a" and normalize_0 [simp]: "normalize 0 = 0" begin text \ Class \<^class>\normalization_semidom\ cultivates the idea that each integral domain can be split into equivalence classes whose representants are associated, i.e. divide each other. \<^const>\normalize\ specifies a canonical representant for each equivalence class. The rationale behind this is that it is easier to reason about equality than equivalences, hence we prefer to think about equality of normalized values rather than associated elements. \ declare unit_factor_is_unit [iff] lemma unit_factor_dvd [simp]: "a \ 0 \ unit_factor a dvd b" by (rule unit_imp_dvd) simp lemma unit_factor_self [simp]: "unit_factor a dvd a" by (cases "a = 0") simp_all lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a" using unit_factor_mult_normalize [of a] by (simp add: ac_simps) lemma normalize_eq_0_iff [simp]: "normalize a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma div_unit_factor [simp]: "a div unit_factor a = normalize a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "unit_factor a \ 0" by simp with nonzero_mult_div_cancel_left have "unit_factor a * normalize a div unit_factor a = normalize a" by blast then show ?thesis by simp qed lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False have "normalize a div a = normalize a div (unit_factor a * normalize a)" by simp also have "\ = 1 div unit_factor a" using False by (subst is_unit_div_mult_cancel_right) simp_all finally show ?thesis . qed lemma is_unit_normalize: assumes "is_unit a" shows "normalize a = 1" proof - from assms have "unit_factor a = a" by (rule is_unit_unit_factor) moreover from assms have "a \ 0" by auto moreover have "normalize a = a div unit_factor a" by simp ultimately show ?thesis by simp qed lemma unit_factor_1 [simp]: "unit_factor 1 = 1" by (rule is_unit_unit_factor) simp lemma normalize_1 [simp]: "normalize 1 = 1" by (rule is_unit_normalize) simp lemma normalize_1_iff: "normalize a = 1 \ is_unit a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (rule is_unit_normalize) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * 1" by simp then have "unit_factor a = a" by simp moreover from \?lhs\ have "a \ 0" by auto then have "is_unit (unit_factor a)" by simp ultimately show ?rhs by simp qed lemma div_normalize [simp]: "a div normalize a = unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "normalize a \ 0" by simp with nonzero_mult_div_cancel_right have "unit_factor a * normalize a div normalize a = unit_factor a" by blast then show ?thesis by simp qed lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b" by (cases "b = 0") simp_all lemma inv_unit_factor_eq_0_iff [simp]: "1 div unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs then have "a * (1 div unit_factor a) = a * 0" by simp then show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a" by (cases "a = 0") (auto intro: is_unit_unit_factor) lemma normalize_unit_factor [simp]: "a \ 0 \ normalize (unit_factor a) = 1" by (rule is_unit_normalize) simp lemma normalize_mult_unit_left [simp]: assumes "a dvd 1" shows "normalize (a * b) = normalize b" proof (cases "b = 0") case False have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)" using assms by (subst unit_factor_mult_unit_left) auto also have "\ = a * b" by simp also have "b = unit_factor b * normalize b" by simp hence "a * b = a * unit_factor b * normalize b" by (simp only: mult_ac) finally show ?thesis using assms False by auto qed auto lemma normalize_mult_unit_right [simp]: assumes "b dvd 1" shows "normalize (a * b) = normalize a" using assms by (subst mult.commute) auto lemma normalize_idem [simp]: "normalize (normalize a) = normalize a" proof (cases "a = 0") case False have "normalize a = normalize (unit_factor a * normalize a)" by simp also from False have "\ = normalize (normalize a)" by (subst normalize_mult_unit_left) auto finally show ?thesis .. qed auto lemma unit_factor_normalize [simp]: assumes "a \ 0" shows "unit_factor (normalize a) = 1" proof - from assms have *: "normalize a \ 0" by simp have "unit_factor (normalize a) * normalize (normalize a) = normalize a" by (simp only: unit_factor_mult_normalize) then have "unit_factor (normalize a) * normalize a = normalize a" by simp with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a" by simp with * show ?thesis by simp qed lemma normalize_dvd_iff [simp]: "normalize a dvd b \ a dvd b" proof - have "normalize a dvd b \ unit_factor a * normalize a dvd b" using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b] by (cases "a = 0") simp_all then show ?thesis by simp qed lemma dvd_normalize_iff [simp]: "a dvd normalize b \ a dvd b" proof - have "a dvd normalize b \ a dvd normalize b * unit_factor b" using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"] by (cases "b = 0") simp_all then show ?thesis by simp qed lemma normalize_idem_imp_unit_factor_eq: assumes "normalize a = a" shows "unit_factor a = of_bool (a \ 0)" proof (cases "a = 0") case True then show ?thesis by simp next case False then show ?thesis using assms unit_factor_normalize [of a] by simp qed lemma normalize_idem_imp_is_unit_iff: assumes "normalize a = a" shows "is_unit a \ a = 1" using assms by (cases "a = 0") (auto dest: is_unit_normalize) lemma coprime_normalize_left_iff [simp]: "coprime (normalize a) b \ coprime a b" - by (rule; rule coprimeI) (auto intro: coprime_common_divisor) + by (rule iffI; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_normalize_right_iff [simp]: "coprime a (normalize b) \ coprime a b" using coprime_normalize_left_iff [of b a] by (simp add: ac_simps) text \ We avoid an explicit definition of associated elements but prefer explicit normalisation instead. In theory we could define an abbreviation like \<^prop>\associated a b \ normalize a = normalize b\ but this is counterproductive without suggestive infix syntax, which we do not want to sacrifice for this purpose here. \ lemma associatedI: assumes "a dvd b" and "b dvd a" shows "normalize a = normalize b" proof (cases "a = 0 \ b = 0") case True with assms show ?thesis by auto next case False from \a dvd b\ obtain c where b: "b = a * c" .. moreover from \b dvd a\ obtain d where a: "a = b * d" .. ultimately have "b * 1 = b * (c * d)" by (simp add: ac_simps) with False have "1 = c * d" unfolding mult_cancel_left by simp then have "is_unit c" and "is_unit d" by auto with a b show ?thesis by (simp add: is_unit_normalize) qed lemma associatedD1: "normalize a = normalize b \ a dvd b" using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric] by simp lemma associatedD2: "normalize a = normalize b \ b dvd a" using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric] by simp lemma associated_unit: "normalize a = normalize b \ is_unit a \ is_unit b" using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2) lemma associated_iff_dvd: "normalize a = normalize b \ a dvd b \ b dvd a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto intro!: associatedI) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * normalize b" by simp then have *: "normalize b * unit_factor a = a" by (simp add: ac_simps) show ?rhs proof (cases "a = 0 \ b = 0") case True with \?lhs\ show ?thesis by auto next case False then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b" by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff) with * show ?thesis by simp qed qed lemma associated_eqI: assumes "a dvd b" and "b dvd a" assumes "normalize a = a" and "normalize b = b" shows "a = b" proof - from assms have "normalize a = normalize b" unfolding associated_iff_dvd by simp with \normalize a = a\ have "a = normalize b" by simp with \normalize b = b\ show "a = b" by simp qed lemma normalize_unit_factor_eqI: assumes "normalize a = normalize b" and "unit_factor a = unit_factor b" shows "a = b" proof - from assms have "unit_factor a * normalize a = unit_factor b * normalize b" by simp then show ?thesis by simp qed lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) end class normalization_semidom_multiplicative = normalization_semidom + assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" begin lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False have "unit_factor (a * b) * normalize (a * b) = a * b" by (rule unit_factor_mult_normalize) then have "normalize (a * b) = a * b div unit_factor (a * b)" by simp also have "\ = a * b div unit_factor (b * a)" by (simp add: ac_simps) also have "\ = a * b div unit_factor b div unit_factor a" using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) also have "\ = a * (b div unit_factor b) div unit_factor a" using False by (subst unit_div_mult_swap) simp_all also have "\ = normalize a * normalize b" using False by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) finally show ?thesis . qed lemma dvd_unit_factor_div: assumes "b dvd a" shows "unit_factor (a div b) = unit_factor a div unit_factor b" proof - from assms have "a = a div b * b" by simp then have "unit_factor a = unit_factor (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: unit_factor_mult) qed lemma dvd_normalize_div: assumes "b dvd a" shows "normalize (a div b) = normalize a div normalize b" proof - from assms have "a = a div b * b" by simp then have "normalize a = normalize (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: normalize_mult) qed end text \Syntactic division remainder operator\ class modulo = dvd + divide + fixes modulo :: "'a \ 'a \ 'a" (infixl "mod" 70) text \Arbitrary quotient and remainder partitions\ class semiring_modulo = comm_semiring_1_cancel + divide + modulo + assumes div_mult_mod_eq: "a div b * b + a mod b = a" begin lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from div_mult_mod_eq have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma mult_div_mod_eq: "b * (a div b) + a mod b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_div_mult_eq: "a mod b + a div b * b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_mult_div_eq: "a mod b + b * (a div b) = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq) lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq) lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b" by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq) lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)" by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq) lemma mod_0_imp_dvd [dest!]: "b dvd a" if "a mod b = 0" proof - have "b dvd (a div b) * b" by simp also have "(a div b) * b = a" using div_mult_mod_eq [of a b] by (simp add: that) finally show ?thesis . qed lemma [nitpick_unfold]: "a mod b = a - a div b * b" by (fact minus_div_mult_eq_mod [symmetric]) end subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin lemma mod_0 [simp]: "0 mod a = 0" using div_mult_mod_eq [of 0 a] by simp lemma mod_by_0 [simp]: "a mod 0 = a" using div_mult_mod_eq [of a 0] by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using div_mult_mod_eq [of a a] by simp lemma dvd_imp_mod_0 [simp]: "b mod a = 0" if "a dvd b" using that minus_div_mult_eq_mod [of b a] by simp lemma mod_eq_0_iff_dvd: "a mod b = 0 \ b dvd a" by (auto intro: mod_0_imp_dvd) lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: "a dvd b \ b mod a = 0" by (simp add: mod_eq_0_iff_dvd) lemma dvd_mod_iff: assumes "c dvd b" shows "c dvd a mod b \ c dvd a" proof - from assms have "(c dvd a mod b) \ (c dvd ((a div b) * b + a mod b))" by (simp add: dvd_add_right_iff) also have "(a div b) * b + a mod b = a" using div_mult_mod_eq [of a b] by simp finally show ?thesis . qed lemma dvd_mod_imp_dvd: assumes "c dvd a mod b" and "c dvd b" shows "c dvd a" using assms dvd_mod_iff [of c b a] by simp lemma dvd_minus_mod [simp]: "b dvd a - a mod b" by (simp add: minus_mod_eq_div_mult) lemma cancel_div_mod_rules: "((a div b) * b + a mod b) + c = a + c" "(b * (a div b) + a mod b) + c = a + c" by (simp_all add: div_mult_mod_eq mult_div_mod_eq) end class idom_modulo = idom + semidom_modulo begin subclass idom_divide .. lemma div_diff [simp]: "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) end subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ ML_file \~~/src/Provers/Arith/cancel_div_mod.ML\ ML \ structure Cancel_Div_Mod_Ring = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val mk_sum = Arith_Data.mk_sum; val dest_sum = Arith_Data.dest_sum; val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ subsection \Ordered semirings and rings\ text \ The theory of partially ordered rings is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" begin lemma mult_mono: "a \ b \ c \ d \ 0 \ b \ 0 \ c \ a * c \ b * d" apply (erule (1) mult_right_mono [THEN order_trans]) apply (erule (1) mult_left_mono) done lemma mult_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a * c \ b * d" by (rule mult_mono) (fast intro: order_trans)+ end lemma mono_mult: fixes a :: "'a::ordered_semiring" shows "a \ 0 \ mono ((*) a)" by (simp add: mono_def mult_left_mono) class ordered_semiring_0 = semiring_0 + ordered_semiring begin lemma mult_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a * b" using mult_left_mono [of 0 b a] by simp lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" using mult_left_mono [of b 0 a] by simp lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" using mult_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_nonpos_nonneg}.\ lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" by (drule mult_right_mono [of b 0]) auto lemma split_mult_neg_le: "(0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b) \ a * b \ 0" by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) end class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass ordered_semiring_0 .. end class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add begin subclass ordered_cancel_semiring .. subclass ordered_cancel_comm_monoid_add .. subclass ordered_ab_semigroup_monoid_add_imp_le .. lemma mult_left_less_imp_less: "c * a < c * b \ 0 \ c \ a < b" by (force simp add: mult_left_mono not_le [symmetric]) lemma mult_right_less_imp_less: "a * c < b * c \ 0 \ c \ a < b" by (force simp add: mult_right_mono not_le [symmetric]) end class zero_less_one = order + zero + one + assumes zero_less_one [simp]: "0 < 1" begin subclass zero_neq_one by standard (simp add: less_imp_neq) lemma zero_le_one [simp]: \0 \ 1\ by (rule less_imp_le) simp end class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one begin lemma convex_bound_le: assumes "x \ a" "y \ a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass linordered_semiring proof fix a b c :: 'a assume *: "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from * show "a * c \ b * c" unfolding le_less using mult_strict_right_mono by (cases "c = 0") auto qed lemma mult_left_le_imp_le: "c * a \ c * b \ 0 < c \ a \ b" by (auto simp add: mult_strict_left_mono _not_less [symmetric]) lemma mult_right_le_imp_le: "a * c \ b * c \ 0 < c \ a \ b" by (auto simp add: mult_strict_right_mono not_less [symmetric]) lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_neg_pos}.\ lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" by (drule mult_strict_right_mono [of b 0]) auto lemma zero_less_mult_pos: assumes "0 < a * b" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b]) qed (auto simp add: le_less not_less) lemma zero_less_mult_pos2: assumes "0 < b * a" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg2 [of a b]) qed (auto simp add: le_less not_less) text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes "a < b" "c < d" "0 < b" "0 \ c" shows "a * c < b * d" proof (cases "c = 0") case True with assms show ?thesis by simp next case False with assms have "a*c < b*c" by (simp add: mult_strict_right_mono [OF \a < b\]) also have "\ < b*d" by (simp add: assms mult_strict_left_mono) finally show ?thesis . qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" - by (rule mult_strict_mono) (insert assms, auto) + using assms by (auto simp add: mult_strict_mono) lemma mult_less_le_imp_less: assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" shows "a * c < b * d" proof - have "a * c < b * c" by (simp add: assms mult_strict_right_mono) also have "... \ b * d" by (intro mult_left_mono) (use assms in auto) finally show ?thesis . qed lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" proof - have "a * c \ b * c" by (simp add: assms mult_right_mono) also have "... < b * d" by (intro mult_strict_left_mono) (use assms in auto) finally show ?thesis . qed end class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one begin subclass linordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" "y < a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add + assumes comm_mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" begin subclass ordered_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" by (rule comm_mult_left_mono) then show "a * c \ b * c" by (simp only: mult.commute) qed end class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add begin subclass comm_semiring_0_cancel .. subclass ordered_comm_semiring .. subclass ordered_cancel_semiring .. end class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass linordered_semiring_strict proof fix a b c :: 'a assume "a < b" "0 < c" then show "c * a < c * b" by (rule comm_mult_strict_left_mono) then show "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end class ordered_ring = ring + ordered_cancel_semiring begin subclass ordered_ab_group_add .. lemma less_add_iff1: "a * e + c < b * e + d \ (a - b) * e + c < d" by (simp add: algebra_simps) lemma less_add_iff2: "a * e + c < b * e + d \ c < (b - a) * e + d" by (simp add: algebra_simps) lemma le_add_iff1: "a * e + c \ b * e + d \ (a - b) * e + c \ d" by (simp add: algebra_simps) lemma le_add_iff2: "a * e + c \ b * e + d \ c \ (b - a) * e + d" by (simp add: algebra_simps) lemma mult_left_mono_neg: "b \ a \ c \ 0 \ c * a \ c * b" by (auto dest: mult_left_mono [of _ _ "- c"]) lemma mult_right_mono_neg: "b \ a \ c \ 0 \ a * c \ b * c" by (auto dest: mult_right_mono [of _ _ "- c"]) lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" using mult_right_mono_neg [of a 0 b] by simp lemma split_mult_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" by (auto simp add: mult_nonpos_nonpos) end class abs_if = minus + uminus + ord + zero + abs + assumes abs_if: "\a\ = (if a < 0 then - a else a)" class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if begin subclass ordered_ring .. subclass ordered_ab_group_add_abs proof fix a b show "\a + b\ \ \a\ + \b\" by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg) qed (auto simp: abs_if) lemma zero_le_square [simp]: "0 \ a * a" using linear [of 0 a] by (auto simp add: mult_nonpos_nonpos) lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: if_split_asm) lemma abs_eq_iff': "\a\ = b \ b \ 0 \ (a = b \ a = - b)" by (cases "a \ 0") auto lemma eq_abs_iff': "a = \b\ \ a \ 0 \ (b = a \ b = - a)" using abs_eq_iff' [of b a] by auto lemma sum_squares_ge_zero: "0 \ x * x + y * y" by (intro add_nonneg_nonneg zero_le_square) lemma not_sum_squares_lt_zero: "\ x * x + y * y < 0" by (simp add: not_less sum_squares_ge_zero) end class linordered_ring_strict = ring + linordered_semiring_strict + ordered_ab_group_add + abs_if begin subclass linordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp subclass ring_no_zero_divisors proof fix a b assume "a \ 0" then have a: "a < 0 \ 0 < a" by (simp add: neq_iff) assume "b \ 0" then have b: "b < 0 \ 0 < b" by (simp add: neq_iff) have "a * b < 0 \ 0 < a * b" proof (cases "a < 0") case True show ?thesis proof (cases "b < 0") case True with \a < 0\ show ?thesis by (auto dest: mult_neg_neg) next case False with b have "0 < b" by auto with \a < 0\ show ?thesis by (auto dest: mult_strict_right_mono) qed next case False with a have "0 < a" by auto show ?thesis proof (cases "b < 0") case True with \0 < a\ show ?thesis by (auto dest: mult_strict_right_mono_neg) next case False with b have "0 < b" by auto with \0 < a\ show ?thesis by auto qed qed then show "a * b \ 0" by (simp add: neq_iff) qed lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]: "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2) lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]: "0 \ a * b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff) lemma mult_less_0_iff [algebra_split_simps, field_split_simps]: "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" using zero_less_mult_iff [of "- a" b] by auto lemma mult_le_0_iff [algebra_split_simps, field_split_simps]: "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" using zero_le_mult_iff [of "- a" b] by auto text \ Cancellation laws for \<^term>\c * a < c * b\ and \<^term>\a * c < b * c\, also with the relations \\\ and equality. \ text \ These ``disjunction'' versions produce two cases when the comparison is an assumption, but effectively four when the comparison is a goal. \ lemma mult_less_cancel_right_disj: "a * c < b * c \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg) qed qed auto lemma mult_less_cancel_left_disj: "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg) qed qed auto text \ The ``conjunction of implication'' lemmas produce two cases when the comparison is a goal, but give four when the comparison is an assumption. \ lemma mult_less_cancel_right: "a * c < b * c \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_right_disj [of a c b] by auto lemma mult_less_cancel_left: "c * a < c * b \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_left_disj [of c a b] by auto lemma mult_le_cancel_right: "a * c \ b * c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_right_disj) lemma mult_le_cancel_left: "c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_left_disj) lemma mult_le_cancel_left_pos: "0 < c \ c * a \ c * b \ a \ b" by (auto simp: mult_le_cancel_left) lemma mult_le_cancel_left_neg: "c < 0 \ c * a \ c * b \ b \ a" by (auto simp: mult_le_cancel_left) lemma mult_less_cancel_left_pos: "0 < c \ c * a < c * b \ a < b" by (auto simp: mult_less_cancel_left) lemma mult_less_cancel_left_neg: "c < 0 \ c * a < c * b \ b < a" by (auto simp: mult_less_cancel_left) end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg class ordered_comm_ring = comm_ring + ordered_comm_semiring begin subclass ordered_ring .. subclass ordered_cancel_comm_semiring .. end class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one + assumes add_mono1: "a < b \ a + 1 < b + 1" begin subclass zero_neq_one - by standard (insert zero_less_one, blast) + by standard subclass comm_semiring_1 by standard (rule mult_1_left) lemma zero_le_one [simp]: "0 \ 1" by (rule zero_less_one [THEN less_imp_le]) lemma not_one_le_zero [simp]: "\ 1 \ 0" by (simp add: not_le) lemma not_one_less_zero [simp]: "\ 1 < 0" by (simp add: not_less) lemma of_bool_less_eq_iff [simp]: \of_bool P \ of_bool Q \ (P \ Q)\ by auto lemma of_bool_less_iff [simp]: \of_bool P < of_bool Q \ \ P \ Q\ by auto lemma mult_left_le: "c \ 1 \ 0 \ a \ a * c \ a" using mult_left_mono[of c 1 a] by simp lemma mult_le_one: "a \ 1 \ 0 \ b \ b \ 1 \ a * b \ 1" using mult_mono[of a 1 b 1] by simp lemma zero_less_two: "0 < 1 + 1" using add_pos_pos[OF zero_less_one zero_less_one] . end class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one + assumes le_add_diff_inverse2 [simp]: "b \ a \ a - b + b = a" begin subclass linordered_nonzero_semiring proof show "a + 1 < b + 1" if "a < b" for a b - proof (rule ccontr, simp add: not_less) - assume "b \ a" - with that show False + proof (rule ccontr) + assume "\ a + 1 < b + 1" + moreover with that have "a + 1 < b + 1" by simp + ultimately show False + by contradiction qed qed lemma zero_less_eq_of_bool [simp]: \0 \ of_bool P\ by simp lemma zero_less_of_bool_iff [simp]: \0 < of_bool P \ P\ by simp lemma of_bool_less_eq_one [simp]: \of_bool P \ 1\ by simp lemma of_bool_less_one_iff [simp]: \of_bool P < 1 \ \ P\ by simp lemma of_bool_or_iff [simp]: \of_bool (P \ Q) = max (of_bool P) (of_bool Q)\ by (simp add: max_def) text \Addition is the inverse of subtraction.\ lemma le_add_diff_inverse [simp]: "b \ a \ b + (a - b) = a" by (frule le_add_diff_inverse2) (simp add: add.commute) lemma add_diff_inverse: "\ a < b \ b + (a - b) = a" by simp lemma add_le_imp_le_diff: assumes "i + k \ n" shows "i \ n - k" proof - have "n - (i + k) + i + k = n" by (simp add: assms add.assoc) with assms add_implies_diff have "i + k \ n - k + k" by fastforce then show ?thesis by simp qed lemma add_le_add_imp_diff_le: assumes 1: "i + k \ n" and 2: "n \ j + k" shows "i + k \ n \ n \ j + k \ n - k \ j" proof - have "n - (i + k) + i + k = n" using 1 by (simp add: add.assoc) moreover have "n - k = n - k - i + i" using 1 by (simp add: add_le_imp_le_diff) ultimately show ?thesis using 2 add_le_imp_le_diff [of "n-k" k "j + k"] by (simp add: add.commute diff_diff_add) qed lemma less_1_mult: "1 < m \ 1 < n \ 1 < m * n" using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one]) end class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" begin subclass linordered_ring_strict .. subclass linordered_semiring_1_strict proof have "0 \ 1 * 1" by (fact zero_le_square) then show "0 < 1" by (simp add: le_less) qed subclass ordered_comm_ring .. subclass idom .. subclass linordered_semidom by standard simp subclass idom_abs_sgn by standard (auto simp add: sgn_if abs_if zero_less_mult_iff) lemma abs_bool_eq [simp]: \\of_bool P\ = of_bool P\ by simp lemma linorder_neqE_linordered_idom: assumes "x \ y" obtains "x < y" | "y < x" using assms by (rule neqE) text \These cancellation simp rules also produce two cases when the comparison is a goal.\ lemma mult_le_cancel_right1: "c \ b * c \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_right [of 1 c b] by simp lemma mult_le_cancel_right2: "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_right [of a c 1] by simp lemma mult_le_cancel_left1: "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_left [of c 1 b] by simp lemma mult_le_cancel_left2: "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_left [of c a 1] by simp lemma mult_less_cancel_right1: "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_right [of 1 c b] by simp lemma mult_less_cancel_right2: "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_right [of a c 1] by simp lemma mult_less_cancel_left1: "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_left [of c 1 b] by simp lemma mult_less_cancel_left2: "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_left [of c a 1] by simp lemma sgn_0_0: "sgn a = 0 \ a = 0" by (fact sgn_eq_0_iff) lemma sgn_1_pos: "sgn a = 1 \ a > 0" unfolding sgn_if by simp lemma sgn_1_neg: "sgn a = - 1 \ a < 0" unfolding sgn_if by auto lemma sgn_pos [simp]: "0 < a \ sgn a = 1" by (simp only: sgn_1_pos) lemma sgn_neg [simp]: "a < 0 \ sgn a = - 1" by (simp only: sgn_1_neg) lemma abs_sgn: "\k\ = k * sgn k" unfolding sgn_if abs_if by auto lemma sgn_greater [simp]: "0 < sgn a \ 0 < a" unfolding sgn_if by auto lemma sgn_less [simp]: "sgn a < 0 \ a < 0" unfolding sgn_if by auto lemma abs_sgn_eq_1 [simp]: "a \ 0 \ \sgn a\ = 1" by simp lemma abs_sgn_eq: "\sgn a\ = (if a = 0 then 0 else 1)" by (simp add: sgn_if) lemma sgn_mult_self_eq [simp]: "sgn a * sgn a = of_bool (a \ 0)" by (cases "a > 0") simp_all lemma abs_mult_self_eq [simp]: "\a\ * \a\ = a * a" by (cases "a > 0") simp_all lemma same_sgn_sgn_add: "sgn (a + b) = sgn a" if "sgn b = sgn a" proof (cases a 0 rule: linorder_cases) case equal with that show ?thesis by simp next case less with that have "b < 0" by (simp add: sgn_1_neg) with \a < 0\ have "a + b < 0" by (rule add_neg_neg) with \a < 0\ show ?thesis by simp next case greater with that have "b > 0" by (simp add: sgn_1_pos) with \a > 0\ have "a + b > 0" by (rule add_pos_pos) with \a > 0\ show ?thesis by simp qed lemma same_sgn_abs_add: "\a + b\ = \a\ + \b\" if "sgn b = sgn a" proof - have "a + b = sgn a * \a\ + sgn b * \b\" by (simp add: sgn_mult_abs) also have "\ = sgn a * (\a\ + \b\)" using that by (simp add: algebra_simps) finally show ?thesis by (auto simp add: abs_mult) qed lemma sgn_not_eq_imp: "sgn a = - sgn b" if "sgn b \ sgn a" and "sgn a \ 0" and "sgn b \ 0" using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg) lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) lemma dvd_abs_iff [simp]: "m dvd \k\ \ m dvd k" by (simp add: abs_if) lemma dvd_if_abs_eq: "\l\ = \k\ \ l dvd k" by (subst abs_dvd_iff [symmetric]) simp text \ The following lemmas can be proven in more general structures, but are dangerous as simp rules in absence of @{thm neg_equal_zero}, @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. \ lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \ a = - 1" by (fact equation_minus_iff) lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \ a = - 1" by (subst minus_equation_iff, auto) lemma le_minus_iff_1 [simp, no_atp]: "1 \ - b \ b \ - 1" by (fact le_minus_iff) lemma minus_le_iff_1 [simp, no_atp]: "- a \ 1 \ - 1 \ a" by (fact minus_le_iff) lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \ b < - 1" by (fact less_minus_iff) lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \ - 1 < a" by (fact minus_less_iff) lemma add_less_zeroD: shows "x+y < 0 \ x<0 \ y<0" by (auto simp: not_less intro: le_less_trans [of _ "x+y"]) end text \Reasoning about inequalities with division\ context linordered_semidom begin lemma less_add_one: "a < a + 1" proof - have "a + 0 < a + 1" by (blast intro: zero_less_one add_strict_left_mono) then show ?thesis by simp qed end context linordered_idom begin lemma mult_right_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ x * y \ x" by (rule mult_left_le) lemma mult_left_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ y * x \ x" by (auto simp add: mult_le_cancel_right2) end text \Absolute Value\ context linordered_idom begin lemma mult_sgn_abs: "sgn x * \x\ = x" by (fact sgn_mult_abs) lemma abs_one: "\1\ = 1" by (fact abs_1) end class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs + assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" context linordered_idom begin subclass ordered_ring_abs by standard (auto simp: abs_if not_less mult_less_0_iff) lemma abs_mult_self: "\a\ * \a\ = a * a" by (fact abs_mult_self_eq) lemma abs_mult_less: assumes ac: "\a\ < c" and bd: "\b\ < d" shows "\a\ * \b\ < c * d" proof - from ac have "0 < c" by (blast intro: le_less_trans abs_ge_zero) with bd show ?thesis by (simp add: ac mult_strict_mono) qed lemma abs_less_iff: "\a\ < b \ a < b \ - a < b" by (simp add: less_le abs_le_iff) (auto simp add: abs_if) lemma abs_mult_pos: "0 \ x \ \y\ * x = \y * x\" by (simp add: abs_mult) lemma abs_mult_pos': "0 \ x \ x * \y\ = \x * y\" by (simp add: abs_mult) lemma abs_diff_less_iff: "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) lemma abs_diff_le_iff: "\x - a\ \ r \ a - r \ x \ x \ a + r" by (auto simp add: diff_le_eq ac_simps abs_le_iff) lemma abs_add_one_gt_zero: "0 < 1 + \x\" by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans) end subsection \Dioids\ text \ Dioids are the alternative extensions of semirings, a semiring can either be a ring or a dioid but never both. \ class dioid = semiring_1 + canonically_ordered_monoid_add begin subclass ordered_semiring by standard (auto simp: le_iff_add distrib_left distrib_right) end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib code_identifier code_module Rings \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Set_Interval.thy b/src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy +++ b/src/HOL/Set_Interval.thy @@ -1,2563 +1,2615 @@ (* Title: HOL/Set_Interval.thy Author: Tobias Nipkow, Clemens Ballarin, Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals Modern convention: Ixy stands for an interval where x and y describe the lower and upper bound and x,y : {c,o,i} where c = closed, o = open, i = infinite. Examples: Ico = {_ ..< _} and Ici = {_ ..} *) section \Set intervals\ theory Set_Interval imports Divides begin (* Belongs in Finite_Set but 2 is not available there *) lemma card_2_iff: "card S = 2 \ (\x y. S = {x,y} \ x \ y)" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_3_iff: "card S = 3 \ (\x y z. S = {x,y,z} \ x \ y \ y \ z \ x \ z)" by (fastforce simp: card_Suc_eq numeral_eq_Suc) context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text\A note of warning when using \<^term>\{.. on type \<^typ>\nat\: it is equivalent to \<^term>\{0::nat.. but some lemmas involving \<^term>\{m.. may not exist in \<^term>\{..-form as well.\ syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) syntax "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) translations "\i\n. A" \ "\i\{..n}. A" "\i "\i\{..i\n. A" \ "\i\{..n}. A" "\i "\i\{..Various equivalences\ lemma (in ord) lessThan_iff [iff]: "(i \ lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \ { b <..} = { max a b <..}" by auto lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" by auto subsection \Logical Equivalences for Set Inclusion and Equality\ lemma atLeast_empty_triv [simp]: "{{}..} = UNIV" by auto lemma atMost_UNIV_triv [simp]: "{..UNIV} = UNIV" by auto lemma atLeast_subset_iff [iff]: "(atLeast x \ atLeast y) = (y \ (x::'a::preorder))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::preorder))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" by auto lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" by auto lemma (in preorder) Ioi_le_Ico: "{a <..} \ {a ..}" by (auto intro: less_imp_le) subsection \Two-sided intervals\ context ord begin lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them alone.\ lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" by auto lemma (in order) atLeastLessThan_eq_atLeastAtMost_diff: "{a..Emptyness, singletons, subset\ context preorder begin lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty_iff[simp]: "{a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ \ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ \ k < l" by auto (blast intro: less_le_trans) lemma atLeastatMost_subset_iff[simp]: "{a..b} \ {c..d} \ (\ a \ b) \ c \ a \ b \ d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((\ a \ b) \ c \ a \ b \ d \ (c < a \ b < d)) \ c \ d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a..b} \ {c ..< d} \ (a \ b \ c \ a \ b < d)" by auto (blast intro: local.order_trans local.le_less_trans elim: )+ lemma Icc_subset_Ici_iff[simp]: "{l..h} \ {l'..} = (\ l\h \ l\l')" by(auto simp: subset_eq intro: order_trans) lemma Icc_subset_Iic_iff[simp]: "{l..h} \ {..h'} = (\ l\h \ h\h')" by(auto simp: subset_eq intro: order_trans) lemma not_Ici_eq_empty[simp]: "{l..} \ {}" by(auto simp: set_eq_iff) lemma not_Iic_eq_empty[simp]: "{..h} \ {}" by(auto simp: set_eq_iff) lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] end context order begin lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: "b \ a \ {a.. k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma Icc_eq_Icc[simp]: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" by (simp add: order_class.order.eq_iff) (auto intro: order_trans) lemma (in linorder) Ico_eq_Ico: "{l.. h=h' \ \ l \ l' a = b \ b = c" proof assume "{a..b} = {c}" hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp with \{a..b} = {c}\ have "c \ a \ b \ c" by auto with * show "a = b \ b = c" by auto qed simp end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_bot begin lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] (* also holds for no_bot but no_top should suffice *) lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" using not_Ici_le_Iic[of l' h] by blast lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] end context no_bot begin lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanLessThan: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end context no_top begin lemma greaterThan_non_empty[simp]: "{x <..} \ {}" using gt_ex[of x] by auto end context no_bot begin lemma lessThan_non_empty[simp]: "{..< x} \ {}" using lt_ex[of x] by auto end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. b \ a \ c\a \ b\d" - apply (auto simp:subset_eq Ball_def not_le) - apply(frule_tac x=a in spec) - apply(erule_tac x=d in allE) - apply auto - done +proof (cases "a < b") + case True + assume assm: "{a.. {c.. a \ a \ d" + using True by (auto simp add: subset_eq Ball_def) + then have 2: "b \ d" + using assm by (auto simp add: subset_eq) + from 1 2 show ?thesis + by simp +qed (auto) lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto lemma (in linorder) Ioc_inj: \{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d\ (is \?P \ ?Q\) proof assume ?Q then show ?P by auto next assume ?P then have \a < x \ x \ b \ c < x \ x \ d\ for x by (simp add: set_eq_iff) from this [of a] this [of b] this [of c] this [of d] show ?Q by auto qed lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" by auto lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" by (auto simp: subset_eq Ball_def) (metis less_le not_less) lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" by (auto simp: set_eq_iff intro: le_bot) lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" by (auto simp: set_eq_iff intro: top_le) lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: "{x..y} = UNIV \ (x = bot \ y = top)" by (auto simp: set_eq_iff intro: top_le le_bot) lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" shows "f m = m'" proof - from f_img have "{m' ..< n'} \ {}" by (metis atLeastLessThan_empty_iff image_is_empty) with f_img have "m' \ f ` {m ..< n}" by auto then obtain k where "f k = m'" "m \ k" by auto moreover have "m' \ f m" using f_img by auto ultimately show "f m = m'" using f_mono by (auto elim: monoE[where x=m and y=k]) qed subsection \Infinite intervals\ context dense_linorder begin lemma infinite_Ioo: assumes "a < b" shows "\ finite {a<.. {}" using \a < b\ by auto ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" using Max_in[of "{a <..< b}"] by auto then obtain x where "Max {a <..< b} < x" "x < b" using dense[of "Max {a<.. {a <..< b}" using \a < Max {a <..< b}\ by auto then have "x \ Max {a <..< b}" using fin by auto with \Max {a <..< b} < x\ show False by auto qed lemma infinite_Icc: "a < b \ \ finite {a .. b}" using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ico: "a < b \ \ finite {a ..< b}" using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioo_iff [simp]: "infinite {a<.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioo) lemma infinite_Icc_iff [simp]: "infinite {a .. b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Icc) lemma infinite_Ico_iff [simp]: "infinite {a.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ico) lemma infinite_Ioc_iff [simp]: "infinite {a<..b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioc) end lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" proof assume "finite {..< a}" then have *: "\x. x < a \ Min {..< a} \ x" by auto obtain x where "x < a" using lt_ex by auto obtain y where "y < Min {..< a}" using lt_ex by auto also have "Min {..< a} \ x" using \x < a\ by fact also note \x < a\ finally have "Min {..< a} \ y" by fact with \y < Min {..< a}\ show False by auto qed lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] by (auto simp: subset_eq less_imp_le) lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" proof assume "finite {a <..}" then have *: "\x. a < x \ x \ Max {a <..}" by auto obtain y where "Max {a <..} < y" using gt_ex by auto obtain x where x: "a < x" using gt_ex by auto also from x have "x \ Max {a <..}" by fact also note \Max {a <..} < y\ finally have "y \ Max { a <..}" by fact with \Max {a <..} < y\ show False by auto qed lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] by (auto simp: subset_eq less_imp_le) subsubsection \Intersection\ context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" by auto end context complete_lattice begin lemma shows Sup_atLeast[simp]: "Sup {x ..} = top" and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" and Sup_atMost[simp]: "Sup {.. y} = y" and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" by (auto intro!: Sup_eqI) lemma shows Inf_atMost[simp]: "Inf {.. x} = bot" and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" and Inf_atLeast[simp]: "Inf {x ..} = x" and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" by (auto intro!: Inf_eqI) end lemma fixes x y :: "'a :: {complete_lattice, dense_linorder}" shows Sup_lessThan[simp]: "Sup {..< y} = y" and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" and Inf_greaterThan[simp]: "Inf {x <..} = x" and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) subsection \Intervals of natural numbers\ subsubsection \The Constant \<^term>\lessThan\\ lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text \The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform \<^term>\{.. to \<^term>\0::nat\ and \<^term>\{..< n}\.\ lemma zero_notin_Suc_image [simp]: "0 \ Suc ` A" by auto lemma lessThan_Suc_eq_insert_0: "{..m::nat. lessThan m) = UNIV" by blast subsubsection \The Constant \<^term>\greaterThan\\ lemma greaterThan_0: "greaterThan 0 = range Suc" unfolding greaterThan_def by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" unfolding greaterThan_def by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(\m::nat. greaterThan m) = {}" by blast subsubsection \The Constant \<^term>\atLeast\\ lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(\m::nat. atLeast m) = UNIV" by blast subsubsection \The Constant \<^term>\atMost\\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(\m::nat. atMost m) = UNIV" by blast subsubsection \The Constant \<^term>\atLeastLessThan\\ text\The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat..The Constant \<^term>\atLeastAtMost\\ lemma Icc_eq_insert_lb_nat: "m \ n \ {m..n} = insert m {Suc m..n}" by auto lemma atLeast0_atMost_Suc: "{0..Suc n} = insert (Suc n) {0..n}" by (simp add: atLeast0AtMost atMost_Suc) lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with \<^term>\Suc\\ text\Not a simprule because the RHS is too messy.\ lemma atLeastLessThanSuc: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by auto lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text \The analogous result is useful on \<^typ>\int\:\ (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m \ 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lemma lessThan_nat_numeral: \ \Evaluation for specific numerals\ "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" by (simp add: numeral_eq_Suc lessThan_Suc) lemma atMost_nat_numeral: \ \Evaluation for specific numerals\ "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" by (simp add: numeral_eq_Suc atMost_Suc) lemma atLeastLessThan_nat_numeral: \ \Evaluation for specific numerals\ "atLeastLessThan m (numeral k :: nat) = (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) else {})" by (simp add: numeral_eq_Suc atLeastLessThanSuc) subsubsection \Image\ context linordered_semidom begin lemma image_add_atLeast[simp]: "plus k ` {i..} = {k + i..}" proof - have "n = k + (n - k)" if "i + k \ n" for n proof - have "n = (n - (k + i)) + (k + i)" using that by (metis add_commute le_add_diff_inverse) then show "n = k + (n - k)" by (metis local.add_diff_cancel_left' add_assoc add_commute) qed then show ?thesis by (fastforce simp: add_le_imp_le_diff add.commute) qed lemma image_add_atLeastAtMost [simp]: "plus k ` {i..j} = {i + k..j + k}" (is "?A = ?B") proof show "?A \ ?B" by (auto simp add: ac_simps) next show "?B \ ?A" proof fix n assume "n \ ?B" then have "i \ n - k" by (simp add: add_le_imp_le_diff) have "n = n - k + k" proof - from \n \ ?B\ have "n = n - (i + k) + (i + k)" by simp also have "\ = n - k - i + i + k" by (simp add: algebra_simps) also have "\ = n - k + k" using \i \ n - k\ by simp finally show ?thesis . qed moreover have "n - k \ {i..j}" using \n \ ?B\ by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) ultimately show "n \ ?A" by (simp add: ac_simps) qed qed lemma image_add_atLeastAtMost' [simp]: "(\n. n + k) ` {i..j} = {i + k..j + k}" by (simp add: add.commute [of _ k]) lemma image_add_atLeastLessThan [simp]: "plus k ` {i..n. n + k) ` {i.. uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" - by (rule imageI) (insert \y \ -x\[THEN le_imp_neg_le], simp) + by (rule imageI) (use \y \ -x\[THEN le_imp_neg_le] in \simp\) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. (-) d ` {a..b}" + proof + fix x + assume "x \ {d - b..d - a}" + then have "d - x \ {a..b}" and "x = d - (d - x)" + by auto + then show "x \ (-) d ` {a..b}" + by (rule rev_image_eqI) + qed +qed(auto) lemma image_diff_atLeastLessThan [simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a.. = {c - b<..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_greaterThanAtMost[simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a<..b} = {c - b.. = {c - b.. = {..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_AtMost[simp]: fixes b c::"'a::linordered_idom" shows "(-) c ` {..b} = {c - b..}" proof - have "(-) c ` {..b} = (+) c ` uminus ` {..b}" unfolding image_image by simp also have "\ = {c - b..}" by simp finally show ?thesis by simp qed lemma image_minus_const_atLeastAtMost' [simp]: "(\t. t-d)`{a..b} = {a-d..b-d}" for d::"'a::linordered_idom" by (metis (no_types, lifting) diff_conv_add_uminus image_add_atLeastAtMost' image_cong) context linordered_field begin lemma image_mult_atLeastAtMost [simp]: "((*) d ` {a..b}) = {d*a..d*b}" if "d>0" using that by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lemma image_divide_atLeastAtMost [simp]: "((\c. c / d) ` {a..b}) = {a/d..b/d}" if "d>0" proof - from that have "inverse d > 0" by simp with image_mult_atLeastAtMost [of "inverse d" a b] have "(*) (inverse d) ` {a..b} = {inverse d * a..inverse d * b}" by blast moreover have "(*) (inverse d) = (\c. c / d)" by (simp add: fun_eq_iff field_simps) ultimately show ?thesis by simp qed lemma image_mult_atLeastAtMost_if: "(*) c ` {x .. y} = (if c > 0 then {c * x .. c * y} else if x \ y then {c * y .. c * x} else {})" proof (cases "c = 0 \ x > y") case True then show ?thesis by auto next case False then have "x \ y" by auto from False consider "c < 0"| "c > 0" by (auto simp add: neq_iff) then show ?thesis proof cases case 1 have "(*) c ` {x..y} = {c * y..c * x}" proof (rule set_eqI) fix d from 1 have "inj (\z. z / c)" by (auto intro: injI) then have "d \ (*) c ` {x..y} \ d / c \ (\z. z div c) ` (*) c ` {x..y}" by (subst inj_image_mem_iff) simp_all also have "\ \ d / c \ {x..y}" using 1 by (simp add: image_image) also have "\ \ d \ {c * y..c * x}" by (auto simp add: field_simps 1) finally show "d \ (*) c ` {x..y} \ d \ {c * y..c * x}" . qed with \x \ y\ show ?thesis by auto qed (simp add: mult_left_mono_neg) qed lemma image_mult_atLeastAtMost_if': "(\x. x * c) ` {x..y} = (if x \ y then if c > 0 then {x * c .. y * c} else {y * c .. x * c} else {})" using image_mult_atLeastAtMost_if [of c x y] by (auto simp add: ac_simps) lemma image_affinity_atLeastAtMost: "((\x. m * x + c) ` {a..b}) = (if {a..b} = {} then {} else if 0 \ m then {m * a + c .. m * b + c} else {m * b + c .. m * a + c})" proof - have *: "(\x. m * x + c) = ((\x. x + c) \ (*) m)" by (simp add: fun_eq_iff) show ?thesis by (simp only: * image_comp [symmetric] image_mult_atLeastAtMost_if) (auto simp add: mult_le_cancel_left) qed lemma image_affinity_atLeastAtMost_diff: "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m*a - c .. m*b - c} else {m*b - c .. m*a - c})" using image_affinity_atLeastAtMost [of m "-c" a b] by simp lemma image_affinity_atLeastAtMost_div: "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m + c .. b/m + c} else {b/m + c .. a/m + c})" using image_affinity_atLeastAtMost [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) lemma image_affinity_atLeastAtMost_div_diff: "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m - c .. b/m - c} else {b/m - c .. a/m - c})" using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) end lemma atLeast1_lessThan_eq_remove0: "{Suc 0..x. x + (l::int)) ` {0..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: if_split_asm) qed qed auto lemma image_int_atLeastLessThan: "int ` {a..Finiteness\ lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (\m. \n\N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast next assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" unfolding finite_nat_set_iff_bounded by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: - "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" -by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) + "\f::nat\nat. (!!n. n \ f n) \ finite {n. f n \ u}" + by (rule finite_subset[of _ "{..u}"]) + (auto intro: order_trans) lemma bounded_Max_nat: fixes P :: "nat \ bool" assumes x: "P x" and M: "\x. P x \ x \ M" obtains m where "P m" "\x. P x \ x \ m" proof - have "finite {x. P x}" using M finite_nat_set_iff_bounded_le by auto then have "Max {x. P x} \ {x. P x}" using Max_in x by auto then show ?thesis by (simp add: \finite {x. P x}\ that) qed text\Any subset of an interval of natural numbers the size of the subset is exactly that interval.\ lemma subset_card_intvl_is_intvl: assumes "A \ {k.. A" by auto with insert have "A \ {k..Proving Inclusions and Equalities between Unions\ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i\n" "x \ M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A \ ?B" by fastforce next show "?B \ ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed qed lemma UN_le_add_shift_strict: "(\ii\{k.. ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k.. M(i)" by auto then have "i - k < n \ x \ M((i-k) + k)" by auto then show "x \ ?A" using UN_le_add_shift by blast qed qed (fastforce) lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" -proof (rule UN_finite_subset, rule) +proof (rule UN_finite_subset, rule subsetI) fix n and a from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) qed lemma UN_finite2_eq: - "(\n::nat. (\i\{0..i\{0.. - (\n. A n) = (\n. B n)" - apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) - apply auto - apply (force simp add: atLeastLessThan_add_Un [of 0])+ - done + assumes "(\n::nat. (\i\{0..i\{0..n. A n) = (\n. B n)" +proof (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) + fix n + show "\ (A ` {0.. (\n. B n)" + using assms by auto +next + fix n + show "\ (B ` {0.. \ (A ` {0..Cardinality\ lemma card_lessThan [simp]: "card {..x. x + l) ` {..x. x + l) ` {.. {l.. (\x. x + l) ` {.. {l.. {..< u -l}" + by auto + then have "(x - l) + l \ (\x. x + l) ` {..< u -l}" + by auto + then show "x \ (\x. x + l) ` {..x. x + l) ` {.. {0.. {0..n}" shows "finite N" using assms finite_atLeastAtMost by (rule finite_subset) lemma ex_bij_betw_nat_finite: "finite M \ \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" shows "(\f. bij_betw f A B) \ (card A = card B)" proof assume "card A = card B" moreover obtain f where "bij_betw f A {0 ..< card A}" using assms ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using assms ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g \ f) A B" by (auto simp: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed (auto simp: bij_betw_same_card) lemma subset_eq_atLeast0_lessThan_card: fixes n :: nat assumes "N \ {0.. n" proof - from assms finite_lessThan have "card N \ card {0..Relational version of @{thm [source] card_inj_on_le}:\ lemma card_le_if_inj_on_rel: assumes "finite B" "\a. a \ A \ \b. b\B \ r a b" "\a1 a2 b. \ a1 \ A; a2 \ A; b \ B; r a1 b; r a2 b \ \ a1 = a2" shows "card A \ card B" proof - let ?P = "\a b. b \ B \ r a b" let ?f = "\a. SOME b. ?P a b" have 1: "?f ` A \ B" by (auto intro: someI2_ex[OF assms(2)]) have "inj_on ?f A" - proof (auto simp: inj_on_def) + unfolding inj_on_def + proof safe fix a1 a2 assume asms: "a1 \ A" "a2 \ A" "?f a1 = ?f a2" have 0: "?f a1 \ B" using "1" \a1 \ A\ by blast have 1: "r a1 (?f a1)" using someI_ex[OF assms(2)[OF \a1 \ A\]] by blast have 2: "r a2 (?f a1)" using someI_ex[OF assms(2)[OF \a2 \ A\]] asms(3) by auto show "a1 = a2" using assms(3)[OF asms(1,2) 0 1 2] . qed with 1 show ?thesis using card_inj_on_le[of ?f A B] assms(1) by simp qed lemma inj_on_funpow_least: \<^marker>\contributor \Lars Noschinski\\ \inj_on (\k. (f ^^ k) s) {0.. if \(f ^^ n) s = s\ \\m. 0 < m \ m < n \ (f ^^ m) s \ s\ proof - { fix k l assume A: "k < n" "l < n" "k \ l" "(f ^^ k) s = (f ^^ l) s" define k' l' where "k' = min k l" and "l' = max k l" with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" by (auto simp: min_def max_def) have "s = (f ^^ ((n - l') + l')) s" using that \l' < n\ by simp also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) finally have "(f ^^ (n - l' + k')) s = s" by simp moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ ultimately have False using that(2) by auto } then show ?thesis by (intro inj_onI) auto qed subsection \Intervals of integers\ lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ -lemma image_atLeastZeroLessThan_int: "0 \ u ==> - {(0::int).. u" + shows "{(0::int).. {y. \x\{x. x < nat u}. y = int x}" + proof + fix x + assume "x \ {0..xa {y. \x\{x. x < nat u}. y = int x}" + by simp + qed +qed (auto) + lemma finite_atLeastZeroLessThan_int: "finite {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int) qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" proof - have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j by (cases j) (use assms in auto) show ?thesis proof (rule card_bij_eq) show "inj_on Suc {k. Suc k \ M \ k < i}" by force show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" by (rule inj_on_diff_nat) (use * in blast) qed (use * in auto) qed lemma card_less_Suc: assumes "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" by simp also have "\ = Suc (card {k \ M - {0}. k < Suc i})" apply (subst card_less_Suc2) using assms by auto also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" by (force intro: arg_cong [where f=card]) also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" by (simp add: card.insert_remove) also have "... = card {k \ M. k < Suc i}" using assms by (force simp add: intro: arg_cong [where f=card]) finally show ?thesis. qed lemma card_le_Suc_Max: "finite S \ card S \ Suc (Max S)" proof (rule classical) assume "finite S" and "\ Suc (Max S) \ card S" then have "Suc (Max S) < card S" by simp with \finite S\ have "S \ {0..Max S}" by auto hence "card S \ card {0..Max S}" by (intro card_mono; auto) thus "card S \ Suc (Max S)" by simp qed subsection \Lemmas useful with the summation operator sum\ text \For examples, see Algebra/poly/UnivPoly2.thy\ subsubsection \Disjoint Unions\ text \Singletons and open intervals\ lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. u ==> {.. u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l..m} Un {m..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch subsubsection \Disjoint Intersections\ text \One- and two-sided intervals\ lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ lemma ivl_disj_int_two: "{l::'a::order<..Some Differences\ lemma ivl_diff[simp]: "i \ n \ {i..Some Subset Conditions\ lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" using linorder_class.le_less_linear[of i n] - apply (auto simp: linorder_not_le) - apply (force intro: leI)+ - done + by safe (force intro: leI)+ subsection \Generic big monoid operation over intervals\ context semiring_char_0 begin lemma inj_on_of_nat [simp]: "inj_on of_nat N" - by rule simp + by (rule inj_onI) simp lemma bij_betw_of_nat [simp]: "bij_betw of_nat N A \ of_nat ` N = A" by (simp add: bij_betw_def) lemma Nats_infinite: "infinite (\ :: 'a set)" by (metis Nats_def finite_imageD infinite_UNIV_char_0 inj_on_of_nat) end context comm_monoid_set begin lemma atLeastLessThan_reindex: "F g {h m.. h) {m.. h) {m..n}" if "bij_betw h {m..n} {h m..h n}" for m n ::nat proof - from that have "inj_on h {m..n}" and "h ` {m..n} = {h m..h n}" by (simp_all add: bij_betw_def) then show ?thesis using reindex [of h "{m..n}" g] by simp qed lemma atLeastLessThan_shift_bounds: "F g {m + k.. plus k) {m.. plus k) {m..n}" for m n k :: nat using atLeastAtMost_reindex [of "plus k" m n g] by (simp add: ac_simps) lemma atLeast_Suc_lessThan_Suc_shift: "F g {Suc m.. Suc) {m.. Suc) {m..n}" using atLeastAtMost_shift_bounds [of _ _ 1] by (simp add: plus_1_eq_Suc) lemma atLeast_atMost_pred_shift: "F (g \ (\n. n - Suc 0)) {Suc m..Suc n} = F g {m..n}" unfolding atLeast_Suc_atMost_Suc_shift by simp lemma atLeast_lessThan_pred_shift: "F (g \ (\n. n - Suc 0)) {Suc m.. int) {m.. int) {m..n}" by (rule atLeastAtMost_reindex) (simp add: image_int_atLeastAtMost) lemma atLeast0_lessThan_Suc: "F g {0..* g n" by (simp add: atLeast0_lessThan_Suc ac_simps) lemma atLeast0_atMost_Suc: "F g {0..Suc n} = F g {0..n} \<^bold>* g (Suc n)" by (simp add: atLeast0_atMost_Suc ac_simps) lemma atLeast0_lessThan_Suc_shift: "F g {0..* F (g \ Suc) {0..* F (g \ Suc) {0..n}" by (simp add: atLeast0_atMost_Suc_eq_insert_0 atLeast_Suc_atMost_Suc_shift) lemma atLeast_Suc_lessThan: "F g {m..* F g {Suc m..* F g {Suc m..n}" if "m \ n" proof - from that have "{m..n} = insert m {Suc m..n}" by auto then show ?thesis by simp qed lemma ivl_cong: "a = c \ b = d \ (\x. c \ x \ x < d \ g x = h x) \ F g {a.. plus m) {0.. n") simp_all lemma atLeastAtMost_shift_0: fixes m n p :: nat assumes "m \ n" shows "F g {m..n} = F (g \ plus m) {0..n - m}" using assms atLeastAtMost_shift_bounds [of g 0 m "n - m"] by simp lemma atLeastLessThan_concat: fixes m n p :: nat shows "m \ n \ n \ p \ F g {m..* F g {n..i. g (m + n - Suc i)) {n..i. g (m + n - i)) {n..m}" by (rule reindex_bij_witness [where i="\i. m + n - i" and j="\i. m + n - i"]) auto lemma atLeastLessThan_rev_at_least_Suc_atMost: "F g {n..i. g (m + n - i)) {Suc n..m}" unfolding atLeastLessThan_rev [of g n m] by (cases m) (simp_all add: atLeast_Suc_atMost_Suc_shift atLeastLessThanSuc_atLeastAtMost) end subsection \Summation indexed over intervals\ syntax (ASCII) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST sum (\x. t) {a..b}" "\x=a..x. t) {a..i\n. t" == "CONST sum (\i. t) {..n}" "\ii. t) {..The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"\x\{a..b}. e"} & \<^term>\\x=a..b. e\ & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..\\x=a.. & @{term[mode=latex_sum]"\x=a..x\{..b}. e"} & \<^term>\\x\b. e\ & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..\\x & @{term[mode=latex_sum]"\xlatex_sum\ (e.g.\ via \mode = latex_sum\ in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on \<^typ>\nat\ it is better to use \<^term>\\x::nat=0.. rather than \\x: \sum\ may not provide all lemmas available for \<^term>\{m.. also in the special form for \<^term>\{...\ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well with the simplifier who adds the unsimplified premise \<^term>\x\B\ to the context.\ context comm_monoid_set begin lemma zero_middle: assumes "1 \ p" "k \ p" shows "F (\j. if j < k then g j else if j = k then \<^bold>1 else h (j - Suc 0)) {..p} = F (\j. if j < k then g j else h j) {..p - Suc 0}" (is "?lhs = ?rhs") proof - have [simp]: "{..p - Suc 0} \ {j. j < k} = {.. - {j. j < k} = {k..p - Suc 0}" using assms by auto have "?lhs = F g {..* F (\j. if j = k then \<^bold>1 else h (j - Suc 0)) {k..p}" using union_disjoint [of "{.. = F g {..* F (\j. h (j - Suc 0)) {Suc k..p}" by (simp add: atLeast_Suc_atMost [of k p] assms) also have "\ = F g {..* F h {k .. p - Suc 0}" using reindex [of Suc "{k..p - Suc 0}"] assms by simp also have "\ = ?rhs" by (simp add: If_cases) finally show ?thesis . qed lemma atMost_Suc [simp]: "F g {..Suc n} = F g {..n} \<^bold>* g (Suc n)" by (simp add: atMost_Suc ac_simps) lemma lessThan_Suc [simp]: "F g {..* g n" by (simp add: lessThan_Suc ac_simps) lemma cl_ivl_Suc [simp]: "F g {m..Suc n} = (if Suc n < m then \<^bold>1 else F g {m..n} \<^bold>* g(Suc n))" by (auto simp: ac_simps atLeastAtMostSuc_conv) lemma op_ivl_Suc [simp]: "F g {m..1 else F g {m..* g(n))" by (auto simp: ac_simps atLeastLessThanSuc) lemma head: fixes n :: nat assumes mn: "m \ n" shows "F g {m..n} = g m \<^bold>* F g {m<..n}" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = F g ({m} \ {m<..n})" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed lemma last_plus: fixes n::nat shows "m \ n \ F g {m..n} = g n \<^bold>* F g {m..1 else F g {m..* g(n))" by (simp add: commute last_plus) lemma ub_add_nat: assumes "(m::nat) \ n + 1" shows "F g {m..n + p} = F g {m..n} \<^bold>* F g {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto thus ?thesis by (auto simp: ivl_disj_int union_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma nat_group: fixes k::nat shows "F (\m. F g {m * k ..< m*k + k}) {.. 0" by auto then show ?thesis by (induct n) (simp_all add: atLeastLessThan_concat add.commute atLeast0LessThan[symmetric]) qed auto lemma triangle_reindex: fixes n :: nat shows "F (\(i,j). g i j) {(i,j). i+j < n} = F (\k. F (\i. g i (k - i)) {..k}) {..(i,j). g i j) {(i,j). i+j \ n} = F (\k. F (\i. g i (k - i)) {..k}) {..n}" using triangle_reindex [of g "Suc n"] by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lemma nat_diff_reindex: "F (\i. g (n - Suc i)) {..i. g(i + k)){m..i. g(i + k)){m..n::nat}" by (rule reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary shift_bounds_cl_Suc_ivl: "F g {Suc m..Suc n} = F (\i. g(Suc i)){m..n}" by (simp add: shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary Suc_reindex_ivl: "m \ n \ F g {m..n} \<^bold>* g (Suc n) = g m \<^bold>* F (\i. g (Suc i)) {m..n}" by (simp add: assoc atLeast_Suc_atMost flip: shift_bounds_cl_Suc_ivl) corollary shift_bounds_Suc_ivl: "F g {Suc m..i. g(Suc i)){m..* F (\i. g (Suc i)) {..n}" proof (induct n) case 0 show ?case by simp next case (Suc n) note IH = this have "F g {..Suc (Suc n)} = F g {..Suc n} \<^bold>* g (Suc (Suc n))" by (rule atMost_Suc) also have "F g {..Suc n} = g 0 \<^bold>* F (\i. g (Suc i)) {..n}" by (rule IH) also have "g 0 \<^bold>* F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = g 0 \<^bold>* (F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)))" by (rule assoc) also have "F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = F (\i. g (Suc i)) {..Suc n}" by (rule atMost_Suc [symmetric]) finally show ?case . qed lemma lessThan_Suc_shift: "F g {..* F (\i. g (Suc i)) {..* F (\i. g (Suc i)) {..i. F (\j. a i j) {0..j. F (\i. a i j) {Suc j..n}) {0..i. F (\j. a i j) {..j. F (\i. a i j) {Suc j..n}) {..k. g (Suc k)) {.. = F (\k. g (Suc k)) {.. b \ F g {a..* g b" by (simp add: atLeastLessThanSuc commute) lemma nat_ivl_Suc': assumes "m \ Suc n" shows "F g {m..Suc n} = g (Suc n) \<^bold>* F g {m..n}" proof - from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto also have "F g \ = g (Suc n) \<^bold>* F g {m..n}" by simp finally show ?thesis . qed lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" proof (induction n) case 0 show ?case by (cases "m=0") auto next case (Suc n) then show ?case by (auto simp: assoc split: if_split_asm) qed lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) end lemma card_sum_le_nat_sum: "\ {0.. \ S" proof (cases "finite S") case True then show ?thesis proof (induction "card S" arbitrary: S) case (Suc x) then have "Max S \ x" using card_le_Suc_Max by fastforce let ?S' = "S - {Max S}" from Suc have "Max S \ S" by (auto intro: Max_in) hence cards: "card S = Suc (card ?S')" using \finite S\ by (intro card.remove; auto) hence "\ {0.. \ ?S'" using Suc by (intro Suc; auto) hence "\ {0.. \ ?S' + Max S" using \Max S \ x\ by simp also have "... = \ S" using sum.remove[OF \finite S\ \Max S \ S\, where g="\x. x"] by simp finally show ?case using cards Suc by auto qed simp qed simp lemma sum_natinterval_diff: fixes f:: "nat \ ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m \ n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma sum_diff_nat_ivl: fixes f :: "nat \ 'a::ab_group_add" shows "\ m \ n; n \ p \ \ sum f {m..x. Q x \ P x \ (\xxxShifting bounds\ context comm_monoid_add begin context fixes f :: "nat \ 'a" assumes "f 0 = 0" begin lemma sum_shift_lb_Suc0_0_upt: "sum f {Suc 0..f 0 = 0\ by simp qed lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis by simp next case (Suc k) moreover have "{0..Suc k} = insert 0 {Suc 0..Suc k}" by auto ultimately show ?thesis using \f 0 = 0\ by simp qed end end lemma sum_Suc_diff: fixes f :: "nat \ 'a::ab_group_add" assumes "m \ Suc n" shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" using assms by (induct n) (auto simp: le_Suc_eq) lemma sum_Suc_diff': fixes f :: "nat \ 'a::ab_group_add" assumes "m \ n" shows "(\i = m..Telescoping\ lemma sum_telescope: fixes f::"nat \ 'a::ab_group_add" shows "sum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" by (induct i) simp_all lemma sum_telescope'': assumes "m \ n" shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) lemma sum_lessThan_telescope: "(\nnThe formula for geometric sums\ lemma sum_power2: "(\i=0.. 1" shows "(\i 0" by simp_all moreover have "(\iy \ 0\) ultimately show ?thesis by simp qed lemma diff_power_eq_sum: fixes y :: "'a::{comm_ring,monoid_mult}" shows "x ^ (Suc n) - y ^ (Suc n) = (x - y) * (\pppp \\COMPLEX_POLYFUN\ in HOL Light\ fixes x :: "'a::{comm_ring,monoid_mult}" shows "x^n - y^n = (x - y) * (\iiiii\n. x^i) = 1 - x^Suc n" by (simp only: one_diff_power_eq lessThan_Suc_atMost) lemma sum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(\i=m..n. x^i) = x^m * (\i\n-m. x^i)" proof - have "(\i=m..n. x^i) = x^m * (\i=m..n. x^(i-m))" by (simp add: sum_distrib_left power_add [symmetric]) also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" using \m \ n\ by (intro sum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto finally show ?thesis . qed lemma sum_gp_multiplied: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" by (metis mult.assoc mult.commute assms sum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis mult.assoc sum_gp_basic) also have "... = x^m - x^Suc n" using assms by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) finally show ?thesis . qed lemma sum_gp: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..n. x^i) = (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" -using sum_gp_multiplied [of m n x] apply auto -by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) +proof (cases "n < m") + case False + assume *: "\ n < m" + then show ?thesis + proof (cases "x = 1") + case False + assume "x \ 1" + then have not_zero: "1 - x \ 0" + by auto + have "(1 - x) * (\i=m..n. x^i) = x ^ m - x * x ^ n" + using sum_gp_multiplied [of m n x] * by auto + then have "(\i=m..n. x^i) = (x ^ m - x * x ^ n) / (1 - x) " + using nonzero_divide_eq_eq mult.commute not_zero + by metis + then show ?thesis + by auto + qed (auto) +qed (auto) subsubsection\Geometric progressions\ lemma sum_gp0: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using sum_gp_basic[of x n] by (simp add: mult.commute field_split_simps) lemma sum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" by (simp add: sum_distrib_left power_add) lemma sum_gp_offset: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..m+n. x^i) = (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" using sum_gp [of x m "m+n"] by (auto simp: power_add algebra_simps) lemma sum_gp_strict: fixes x :: "'a::{comm_ring,division_ring}" shows "(\iThe formulae for arithmetic sums\ context comm_semiring_1 begin lemma double_gauss_sum: "2 * (\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1)" by (induct n) (simp_all add: sum.atLeast0_atMost_Suc algebra_simps left_add_twice) lemma double_gauss_sum_from_Suc_0: "2 * (\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1)" proof - have "sum of_nat {Suc 0..n} = sum of_nat (insert 0 {Suc 0..n})" by simp also have "\ = sum of_nat {0..n}" by (cases n) (simp_all add: atLeast0_atMost_Suc_eq_insert_0) finally show ?thesis by (simp add: double_gauss_sum) qed lemma double_arith_series: "2 * (\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d)" proof - have "(\i = 0..n. a + of_nat i * d) = ((\i = 0..n. a) + (\i = 0..n. of_nat i * d))" by (rule sum.distrib) also have "\ = (of_nat (Suc n) * a + d * (\i = 0..n. of_nat i))" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis by (simp add: algebra_simps double_gauss_sum left_add_twice) qed end context unique_euclidean_semiring_with_nat begin lemma gauss_sum: "(\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum [of n, symmetric] by simp lemma gauss_sum_from_Suc_0: "(\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum_from_Suc_0 [of n, symmetric] by simp lemma arith_series: "(\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d) div 2" using double_arith_series [of a d n, symmetric] by simp end lemma gauss_sum_nat: "\{0..n} = (n * Suc n) div 2" using gauss_sum [of n, where ?'a = nat] by simp lemma arith_series_nat: "(\i = 0..n. a + i * d) = Suc n * (2 * a + n * d) div 2" using arith_series [of a d n] by simp lemma Sum_Icc_int: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" if "m \ n" for m n :: int using that proof (induct i \ "nat (n - m)" arbitrary: m n) case 0 then have "m = n" by arith then show ?case by (simp add: algebra_simps mult_2 [symmetric]) next case (Suc i) have 0: "i = nat((n-1) - m)" "m \ n-1" using Suc(2,3) by arith+ have "\ {m..n} = \ {m..1+(n-1)}" by simp also have "\ = \ {m..n-1} + n" using \m \ n\ by(subst atLeastAtMostPlus1_int_conv) simp_all also have "\ = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n" by(simp add: Suc(1)[OF 0]) also have "\ = ((n-1)*(n-1+1) - m*(m-1) + 2*n) div 2" by simp also have "\ = (n*(n+1) - m*(m-1)) div 2" by (simp add: algebra_simps mult_2_right) finally show ?case . qed lemma Sum_Icc_nat: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" for m n :: nat proof (cases "m \ n") case True then have *: "m * (m - 1) \ n * (n + 1)" by (meson diff_le_self order_trans le_add1 mult_le_mono) have "int (\{m..n}) = (\{int m..int n})" by (simp add: sum.atLeast_int_atMost_int_shift) also have "\ = (int n * (int n + 1) - int m * (int m - 1)) div 2" using \m \ n\ by (simp add: Sum_Icc_int) also have "\ = int ((n * (n + 1) - m * (m - 1)) div 2)" using le_square * by (simp add: algebra_simps of_nat_div of_nat_diff) finally show ?thesis by (simp only: of_nat_eq_iff) next case False then show ?thesis by (auto dest: less_imp_Suc_add simp add: not_le algebra_simps) qed lemma Sum_Ico_nat: "\{m..Division remainder\ lemma range_mod: fixes n :: nat assumes "n > 0" shows "range (\m. m mod n) = {0.. ?A \ m \ ?B" proof assume "m \ ?A" with assms show "m \ ?B" by auto next assume "m \ ?B" moreover have "m mod n \ ?A" by (rule rangeI) ultimately show "m \ ?A" by simp qed qed subsection \Products indexed over intervals\ syntax (ASCII) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" \ "CONST prod (\x. t) {a..b}" "\x=a.. "CONST prod (\x. t) {a..i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lemma prod_int_eq: "prod int {i..j} = \{int i..int j}" proof (cases "i \ j") case True then show ?thesis by (metis le_iff_add prod_int_plus_eq) next case False then show ?thesis by auto qed subsection \Efficient folding over intervals\ function fold_atLeastAtMost_nat where [simp del]: "fold_atLeastAtMost_nat f a (b::nat) acc = (if a > b then acc else fold_atLeastAtMost_nat f (a+1) b (f a acc))" by pat_completeness auto termination by (relation "measure (\(_,a,b,_). Suc b - a)") auto lemma fold_atLeastAtMost_nat: assumes "comp_fun_commute f" shows "fold_atLeastAtMost_nat f a b acc = Finite_Set.fold f acc {a..b}" using assms proof (induction f a b acc rule: fold_atLeastAtMost_nat.induct, goal_cases) case (1 f a b acc) interpret comp_fun_commute f by fact show ?case proof (cases "a > b") case True thus ?thesis by (subst fold_atLeastAtMost_nat.simps) auto next case False with 1 show ?thesis by (subst fold_atLeastAtMost_nat.simps) (auto simp: atLeastAtMost_insertL[symmetric] fold_fun_left_comm) qed qed lemma sum_atLeastAtMost_code: "sum f {a..b} = fold_atLeastAtMost_nat (\a acc. f a + acc) a b 0" proof - have "comp_fun_commute (\a. (+) (f a))" by unfold_locales (auto simp: o_def add_ac) thus ?thesis by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def) qed lemma prod_atLeastAtMost_code: "prod f {a..b} = fold_atLeastAtMost_nat (\a acc. f a * acc) a b 1" proof - have "comp_fun_commute (\a. (*) (f a))" by unfold_locales (auto simp: o_def mult_ac) thus ?thesis by (simp add: prod.eq_fold fold_atLeastAtMost_nat o_def) qed (* TODO: Add support for folding over more kinds of intervals here *) end diff --git a/src/HOL/Transfer.thy b/src/HOL/Transfer.thy --- a/src/HOL/Transfer.thy +++ b/src/HOL/Transfer.thy @@ -1,692 +1,697 @@ (* Title: HOL/Transfer.thy Author: Brian Huffman, TU Muenchen Author: Ondrej Kuncar, TU Muenchen *) section \Generic theorem transfer using relations\ theory Transfer imports Basic_BNF_LFPs Hilbert_Choice Metis begin subsection \Relator for function space\ bundle lifting_syntax begin notation rel_fun (infixr "===>" 55) notation map_fun (infixr "--->" 55) end context includes lifting_syntax begin lemma rel_funD2: assumes "rel_fun A B f g" and "A x x" shows "B (f x) (g x)" using assms by (rule rel_funD) lemma rel_funE: assumes "rel_fun A B f g" and "A x y" obtains "B (f x) (g y)" using assms by (simp add: rel_fun_def) lemmas rel_fun_eq = fun.rel_eq lemma rel_fun_eq_rel: shows "rel_fun (=) R = (\f g. \x. R (f x) (g x))" by (simp add: rel_fun_def) subsection \Transfer method\ text \Explicit tag for relation membership allows for backward proof methods.\ definition Rel :: "('a \ 'b \ bool) \ 'a \ 'b \ bool" where "Rel r \ r" text \Handling of equality relations\ definition is_equality :: "('a \ 'a \ bool) \ bool" where "is_equality R \ R = (=)" lemma is_equality_eq: "is_equality (=)" unfolding is_equality_def by simp text \Reverse implication for monotonicity rules\ definition rev_implies where "rev_implies x y \ (y \ x)" text \Handling of meta-logic connectives\ definition transfer_forall where "transfer_forall \ All" definition transfer_implies where "transfer_implies \ (\)" definition transfer_bforall :: "('a \ bool) \ ('a \ bool) \ bool" where "transfer_bforall \ (\P Q. \x. P x \ Q x)" lemma transfer_forall_eq: "(\x. P x) \ Trueprop (transfer_forall (\x. P x))" unfolding atomize_all transfer_forall_def .. lemma transfer_implies_eq: "(A \ B) \ Trueprop (transfer_implies A B)" unfolding atomize_imp transfer_implies_def .. lemma transfer_bforall_unfold: "Trueprop (transfer_bforall P (\x. Q x)) \ (\x. P x \ Q x)" unfolding transfer_bforall_def atomize_imp atomize_all .. lemma transfer_start: "\P; Rel (=) P Q\ \ Q" unfolding Rel_def by simp lemma transfer_start': "\P; Rel (\) P Q\ \ Q" unfolding Rel_def by simp lemma transfer_prover_start: "\x = x'; Rel R x' y\ \ Rel R x y" by simp lemma untransfer_start: "\Q; Rel (=) P Q\ \ P" unfolding Rel_def by simp lemma Rel_eq_refl: "Rel (=) x x" unfolding Rel_def .. lemma Rel_app: assumes "Rel (A ===> B) f g" and "Rel A x y" shows "Rel B (f x) (g y)" using assms unfolding Rel_def rel_fun_def by fast lemma Rel_abs: assumes "\x y. Rel A x y \ Rel B (f x) (g y)" shows "Rel (A ===> B) (\x. f x) (\y. g y)" using assms unfolding Rel_def rel_fun_def by fast subsection \Predicates on relations, i.e. ``class constraints''\ definition left_total :: "('a \ 'b \ bool) \ bool" where "left_total R \ (\x. \y. R x y)" definition left_unique :: "('a \ 'b \ bool) \ bool" where "left_unique R \ (\x y z. R x z \ R y z \ x = y)" definition right_total :: "('a \ 'b \ bool) \ bool" where "right_total R \ (\y. \x. R x y)" definition right_unique :: "('a \ 'b \ bool) \ bool" where "right_unique R \ (\x y z. R x y \ R x z \ y = z)" definition bi_total :: "('a \ 'b \ bool) \ bool" where "bi_total R \ (\x. \y. R x y) \ (\y. \x. R x y)" definition bi_unique :: "('a \ 'b \ bool) \ bool" where "bi_unique R \ (\x y z. R x y \ R x z \ y = z) \ (\x y z. R x z \ R y z \ x = y)" lemma left_unique_iff: "left_unique R \ (\z. \\<^sub>\\<^sub>1x. R x z)" unfolding Uniq_def left_unique_def by force lemma left_uniqueI: "(\x y z. \ A x z; A y z \ \ x = y) \ left_unique A" unfolding left_unique_def by blast lemma left_uniqueD: "\ left_unique A; A x z; A y z \ \ x = y" unfolding left_unique_def by blast lemma left_totalI: "(\x. \y. R x y) \ left_total R" unfolding left_total_def by blast lemma left_totalE: assumes "left_total R" obtains "(\x. \y. R x y)" using assms unfolding left_total_def by blast lemma bi_uniqueDr: "\ bi_unique A; A x y; A x z \ \ y = z" by(simp add: bi_unique_def) lemma bi_uniqueDl: "\ bi_unique A; A x y; A z y \ \ x = z" by(simp add: bi_unique_def) lemma bi_unique_iff: "bi_unique R \ (\z. \\<^sub>\\<^sub>1x. R x z) \ (\z. \\<^sub>\\<^sub>1x. R z x)" unfolding Uniq_def bi_unique_def by force lemma right_unique_iff: "right_unique R \ (\z. \\<^sub>\\<^sub>1x. R z x)" unfolding Uniq_def right_unique_def by force lemma right_uniqueI: "(\x y z. \ A x y; A x z \ \ y = z) \ right_unique A" unfolding right_unique_def by fast lemma right_uniqueD: "\ right_unique A; A x y; A x z \ \ y = z" unfolding right_unique_def by fast lemma right_totalI: "(\y. \x. A x y) \ right_total A" by(simp add: right_total_def) lemma right_totalE: assumes "right_total A" obtains x where "A x y" using assms by(auto simp add: right_total_def) lemma right_total_alt_def2: "right_total R \ ((R ===> (\)) ===> (\)) All All" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding right_total_def rel_fun_def by blast next assume \
: ?rhs show ?lhs using \
[unfolded rel_fun_def, rule_format, of "\x. True" "\y. \x. R x y"] unfolding right_total_def by blast qed lemma right_unique_alt_def2: "right_unique R \ (R ===> R ===> (\)) (=) (=)" unfolding right_unique_def rel_fun_def by auto lemma bi_total_alt_def2: "bi_total R \ ((R ===> (=)) ===> (=)) All All" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding bi_total_def rel_fun_def by blast next assume \
: ?rhs show ?lhs using \
[unfolded rel_fun_def, rule_format, of "\x. \y. R x y" "\y. True"] using \
[unfolded rel_fun_def, rule_format, of "\x. True" "\y. \x. R x y"] by (auto simp: bi_total_def) qed lemma bi_unique_alt_def2: "bi_unique R \ (R ===> R ===> (=)) (=) (=)" unfolding bi_unique_def rel_fun_def by auto lemma [simp]: shows left_unique_conversep: "left_unique A\\ \ right_unique A" and right_unique_conversep: "right_unique A\\ \ left_unique A" by(auto simp add: left_unique_def right_unique_def) lemma [simp]: shows left_total_conversep: "left_total A\\ \ right_total A" and right_total_conversep: "right_total A\\ \ left_total A" by(simp_all add: left_total_def right_total_def) lemma bi_unique_conversep [simp]: "bi_unique R\\ = bi_unique R" by(auto simp add: bi_unique_def) lemma bi_total_conversep [simp]: "bi_total R\\ = bi_total R" by(auto simp add: bi_total_def) lemma right_unique_alt_def: "right_unique R = (conversep R OO R \ (=))" unfolding right_unique_def by blast lemma left_unique_alt_def: "left_unique R = (R OO (conversep R) \ (=))" unfolding left_unique_def by blast lemma right_total_alt_def: "right_total R = (conversep R OO R \ (=))" unfolding right_total_def by blast lemma left_total_alt_def: "left_total R = (R OO conversep R \ (=))" unfolding left_total_def by blast lemma bi_total_alt_def: "bi_total A = (left_total A \ right_total A)" unfolding left_total_def right_total_def bi_total_def by blast lemma bi_unique_alt_def: "bi_unique A = (left_unique A \ right_unique A)" unfolding left_unique_def right_unique_def bi_unique_def by blast lemma bi_totalI: "left_total R \ right_total R \ bi_total R" unfolding bi_total_alt_def .. lemma bi_uniqueI: "left_unique R \ right_unique R \ bi_unique R" unfolding bi_unique_alt_def .. end lemma is_equality_lemma: "(\R. is_equality R \ PROP (P R)) \ PROP (P (=))" unfolding is_equality_def proof (rule equal_intr_rule) show "(\R. R = (=) \ PROP P R) \ PROP P (=)" apply (drule meta_spec) apply (erule meta_mp [OF _ refl]) done qed simp lemma Domainp_lemma: "(\R. Domainp T = R \ PROP (P R)) \ PROP (P (Domainp T))" proof (rule equal_intr_rule) show "(\R. Domainp T = R \ PROP P R) \ PROP P (Domainp T)" apply (drule meta_spec) apply (erule meta_mp [OF _ refl]) done qed simp ML_file \Tools/Transfer/transfer.ML\ declare refl [transfer_rule] hide_const (open) Rel context includes lifting_syntax begin text \Handling of domains\ lemma Domainp_iff: "Domainp T x \ (\y. T x y)" by auto lemma Domainp_refl[transfer_domain_rule]: "Domainp T = Domainp T" .. lemma Domain_eq_top[transfer_domain_rule]: "Domainp (=) = top" by auto lemma Domainp_pred_fun_eq[relator_domain]: assumes "left_unique T" shows "Domainp (T ===> S) = pred_fun (Domainp T) (Domainp S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix x assume "?lhs x" then show "?rhs x" using assms unfolding rel_fun_def pred_fun_def by blast next fix x assume "?rhs x" then have "\g. \y xa. T xa y \ S (x xa) (g y)" using assms unfolding Domainp_iff left_unique_def pred_fun_def by (intro choice) blast then show "?lhs x" by blast qed text \Properties are preserved by relation composition.\ lemma OO_def: "R OO S = (\x z. \y. R x y \ S y z)" by auto lemma bi_total_OO: "\bi_total A; bi_total B\ \ bi_total (A OO B)" unfolding bi_total_def OO_def by fast lemma bi_unique_OO: "\bi_unique A; bi_unique B\ \ bi_unique (A OO B)" unfolding bi_unique_def OO_def by blast lemma right_total_OO: "\right_total A; right_total B\ \ right_total (A OO B)" unfolding right_total_def OO_def by fast lemma right_unique_OO: "\right_unique A; right_unique B\ \ right_unique (A OO B)" unfolding right_unique_def OO_def by fast lemma left_total_OO: "left_total R \ left_total S \ left_total (R OO S)" unfolding left_total_def OO_def by fast lemma left_unique_OO: "left_unique R \ left_unique S \ left_unique (R OO S)" unfolding left_unique_def OO_def by blast subsection \Properties of relators\ lemma left_total_eq[transfer_rule]: "left_total (=)" unfolding left_total_def by blast lemma left_unique_eq[transfer_rule]: "left_unique (=)" unfolding left_unique_def by blast lemma right_total_eq [transfer_rule]: "right_total (=)" unfolding right_total_def by simp lemma right_unique_eq [transfer_rule]: "right_unique (=)" unfolding right_unique_def by simp lemma bi_total_eq[transfer_rule]: "bi_total (=)" unfolding bi_total_def by simp lemma bi_unique_eq[transfer_rule]: "bi_unique (=)" unfolding bi_unique_def by simp lemma left_total_fun[transfer_rule]: assumes "left_unique A" "left_total B" shows "left_total (A ===> B)" unfolding left_total_def proof fix f show "Ex ((A ===> B) f)" unfolding rel_fun_def proof (intro exI strip) fix x y assume A: "A x y" have "(THE x. A x y) = x" using A assms by (simp add: left_unique_def the_equality) then show "B (f x) (SOME z. B (f (THE x. A x y)) z)" using assms by (force simp: left_total_def intro: someI_ex) qed qed lemma left_unique_fun[transfer_rule]: "\left_total A; left_unique B\ \ left_unique (A ===> B)" unfolding left_total_def left_unique_def rel_fun_def by (clarify, rule ext, fast) lemma right_total_fun [transfer_rule]: assumes "right_unique A" "right_total B" shows "right_total (A ===> B)" unfolding right_total_def proof fix g show "\x. (A ===> B) x g" unfolding rel_fun_def proof (intro exI strip) fix x y assume A: "A x y" have "(THE y. A x y) = y" using A assms by (simp add: right_unique_def the_equality) then show "B (SOME z. B z (g (THE y. A x y))) (g y)" using assms by (force simp: right_total_def intro: someI_ex) qed qed lemma right_unique_fun [transfer_rule]: "\right_total A; right_unique B\ \ right_unique (A ===> B)" unfolding right_total_def right_unique_def rel_fun_def by (clarify, rule ext, fast) lemma bi_total_fun[transfer_rule]: "\bi_unique A; bi_total B\ \ bi_total (A ===> B)" unfolding bi_unique_alt_def bi_total_alt_def by (blast intro: right_total_fun left_total_fun) lemma bi_unique_fun[transfer_rule]: "\bi_total A; bi_unique B\ \ bi_unique (A ===> B)" unfolding bi_unique_alt_def bi_total_alt_def by (blast intro: right_unique_fun left_unique_fun) end lemma if_conn: "(if P \ Q then t else e) = (if P then if Q then t else e else e)" "(if P \ Q then t else e) = (if P then t else if Q then t else e)" "(if P \ Q then t else e) = (if P then if Q then t else e else t)" "(if \ P then t else e) = (if P then e else t)" by auto ML_file \Tools/Transfer/transfer_bnf.ML\ ML_file \Tools/BNF/bnf_fp_rec_sugar_transfer.ML\ declare pred_fun_def [simp] declare rel_fun_eq [relator_eq] (* Delete the automated generated rule from the bnf command; we have a more general rule (Domainp_pred_fun_eq) that subsumes it. *) declare fun.Domainp_rel[relator_domain del] subsection \Transfer rules\ context includes lifting_syntax begin lemma Domainp_forall_transfer [transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> (=)) (transfer_bforall (Domainp A)) transfer_forall" using assms unfolding right_total_def unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff by fast text \Transfer rules using implication instead of equality on booleans.\ lemma transfer_forall_transfer [transfer_rule]: "bi_total A \ ((A ===> (=)) ===> (=)) transfer_forall transfer_forall" "right_total A \ ((A ===> (=)) ===> implies) transfer_forall transfer_forall" "right_total A \ ((A ===> implies) ===> implies) transfer_forall transfer_forall" "bi_total A \ ((A ===> (=)) ===> rev_implies) transfer_forall transfer_forall" "bi_total A \ ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall" unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def by fast+ lemma transfer_implies_transfer [transfer_rule]: "((=) ===> (=) ===> (=) ) transfer_implies transfer_implies" "(rev_implies ===> implies ===> implies ) transfer_implies transfer_implies" "(rev_implies ===> (=) ===> implies ) transfer_implies transfer_implies" "((=) ===> implies ===> implies ) transfer_implies transfer_implies" "((=) ===> (=) ===> implies ) transfer_implies transfer_implies" "(implies ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" "(implies ===> (=) ===> rev_implies) transfer_implies transfer_implies" "((=) ===> rev_implies ===> rev_implies) transfer_implies transfer_implies" "((=) ===> (=) ===> rev_implies) transfer_implies transfer_implies" unfolding transfer_implies_def rev_implies_def rel_fun_def by auto lemma eq_imp_transfer [transfer_rule]: "right_unique A \ (A ===> A ===> (\)) (=) (=)" unfolding right_unique_alt_def2 . text \Transfer rules using equality.\ lemma left_unique_transfer [transfer_rule]: assumes "right_total A" assumes "right_total B" assumes "bi_unique A" shows "((A ===> B ===> (=)) ===> implies) left_unique left_unique" using assms unfolding left_unique_def right_total_def bi_unique_def rel_fun_def by metis lemma eq_transfer [transfer_rule]: assumes "bi_unique A" shows "(A ===> A ===> (=)) (=) (=)" using assms unfolding bi_unique_def rel_fun_def by auto lemma right_total_Ex_transfer[transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> (=)) (Bex (Collect (Domainp A))) Ex" using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff by fast lemma right_total_All_transfer[transfer_rule]: assumes "right_total A" shows "((A ===> (=)) ===> (=)) (Ball (Collect (Domainp A))) All" using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff by fast context includes lifting_syntax begin lemma right_total_fun_eq_transfer: assumes [transfer_rule]: "right_total A" "bi_unique B" shows "((A ===> B) ===> (A ===> B) ===> (=)) (\f g. \x\Collect(Domainp A). f x = g x) (=)" unfolding fun_eq_iff by transfer_prover end lemma All_transfer [transfer_rule]: assumes "bi_total A" shows "((A ===> (=)) ===> (=)) All All" using assms unfolding bi_total_def rel_fun_def by fast lemma Ex_transfer [transfer_rule]: assumes "bi_total A" shows "((A ===> (=)) ===> (=)) Ex Ex" using assms unfolding bi_total_def rel_fun_def by fast lemma Ex1_parametric [transfer_rule]: assumes [transfer_rule]: "bi_unique A" "bi_total A" shows "((A ===> (=)) ===> (=)) Ex1 Ex1" unfolding Ex1_def by transfer_prover declare If_transfer [transfer_rule] lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let" unfolding rel_fun_def by simp declare id_transfer [transfer_rule] declare comp_transfer [transfer_rule] lemma curry_transfer [transfer_rule]: "((rel_prod A B ===> C) ===> A ===> B ===> C) curry curry" unfolding curry_def by transfer_prover lemma fun_upd_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd" unfolding fun_upd_def by transfer_prover lemma case_nat_transfer [transfer_rule]: "(A ===> ((=) ===> A) ===> (=) ===> A) case_nat case_nat" unfolding rel_fun_def by (simp split: nat.split) lemma rec_nat_transfer [transfer_rule]: "(A ===> ((=) ===> A ===> A) ===> (=) ===> A) rec_nat rec_nat" - unfolding rel_fun_def by (clarsimp, rename_tac n, induct_tac n, simp_all) + unfolding rel_fun_def + apply safe + subgoal for _ _ _ _ _ n + by (induction n) simp_all + done + lemma funpow_transfer [transfer_rule]: "((=) ===> (A ===> A) ===> (A ===> A)) compow compow" unfolding funpow_def by transfer_prover lemma mono_transfer[transfer_rule]: assumes [transfer_rule]: "bi_total A" assumes [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" assumes [transfer_rule]: "(B ===> B ===> (=)) (\) (\)" shows "((A ===> B) ===> (=)) mono mono" unfolding mono_def by transfer_prover lemma right_total_relcompp_transfer[transfer_rule]: assumes [transfer_rule]: "right_total B" shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (\R S x z. \y\Collect (Domainp B). R x y \ S y z) (OO)" unfolding OO_def by transfer_prover lemma relcompp_transfer[transfer_rule]: assumes [transfer_rule]: "bi_total B" shows "((A ===> B ===> (=)) ===> (B ===> C ===> (=)) ===> A ===> C ===> (=)) (OO) (OO)" unfolding OO_def by transfer_prover lemma right_total_Domainp_transfer[transfer_rule]: assumes [transfer_rule]: "right_total B" shows "((A ===> B ===> (=)) ===> A ===> (=)) (\T x. \y\Collect(Domainp B). T x y) Domainp" apply(subst(2) Domainp_iff[abs_def]) by transfer_prover lemma Domainp_transfer[transfer_rule]: assumes [transfer_rule]: "bi_total B" shows "((A ===> B ===> (=)) ===> A ===> (=)) Domainp Domainp" unfolding Domainp_iff by transfer_prover lemma reflp_transfer[transfer_rule]: "bi_total A \ ((A ===> A ===> (=)) ===> (=)) reflp reflp" "right_total A \ ((A ===> A ===> implies) ===> implies) reflp reflp" "right_total A \ ((A ===> A ===> (=)) ===> implies) reflp reflp" "bi_total A \ ((A ===> A ===> rev_implies) ===> rev_implies) reflp reflp" "bi_total A \ ((A ===> A ===> (=)) ===> rev_implies) reflp reflp" unfolding reflp_def rev_implies_def bi_total_def right_total_def rel_fun_def by fast+ lemma right_unique_transfer [transfer_rule]: "\ right_total A; right_total B; bi_unique B \ \ ((A ===> B ===> (=)) ===> implies) right_unique right_unique" unfolding right_unique_def right_total_def bi_unique_def rel_fun_def by metis lemma left_total_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) left_total left_total" unfolding left_total_def by transfer_prover lemma right_total_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) right_total right_total" unfolding right_total_def by transfer_prover lemma left_unique_parametric [transfer_rule]: assumes [transfer_rule]: "bi_unique A" "bi_total A" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) left_unique left_unique" unfolding left_unique_def by transfer_prover lemma prod_pred_parametric [transfer_rule]: "((A ===> (=)) ===> (B ===> (=)) ===> rel_prod A B ===> (=)) pred_prod pred_prod" unfolding prod.pred_set Basic_BNFs.fsts_def Basic_BNFs.snds_def fstsp.simps sndsp.simps by simp transfer_prover lemma apfst_parametric [transfer_rule]: "((A ===> B) ===> rel_prod A C ===> rel_prod B C) apfst apfst" unfolding apfst_def by transfer_prover lemma rel_fun_eq_eq_onp: "((=) ===> eq_onp P) = eq_onp (\f. \x. P(f x))" unfolding eq_onp_def rel_fun_def by auto lemma rel_fun_eq_onp_rel: shows "((eq_onp R) ===> S) = (\f g. \x. R x \ S (f x) (g x))" by (auto simp add: eq_onp_def rel_fun_def) lemma eq_onp_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "((A ===> (=)) ===> A ===> A ===> (=)) eq_onp eq_onp" unfolding eq_onp_def by transfer_prover lemma rtranclp_parametric [transfer_rule]: assumes "bi_unique A" "bi_total A" shows "((A ===> A ===> (=)) ===> A ===> A ===> (=)) rtranclp rtranclp" proof(rule rel_funI iffI)+ fix R :: "'a \ 'a \ bool" and R' x y x' y' assume R: "(A ===> A ===> (=)) R R'" and "A x x'" { assume "R\<^sup>*\<^sup>* x y" "A y y'" thus "R'\<^sup>*\<^sup>* x' y'" proof(induction arbitrary: y') case base with \bi_unique A\ \A x x'\ have "x' = y'" by(rule bi_uniqueDr) thus ?case by simp next case (step y z z') from \bi_total A\ obtain y' where "A y y'" unfolding bi_total_def by blast hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH) moreover from R \A y y'\ \A z z'\ \R y z\ have "R' y' z'" by(auto dest: rel_funD) ultimately show ?case .. qed next assume "R'\<^sup>*\<^sup>* x' y'" "A y y'" thus "R\<^sup>*\<^sup>* x y" proof(induction arbitrary: y) case base with \bi_unique A\ \A x x'\ have "x = y" by(rule bi_uniqueDl) thus ?case by simp next case (step y' z' z) from \bi_total A\ obtain y where "A y y'" unfolding bi_total_def by blast hence "R\<^sup>*\<^sup>* x y" by(rule step.IH) moreover from R \A y y'\ \A z z'\ \R' y' z'\ have "R y z" by(auto dest: rel_funD) ultimately show ?case .. qed } qed lemma right_unique_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_unique B" "bi_total B" shows "((A ===> B ===> (=)) ===> (=)) right_unique right_unique" unfolding right_unique_def by transfer_prover lemma map_fun_parametric [transfer_rule]: "((A ===> B) ===> (C ===> D) ===> (B ===> C) ===> A ===> D) map_fun map_fun" unfolding map_fun_def by transfer_prover end subsection \\<^const>\of_bool\ and \<^const>\of_nat\\ context includes lifting_syntax begin lemma transfer_rule_of_bool: \((\) ===> (\)) of_bool of_bool\ if [transfer_rule]: \0 \ 0\ \1 \ 1\ for R :: \'a::zero_neq_one \ 'b::zero_neq_one \ bool\ (infix \\\ 50) unfolding of_bool_def by transfer_prover lemma transfer_rule_of_nat: "((=) ===> (\)) of_nat of_nat" if [transfer_rule]: \0 \ 0\ \1 \ 1\ \((\) ===> (\) ===> (\)) (+) (+)\ for R :: \'a::semiring_1 \ 'b::semiring_1 \ bool\ (infix \\\ 50) unfolding of_nat_def by transfer_prover end end diff --git a/src/HOL/Transitive_Closure.thy b/src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy +++ b/src/HOL/Transitive_Closure.thy @@ -1,1362 +1,1363 @@ (* Title: HOL/Transitive_Closure.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section \Reflexive and Transitive closure of a relation\ theory Transitive_Closure imports Relation abbrevs "^*" = "\<^sup>*" "\<^sup>*\<^sup>*" and "^+" = "\<^sup>+" "\<^sup>+\<^sup>+" and "^=" = "\<^sup>=" "\<^sup>=\<^sup>=" begin ML_file \~~/src/Provers/trancl.ML\ text \ \rtrancl\ is reflexive/transitive closure, \trancl\ is transitive closure, \reflcl\ is reflexive closure. These postfix operators have \<^emph>\maximum priority\, forcing their operands to be atomic. \ context notes [[inductive_internals]] begin inductive_set rtrancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>*)" [1000] 999) for r :: "('a \ 'a) set" where rtrancl_refl [intro!, Pure.intro!, simp]: "(a, a) \ r\<^sup>*" | rtrancl_into_rtrancl [Pure.intro]: "(a, b) \ r\<^sup>* \ (b, c) \ r \ (a, c) \ r\<^sup>*" inductive_set trancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>+)" [1000] 999) for r :: "('a \ 'a) set" where r_into_trancl [intro, Pure.intro]: "(a, b) \ r \ (a, b) \ r\<^sup>+" | trancl_into_trancl [Pure.intro]: "(a, b) \ r\<^sup>+ \ (b, c) \ r \ (a, c) \ r\<^sup>+" notation rtranclp ("(_\<^sup>*\<^sup>*)" [1000] 1000) and tranclp ("(_\<^sup>+\<^sup>+)" [1000] 1000) declare rtrancl_def [nitpick_unfold del] rtranclp_def [nitpick_unfold del] trancl_def [nitpick_unfold del] tranclp_def [nitpick_unfold del] end abbreviation reflcl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>=)" [1000] 999) where "r\<^sup>= \ r \ Id" abbreviation reflclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" ("(_\<^sup>=\<^sup>=)" [1000] 1000) where "r\<^sup>=\<^sup>= \ sup r (=)" notation (ASCII) rtrancl ("(_^*)" [1000] 999) and trancl ("(_^+)" [1000] 999) and reflcl ("(_^=)" [1000] 999) and rtranclp ("(_^**)" [1000] 1000) and tranclp ("(_^++)" [1000] 1000) and reflclp ("(_^==)" [1000] 1000) subsection \Reflexive closure\ lemma refl_reflcl[simp]: "refl (r\<^sup>=)" by (simp add: refl_on_def) lemma antisym_reflcl[simp]: "antisym (r\<^sup>=) = antisym r" by (simp add: antisym_def) lemma trans_reflclI[simp]: "trans r \ trans (r\<^sup>=)" unfolding trans_def by blast lemma reflclp_idemp [simp]: "(P\<^sup>=\<^sup>=)\<^sup>=\<^sup>= = P\<^sup>=\<^sup>=" by blast subsection \Reflexive-transitive closure\ lemma reflcl_set_eq [pred_set_conv]: "(sup (\x y. (x, y) \ r) (=)) = (\x y. (x, y) \ r \ Id)" by (auto simp: fun_eq_iff) lemma r_into_rtrancl [intro]: "\p. p \ r \ p \ r\<^sup>*" \ \\rtrancl\ of \r\ contains \r\\ by (simp add: split_tupled_all rtrancl_refl [THEN rtrancl_into_rtrancl]) lemma r_into_rtranclp [intro]: "r x y \ r\<^sup>*\<^sup>* x y" \ \\rtrancl\ of \r\ contains \r\\ by (erule rtranclp.rtrancl_refl [THEN rtranclp.rtrancl_into_rtrancl]) lemma rtranclp_mono: "r \ s \ r\<^sup>*\<^sup>* \ s\<^sup>*\<^sup>*" \ \monotonicity of \rtrancl\\ proof (rule predicate2I) show "s\<^sup>*\<^sup>* x y" if "r \ s" "r\<^sup>*\<^sup>* x y" for x y using \r\<^sup>*\<^sup>* x y\ \r \ s\ by (induction rule: rtranclp.induct) (blast intro: rtranclp.rtrancl_into_rtrancl)+ qed lemma mono_rtranclp[mono]: "(\a b. x a b \ y a b) \ x\<^sup>*\<^sup>* a b \ y\<^sup>*\<^sup>* a b" using rtranclp_mono[of x y] by auto lemmas rtrancl_mono = rtranclp_mono [to_set] theorem rtranclp_induct [consumes 1, case_names base step, induct set: rtranclp]: assumes a: "r\<^sup>*\<^sup>* a b" and cases: "P a" "\y z. r\<^sup>*\<^sup>* a y \ r y z \ P y \ P z" shows "P b" using a by (induct x\a b) (rule cases)+ lemmas rtrancl_induct [induct set: rtrancl] = rtranclp_induct [to_set] lemmas rtranclp_induct2 = rtranclp_induct[of _ "(ax,ay)" "(bx,by)", split_rule, consumes 1, case_names refl step] lemmas rtrancl_induct2 = rtrancl_induct[of "(ax,ay)" "(bx,by)", split_format (complete), consumes 1, case_names refl step] lemma refl_rtrancl: "refl (r\<^sup>*)" unfolding refl_on_def by fast text \Transitivity of transitive closure.\ lemma trans_rtrancl: "trans (r\<^sup>*)" proof (rule transI) fix x y z assume "(x, y) \ r\<^sup>*" assume "(y, z) \ r\<^sup>*" then show "(x, z) \ r\<^sup>*" proof induct case base show "(x, y) \ r\<^sup>*" by fact next case (step u v) from \(x, u) \ r\<^sup>*\ and \(u, v) \ r\ show "(x, v) \ r\<^sup>*" .. qed qed lemmas rtrancl_trans = trans_rtrancl [THEN transD] lemma rtranclp_trans: assumes "r\<^sup>*\<^sup>* x y" and "r\<^sup>*\<^sup>* y z" shows "r\<^sup>*\<^sup>* x z" using assms(2,1) by induct iprover+ lemma rtranclE [cases set: rtrancl]: fixes a b :: 'a assumes major: "(a, b) \ r\<^sup>*" obtains (base) "a = b" | (step) y where "(a, y) \ r\<^sup>*" and "(y, b) \ r" \ \elimination of \rtrancl\ -- by induction on a special formula\ proof - have "a = b \ (\y. (a, y) \ r\<^sup>* \ (y, b) \ r)" by (rule major [THEN rtrancl_induct]) blast+ then show ?thesis by (auto intro: base step) qed lemma rtrancl_Int_subset: "Id \ s \ (r\<^sup>* \ s) O r \ s \ r\<^sup>* \ s" by (fastforce elim: rtrancl_induct) lemma converse_rtranclp_into_rtranclp: "r a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>*\<^sup>* a c" by (rule rtranclp_trans) iprover+ lemmas converse_rtrancl_into_rtrancl = converse_rtranclp_into_rtranclp [to_set] text \\<^medskip> More \<^term>\r\<^sup>*\ equations and inclusions.\ lemma rtranclp_idemp [simp]: "(r\<^sup>*\<^sup>*)\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*" proof - have "r\<^sup>*\<^sup>*\<^sup>*\<^sup>* x y \ r\<^sup>*\<^sup>* x y" for x y by (induction rule: rtranclp_induct) (blast intro: rtranclp_trans)+ then show ?thesis by (auto intro!: order_antisym) qed lemmas rtrancl_idemp [simp] = rtranclp_idemp [to_set] lemma rtrancl_idemp_self_comp [simp]: "R\<^sup>* O R\<^sup>* = R\<^sup>*" by (force intro: rtrancl_trans) lemma rtrancl_subset_rtrancl: "r \ s\<^sup>* \ r\<^sup>* \ s\<^sup>*" by (drule rtrancl_mono, simp) lemma rtranclp_subset: "R \ S \ S \ R\<^sup>*\<^sup>* \ S\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*" by (fastforce dest: rtranclp_mono) lemmas rtrancl_subset = rtranclp_subset [to_set] lemma rtranclp_sup_rtranclp: "(sup (R\<^sup>*\<^sup>*) (S\<^sup>*\<^sup>*))\<^sup>*\<^sup>* = (sup R S)\<^sup>*\<^sup>*" by (blast intro!: rtranclp_subset intro: rtranclp_mono [THEN predicate2D]) lemmas rtrancl_Un_rtrancl = rtranclp_sup_rtranclp [to_set] lemma rtranclp_reflclp [simp]: "(R\<^sup>=\<^sup>=)\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*" by (blast intro!: rtranclp_subset) lemmas rtrancl_reflcl [simp] = rtranclp_reflclp [to_set] lemma rtrancl_r_diff_Id: "(r - Id)\<^sup>* = r\<^sup>*" by (rule rtrancl_subset [symmetric]) auto lemma rtranclp_r_diff_Id: "(inf r (\))\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*" by (rule rtranclp_subset [symmetric]) auto theorem rtranclp_converseD: assumes "(r\\)\<^sup>*\<^sup>* x y" shows "r\<^sup>*\<^sup>* y x" using assms by induct (iprover intro: rtranclp_trans dest!: conversepD)+ lemmas rtrancl_converseD = rtranclp_converseD [to_set] theorem rtranclp_converseI: assumes "r\<^sup>*\<^sup>* y x" shows "(r\\)\<^sup>*\<^sup>* x y" using assms by induct (iprover intro: rtranclp_trans conversepI)+ lemmas rtrancl_converseI = rtranclp_converseI [to_set] lemma rtrancl_converse: "(r\)\<^sup>* = (r\<^sup>*)\" by (fast dest!: rtrancl_converseD intro!: rtrancl_converseI) lemma sym_rtrancl: "sym r \ sym (r\<^sup>*)" by (simp only: sym_conv_converse_eq rtrancl_converse [symmetric]) theorem converse_rtranclp_induct [consumes 1, case_names base step]: assumes major: "r\<^sup>*\<^sup>* a b" and cases: "P b" "\y z. r y z \ r\<^sup>*\<^sup>* z b \ P z \ P y" shows "P a" using rtranclp_converseI [OF major] by induct (iprover intro: cases dest!: conversepD rtranclp_converseD)+ lemmas converse_rtrancl_induct = converse_rtranclp_induct [to_set] lemmas converse_rtranclp_induct2 = converse_rtranclp_induct [of _ "(ax, ay)" "(bx, by)", split_rule, consumes 1, case_names refl step] lemmas converse_rtrancl_induct2 = converse_rtrancl_induct [of "(ax, ay)" "(bx, by)", split_format (complete), consumes 1, case_names refl step] lemma converse_rtranclpE [consumes 1, case_names base step]: assumes major: "r\<^sup>*\<^sup>* x z" and cases: "x = z \ P" "\y. r x y \ r\<^sup>*\<^sup>* y z \ P" shows P proof - have "x = z \ (\y. r x y \ r\<^sup>*\<^sup>* y z)" - by (rule_tac major [THEN converse_rtranclp_induct]) iprover+ + by (rule major [THEN converse_rtranclp_induct]) iprover+ then show ?thesis by (auto intro: cases) qed lemmas converse_rtranclE = converse_rtranclpE [to_set] lemmas converse_rtranclpE2 = converse_rtranclpE [of _ "(xa,xb)" "(za,zb)", split_rule] lemmas converse_rtranclE2 = converse_rtranclE [of "(xa,xb)" "(za,zb)", split_rule] lemma r_comp_rtrancl_eq: "r O r\<^sup>* = r\<^sup>* O r" by (blast elim: rtranclE converse_rtranclE intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl) lemma rtrancl_unfold: "r\<^sup>* = Id \ r\<^sup>* O r" by (auto intro: rtrancl_into_rtrancl elim: rtranclE) lemma rtrancl_Un_separatorE: "(a, b) \ (P \ Q)\<^sup>* \ \x y. (a, x) \ P\<^sup>* \ (x, y) \ Q \ x = y \ (a, b) \ P\<^sup>*" proof (induct rule: rtrancl.induct) case rtrancl_refl then show ?case by blast next case rtrancl_into_rtrancl then show ?case by (blast intro: rtrancl_trans) qed lemma rtrancl_Un_separator_converseE: "(a, b) \ (P \ Q)\<^sup>* \ \x y. (x, b) \ P\<^sup>* \ (y, x) \ Q \ y = x \ (a, b) \ P\<^sup>*" proof (induct rule: converse_rtrancl_induct) case base then show ?case by blast next case step then show ?case by (blast intro: rtrancl_trans) qed lemma Image_closed_trancl: assumes "r `` X \ X" shows "r\<^sup>* `` X = X" proof - from assms have **: "{y. \x\X. (x, y) \ r} \ X" by auto have "x \ X" if 1: "(y, x) \ r\<^sup>*" and 2: "y \ X" for x y proof - from 1 show "x \ X" proof induct case base show ?case by (fact 2) next case step with ** show ?case by auto qed qed then show ?thesis by auto qed subsection \Transitive closure\ lemma totalp_on_tranclp: "totalp_on A R \ totalp_on A (tranclp R)" by (auto intro: totalp_onI dest: totalp_onD) lemma total_on_trancl: "total_on A r \ total_on A (trancl r)" by (rule totalp_on_tranclp[to_set]) lemma trancl_mono: assumes "p \ r\<^sup>+" "r \ s" shows "p \ s\<^sup>+" proof - have "\(a, b) \ r\<^sup>+; r \ s\ \ (a, b) \ s\<^sup>+" for a b by (induction rule: trancl.induct) (iprover dest: subsetD)+ with assms show ?thesis by (cases p) force qed lemma r_into_trancl': "\p. p \ r \ p \ r\<^sup>+" by (simp only: split_tupled_all) (erule r_into_trancl) text \\<^medskip> Conversions between \trancl\ and \rtrancl\.\ lemma tranclp_into_rtranclp: "r\<^sup>+\<^sup>+ a b \ r\<^sup>*\<^sup>* a b" by (erule tranclp.induct) iprover+ lemmas trancl_into_rtrancl = tranclp_into_rtranclp [to_set] lemma rtranclp_into_tranclp1: assumes "r\<^sup>*\<^sup>* a b" shows "r b c \ r\<^sup>+\<^sup>+ a c" using assms by (induct arbitrary: c) iprover+ lemmas rtrancl_into_trancl1 = rtranclp_into_tranclp1 [to_set] lemma rtranclp_into_tranclp2: assumes "r a b" "r\<^sup>*\<^sup>* b c" shows "r\<^sup>+\<^sup>+ a c" \ \intro rule from \r\ and \rtrancl\\ using \r\<^sup>*\<^sup>* b c\ proof (cases rule: rtranclp.cases) case rtrancl_refl with assms show ?thesis by iprover next case rtrancl_into_rtrancl with assms show ?thesis by (auto intro: rtranclp_trans [THEN rtranclp_into_tranclp1]) qed lemmas rtrancl_into_trancl2 = rtranclp_into_tranclp2 [to_set] text \Nice induction rule for \trancl\\ lemma tranclp_induct [consumes 1, case_names base step, induct pred: tranclp]: assumes a: "r\<^sup>+\<^sup>+ a b" and cases: "\y. r a y \ P y" "\y z. r\<^sup>+\<^sup>+ a y \ r y z \ P y \ P z" shows "P b" using a by (induct x\a b) (iprover intro: cases)+ lemmas trancl_induct [induct set: trancl] = tranclp_induct [to_set] lemmas tranclp_induct2 = tranclp_induct [of _ "(ax, ay)" "(bx, by)", split_rule, consumes 1, case_names base step] lemmas trancl_induct2 = trancl_induct [of "(ax, ay)" "(bx, by)", split_format (complete), consumes 1, case_names base step] lemma tranclp_trans_induct: assumes major: "r\<^sup>+\<^sup>+ x y" and cases: "\x y. r x y \ P x y" "\x y z. r\<^sup>+\<^sup>+ x y \ P x y \ r\<^sup>+\<^sup>+ y z \ P y z \ P x z" shows "P x y" \ \Another induction rule for trancl, incorporating transitivity\ by (iprover intro: major [THEN tranclp_induct] cases) lemmas trancl_trans_induct = tranclp_trans_induct [to_set] lemma tranclE [cases set: trancl]: assumes "(a, b) \ r\<^sup>+" obtains (base) "(a, b) \ r" | (step) c where "(a, c) \ r\<^sup>+" and "(c, b) \ r" using assms by cases simp_all lemma trancl_Int_subset: "r \ s \ (r\<^sup>+ \ s) O r \ s \ r\<^sup>+ \ s" by (fastforce simp add: elim: trancl_induct) lemma trancl_unfold: "r\<^sup>+ = r \ r\<^sup>+ O r" by (auto intro: trancl_into_trancl elim: tranclE) text \Transitivity of \<^term>\r\<^sup>+\\ lemma trans_trancl [simp]: "trans (r\<^sup>+)" proof (rule transI) fix x y z assume "(x, y) \ r\<^sup>+" assume "(y, z) \ r\<^sup>+" then show "(x, z) \ r\<^sup>+" proof induct case (base u) from \(x, y) \ r\<^sup>+\ and \(y, u) \ r\ show "(x, u) \ r\<^sup>+" .. next case (step u v) from \(x, u) \ r\<^sup>+\ and \(u, v) \ r\ show "(x, v) \ r\<^sup>+" .. qed qed lemmas trancl_trans = trans_trancl [THEN transD] lemma tranclp_trans: assumes "r\<^sup>+\<^sup>+ x y" and "r\<^sup>+\<^sup>+ y z" shows "r\<^sup>+\<^sup>+ x z" using assms(2,1) by induct iprover+ lemma trancl_id [simp]: "trans r \ r\<^sup>+ = r" unfolding trans_def by (fastforce simp add: elim: trancl_induct) lemma rtranclp_tranclp_tranclp: assumes "r\<^sup>*\<^sup>* x y" shows "\z. r\<^sup>+\<^sup>+ y z \ r\<^sup>+\<^sup>+ x z" using assms by induct (iprover intro: tranclp_trans)+ lemmas rtrancl_trancl_trancl = rtranclp_tranclp_tranclp [to_set] lemma tranclp_into_tranclp2: "r a b \ r\<^sup>+\<^sup>+ b c \ r\<^sup>+\<^sup>+ a c" by (erule tranclp_trans [OF tranclp.r_into_trancl]) lemmas trancl_into_trancl2 = tranclp_into_tranclp2 [to_set] lemma tranclp_converseI: assumes "(r\<^sup>+\<^sup>+)\\ x y" shows "(r\\)\<^sup>+\<^sup>+ x y" using conversepD [OF assms] proof (induction rule: tranclp_induct) case (base y) then show ?case by (iprover intro: conversepI) next case (step y z) then show ?case by (iprover intro: conversepI tranclp_trans) qed lemmas trancl_converseI = tranclp_converseI [to_set] lemma tranclp_converseD: assumes "(r\\)\<^sup>+\<^sup>+ x y" shows "(r\<^sup>+\<^sup>+)\\ x y" proof - have "r\<^sup>+\<^sup>+ y x" using assms by (induction rule: tranclp_induct) (iprover dest: conversepD intro: tranclp_trans)+ then show ?thesis by (rule conversepI) qed lemmas trancl_converseD = tranclp_converseD [to_set] lemma tranclp_converse: "(r\\)\<^sup>+\<^sup>+ = (r\<^sup>+\<^sup>+)\\" by (fastforce simp add: fun_eq_iff intro!: tranclp_converseI dest!: tranclp_converseD) lemmas trancl_converse = tranclp_converse [to_set] lemma sym_trancl: "sym r \ sym (r\<^sup>+)" by (simp only: sym_conv_converse_eq trancl_converse [symmetric]) lemma converse_tranclp_induct [consumes 1, case_names base step]: assumes major: "r\<^sup>+\<^sup>+ a b" and cases: "\y. r y b \ P y" "\y z. r y z \ r\<^sup>+\<^sup>+ z b \ P z \ P y" shows "P a" proof - have "r\\\<^sup>+\<^sup>+ b a" by (intro tranclp_converseI conversepI major) then show ?thesis by (induction rule: tranclp_induct) (blast intro: cases dest: tranclp_converseD)+ qed lemmas converse_trancl_induct = converse_tranclp_induct [to_set] lemma tranclpD: "R\<^sup>+\<^sup>+ x y \ \z. R x z \ R\<^sup>*\<^sup>* z y" proof (induction rule: converse_tranclp_induct) case (step u v) then show ?case by (blast intro: rtranclp_trans) qed auto lemmas tranclD = tranclpD [to_set] lemma converse_tranclpE: assumes major: "tranclp r x z" and base: "r x z \ P" and step: "\y. r x y \ tranclp r y z \ P" shows P proof - from tranclpD [OF major] obtain y where "r x y" and "rtranclp r y z" by iprover from this(2) show P proof (cases rule: rtranclp.cases) case rtrancl_refl with \r x y\ base show P by iprover next case rtrancl_into_rtrancl then have "tranclp r y z" by (iprover intro: rtranclp_into_tranclp1) with \r x y\ step show P by iprover qed qed lemmas converse_tranclE = converse_tranclpE [to_set] lemma tranclD2: "(x, y) \ R\<^sup>+ \ \z. (x, z) \ R\<^sup>* \ (z, y) \ R" by (blast elim: tranclE intro: trancl_into_rtrancl) lemma irrefl_tranclI: "r\ \ r\<^sup>* = {} \ (x, x) \ r\<^sup>+" by (blast elim: tranclE dest: trancl_into_rtrancl) lemma irrefl_trancl_rD: "\x. (x, x) \ r\<^sup>+ \ (x, y) \ r \ x \ y" by (blast dest: r_into_trancl) lemma trancl_subset_Sigma_aux: "(a, b) \ r\<^sup>* \ r \ A \ A \ a = b \ a \ A" by (induct rule: rtrancl_induct) auto lemma trancl_subset_Sigma: assumes "r \ A \ A" shows "r\<^sup>+ \ A \ A" proof (rule trancl_Int_subset [OF assms]) show "(r\<^sup>+ \ A \ A) O r \ A \ A" using assms by auto qed lemma reflclp_tranclp [simp]: "(r\<^sup>+\<^sup>+)\<^sup>=\<^sup>= = r\<^sup>*\<^sup>*" by (fast elim: rtranclp.cases tranclp_into_rtranclp dest: rtranclp_into_tranclp1) lemmas reflcl_trancl [simp] = reflclp_tranclp [to_set] lemma trancl_reflcl [simp]: "(r\<^sup>=)\<^sup>+ = r\<^sup>*" proof - have "(a, b) \ (r\<^sup>=)\<^sup>+ \ (a, b) \ r\<^sup>*" for a b by (force dest: trancl_into_rtrancl) moreover have "(a, b) \ (r\<^sup>=)\<^sup>+" if "(a, b) \ r\<^sup>*" for a b using that proof (cases a b rule: rtranclE) case step show ?thesis by (rule rtrancl_into_trancl1) (use step in auto) qed auto ultimately show ?thesis by auto qed lemma rtrancl_trancl_reflcl [code]: "r\<^sup>* = (r\<^sup>+)\<^sup>=" by simp lemma trancl_empty [simp]: "{}\<^sup>+ = {}" by (auto elim: trancl_induct) lemma rtrancl_empty [simp]: "{}\<^sup>* = Id" by (rule subst [OF reflcl_trancl]) simp lemma rtranclpD: "R\<^sup>*\<^sup>* a b \ a = b \ a \ b \ R\<^sup>+\<^sup>+ a b" by (force simp: reflclp_tranclp [symmetric] simp del: reflclp_tranclp) lemmas rtranclD = rtranclpD [to_set] lemma rtrancl_eq_or_trancl: "(x,y) \ R\<^sup>* \ x = y \ x \ y \ (x, y) \ R\<^sup>+" by (fast elim: trancl_into_rtrancl dest: rtranclD) lemma trancl_unfold_right: "r\<^sup>+ = r\<^sup>* O r" by (auto dest: tranclD2 intro: rtrancl_into_trancl1) lemma trancl_unfold_left: "r\<^sup>+ = r O r\<^sup>*" by (auto dest: tranclD intro: rtrancl_into_trancl2) lemma trancl_insert: "(insert (y, x) r)\<^sup>+ = r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*}" \ \primitive recursion for \trancl\ over finite relations\ proof - have "\a b. (a, b) \ (insert (y, x) r)\<^sup>+ \ (a, b) \ r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*}" by (erule trancl_induct) (blast intro: rtrancl_into_trancl1 trancl_into_rtrancl trancl_trans)+ moreover have "r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*} \ (insert (y, x) r)\<^sup>+" by (blast intro: trancl_mono rtrancl_mono [THEN [2] rev_subsetD] rtrancl_trancl_trancl rtrancl_into_trancl2) ultimately show ?thesis by auto qed lemma trancl_insert2: "(insert (a, b) r)\<^sup>+ = r\<^sup>+ \ {(x, y). ((x, a) \ r\<^sup>+ \ x = a) \ ((b, y) \ r\<^sup>+ \ y = b)}" by (auto simp: trancl_insert rtrancl_eq_or_trancl) lemma rtrancl_insert: "(insert (a,b) r)\<^sup>* = r\<^sup>* \ {(x, y). (x, a) \ r\<^sup>* \ (b, y) \ r\<^sup>*}" using trancl_insert[of a b r] by (simp add: rtrancl_trancl_reflcl del: reflcl_trancl) blast text \Simplifying nested closures\ lemma rtrancl_trancl_absorb[simp]: "(R\<^sup>*)\<^sup>+ = R\<^sup>*" by (simp add: trans_rtrancl) lemma trancl_rtrancl_absorb[simp]: "(R\<^sup>+)\<^sup>* = R\<^sup>*" by (subst reflcl_trancl[symmetric]) simp lemma rtrancl_reflcl_absorb[simp]: "(R\<^sup>*)\<^sup>= = R\<^sup>*" by auto text \\Domain\ and \Range\\ lemma Domain_rtrancl [simp]: "Domain (R\<^sup>*) = UNIV" by blast lemma Range_rtrancl [simp]: "Range (R\<^sup>*) = UNIV" by blast lemma rtrancl_Un_subset: "(R\<^sup>* \ S\<^sup>*) \ (R \ S)\<^sup>*" by (rule rtrancl_Un_rtrancl [THEN subst]) fast lemma in_rtrancl_UnI: "x \ R\<^sup>* \ x \ S\<^sup>* \ x \ (R \ S)\<^sup>*" by (blast intro: subsetD [OF rtrancl_Un_subset]) lemma trancl_domain [simp]: "Domain (r\<^sup>+) = Domain r" by (unfold Domain_unfold) (blast dest: tranclD) lemma trancl_range [simp]: "Range (r\<^sup>+) = Range r" unfolding Domain_converse [symmetric] by (simp add: trancl_converse [symmetric]) lemma Not_Domain_rtrancl: assumes "x \ Domain R" shows "(x, y) \ R\<^sup>* \ x = y" proof - have "(x, y) \ R\<^sup>* \ x = y" by (erule rtrancl_induct) (use assms in auto) then show ?thesis by auto qed lemma trancl_subset_Field2: "r\<^sup>+ \ Field r \ Field r" by (rule trancl_Int_subset) (auto simp: Field_def) lemma finite_trancl[simp]: "finite (r\<^sup>+) = finite r" proof show "finite (r\<^sup>+) \ finite r" by (blast intro: r_into_trancl' finite_subset) show "finite r \ finite (r\<^sup>+)" by (auto simp: finite_Field trancl_subset_Field2 [THEN finite_subset]) qed lemma finite_rtrancl_Image[simp]: assumes "finite R" "finite A" shows "finite (R\<^sup>* `` A)" proof (rule ccontr) assume "infinite (R\<^sup>* `` A)" with assms show False by(simp add: rtrancl_trancl_reflcl Un_Image del: reflcl_trancl) qed text \More about converse \rtrancl\ and \trancl\, should be merged with main body.\ lemma single_valued_confluent: assumes "single_valued r" and xy: "(x, y) \ r\<^sup>*" and xz: "(x, z) \ r\<^sup>*" shows "(y, z) \ r\<^sup>* \ (z, y) \ r\<^sup>*" using xy proof (induction rule: rtrancl_induct) case base show ?case by (simp add: assms) next case (step y z) with xz \single_valued r\ show ?case by (auto elim: converse_rtranclE dest: single_valuedD intro: rtrancl_trans) qed lemma r_r_into_trancl: "(a, b) \ R \ (b, c) \ R \ (a, c) \ R\<^sup>+" by (fast intro: trancl_trans) lemma trancl_into_trancl: "(a, b) \ r\<^sup>+ \ (b, c) \ r \ (a, c) \ r\<^sup>+" by (induct rule: trancl_induct) (fast intro: r_r_into_trancl trancl_trans)+ lemma tranclp_rtranclp_tranclp: assumes "r\<^sup>+\<^sup>+ a b" "r\<^sup>*\<^sup>* b c" shows "r\<^sup>+\<^sup>+ a c" proof - obtain z where "r a z" "r\<^sup>*\<^sup>* z c" using assms by (iprover dest: tranclpD rtranclp_trans) then show ?thesis by (blast dest: rtranclp_into_tranclp2) qed lemma rtranclp_conversep: "r\\\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*\\" by(auto simp add: fun_eq_iff intro: rtranclp_converseI rtranclp_converseD) lemmas symp_rtranclp = sym_rtrancl[to_pred] lemmas symp_conv_conversep_eq = sym_conv_converse_eq[to_pred] lemmas rtranclp_tranclp_absorb [simp] = rtrancl_trancl_absorb[to_pred] lemmas tranclp_rtranclp_absorb [simp] = trancl_rtrancl_absorb[to_pred] lemmas rtranclp_reflclp_absorb [simp] = rtrancl_reflcl_absorb[to_pred] lemmas trancl_rtrancl_trancl = tranclp_rtranclp_tranclp [to_set] lemmas transitive_closure_trans [trans] = r_r_into_trancl trancl_trans rtrancl_trans trancl.trancl_into_trancl trancl_into_trancl2 rtrancl.rtrancl_into_rtrancl converse_rtrancl_into_rtrancl rtrancl_trancl_trancl trancl_rtrancl_trancl lemmas transitive_closurep_trans' [trans] = tranclp_trans rtranclp_trans tranclp.trancl_into_trancl tranclp_into_tranclp2 rtranclp.rtrancl_into_rtrancl converse_rtranclp_into_rtranclp rtranclp_tranclp_tranclp tranclp_rtranclp_tranclp declare trancl_into_rtrancl [elim] subsection \Symmetric closure\ definition symclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" where "symclp r x y \ r x y \ r y x" lemma symclpI [simp, intro?]: shows symclpI1: "r x y \ symclp r x y" and symclpI2: "r y x \ symclp r x y" by(simp_all add: symclp_def) lemma symclpE [consumes 1, cases pred]: assumes "symclp r x y" obtains (base) "r x y" | (sym) "r y x" using assms by(auto simp add: symclp_def) lemma symclp_pointfree: "symclp r = sup r r\\" by(auto simp add: symclp_def fun_eq_iff) lemma symclp_greater: "r \ symclp r" by(simp add: symclp_pointfree) lemma symclp_conversep [simp]: "symclp r\\ = symclp r" by(simp add: symclp_pointfree sup.commute) lemma symp_symclp [simp]: "symp (symclp r)" by(auto simp add: symp_def elim: symclpE intro: symclpI) lemma symp_symclp_eq: "symp r \ symclp r = r" by(simp add: symclp_pointfree symp_conv_conversep_eq) lemma symp_rtranclp_symclp [simp]: "symp (symclp r)\<^sup>*\<^sup>*" by(simp add: symp_rtranclp) lemma rtranclp_symclp_sym [sym]: "(symclp r)\<^sup>*\<^sup>* x y \ (symclp r)\<^sup>*\<^sup>* y x" by(rule sympD[OF symp_rtranclp_symclp]) lemma symclp_idem [simp]: "symclp (symclp r) = symclp r" by(simp add: symclp_pointfree sup_commute converse_join) lemma reflp_rtranclp [simp]: "reflp R\<^sup>*\<^sup>*" using refl_rtrancl[to_pred, of R] reflp_refl_eq[of "{(x, y). R\<^sup>*\<^sup>* x y}"] by simp subsection \The power operation on relations\ text \\R ^^ n = R O \ O R\, the n-fold composition of \R\\ overloading relpow \ "compow :: nat \ ('a \ 'a) set \ ('a \ 'a) set" relpowp \ "compow :: nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" begin primrec relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where "relpow 0 R = Id" | "relpow (Suc n) R = (R ^^ n) O R" primrec relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" where "relpowp 0 R = HOL.eq" | "relpowp (Suc n) R = (R ^^ n) OO R" end lemma relpowp_relpow_eq [pred_set_conv]: "(\x y. (x, y) \ R) ^^ n = (\x y. (x, y) \ R ^^ n)" for R :: "'a rel" by (induct n) (simp_all add: relcompp_relcomp_eq) text \For code generation:\ definition relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where relpow_code_def [code_abbrev]: "relpow = compow" definition relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" where relpowp_code_def [code_abbrev]: "relpowp = compow" lemma [code]: "relpow (Suc n) R = (relpow n R) O R" "relpow 0 R = Id" by (simp_all add: relpow_code_def) lemma [code]: "relpowp (Suc n) R = (R ^^ n) OO R" "relpowp 0 R = HOL.eq" by (simp_all add: relpowp_code_def) hide_const (open) relpow hide_const (open) relpowp lemma relpow_1 [simp]: "R ^^ 1 = R" for R :: "('a \ 'a) set" by simp lemma relpowp_1 [simp]: "P ^^ 1 = P" for P :: "'a \ 'a \ bool" by (fact relpow_1 [to_pred]) lemma relpow_0_I: "(x, x) \ R ^^ 0" by simp lemma relpowp_0_I: "(P ^^ 0) x x" by (fact relpow_0_I [to_pred]) lemma relpow_Suc_I: "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" by auto lemma relpowp_Suc_I: "(P ^^ n) x y \ P y z \ (P ^^ Suc n) x z" by (fact relpow_Suc_I [to_pred]) lemma relpow_Suc_I2: "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" by (induct n arbitrary: z) (simp, fastforce) lemma relpowp_Suc_I2: "P x y \ (P ^^ n) y z \ (P ^^ Suc n) x z" by (fact relpow_Suc_I2 [to_pred]) lemma relpow_0_E: "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" by simp lemma relpowp_0_E: "(P ^^ 0) x y \ (x = y \ Q) \ Q" by (fact relpow_0_E [to_pred]) lemma relpow_Suc_E: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" by auto lemma relpowp_Suc_E: "(P ^^ Suc n) x z \ (\y. (P ^^ n) x y \ P y z \ Q) \ Q" by (fact relpow_Suc_E [to_pred]) lemma relpow_E: "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) \ (\y m. n = Suc m \ (x, y) \ R ^^ m \ (y, z) \ R \ P) \ P" by (cases n) auto lemma relpowp_E: "(P ^^ n) x z \ (n = 0 \ x = z \ Q) \ (\y m. n = Suc m \ (P ^^ m) x y \ P y z \ Q) \ Q" by (fact relpow_E [to_pred]) lemma relpow_Suc_D2: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" by (induct n arbitrary: x z) (blast intro: relpow_0_I relpow_Suc_I elim: relpow_0_E relpow_Suc_E)+ lemma relpowp_Suc_D2: "(P ^^ Suc n) x z \ \y. P x y \ (P ^^ n) y z" by (fact relpow_Suc_D2 [to_pred]) lemma relpow_Suc_E2: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n \ P) \ P" by (blast dest: relpow_Suc_D2) lemma relpowp_Suc_E2: "(P ^^ Suc n) x z \ (\y. P x y \ (P ^^ n) y z \ Q) \ Q" by (fact relpow_Suc_E2 [to_pred]) lemma relpow_Suc_D2': "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" by (induct n) (simp_all, blast) lemma relpowp_Suc_D2': "\x y z. (P ^^ n) x y \ P y z \ (\w. P x w \ (P ^^ n) w z)" by (fact relpow_Suc_D2' [to_pred]) lemma relpow_E2: assumes "(x, z) \ R ^^ n" "n = 0 \ x = z \ P" "\y m. n = Suc m \ (x, y) \ R \ (y, z) \ R ^^ m \ P" shows "P" proof (cases n) case 0 with assms show ?thesis by simp next case (Suc m) with assms relpow_Suc_D2' [of m R] show ?thesis by force qed lemma relpowp_E2: "(P ^^ n) x z \ (n = 0 \ x = z \ Q) \ (\y m. n = Suc m \ P x y \ (P ^^ m) y z \ Q) \ Q" by (fact relpow_E2 [to_pred]) lemma relpow_add: "R ^^ (m + n) = R^^m O R^^n" by (induct n) auto lemma relpowp_add: "P ^^ (m + n) = P ^^ m OO P ^^ n" by (fact relpow_add [to_pred]) lemma relpow_commute: "R O R ^^ n = R ^^ n O R" by (induct n) (simp_all add: O_assoc [symmetric]) lemma relpowp_commute: "P OO P ^^ n = P ^^ n OO P" by (fact relpow_commute [to_pred]) lemma relpow_empty: "0 < n \ ({} :: ('a \ 'a) set) ^^ n = {}" by (cases n) auto lemma relpowp_bot: "0 < n \ (\ :: 'a \ 'a \ bool) ^^ n = \" by (fact relpow_empty [to_pred]) lemma rtrancl_imp_UN_relpow: assumes "p \ R\<^sup>*" shows "p \ (\n. R ^^ n)" proof (cases p) case (Pair x y) with assms have "(x, y) \ R\<^sup>*" by simp then have "(x, y) \ (\n. R ^^ n)" proof induct case base show ?case by (blast intro: relpow_0_I) next case step then show ?case by (blast intro: relpow_Suc_I) qed with Pair show ?thesis by simp qed lemma rtranclp_imp_Sup_relpowp: assumes "(P\<^sup>*\<^sup>*) x y" shows "(\n. P ^^ n) x y" using assms and rtrancl_imp_UN_relpow [of "(x, y)", to_pred] by simp lemma relpow_imp_rtrancl: assumes "p \ R ^^ n" shows "p \ R\<^sup>*" proof (cases p) case (Pair x y) with assms have "(x, y) \ R ^^ n" by simp then have "(x, y) \ R\<^sup>*" proof (induct n arbitrary: x y) case 0 then show ?case by simp next case Suc then show ?case by (blast elim: relpow_Suc_E intro: rtrancl_into_rtrancl) qed with Pair show ?thesis by simp qed lemma relpowp_imp_rtranclp: "(P ^^ n) x y \ (P\<^sup>*\<^sup>*) x y" using relpow_imp_rtrancl [of "(x, y)", to_pred] by simp lemma rtrancl_is_UN_relpow: "R\<^sup>* = (\n. R ^^ n)" by (blast intro: rtrancl_imp_UN_relpow relpow_imp_rtrancl) lemma rtranclp_is_Sup_relpowp: "P\<^sup>*\<^sup>* = (\n. P ^^ n)" using rtrancl_is_UN_relpow [to_pred, of P] by auto lemma rtrancl_power: "p \ R\<^sup>* \ (\n. p \ R ^^ n)" by (simp add: rtrancl_is_UN_relpow) lemma rtranclp_power: "(P\<^sup>*\<^sup>*) x y \ (\n. (P ^^ n) x y)" by (simp add: rtranclp_is_Sup_relpowp) lemma trancl_power: "p \ R\<^sup>+ \ (\n > 0. p \ R ^^ n)" proof - have "(a, b) \ R\<^sup>+ \ (\n>0. (a, b) \ R ^^ n)" for a b proof safe show "(a, b) \ R\<^sup>+ \ \n>0. (a, b) \ R ^^ n" by (fastforce simp: rtrancl_is_UN_relpow relcomp_unfold dest: tranclD2) show "(a, b) \ R\<^sup>+" if "n > 0" "(a, b) \ R ^^ n" for n proof (cases n) case (Suc m) with that show ?thesis by (auto simp: dest: relpow_imp_rtrancl rtrancl_into_trancl1) qed (use that in auto) qed then show ?thesis by (cases p) auto qed lemma tranclp_power: "(P\<^sup>+\<^sup>+) x y \ (\n > 0. (P ^^ n) x y)" using trancl_power [to_pred, of P "(x, y)"] by simp lemma rtrancl_imp_relpow: "p \ R\<^sup>* \ \n. p \ R ^^ n" by (auto dest: rtrancl_imp_UN_relpow) lemma rtranclp_imp_relpowp: "(P\<^sup>*\<^sup>*) x y \ \n. (P ^^ n) x y" by (auto dest: rtranclp_imp_Sup_relpowp) text \By Sternagel/Thiemann:\ lemma relpow_fun_conv: "(a, b) \ R ^^ n \ (\f. f 0 = a \ f n = b \ (\i R))" proof (induct n arbitrary: b) case 0 show ?case by auto next case (Suc n) show ?case - proof (simp add: relcomp_unfold Suc) - show "(\y. (\f. f 0 = a \ f n = y \ (\i R)) \ (y,b) \ R) \ + proof - + have "(\y. (\f. f 0 = a \ f n = y \ (\i R)) \ (y,b) \ R) \ (\f. f 0 = a \ f(Suc n) = b \ (\i R))" - (is "?l = ?r") + (is "?l \ ?r") proof assume ?l then obtain c f where 1: "f 0 = a" "f n = c" "\i. i < n \ (f i, f (Suc i)) \ R" "(c,b) \ R" by auto let ?g = "\ m. if m = Suc n then b else f m" show ?r by (rule exI[of _ ?g]) (simp add: 1) next assume ?r then obtain f where 1: "f 0 = a" "b = f (Suc n)" "\i. i < Suc n \ (f i, f (Suc i)) \ R" by auto - show ?l by (rule exI[of _ "f n"], rule conjI, rule exI[of _ f], insert 1, auto) + show ?l by (rule exI[of _ "f n"], rule conjI, rule exI[of _ f], auto simp add: 1) qed + then show ?thesis by (simp add: relcomp_unfold Suc) qed qed lemma relpowp_fun_conv: "(P ^^ n) x y \ (\f. f 0 = x \ f n = y \ (\i 'a) set" assumes "finite R" and "k > 0" shows "R^^k \ (\n\{n. 0 < n \ n \ card R}. R^^n)" (is "_ \ ?r") proof - have "(a, b) \ R^^(Suc k) \ \n. 0 < n \ n \ card R \ (a, b) \ R^^n" for a b k proof (induct k arbitrary: b) case 0 then have "R \ {}" by auto with card_0_eq[OF \finite R\] have "card R \ Suc 0" by auto then show ?case using 0 by force next case (Suc k) then obtain a' where "(a, a') \ R^^(Suc k)" and "(a', b) \ R" by auto from Suc(1)[OF \(a, a') \ R^^(Suc k)\] obtain n where "n \ card R" and "(a, a') \ R ^^ n" by auto have "(a, b) \ R^^(Suc n)" using \(a, a') \ R^^n\ and \(a', b)\ R\ by auto from \n \ card R\ consider "n < card R" | "n = card R" by force then show ?case proof cases case 1 then show ?thesis using \(a, b) \ R^^(Suc n)\ Suc_leI[OF \n < card R\] by blast next case 2 from \(a, b) \ R ^^ (Suc n)\ [unfolded relpow_fun_conv] obtain f where "f 0 = a" and "f (Suc n) = b" and steps: "\i. i \ n \ (f i, f (Suc i)) \ R" by auto let ?p = "\i. (f i, f(Suc i))" let ?N = "{i. i \ n}" have "?p ` ?N \ R" using steps by auto from card_mono[OF assms(1) this] have "card (?p ` ?N) \ card R" . also have "\ < card ?N" using \n = card R\ by simp finally have "\ inj_on ?p ?N" by (rule pigeonhole) then obtain i j where i: "i \ n" and j: "j \ n" and ij: "i \ j" and pij: "?p i = ?p j" by (auto simp: inj_on_def) let ?i = "min i j" let ?j = "max i j" have i: "?i \ n" and j: "?j \ n" and pij: "?p ?i = ?p ?j" and ij: "?i < ?j" using i j ij pij unfolding min_def max_def by auto from i j pij ij obtain i j where i: "i \ n" and j: "j \ n" and ij: "i < j" and pij: "?p i = ?p j" by blast let ?g = "\l. if l \ i then f l else f (l + (j - i))" let ?n = "Suc (n - (j - i))" have abl: "(a, b) \ R ^^ ?n" unfolding relpow_fun_conv proof (rule exI[of _ ?g], intro conjI impI allI) show "?g ?n = b" using \f(Suc n) = b\ j ij by auto next fix k assume "k < ?n" show "(?g k, ?g (Suc k)) \ R" proof (cases "k < i") case True with i have "k \ n" by auto from steps[OF this] show ?thesis using True by simp next case False then have "i \ k" by auto show ?thesis proof (cases "k = i") case True then show ?thesis using ij pij steps[OF i] by simp next case False with \i \ k\ have "i < k" by auto then have small: "k + (j - i) \ n" using \k by arith show ?thesis using steps[OF small] \i by auto qed qed qed (simp add: \f 0 = a\) moreover have "?n \ n" using i j ij by arith ultimately show ?thesis using \n = card R\ by blast qed qed then show ?thesis using gr0_implies_Suc[OF \k > 0\] by auto qed lemma relpow_finite_bounded: fixes R :: "('a \ 'a) set" assumes "finite R" shows "R^^k \ (\n\{n. n \ card R}. R^^n)" proof (cases k) case (Suc k') then show ?thesis using relpow_finite_bounded1[OF assms, of k] by auto qed force lemma rtrancl_finite_eq_relpow: "finite R \ R\<^sup>* = (\n\{n. n \ card R}. R^^n)" by (fastforce simp: rtrancl_power dest: relpow_finite_bounded) lemma trancl_finite_eq_relpow: assumes "finite R" shows "R\<^sup>+ = (\n\{n. 0 < n \ n \ card R}. R^^n)" proof - have "\a b n. \0 < n; (a, b) \ R ^^ n\ \ \x>0. x \ card R \ (a, b) \ R ^^ x" using assms by (auto dest: relpow_finite_bounded1) then show ?thesis by (auto simp: trancl_power) qed lemma finite_relcomp[simp,intro]: assumes "finite R" and "finite S" shows "finite (R O S)" proof- have "R O S = (\(x, y)\R. \(u, v)\S. if u = y then {(x, v)} else {})" by (force simp: split_def image_constant_conv split: if_splits) then show ?thesis using assms by clarsimp qed lemma finite_relpow [simp, intro]: fixes R :: "('a \ 'a) set" assumes "finite R" shows "n > 0 \ finite (R^^n)" proof (induct n) case 0 then show ?case by simp next case (Suc n) then show ?case by (cases n) (use assms in simp_all) qed lemma single_valued_relpow: fixes R :: "('a \ 'a) set" shows "single_valued R \ single_valued (R ^^ n)" proof (induct n arbitrary: R) case 0 then show ?case by simp next case (Suc n) show ?case by (rule single_valuedI) (use Suc in \fast dest: single_valuedD elim: relpow_Suc_E\) qed subsection \Bounded transitive closure\ definition ntrancl :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where "ntrancl n R = (\i\{i. 0 < i \ i \ Suc n}. R ^^ i)" lemma ntrancl_Zero [simp, code]: "ntrancl 0 R = R" proof show "R \ ntrancl 0 R" unfolding ntrancl_def by fastforce have "0 < i \ i \ Suc 0 \ i = 1" for i by auto then show "ntrancl 0 R \ R" unfolding ntrancl_def by auto qed lemma ntrancl_Suc [simp]: "ntrancl (Suc n) R = ntrancl n R O (Id \ R)" proof have "(a, b) \ ntrancl n R O (Id \ R)" if "(a, b) \ ntrancl (Suc n) R" for a b proof - from that obtain i where "0 < i" "i \ Suc (Suc n)" "(a, b) \ R ^^ i" unfolding ntrancl_def by auto show ?thesis proof (cases "i = 1") case True with \(a, b) \ R ^^ i\ show ?thesis by (auto simp: ntrancl_def) next case False with \0 < i\ obtain j where j: "i = Suc j" "0 < j" by (cases i) auto with \(a, b) \ R ^^ i\ obtain c where c1: "(a, c) \ R ^^ j" and c2: "(c, b) \ R" by auto from c1 j \i \ Suc (Suc n)\ have "(a, c) \ ntrancl n R" by (fastforce simp: ntrancl_def) with c2 show ?thesis by fastforce qed qed then show "ntrancl (Suc n) R \ ntrancl n R O (Id \ R)" by auto show "ntrancl n R O (Id \ R) \ ntrancl (Suc n) R" by (fastforce simp: ntrancl_def) qed lemma [code]: "ntrancl (Suc n) r = (let r' = ntrancl n r in r' \ r' O r)" by (auto simp: Let_def) lemma finite_trancl_ntranl: "finite R \ trancl R = ntrancl (card R - 1) R" by (cases "card R") (auto simp: trancl_finite_eq_relpow relpow_empty ntrancl_def) subsection \Acyclic relations\ definition acyclic :: "('a \ 'a) set \ bool" where "acyclic r \ (\x. (x,x) \ r\<^sup>+)" abbreviation acyclicP :: "('a \ 'a \ bool) \ bool" where "acyclicP r \ acyclic {(x, y). r x y}" lemma acyclic_irrefl [code]: "acyclic r \ irrefl (r\<^sup>+)" by (simp add: acyclic_def irrefl_def) lemma acyclicI: "\x. (x, x) \ r\<^sup>+ \ acyclic r" by (simp add: acyclic_def) lemma (in preorder) acyclicI_order: assumes *: "\a b. (a, b) \ r \ f b < f a" shows "acyclic r" proof - have "f b < f a" if "(a, b) \ r\<^sup>+" for a b using that by induct (auto intro: * less_trans) then show ?thesis by (auto intro!: acyclicI) qed lemma acyclic_insert [iff]: "acyclic (insert (y, x) r) \ acyclic r \ (x, y) \ r\<^sup>*" by (simp add: acyclic_def trancl_insert) (blast intro: rtrancl_trans) lemma acyclic_converse [iff]: "acyclic (r\) \ acyclic r" by (simp add: acyclic_def trancl_converse) lemmas acyclicP_converse [iff] = acyclic_converse [to_pred] lemma acyclic_impl_antisym_rtrancl: "acyclic r \ antisym (r\<^sup>*)" by (simp add: acyclic_def antisym_def) (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl) (* Other direction: acyclic = no loops antisym = only self loops Goalw [acyclic_def,antisym_def] "antisym( r\<^sup>* ) \ acyclic(r - Id) \ antisym( r\<^sup>* ) = acyclic(r - Id)"; *) lemma acyclic_subset: "acyclic s \ r \ s \ acyclic r" unfolding acyclic_def by (blast intro: trancl_mono) subsection \Setup of transitivity reasoner\ ML \ structure Trancl_Tac = Trancl_Tac ( val r_into_trancl = @{thm trancl.r_into_trancl}; val trancl_trans = @{thm trancl_trans}; val rtrancl_refl = @{thm rtrancl.rtrancl_refl}; val r_into_rtrancl = @{thm r_into_rtrancl}; val trancl_into_rtrancl = @{thm trancl_into_rtrancl}; val rtrancl_trancl_trancl = @{thm rtrancl_trancl_trancl}; val trancl_rtrancl_trancl = @{thm trancl_rtrancl_trancl}; val rtrancl_trans = @{thm rtrancl_trans}; fun decomp \<^Const_>\Trueprop for t\ = let fun dec \<^Const_>\Set.member _ for \<^Const_>\Pair _ _ for a b\ rel\ = let fun decr \<^Const_>\rtrancl _ for r\ = (r,"r*") | decr \<^Const_>\trancl _ for r\ = (r,"r+") | decr r = (r,"r"); val (rel,r) = decr (Envir.beta_eta_contract rel); in SOME (a,b,rel,r) end | dec _ = NONE in dec t end | decomp _ = NONE; ); structure Tranclp_Tac = Trancl_Tac ( val r_into_trancl = @{thm tranclp.r_into_trancl}; val trancl_trans = @{thm tranclp_trans}; val rtrancl_refl = @{thm rtranclp.rtrancl_refl}; val r_into_rtrancl = @{thm r_into_rtranclp}; val trancl_into_rtrancl = @{thm tranclp_into_rtranclp}; val rtrancl_trancl_trancl = @{thm rtranclp_tranclp_tranclp}; val trancl_rtrancl_trancl = @{thm tranclp_rtranclp_tranclp}; val rtrancl_trans = @{thm rtranclp_trans}; fun decomp \<^Const_>\Trueprop for t\ = let fun dec (rel $ a $ b) = let fun decr \<^Const_>\rtranclp _ for r\ = (r,"r*") | decr \<^Const_>\tranclp _ for r\ = (r,"r+") | decr r = (r,"r"); val (rel,r) = decr rel; in SOME (a, b, rel, r) end | dec _ = NONE in dec t end | decomp _ = NONE; ); \ setup \ map_theory_simpset (fn ctxt => ctxt addSolver (mk_solver "Trancl" Trancl_Tac.trancl_tac) addSolver (mk_solver "Rtrancl" Trancl_Tac.rtrancl_tac) addSolver (mk_solver "Tranclp" Tranclp_Tac.trancl_tac) addSolver (mk_solver "Rtranclp" Tranclp_Tac.rtrancl_tac)) \ lemma transp_rtranclp [simp]: "transp R\<^sup>*\<^sup>*" by(auto simp add: transp_def) text \Optional methods.\ method_setup trancl = \Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.trancl_tac)\ \simple transitivity reasoner\ method_setup rtrancl = \Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.rtrancl_tac)\ \simple transitivity reasoner\ method_setup tranclp = \Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.trancl_tac)\ \simple transitivity reasoner (predicate version)\ method_setup rtranclp = \Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.rtrancl_tac)\ \simple transitivity reasoner (predicate version)\ end diff --git a/src/HOL/Wellfounded.thy b/src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy +++ b/src/HOL/Wellfounded.thy @@ -1,981 +1,992 @@ (* Title: HOL/Wellfounded.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Konrad Slind Author: Alexander Krauss Author: Andrei Popescu, TU Muenchen *) section \Well-founded Recursion\ theory Wellfounded imports Transitive_Closure begin subsection \Basic Definitions\ definition wf :: "('a \ 'a) set \ bool" where "wf r \ (\P. (\x. (\y. (y, x) \ r \ P y) \ P x) \ (\x. P x))" definition wfP :: "('a \ 'a \ bool) \ bool" where "wfP r \ wf {(x, y). r x y}" lemma wfP_wf_eq [pred_set_conv]: "wfP (\x y. (x, y) \ r) = wf r" by (simp add: wfP_def) lemma wfUNIVI: "(\P x. (\x. (\y. (y, x) \ r \ P y) \ P x) \ P x) \ wf r" unfolding wf_def by blast lemmas wfPUNIVI = wfUNIVI [to_pred] text \Restriction to domain \A\ and range \B\. If \r\ is well-founded over their intersection, then \wf r\.\ lemma wfI: assumes "r \ A \ B" and "\x P. \\x. (\y. (y, x) \ r \ P y) \ P x; x \ A; x \ B\ \ P x" shows "wf r" using assms unfolding wf_def by blast lemma wf_induct: assumes "wf r" and "\x. \y. (y, x) \ r \ P y \ P x" shows "P a" using assms unfolding wf_def by blast lemmas wfP_induct = wf_induct [to_pred] lemmas wf_induct_rule = wf_induct [rule_format, consumes 1, case_names less, induct set: wf] lemmas wfP_induct_rule = wf_induct_rule [to_pred, induct set: wfP] lemma wf_not_sym: "wf r \ (a, x) \ r \ (x, a) \ r" by (induct a arbitrary: x set: wf) blast lemma wf_asym: assumes "wf r" "(a, x) \ r" obtains "(x, a) \ r" by (drule wf_not_sym[OF assms]) lemma wf_imp_asym: "wf r \ asym r" by (auto intro: asymI elim: wf_asym) lemma wfP_imp_asymp: "wfP r \ asymp r" by (rule wf_imp_asym[to_pred]) lemma wf_not_refl [simp]: "wf r \ (a, a) \ r" by (blast elim: wf_asym) lemma wf_irrefl: assumes "wf r" obtains "(a, a) \ r" by (drule wf_not_refl[OF assms]) lemma wf_imp_irrefl: assumes "wf r" shows "irrefl r" using wf_irrefl [OF assms] by (auto simp add: irrefl_def) lemma wfP_imp_irreflp: "wfP r \ irreflp r" by (rule wf_imp_irrefl[to_pred]) lemma wf_wellorderI: assumes wf: "wf {(x::'a::ord, y). x < y}" and lin: "OFCLASS('a::ord, linorder_class)" shows "OFCLASS('a::ord, wellorder_class)" apply (rule wellorder_class.intro [OF lin]) apply (simp add: wellorder_class.intro class.wellorder_axioms.intro wf_induct_rule [OF wf]) done lemma (in wellorder) wf: "wf {(x, y). x < y}" unfolding wf_def by (blast intro: less_induct) lemma (in wellorder) wfP_less[simp]: "wfP (<)" by (simp add: wf wfP_def) subsection \Basic Results\ text \Point-free characterization of well-foundedness\ lemma wfE_pf: assumes wf: "wf R" and a: "A \ R `` A" shows "A = {}" proof - from wf have "x \ A" for x proof induct fix x assume "\y. (y, x) \ R \ y \ A" then have "x \ R `` A" by blast with a show "x \ A" by blast qed then show ?thesis by auto qed lemma wfI_pf: assumes a: "\A. A \ R `` A \ A = {}" shows "wf R" proof (rule wfUNIVI) fix P :: "'a \ bool" and x let ?A = "{x. \ P x}" assume "\x. (\y. (y, x) \ R \ P y) \ P x" then have "?A \ R `` ?A" by blast with a show "P x" by blast qed subsubsection \Minimal-element characterization of well-foundedness\ lemma wfE_min: assumes wf: "wf R" and Q: "x \ Q" obtains z where "z \ Q" "\y. (y, z) \ R \ y \ Q" using Q wfE_pf[OF wf, of Q] by blast lemma wfE_min': "wf R \ Q \ {} \ (\z. z \ Q \ (\y. (y, z) \ R \ y \ Q) \ thesis) \ thesis" using wfE_min[of R _ Q] by blast lemma wfI_min: assumes a: "\x Q. x \ Q \ \z\Q. \y. (y, z) \ R \ y \ Q" shows "wf R" proof (rule wfI_pf) fix A assume b: "A \ R `` A" have False if "x \ A" for x using a[OF that] b by blast then show "A = {}" by blast qed lemma wf_eq_minimal: "wf r \ (\Q x. x \ Q \ (\z\Q. \y. (y, z) \ r \ y \ Q))" apply (rule iffI) apply (blast intro: elim!: wfE_min) by (rule wfI_min) auto lemmas wfP_eq_minimal = wf_eq_minimal [to_pred] subsubsection \Well-foundedness of transitive closure\ lemma wf_trancl: assumes "wf r" shows "wf (r\<^sup>+)" proof - have "P x" if induct_step: "\x. (\y. (y, x) \ r\<^sup>+ \ P y) \ P x" for P x proof (rule induct_step) show "P y" if "(y, x) \ r\<^sup>+" for y using \wf r\ and that proof (induct x arbitrary: y) case (less x) note hyp = \\x' y'. (x', x) \ r \ (y', x') \ r\<^sup>+ \ P y'\ from \(y, x) \ r\<^sup>+\ show "P y" proof cases case base show "P y" proof (rule induct_step) fix y' assume "(y', y) \ r\<^sup>+" with \(y, x) \ r\ show "P y'" by (rule hyp [of y y']) qed next case step then obtain x' where "(x', x) \ r" and "(y, x') \ r\<^sup>+" by simp then show "P y" by (rule hyp [of x' y]) qed qed qed then show ?thesis unfolding wf_def by blast qed lemmas wfP_trancl = wf_trancl [to_pred] lemma wf_converse_trancl: "wf (r\) \ wf ((r\<^sup>+)\)" apply (subst trancl_converse [symmetric]) apply (erule wf_trancl) done text \Well-foundedness of subsets\ lemma wf_subset: "wf r \ p \ r \ wf p" by (simp add: wf_eq_minimal) fast lemmas wfP_subset = wf_subset [to_pred] text \Well-foundedness of the empty relation\ lemma wf_empty [iff]: "wf {}" by (simp add: wf_def) lemma wfP_empty [iff]: "wfP (\x y. False)" proof - have "wfP bot" by (fact wf_empty[to_pred bot_empty_eq2]) then show ?thesis by (simp add: bot_fun_def) qed lemma wf_Int1: "wf r \ wf (r \ r')" by (erule wf_subset) (rule Int_lower1) lemma wf_Int2: "wf r \ wf (r' \ r)" by (erule wf_subset) (rule Int_lower2) text \Exponentiation.\ lemma wf_exp: assumes "wf (R ^^ n)" shows "wf R" proof (rule wfI_pf) fix A assume "A \ R `` A" then have "A \ (R ^^ n) `` A" by (induct n) force+ with \wf (R ^^ n)\ show "A = {}" by (rule wfE_pf) qed text \Well-foundedness of \insert\.\ lemma wf_insert [iff]: "wf (insert (y,x) r) \ wf r \ (x,y) \ r\<^sup>*" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (blast elim: wf_trancl [THEN wf_irrefl] intro: rtrancl_into_trancl1 wf_subset rtrancl_mono [THEN subsetD]) next assume R: ?rhs then have R': "Q \ {} \ (\z\Q. \y. (y, z) \ r \ y \ Q)" for Q by (auto simp: wf_eq_minimal) show ?lhs unfolding wf_eq_minimal proof clarify fix Q :: "'a set" and q assume "q \ Q" then obtain a where "a \ Q" and a: "\y. (y, a) \ r \ y \ Q" using R by (auto simp: wf_eq_minimal) show "\z\Q. \y'. (y', z) \ insert (y, x) r \ y' \ Q" proof (cases "a=x") case True show ?thesis proof (cases "y \ Q") case True then obtain z where "z \ Q" "(z, y) \ r\<^sup>*" "\z'. (z', z) \ r \ z' \ Q \ (z', y) \ r\<^sup>*" using R' [of "{z \ Q. (z,y) \ r\<^sup>*}"] by auto - with R show ?thesis - by (rule_tac x="z" in bexI) (blast intro: rtrancl_trans) + then have "\y'. (y', z) \ insert (y, x) r \ y' \ Q" + using R by(blast intro: rtrancl_trans)+ + then show ?thesis + by (rule bexI) fact next case False then show ?thesis using a \a \ Q\ by blast qed next case False with a \a \ Q\ show ?thesis by blast qed qed qed subsubsection \Well-foundedness of image\ lemma wf_map_prod_image_Dom_Ran: fixes r:: "('a \ 'a) set" and f:: "'a \ 'b" assumes wf_r: "wf r" and inj: "\ a a'. a \ Domain r \ a' \ Range r \ f a = f a' \ a = a'" shows "wf (map_prod f f ` r)" proof (unfold wf_eq_minimal, clarify) fix B :: "'b set" and b::"'b" assume "b \ B" define A where "A = f -` B \ Domain r" show "\z\B. \y. (y, z) \ map_prod f f ` r \ y \ B" proof (cases "A = {}") case False then obtain a0 where "a0 \ A" and "\a. (a, a0) \ r \ a \ A" using wfE_min[OF wf_r] by auto thus ?thesis using inj unfolding A_def by (intro bexI[of _ "f a0"]) auto - qed (insert \b \ B\, unfold A_def, auto) + qed (use \b \ B\ in \unfold A_def, auto\) qed lemma wf_map_prod_image: "wf r \ inj f \ wf (map_prod f f ` r)" by(rule wf_map_prod_image_Dom_Ran) (auto dest: inj_onD) subsection \Well-Foundedness Results for Unions\ lemma wf_union_compatible: assumes "wf R" "wf S" assumes "R O S \ R" shows "wf (R \ S)" proof (rule wfI_min) fix x :: 'a and Q let ?Q' = "{x \ Q. \y. (y, x) \ R \ y \ Q}" assume "x \ Q" obtain a where "a \ ?Q'" by (rule wfE_min [OF \wf R\ \x \ Q\]) blast with \wf S\ obtain z where "z \ ?Q'" and zmin: "\y. (y, z) \ S \ y \ ?Q'" by (erule wfE_min) have "y \ Q" if "(y, z) \ S" for y proof from that have "y \ ?Q'" by (rule zmin) assume "y \ Q" with \y \ ?Q'\ obtain w where "(w, y) \ R" and "w \ Q" by auto from \(w, y) \ R\ \(y, z) \ S\ have "(w, z) \ R O S" by (rule relcompI) with \R O S \ R\ have "(w, z) \ R" .. with \z \ ?Q'\ have "w \ Q" by blast with \w \ Q\ show False by contradiction qed with \z \ ?Q'\ show "\z\Q. \y. (y, z) \ R \ S \ y \ Q" by blast qed text \Well-foundedness of indexed union with disjoint domains and ranges.\ lemma wf_UN: assumes r: "\i. i \ I \ wf (r i)" and disj: "\i j. \i \ I; j \ I; r i \ r j\ \ Domain (r i) \ Range (r j) = {}" shows "wf (\i\I. r i)" unfolding wf_eq_minimal proof clarify fix A and a :: "'b" assume "a \ A" show "\z\A. \y. (y, z) \ \(r ` I) \ y \ A" proof (cases "\i\I. \a\A. \b\A. (b, a) \ r i") case True then obtain i b c where ibc: "i \ I" "b \ A" "c \ A" "(c,b) \ r i" by blast have ri: "\Q. Q \ {} \ \z\Q. \y. (y, z) \ r i \ y \ Q" using r [OF \i \ I\] unfolding wf_eq_minimal by auto show ?thesis using ri [of "{a. a \ A \ (\b\A. (b, a) \ r i) }"] ibc disj by blast next case False with \a \ A\ show ?thesis by blast qed qed lemma wfP_SUP: "\i. wfP (r i) \ \i j. r i \ r j \ inf (Domainp (r i)) (Rangep (r j)) = bot \ wfP (\(range r))" by (rule wf_UN[to_pred]) simp_all lemma wf_Union: assumes "\r\R. wf r" and "\r\R. \s\R. r \ s \ Domain r \ Range s = {}" shows "wf (\R)" using assms wf_UN[of R "\i. i"] by simp text \ Intuition: We find an \R \ S\-min element of a nonempty subset \A\ by case distinction. \<^enum> There is a step \a \R\ b\ with \a, b \ A\. Pick an \R\-min element \z\ of the (nonempty) set \{a\A | \b\A. a \R\ b}\. By definition, there is \z' \ A\ s.t. \z \R\ z'\. Because \z\ is \R\-min in the subset, \z'\ must be \R\-min in \A\. Because \z'\ has an \R\-predecessor, it cannot have an \S\-successor and is thus \S\-min in \A\ as well. \<^enum> There is no such step. Pick an \S\-min element of \A\. In this case it must be an \R\-min element of \A\ as well. \ lemma wf_Un: "wf r \ wf s \ Domain r \ Range s = {} \ wf (r \ s)" using wf_union_compatible[of s r] by (auto simp: Un_ac) lemma wf_union_merge: "wf (R \ S) = wf (R O R \ S O R \ S)" (is "wf ?A = wf ?B") proof assume "wf ?A" with wf_trancl have wfT: "wf (?A\<^sup>+)" . moreover have "?B \ ?A\<^sup>+" by (subst trancl_unfold, subst trancl_unfold) blast ultimately show "wf ?B" by (rule wf_subset) next assume "wf ?B" show "wf ?A" proof (rule wfI_min) fix Q :: "'a set" and x assume "x \ Q" with \wf ?B\ obtain z where "z \ Q" and "\y. (y, z) \ ?B \ y \ Q" by (erule wfE_min) then have 1: "\y. (y, z) \ R O R \ y \ Q" and 2: "\y. (y, z) \ S O R \ y \ Q" and 3: "\y. (y, z) \ S \ y \ Q" by auto show "\z\Q. \y. (y, z) \ ?A \ y \ Q" proof (cases "\y. (y, z) \ R \ y \ Q") case True with \z \ Q\ 3 show ?thesis by blast next case False then obtain z' where "z'\Q" "(z', z) \ R" by blast have "\y. (y, z') \ ?A \ y \ Q" proof (intro allI impI) fix y assume "(y, z') \ ?A" then show "y \ Q" proof assume "(y, z') \ R" then have "(y, z) \ R O R" using \(z', z) \ R\ .. with 1 show "y \ Q" . next assume "(y, z') \ S" then have "(y, z) \ S O R" using \(z', z) \ R\ .. with 2 show "y \ Q" . qed qed with \z' \ Q\ show ?thesis .. qed qed qed lemma wf_comp_self: "wf R \ wf (R O R)" \ \special case\ by (rule wf_union_merge [where S = "{}", simplified]) subsection \Well-Foundedness of Composition\ text \Bachmair and Dershowitz 1986, Lemma 2. [Provided by Tjark Weber]\ lemma qc_wf_relto_iff: assumes "R O S \ (R \ S)\<^sup>* O R" \ \R quasi-commutes over S\ shows "wf (S\<^sup>* O R O S\<^sup>*) \ wf R" (is "wf ?S \ _") proof show "wf R" if "wf ?S" proof - have "R \ ?S" by auto with wf_subset [of ?S] that show "wf R" by auto qed next show "wf ?S" if "wf R" proof (rule wfI_pf) fix A assume A: "A \ ?S `` A" let ?X = "(R \ S)\<^sup>* `` A" have *: "R O (R \ S)\<^sup>* \ (R \ S)\<^sup>* O R" proof - have "(x, z) \ (R \ S)\<^sup>* O R" if "(y, z) \ (R \ S)\<^sup>*" and "(x, y) \ R" for x y z using that proof (induct y z) case rtrancl_refl then show ?case by auto next case (rtrancl_into_rtrancl a b c) then have "(x, c) \ ((R \ S)\<^sup>* O (R \ S)\<^sup>*) O R" using assms by blast then show ?case by simp qed then show ?thesis by auto qed then have "R O S\<^sup>* \ (R \ S)\<^sup>* O R" using rtrancl_Un_subset by blast then have "?S \ (R \ S)\<^sup>* O (R \ S)\<^sup>* O R" by (simp add: relcomp_mono rtrancl_mono) also have "\ = (R \ S)\<^sup>* O R" by (simp add: O_assoc[symmetric]) finally have "?S O (R \ S)\<^sup>* \ (R \ S)\<^sup>* O R O (R \ S)\<^sup>*" by (simp add: O_assoc[symmetric] relcomp_mono) also have "\ \ (R \ S)\<^sup>* O (R \ S)\<^sup>* O R" using * by (simp add: relcomp_mono) finally have "?S O (R \ S)\<^sup>* \ (R \ S)\<^sup>* O R" by (simp add: O_assoc[symmetric]) then have "(?S O (R \ S)\<^sup>*) `` A \ ((R \ S)\<^sup>* O R) `` A" by (simp add: Image_mono) moreover have "?X \ (?S O (R \ S)\<^sup>*) `` A" using A by (auto simp: relcomp_Image) ultimately have "?X \ R `` ?X" by (auto simp: relcomp_Image) then have "?X = {}" using \wf R\ by (simp add: wfE_pf) moreover have "A \ ?X" by auto ultimately show "A = {}" by simp qed qed corollary wf_relcomp_compatible: assumes "wf R" and "R O S \ S O R" shows "wf (S O R)" proof - have "R O S \ (R \ S)\<^sup>* O R" using assms by blast then have "wf (S\<^sup>* O R O S\<^sup>*)" by (simp add: assms qc_wf_relto_iff) then show ?thesis by (rule Wellfounded.wf_subset) blast qed subsection \Acyclic relations\ lemma wf_acyclic: "wf r \ acyclic r" by (simp add: acyclic_def) (blast elim: wf_trancl [THEN wf_irrefl]) lemmas wfP_acyclicP = wf_acyclic [to_pred] subsubsection \Wellfoundedness of finite acyclic relations\ lemma finite_acyclic_wf: assumes "finite r" "acyclic r" shows "wf r" using assms proof (induction r rule: finite_induct) case (insert x r) then show ?case by (cases x) simp qed simp lemma finite_acyclic_wf_converse: "finite r \ acyclic r \ wf (r\)" apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf]) apply (erule acyclic_converse [THEN iffD2]) done text \ Observe that the converse of an irreflexive, transitive, and finite relation is again well-founded. Thus, we may employ it for well-founded induction. \ lemma wf_converse: assumes "irrefl r" and "trans r" and "finite r" shows "wf (r\)" proof - have "acyclic r" using \irrefl r\ and \trans r\ by (simp add: irrefl_def acyclic_irrefl) with \finite r\ show ?thesis by (rule finite_acyclic_wf_converse) qed lemma wf_iff_acyclic_if_finite: "finite r \ wf r = acyclic r" by (blast intro: finite_acyclic_wf wf_acyclic) subsection \\<^typ>\nat\ is well-founded\ lemma less_nat_rel: "(<) = (\m n. n = Suc m)\<^sup>+\<^sup>+" proof (rule ext, rule ext, rule iffI) fix n m :: nat show "(\m n. n = Suc m)\<^sup>+\<^sup>+ m n" if "m < n" using that proof (induct n) case 0 then show ?case by auto next case (Suc n) then show ?case by (auto simp add: less_Suc_eq_le le_less intro: tranclp.trancl_into_trancl) qed show "m < n" if "(\m n. n = Suc m)\<^sup>+\<^sup>+ m n" using that by (induct n) (simp_all add: less_Suc_eq_le reflexive le_less) qed definition pred_nat :: "(nat \ nat) set" where "pred_nat = {(m, n). n = Suc m}" definition less_than :: "(nat \ nat) set" where "less_than = pred_nat\<^sup>+" lemma less_eq: "(m, n) \ pred_nat\<^sup>+ \ m < n" unfolding less_nat_rel pred_nat_def trancl_def by simp lemma pred_nat_trancl_eq_le: "(m, n) \ pred_nat\<^sup>* \ m \ n" unfolding less_eq rtrancl_eq_or_trancl by auto lemma wf_pred_nat: "wf pred_nat" - apply (unfold wf_def pred_nat_def) - apply clarify - apply (induct_tac x) - apply blast+ - done + unfolding wf_def +proof clarify + fix P x + assume "\x'. (\y. (y, x') \ pred_nat \ P y) \ P x'" + then show "P x" + unfolding pred_nat_def by (induction x) blast+ +qed lemma wf_less_than [iff]: "wf less_than" by (simp add: less_than_def wf_pred_nat [THEN wf_trancl]) lemma trans_less_than [iff]: "trans less_than" by (simp add: less_than_def) lemma less_than_iff [iff]: "((x,y) \ less_than) = (xAccessible Part\ text \ Inductive definition of the accessible part \acc r\ of a relation; see also @{cite "paulin-tlca"}. \ inductive_set acc :: "('a \ 'a) set \ 'a set" for r :: "('a \ 'a) set" where accI: "(\y. (y, x) \ r \ y \ acc r) \ x \ acc r" abbreviation termip :: "('a \ 'a \ bool) \ 'a \ bool" where "termip r \ accp (r\\)" abbreviation termi :: "('a \ 'a) set \ 'a set" where "termi r \ acc (r\)" lemmas accpI = accp.accI lemma accp_eq_acc [code]: "accp r = (\x. x \ Wellfounded.acc {(x, y). r x y})" by (simp add: acc_def) text \Induction rules\ theorem accp_induct: assumes major: "accp r a" assumes hyp: "\x. accp r x \ \y. r y x \ P y \ P x" shows "P a" apply (rule major [THEN accp.induct]) apply (rule hyp) apply (rule accp.accI) apply auto done lemmas accp_induct_rule = accp_induct [rule_format, induct set: accp] theorem accp_downward: "accp r b \ r a b \ accp r a" by (cases rule: accp.cases) lemma not_accp_down: assumes na: "\ accp R x" obtains z where "R z x" and "\ accp R z" proof - assume a: "\z. R z x \ \ accp R z \ thesis" show thesis proof (cases "\z. R z x \ accp R z") case True then have "\z. R z x \ accp R z" by auto then have "accp R x" by (rule accp.accI) with na show thesis .. next case False then obtain z where "R z x" and "\ accp R z" by auto with a show thesis . qed qed lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a \ accp r a \ accp r b" by (erule rtranclp_induct) (blast dest: accp_downward)+ theorem accp_downwards: "accp r a \ r\<^sup>*\<^sup>* b a \ accp r b" by (blast dest: accp_downwards_aux) theorem accp_wfPI: "\x. accp r x \ wfP r" - apply (rule wfPUNIVI) - apply (rule_tac P = P in accp_induct) - apply blast+ - done +proof (rule wfPUNIVI) + fix P x + assume "\x. accp r x" "\x. (\y. r y x \ P y) \ P x" + then show "P x" + using accp_induct[where P = P] by blast +qed theorem accp_wfPD: "wfP r \ accp r x" apply (erule wfP_induct_rule) apply (rule accp.accI) apply blast done theorem wfP_accp_iff: "wfP r = (\x. accp r x)" by (blast intro: accp_wfPI dest: accp_wfPD) text \Smaller relations have bigger accessible parts:\ lemma accp_subset: assumes "R1 \ R2" shows "accp R2 \ accp R1" proof (rule predicate1I) fix x assume "accp R2 x" then show "accp R1 x" proof (induct x) fix x assume "\y. R2 y x \ accp R1 y" with assms show "accp R1 x" by (blast intro: accp.accI) qed qed text \This is a generalized induction theorem that works on subsets of the accessible part.\ lemma accp_subset_induct: assumes subset: "D \ accp R" and dcl: "\x z. D x \ R z x \ D z" and "D x" and istep: "\x. D x \ (\z. R z x \ P z) \ P x" shows "P x" proof - from subset and \D x\ have "accp R x" .. then show "P x" using \D x\ proof (induct x) fix x assume "D x" and "\y. R y x \ D y \ P y" with dcl and istep show "P x" by blast qed qed text \Set versions of the above theorems\ lemmas acc_induct = accp_induct [to_set] lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc] lemmas acc_downward = accp_downward [to_set] lemmas not_acc_down = not_accp_down [to_set] lemmas acc_downwards_aux = accp_downwards_aux [to_set] lemmas acc_downwards = accp_downwards [to_set] lemmas acc_wfI = accp_wfPI [to_set] lemmas acc_wfD = accp_wfPD [to_set] lemmas wf_acc_iff = wfP_accp_iff [to_set] lemmas acc_subset = accp_subset [to_set] lemmas acc_subset_induct = accp_subset_induct [to_set] subsection \Tools for building wellfounded relations\ text \Inverse Image\ lemma wf_inv_image [simp,intro!]: fixes f :: "'a \ 'b" assumes "wf r" shows "wf (inv_image r f)" -proof (clarsimp simp: inv_image_def wf_eq_minimal) - fix P and x::'a - assume "x \ P" - then obtain w where w: "w \ {w. \x::'a. x \ P \ f x = w}" - by auto - have *: "\Q u. u \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" - using assms by (auto simp add: wf_eq_minimal) - show "\z\P. \y. (f y, f z) \ r \ y \ P" - using * [OF w] by auto +proof - + have "\x P. x \ P \ \z\P. \y. (f y, f z) \ r \ y \ P" + proof - + fix P and x::'a + assume "x \ P" + then obtain w where w: "w \ {w. \x::'a. x \ P \ f x = w}" + by auto + have *: "\Q u. u \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" + using assms by (auto simp add: wf_eq_minimal) + show "\z\P. \y. (f y, f z) \ r \ y \ P" + using * [OF w] by auto + qed + then show ?thesis + by (clarsimp simp: inv_image_def wf_eq_minimal) qed text \Measure functions into \<^typ>\nat\\ definition measure :: "('a \ nat) \ ('a \ 'a) set" where "measure = inv_image less_than" lemma in_measure[simp, code_unfold]: "(x, y) \ measure f \ f x < f y" by (simp add:measure_def) lemma wf_measure [iff]: "wf (measure f)" unfolding measure_def by (rule wf_less_than [THEN wf_inv_image]) lemma wf_if_measure: "(\x. P x \ f(g x) < f x) \ wf {(y,x). P x \ y = g x}" for f :: "'a \ nat" using wf_measure[of f] unfolding measure_def inv_image_def less_than_def less_eq by (rule wf_subset) auto subsubsection \Lexicographic combinations\ definition lex_prod :: "('a \'a) set \ ('b \ 'b) set \ (('a \ 'b) \ ('a \ 'b)) set" (infixr "<*lex*>" 80) where "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \ ra \ a = a' \ (b, b') \ rb}" lemma in_lex_prod[simp]: "((a, b), (a', b')) \ r <*lex*> s \ (a, a') \ r \ a = a' \ (b, b') \ s" by (auto simp:lex_prod_def) lemma wf_lex_prod [intro!]: assumes "wf ra" "wf rb" shows "wf (ra <*lex*> rb)" proof (rule wfI) fix z :: "'a \ 'b" and P assume * [rule_format]: "\u. (\v. (v, u) \ ra <*lex*> rb \ P v) \ P u" obtain x y where zeq: "z = (x,y)" by fastforce have "P(x,y)" using \wf ra\ proof (induction x arbitrary: y rule: wf_induct_rule) case (less x) note lessx = less show ?case using \wf rb\ less proof (induction y rule: wf_induct_rule) case (less y) show ?case by (force intro: * less.IH lessx) qed qed then show "P z" by (simp add: zeq) qed auto text \\<*lex*>\ preserves transitivity\ lemma trans_lex_prod [simp,intro!]: "trans R1 \ trans R2 \ trans (R1 <*lex*> R2)" unfolding trans_def lex_prod_def by blast lemma total_on_lex_prod [simp]: "total_on A r \ total_on B s \ total_on (A \ B) (r <*lex*> s)" by (auto simp: total_on_def) lemma asym_lex_prod: "\asym R; asym S\ \ asym (R <*lex*> S)" by (auto simp add: asym_iff lex_prod_def) lemma total_lex_prod [simp]: "total r \ total s \ total (r <*lex*> s)" by (auto simp: total_on_def) text \lexicographic combinations with measure functions\ definition mlex_prod :: "('a \ nat) \ ('a \ 'a) set \ ('a \ 'a) set" (infixr "<*mlex*>" 80) where "f <*mlex*> R = inv_image (less_than <*lex*> R) (\x. (f x, x))" lemma wf_mlex: "wf R \ wf (f <*mlex*> R)" and mlex_less: "f x < f y \ (x, y) \ f <*mlex*> R" and mlex_leq: "f x \ f y \ (x, y) \ R \ (x, y) \ f <*mlex*> R" and mlex_iff: "(x, y) \ f <*mlex*> R \ f x < f y \ f x = f y \ (x, y) \ R" by (auto simp: mlex_prod_def) text \Proper subset relation on finite sets.\ definition finite_psubset :: "('a set \ 'a set) set" where "finite_psubset = {(A, B). A \ B \ finite B}" lemma wf_finite_psubset[simp]: "wf finite_psubset" apply (unfold finite_psubset_def) apply (rule wf_measure [THEN wf_subset]) apply (simp add: measure_def inv_image_def less_than_def less_eq) apply (fast elim!: psubset_card_mono) done lemma trans_finite_psubset: "trans finite_psubset" by (auto simp: finite_psubset_def less_le trans_def) lemma in_finite_psubset[simp]: "(A, B) \ finite_psubset \ A \ B \ finite B" unfolding finite_psubset_def by auto text \max- and min-extension of order to finite sets\ inductive_set max_ext :: "('a \ 'a) set \ ('a set \ 'a set) set" for R :: "('a \ 'a) set" where max_extI[intro]: "finite X \ finite Y \ Y \ {} \ (\x. x \ X \ \y\Y. (x, y) \ R) \ (X, Y) \ max_ext R" lemma max_ext_wf: assumes wf: "wf r" shows "wf (max_ext r)" proof (rule acc_wfI, intro allI) show "M \ acc (max_ext r)" (is "_ \ ?W") for M proof (induct M rule: infinite_finite_induct) case empty show ?case by (rule accI) (auto elim: max_ext.cases) next case (insert a M) from wf \M \ ?W\ \finite M\ show "insert a M \ ?W" proof (induct arbitrary: M) fix M a assume "M \ ?W" assume [intro]: "finite M" assume hyp: "\b M. (b, a) \ r \ M \ ?W \ finite M \ insert b M \ ?W" have add_less: "M \ ?W \ (\y. y \ N \ (y, a) \ r) \ N \ M \ ?W" if "finite N" "finite M" for N M :: "'a set" using that by (induct N arbitrary: M) (auto simp: hyp) show "insert a M \ ?W" proof (rule accI) fix N assume Nless: "(N, insert a M) \ max_ext r" then have *: "\x. x \ N \ (x, a) \ r \ (\y \ M. (x, y) \ r)" by (auto elim!: max_ext.cases) let ?N1 = "{n \ N. (n, a) \ r}" let ?N2 = "{n \ N. (n, a) \ r}" have N: "?N1 \ ?N2 = N" by (rule set_eqI) auto from Nless have "finite N" by (auto elim: max_ext.cases) then have finites: "finite ?N1" "finite ?N2" by auto have "?N2 \ ?W" proof (cases "M = {}") case [simp]: True have Mw: "{} \ ?W" by (rule accI) (auto elim: max_ext.cases) from * have "?N2 = {}" by auto with Mw show "?N2 \ ?W" by (simp only:) next case False from * finites have N2: "(?N2, M) \ max_ext r" - by (rule_tac max_extI[OF _ _ \M \ {}\]) auto + using max_extI[OF _ _ \M \ {}\, where ?X = ?N2] by auto with \M \ ?W\ show "?N2 \ ?W" by (rule acc_downward) qed with finites have "?N1 \ ?N2 \ ?W" by (rule add_less) simp then show "N \ ?W" by (simp only: N) qed qed next case infinite show ?case by (rule accI) (auto elim: max_ext.cases simp: infinite) qed qed lemma max_ext_additive: "(A, B) \ max_ext R \ (C, D) \ max_ext R \ (A \ C, B \ D) \ max_ext R" by (force elim!: max_ext.cases) definition min_ext :: "('a \ 'a) set \ ('a set \ 'a set) set" where "min_ext r = {(X, Y) | X Y. X \ {} \ (\y \ Y. (\x \ X. (x, y) \ r))}" lemma min_ext_wf: assumes "wf r" shows "wf (min_ext r)" proof (rule wfI_min) show "\m \ Q. (\n. (n, m) \ min_ext r \ n \ Q)" if nonempty: "x \ Q" for Q :: "'a set set" and x proof (cases "Q = {{}}") case True then show ?thesis by (simp add: min_ext_def) next case False with nonempty obtain e x where "x \ Q" "e \ x" by force then have eU: "e \ \Q" by auto with \wf r\ obtain z where z: "z \ \Q" "\y. (y, z) \ r \ y \ \Q" by (erule wfE_min) from z obtain m where "m \ Q" "z \ m" by auto from \m \ Q\ show ?thesis proof (intro rev_bexI allI impI) fix n assume smaller: "(n, m) \ min_ext r" with \z \ m\ obtain y where "y \ n" "(y, z) \ r" by (auto simp: min_ext_def) with z(2) show "n \ Q" by auto qed qed qed subsubsection \Bounded increase must terminate\ lemma wf_bounded_measure: fixes ub :: "'a \ nat" and f :: "'a \ nat" assumes "\a b. (b, a) \ r \ ub b \ ub a \ ub a \ f b \ f b > f a" shows "wf r" by (rule wf_subset[OF wf_measure[of "\a. ub a - f a"]]) (auto dest: assms) lemma wf_bounded_set: fixes ub :: "'a \ 'b set" and f :: "'a \ 'b set" assumes "\a b. (b,a) \ r \ finite (ub a) \ ub b \ ub a \ ub a \ f b \ f b \ f a" shows "wf r" apply (rule wf_bounded_measure[of r "\a. card (ub a)" "\a. card (f a)"]) apply (drule assms) apply (blast intro: card_mono finite_subset psubset_card_mono dest: psubset_eq[THEN iffD2]) done lemma finite_subset_wf: assumes "finite A" shows "wf {(X, Y). X \ Y \ Y \ A}" by (rule wf_subset[OF wf_finite_psubset[unfolded finite_psubset_def]]) (auto intro: finite_subset[OF _ assms]) hide_const (open) acc accp end diff --git a/src/HOL/Wfrec.thy b/src/HOL/Wfrec.thy --- a/src/HOL/Wfrec.thy +++ b/src/HOL/Wfrec.thy @@ -1,115 +1,120 @@ (* Title: HOL/Wfrec.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Konrad Slind *) section \Well-Founded Recursion Combinator\ theory Wfrec imports Wellfounded begin inductive wfrec_rel :: "('a \ 'a) set \ (('a \ 'b) \ ('a \ 'b)) \ 'a \ 'b \ bool" for R F where wfrecI: "(\z. (z, x) \ R \ wfrec_rel R F z (g z)) \ wfrec_rel R F x (F g x)" definition cut :: "('a \ 'b) \ ('a \ 'a) set \ 'a \ 'a \ 'b" where "cut f R x = (\y. if (y, x) \ R then f y else undefined)" definition adm_wf :: "('a \ 'a) set \ (('a \ 'b) \ ('a \ 'b)) \ bool" where "adm_wf R F \ (\f g x. (\z. (z, x) \ R \ f z = g z) \ F f x = F g x)" definition wfrec :: "('a \ 'a) set \ (('a \ 'b) \ ('a \ 'b)) \ ('a \ 'b)" where "wfrec R F = (\x. THE y. wfrec_rel R (\f x. F (cut f R x) x) x y)" lemma cuts_eq: "(cut f R x = cut g R x) \ (\y. (y, x) \ R \ f y = g y)" by (simp add: fun_eq_iff cut_def) lemma cut_apply: "(x, a) \ R \ cut f R a x = f x" by (simp add: cut_def) text \ Inductive characterization of \wfrec\ combinator; for details see: John Harrison, "Inductive definitions: automation and application". \ lemma theI_unique: "\!x. P x \ P x \ x = The P" by (auto intro: the_equality[symmetric] theI) lemma wfrec_unique: assumes "adm_wf R F" "wf R" shows "\!y. wfrec_rel R F x y" using \wf R\ proof induct define f where "f y = (THE z. wfrec_rel R F y z)" for y case (less x) then have "\y z. (y, x) \ R \ wfrec_rel R F y z \ z = f y" unfolding f_def by (rule theI_unique) with \adm_wf R F\ show ?case by (subst wfrec_rel.simps) (auto simp: adm_wf_def) qed lemma adm_lemma: "adm_wf R (\f x. F (cut f R x) x)" by (auto simp: adm_wf_def intro!: arg_cong[where f="\x. F x y" for y] cuts_eq[THEN iffD2]) lemma wfrec: "wf R \ wfrec R F a = F (cut (wfrec R F) R a) a" apply (simp add: wfrec_def) apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality]) apply assumption apply (rule wfrec_rel.wfrecI) apply (erule adm_lemma [THEN wfrec_unique, THEN theI']) done text \This form avoids giant explosions in proofs. NOTE USE OF \\\.\ lemma def_wfrec: "f \ wfrec R F \ wf R \ f a = F (cut f R a) a" by (auto intro: wfrec) subsubsection \Well-founded recursion via genuine fixpoints\ lemma wfrec_fixpoint: assumes wf: "wf R" and adm: "adm_wf R F" shows "wfrec R F = F (wfrec R F)" proof (rule ext) fix x have "wfrec R F x = F (cut (wfrec R F) R x) x" using wfrec[of R F] wf by simp also have "\y. (y, x) \ R \ cut (wfrec R F) R x y = wfrec R F y" by (auto simp add: cut_apply) then have "F (cut (wfrec R F) R x) x = F (wfrec R F) x" using adm adm_wf_def[of R F] by auto finally show "wfrec R F x = F (wfrec R F) x" . qed lemma wfrec_def_adm: "f \ wfrec R F \ wf R \ adm_wf R F \ f = F f" using wfrec_fixpoint by simp subsection \Wellfoundedness of \same_fst\\ definition same_fst :: "('a \ bool) \ ('a \ ('b \ 'b) set) \ (('a \ 'b) \ ('a \ 'b)) set" where "same_fst P R = {((x', y'), (x, y)) . x' = x \ P x \ (y',y) \ R x}" \ \For \<^const>\wfrec\ declarations where the first n parameters stay unchanged in the recursive call.\ lemma same_fstI [intro!]: "P x \ (y', y) \ R x \ ((x, y'), (x, y)) \ same_fst P R" by (simp add: same_fst_def) lemma wf_same_fst: assumes "\x. P x \ wf (R x)" shows "wf (same_fst P R)" -proof (clarsimp simp add: wf_def same_fst_def) - fix Q a b - assume *: "\a b. (\x. P a \ (x,b) \ R a \ Q (a,x)) \ Q (a,b)" - show "Q(a,b)" - proof (cases "wf (R a)") - case True - then show ?thesis - by (induction b rule: wf_induct_rule) (use * in blast) - qed (use * assms in blast) +proof - + have "\a b Q. \a b. (\x. P a \ (x, b) \ R a \ Q (a, x)) \ Q (a, b) \ Q (a, b)" + proof - + fix Q a b + assume *: "\a b. (\x. P a \ (x,b) \ R a \ Q (a,x)) \ Q (a,b)" + show "Q(a,b)" + proof (cases "wf (R a)") + case True + then show ?thesis + by (induction b rule: wf_induct_rule) (use * in blast) + qed (use * assms in blast) + qed + then show ?thesis + by (clarsimp simp add: wf_def same_fst_def) qed end